I'm trying to proof the theorem that if n > 0 then g n b = True (see below). This is the case, because g (Suc n) b only ever calls g 0 True. Unfortunately, I don't have that fact in my induction when I try to proof g 0 b. How can I finish the proof (what do I have to replace the sorry with)?
fun g :: "nat ⇒ bool ⇒ bool" where
"g (Suc n) b = g n True" |
"g 0 b = b"
theorem
fixes n::nat and b::bool
assumes "n > 0"
shows "g n b"
proof (induct n b rule: g.induct)
fix n
fix b
assume "g n True"
thus "g (Suc n) b" by (metis g.simps(1))
next
fix b
show "g 0 b" sorry
qed
You forgot to use the assumption n > 0 in your induction.
E.g., you may write
theorem
fixes n::nat and b::bool
assumes "n > 0"
shows "g n b"
using assms (* this is important *)
proof (induct n b rule: g.induct)
case (1 n b)
thus ?case by (cases n) auto
next
case (2 b)
thus ?case by auto
qed
Alternatively you may immediately start your theorem like this
and shorten it further:
theorem "n > 0 ==> g n b"
proof (induct n b rule: g.induct)
case (1 n b)
thus ?case by (cases n) auto
qed auto
Related
I was trying to recreate a simplified version of the natural numbers, for learning purposes (as it involves inductive definitions, recursive functions, etc...). In that process however, I got stuck in something that I thought would be very trivial.
Basically, I have a definition for natural numbers 'natt' and a definition for the '<' relation:
datatype natt = Zero | Succ natt
primrec natt_less :: "natt ⇒ natt ⇒ bool" (infixl "<" 75) where
"natt_less n Zero = False"
| "natt_less n (Succ m') = (case n of Zero ⇒ True | Succ n' ⇒ natt_less n' m')"
and from these, I tried to prove 3 basic properties of the < relation:
Non-reflexivity: ~ (a < a)
Non-symmetry: a < b ⟹ ~ (b < a)
Transitivity: a < b ⟹ b < c ⟹ a < c
I was able to prove the first, but not the others. What took me even more by surprise, is that there are some sub-lemmas that would aid on these, such as Succ a < b ⟹ a < b, a < b ⟹ a < Succ b or a < b ∨ a = b ∨ b < a, which seem even more trivial, but nonetheless I was also not able to prove, even after many attempts. It seems like only one of these (including 2. and 3.) is enough to prove the rest, but I wasn't able to prove any of them.
I'm mostly trying to use induction. Together with the fact that I've made the definitions myself, there are two possibilities - Either my definitions are wrong, and do not have the desired properties, or I'm missing some method/argument. So, I have two questions:
Is my definition wrong (i.e. it does not accurately represent < and lacks the desired properties)? If so, how may I fix it?
If not, how can I prove these seemingly trivial properties?
For context, my current attempts are by induction, which I can prove the base case, but always get stuck in the induction case, not really knowing where to go with the assumptions, such as in this example:
lemma less_Succ_1: "Succ a < b ⟹ a < b"
proof (induction b)
case Zero
assume "Succ a < Zero"
then have "False" by simp
then show ?case by simp
next
case (Succ b)
assume "(Succ a < b ⟹ a < b)" "Succ a < Succ b"
then show "a < Succ b" oops
After some tips from user9716869, it's clear that my main problem was the lack of knowledge about the arbitrary option in induction. Using (induction _ arbitrary: _) and (cases _) (see the reference manual for details), the proofs are quite straight forward.
Since these are made for educational purposes, the following proofs are not meant to be concise, but to make every step very clear. Most of these could be vastly reduced if more automation is desired, and some can be done in one line (which I left as a comment below the lemma).
Note: In these proofs, we are using an implicit lemma about inductive types, their injectivity (which implies (Succ a = Succ b) ≡ (a = b) and Zero ≠ Succ a). Furthermore, (Succ a < Succ b) ≡ (a < b) by definition.
First, we prove 2 useful lemmas:
a < b ⟹ b ≠ Zero
b ≠ Zero ⟷ (∃ b'. b = Succ b')
lemma greater_not_Zero [simp]: "a < b ⟹ b ≠ Zero"
(*by clarsimp*)
proof
assume "a < b" "b = Zero"
then have "a < Zero" by simp
then show "False" by simp
qed
lemma not_Zero_is_Succ: "b ≠ Zero ⟷ (∃ b'. b = Succ b')"
(*by (standard, cases b) auto*)
proof
show "b ≠ Zero ⟹ ∃ b'. b = Succ b'"
proof (cases b)
case Zero
assume ‹b ≠ Zero›
moreover note ‹b = Zero›
ultimately show "∃b'. b = Succ b'" by contradiction
next
case (Succ b')
assume ‹b ≠ Zero›
note ‹b = Succ b'›
then show "∃b'. b = Succ b'" by simp
qed
next
assume "∃ b'. b = Succ b'"
then obtain b'::natt where "b = Succ b'" by clarsimp
then show "b ≠ Zero" by simp
qed
Armed with these, we can prove the 3 main statements:
Non-reflexivity: ~ (a < a)
Non-symmetry: a < b ⟹ ~ (b < a)
Transitivity: a < b ⟹ b < c ⟹ a < c
lemma less_not_refl [simp]: "¬ a < a"
(*by (induction a) auto*)
proof (induction a)
case Zero
show "¬ Zero < Zero" by simp
next
case (Succ a)
note IH = ‹¬ a < a›
show "¬ Succ a < Succ a"
proof
assume "Succ a < Succ a"
then have "a < a" by simp
then show "False" using IH by contradiction
qed
qed
lemma less_not_sym: "a < b ⟹ ¬ b < a"
proof (induction a arbitrary: b)
case Zero
then show "¬ b < Zero" by simp
next
case (Succ a)
note IH = ‹⋀b. a < b ⟹ ¬ b < a›
and IH_prems = ‹Succ a < b›
show "¬ b < Succ a"
proof
assume asm:"b < Succ a"
have "b ≠ Zero" using IH_prems by simp
then obtain b'::natt where eq: "b = Succ (b')"
using not_Zero_is_Succ by clarsimp
then have "b' < a" using asm by simp
then have "¬ a < b'" using IH by clarsimp
moreover have "a < b'" using IH_prems eq by simp
ultimately show "False" by contradiction
qed
qed
lemma less_trans [trans]: "a < b ⟹ b < c ⟹ a < c"
proof (induction c arbitrary: a b)
case Zero
note ‹b < Zero›
then have "False" by simp
then show ?case by simp
next
case (Succ c)
note IH = ‹⋀a b. a < b ⟹ b < c ⟹ a < c›
and IH_prems = ‹a < b› ‹b < Succ c›
show "a < Succ c"
proof (cases a)
case Zero
note ‹a = Zero›
then show "a < Succ c" by simp
next
case (Succ a')
note cs_prem = ‹a = Succ a'›
have "b ≠ Zero" using IH_prems by simp
then obtain b' where b_eq: "b = Succ b'"
using not_Zero_is_Succ by clarsimp
then have "a' < b'" using IH_prems cs_prem b_eq by simp
moreover have "b' < c" using IH_prems b_eq by simp
ultimately have "a' < c" using IH by simp
then show "a < Succ c" using cs_prem by simp
qed
qed
I find myself solving a goal that with safe splits to 32 subgoals. It is a quite algebraic goal so overall I need to use argo, algebra and auto. I was wondering if there is a way to specify that auto should be applied say 2 times, then algebra 10 times etc. Where should I look for this syntax in the future? Is it part of eisbach?
There is the REPEAT_DETERM_N tactical in $ISABELLE_HOME/src/Pure/tactical.ML I never used it so I'm not 100% sure it's what you need.
Alternatively your functionality can be done somewhat like that:
theory NTimes
imports
Main
"~~/src/HOL/Eisbach/Eisbach"
begin
ML ‹
infixr 2 TIMES
fun 0 TIMES _ = all_tac
| n TIMES tac = tac THEN (n - 1) TIMES tac
›
notepad
begin
fix A B C D
have test1: "A ∧ B ∧ C ∧ D ⟹ True"
apply (tactic ‹3 TIMES eresolve_tac #{context} [#{thm conjE}] 1›)
apply (rule TrueI)
done
fix E
have test2: "A ∧ B ∧ C ∧ D ∧ E ⟹ True"
apply (tactic ‹2 TIMES 2 TIMES eresolve_tac #{context} [#{thm conjE}] 1›)
apply (rule TrueI)
done
end
(* For good examples for working
with higher order methods in ML see $ISABELLE_HOME/src/HOL/Eisbach/Eisbach.thy *)
method_setup ntimes = ‹
Scan.lift Parse.nat -- Method.text_closure >>
(fn (n, closure) => fn ctxt => fn facts =>
let
val tac = method_evaluate closure ctxt facts
in
SIMPLE_METHOD (n TIMES tac) facts
end)
›
notepad
begin
fix A B C D
have test1: "A ∧ B ∧ C ∧ D ⟹ True"
apply (ntimes 3 ‹erule conjE›)
apply (rule TrueI)
done
fix E
have test2: "A ∧ B ∧ C ∧ D ∧ E ⟹ True"
apply (ntimes 2 ‹ntimes 2 ‹erule conjE››)
apply (rule TrueI)
done
have test3: "A ∧ B ∧ C ∧ D ∧ E ⟹ True"
apply (ntimes 3 ‹erule conjE›)
apply (rule TrueI)
done
have test4: "A = A" "B = B" "C = C"
apply -
apply (ntimes 2 ‹fastforce›)
apply (rule refl)
done
(* in some examples one can instead use subgoal ranges *)
have test5: "A = A" "B = B" "C = C"
apply -
apply (fastforce+)[2]
apply (rule refl)
done
end
end
I'm not an expert in Isabelle/ML Programming so this code is likely of low quality, but I hope it's a good starting point for you!
Consider as an example the following definition of inequality of natural numbers in Isabelle:
inductive unequal :: "nat ⇒ nat ⇒ bool" where
zero_suc: "unequal 0 (Suc _)" |
suc_zero: "unequal (Suc _) 0" |
suc_suc: "unequal n m ⟹ unequal (Suc n) (Suc m)"
I want to prove irreflexivity of unequal, that is, ¬ unequal n n. For illustration purposes let me first prove the contrived lemma ¬ unequal (n + m) (n + m):
lemma "¬ unequal (n + m) (n + m)"
proof
assume "unequal (n + m) (n + m)"
then show False
proof (induction "n + m" "n + m" arbitrary: n m)
case zero_suc
then show False by simp
next
case suc_zero
then show False by simp
next
case suc_suc
then show False by presburger
qed
qed
In the first two cases, False must be deduced from the assumptions 0 = n + m and Suc _ = n + m, which is trivial.
I would expect that the proof of ¬ unequal n n can be done in an analogous way, that is, according to the following pattern:
lemma "¬ unequal n n"
proof
assume "unequal n n"
then show False
proof (induction n n arbitrary: n)
case zero_suc
then show False sorry
next
case suc_zero
then show False sorry
next
case suc_suc
then show False sorry
qed
qed
In particular, I would expect that in the first two cases, I get the assumptions 0 = n and Suc _ = n. However, I get no assumptions at all, meaning that I am asked to prove False from nothing. Why is this and how can I conduct the proof of inequality?
You are inducting over unequal. Instead, you should induct over n, like this:
lemma "¬ (unequal n n)"
proof (induct n)
case 0
then show ?case sorry
next
case (Suc n)
then show ?case sorry
qed
Then we can use Sledgehammer on each of the subgoals marked with sorry. Sledgehammer (with CVC4) recommends us to complete the proof as follows:
lemma "¬ (unequal n n)"
proof (induct n)
case 0
then show ?case using unequal.cases by blast
next
case (Suc n)
then show ?case using unequal.cases by blast
qed
The induction method handles variable instantiations and non-variable instantiations differently. A non-variable instantiation t is a shorthand for x ≡ t where x is a fresh variable. As a result, induction is done on x, and the context additionally contains the definition x ≡ t.
Therefore, (induction "n + m" "n + m" arbitrary: n m) in the first proof is equivalent to (induction k ≡ "n + m" l ≡ "n + m" arbitrary: n m) with the effect described above. To get this effect for the second proof, you have to replace (induction n n arbitrary: n) with (induction k ≡ n l ≡ n arbitrary: n). The assumptions will actually become so simple that the pre-simplifier, which is run by the induction method, can derive False from them. As a result, there will be no cases left to prove, and you can replace the whole inner proof–qed block with by (induction k ≡ n l ≡ n arbitrary: n).
I have a function called feedback which calculates the power of 3 (i.e)
feedback(t) = 3^t
primrec feedback :: "nat ⇒ nat" where
"feedback 0 = Suc(0)"|
"feedback (Suc t) = (feedback t)*3"
I want to prove that
if t > 5 then feedback(t) > 200
using Induction
lemma th2: "¬(t>5) ∨ ((feedback t) > 200)" (is "?H(t)" is "?P(t)∨?Q(t)" is "(?P(t))∨(?F(t) > 200)")
proof(induct t)
case 0 show "?P 0 ∨ ?Q 0" by simp
next
assume a:" ?F(t) > 200"
assume d: "?P(t) = False"
have b: "?F (Suc(t)) ≥ ?F(t)" by simp
from b and a have c: "?F(Suc(t)) > 200" by simp
from c have e: "?Q(Suc(t))" by simp
from d have f:"?P(Suc(t)) = False" by simp
from f and e have g: "?P(Suc(t))∨?Q(Suc(t))" by simp
from a and d and g have h: "?P(t)∨?Q(t) ⟹ ?P(Suc(t))∨?Q(Suc(t))" by simp
from a and d have "?H(Suc(t))" by simp
qed
First I prove that
feedback(t+1) >= feedback(t)
then assume feedback(t) > 200, so feedback(t+1)>200
Assume t>5
this implies (t+1) > 5
Also ~((t+1)>5) V (feedback (t+1) > 200) is True
Thus if P(t) is true then P(t+1) is true
But this is not working. I have no idea what the problem is
Well, first of all, you cannot simply assume arbitrary things in Isar. Or rather, you can do that, but you won't be able to show your goal after you've done that. The things that Isar allows you to assume are quite rigid; in your case, it's ¬ 5 < t ∨ 200 < feedback t.
I recommend using the case command, which assumes the right things for you. Then you can do a case distinction about that disjunction and then another one about whether t = 5:
lemma th2: "¬(t>5) ∨ ((feedback t) > 200)"
proof (induct t)
case 0
show ?case by simp
next
case (Suc t)
thus ?case
proof
assume "¬t > 5"
moreover have "feedback 6 = 729" by code_simp
-- ‹"simp add: eval_nat_numeral" would also work›
ultimately show ?thesis
by (cases "t = 5") auto
next
assume "feedback t > 200"
thus ?thesis by simp
qed
qed
Or, more compactly:
lemma th2: "¬(t>5) ∨ ((feedback t) > 200)"
proof (induct t)
case (Suc t)
moreover have "feedback 6 = 729" by code_simp
ultimately show ?case by (cases "t = 5") auto
qed simp_all
If your feedback function is actually monotonic, I would recommend proving that first, then the proof becomes a little less tedious.
I have a datatype and an inductive predicate over it (which is actually a small-step semantics of some transition system):
datatype dtype = E | A | B dtype
inductive dsem :: "dtype ⇒ dtype ⇒ bool" where
"dsem A E"
| "dsem (B E) E"
| "dsem d d' ⟹ dsem (B d) (B d')"
and also a function which is computed by case distinction:
fun f :: "dtype ⇒ nat" where
"f E = 0"
| "f A = 1"
| "f (B _) = 2"
I'm trying to prove some property about the inductive predicate, and assumptions also involve computing the value of f which doesn't participate in induction.
lemma
assumes d: "dsem s s'"
and h: "h s v"
and v: "v = f s"
shows "P v"
using d h
proof (induct rule: dsem.induct)
For the 3rd semantics rule Isabelle computes the subgoal
⋀d d'. dsem d d' ⟹ (h d v ⟹ P v) ⟹ h (B d) v ⟹ P v
where the value of s is lost so it is impossible to compute the value v.
I can neither include v into the induction assumptions because then Isabelle generates the subgoal
⋀d d'. dsem d d' ⟹ (h d v ⟹ v = f d ⟹ P v) ⟹ h (B d) v ⟹ v = f (B d) ⟹ P v
where the induction hypothesis says v = f d which is incorrect since v = f (B d) in this case. Nor can I put v into arbitrary: ... because the value of v must be fixed throughout the proof.
It would be nice to have an explicit binding s = B d in the generated subgoal; unfortunately, the rule dsem.induct doesn't provide it.
Does anybody know a workaround for computing the value v in this case?
It seems strange to me that v should be at the same time fixed and computed from s and that is what chris is saying in the comments.
If the solution Brian gives in the comments is what you want, it could duplicate the expression f s which could be big (and use s several times) and perhaps the point of the assumption v = f s was to avoid this.
A first workaround (that was possibly what Brian implicitly proposed) is to make Isabelle do the unfolding:
lemma
assumes d: "dsem s s'"
and h: "h s v"
and v: "v = big_f s s"
shows "P v"
using d h
unfolding v -- {* <<<< *}
proof (induct rule: dsem.induct)
A second workaround could be to abbreviate big_f instead of big_f s s:
lemma
assumes d: "dsem s s'"
and h: "h s (f s)"
and v: "f = (λs. big_f s s)" -- {* <<<< *}
shows "P (f s)"
using d h
proof (induct rule: dsem.induct)