Proving basic properties of recursive "less than" definition for naturals in Isabelle - isabelle

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

Related

Focussing on new subgoals in Eisbach

In Eisbach I can use ; to apply a method to all new subgoals created by a method.
However, I often know how many subgoals are created and would like to apply different methods to the new subgoals.
Is there a way to say something like "apply method X to the first new subgoal and method Y to the second new subgoal"?
Here is a simple use case:
I want to develop a method that works on 2 conjunctions of arbitrary length but with the same structure.
The method should be usable to show that conjunction 1 implies conjunction 2 by showing that the implication holds for each component.
It should be usable like this:
lemma example:
assumes c: "a 0 ∧ a 1 ∧ a 2 ∧ a 3"
and imp: "⋀i. a i ⟹ a' i"
shows "a' 0 ∧ a' 1 ∧ a' 2 ∧ a' 3"
proof (conj_one_by_one pre: c)
show "a 0 ⟹ a' 0" by (rule imp)
show "a 1 ⟹ a' 1" by (rule imp)
show "a 2 ⟹ a' 2" by (rule imp)
show "a 3 ⟹ a' 3" by (rule imp)
qed
When implementing this method in Eisbach, I have a problem after using rule conjI.
I get two subgoals that I want to recursively work on, but I want to use different facts for the two cases.
I came up with the following workaround, which uses artificial markers for the two subgoals and is kind of ugly:
definition "marker_L x ≡ x"
definition "marker_R x ≡ x"
lemma conjI_marked:
assumes "marker_L P" and "marker_R Q"
shows "P ∧ Q"
using assms unfolding marker_L_def marker_R_def by simp
method conj_one_by_one uses pre = (
match pre in
p: "?P ∧ ?Q" ⇒ ‹
(unfold marker_L_def marker_R_def)?,
rule conjI_marked;(
(match conclusion in "marker_L _" ⇒ ‹(conj_one_by_one pre: p[THEN conjunct1])?›)
| (match conclusion in "marker_R _" ⇒ ‹(conj_one_by_one pre: p[THEN conjunct2])?›))›)
| ((unfold marker_L_def marker_R_def)?, insert pre)
This is not a complete answer, but you might be able to derive some useful information from what is stated here.
In Eisbach I can use ; to apply a method to all new subgoals created
by a method. However, I often know how many subgoals are created and
would like to apply different methods to the new subgoals. Is there a
way to say something like "apply method X to the first new subgoal and
method Y to the second new subgoal"?
You can use the standard tactical RANGE to define your own tactic that you can apply to consecutive subgoals. I provide a very specialized and significantly simplified use case below:
ML‹
fun mytac ctxt thms = thms
|> map (fn thm => resolve_tac ctxt (single thm))
|> RANGE
›
lemma
assumes A: A and B: B and C: C
shows "A ∧ B ∧ C"
apply(intro conjI)
apply(tactic‹mytac #{context} [#{thm A}, #{thm B}, #{thm C}] 1›)
done
Hopefully, it should be reasonably easy to extend it to more complicated use cases (while being more careful than I am about subgoal indexing: you might also need SELECT_GOAL to ensure that the implementation is safe). While in the example above mytac accepts a list of theorems, it should be easy to see how these theorems can be replaced by tactics and with some further work, the tactic can be wrapped as a higher-order method.
I want to develop a method that works on 2 conjunctions of arbitrary
length but with the same structure. The method should be usable to
show that conjunction 1 implies conjunction 2 by showing that the
implication holds for each component. It should be usable like this:
UPDATE
Having had another look at the problem, it seems that there exists a substantially more natural solution. The solution follows the outline from the original answer, but the meta implication is replaced with the HOL's object logic implication (the 'to and fro' conversion can be achieved using atomize (full) and intro impI):
lemma arg_imp2: "(a ⟶ b) ⟹ (c ⟶ d) ⟹ ((a ∧ c) ⟶ (b ∧ d))" by auto
lemma example:
assumes "a 0 ∧ a 1 ∧ a 2 ∧ a 3"
and imp: "⋀i. a i ⟹ a' i"
shows "a' 0 ∧ a' 1 ∧ a' 2 ∧ a' 3"
apply(insert assms(1), atomize (full))
apply(intro arg_imp2; intro impI; intro imp; assumption)
done
LEGACY (this was part of the original answer, but is almost irrelevant due to the UPDATE suggested above)
If this is the only application that you have in mind, perhaps, there is a reasonably natural solution based on the following iterative procedure:
lemma arg_imp2: "(a ⟹ b) ⟹ (c ⟹ d) ⟹ ((a ∧ c) ⟹ (b ∧ d))" by auto
lemma example:
assumes c: "a 0 ∧ a 1 ∧ a 2 ∧ a 3"
and imp: "⋀i. a i ⟹ a' i"
shows "a' 0 ∧ a' 1 ∧ a' 2 ∧ a' 3"
using c
apply(intro arg_imp2[of ‹a 0› ‹a' 0› ‹a 1 ∧ a 2 ∧ a 3› ‹a' 1 ∧ a' 2 ∧ a' 3›])
apply(rule imp)
apply(assumption)
apply(intro arg_imp2[of ‹a 1› ‹a' 1› ‹a 2 ∧ a 3› ‹a' 2 ∧ a' 3›])
apply(rule imp)
apply(assumption)
apply(intro arg_imp2[of ‹a 2› ‹a' 2› ‹a 3› ‹a' 3›])
apply(rule imp)
apply(assumption)
apply(rule imp)
apply(assumption+)
done
I am not certain how easy it would be to express this in Eisbach, but it should be reasonably easy to express this in Isabelle/ML.
Using the pointers from user9716869, I was able to write a method that does what I want:
ML‹
fun split_with_tac (tac1: int -> tactic) (ts: (int -> tactic) list) (i: int) (st: thm): thm Seq.seq =
let
val st's = tac1 i st
fun next st' =
let
val new_subgoals_count = 1 + Thm.nprems_of st' - Thm.nprems_of st
in
if new_subgoals_count <> length ts then Seq.empty
else
RANGE ts i st'
end
in
st's |> Seq.maps next
end
fun tok_to_method_text ctxt tok =
case Token.get_value tok of
SOME (Token.Source src) => Method.read ctxt src
| _ =>
let
val (text, src) = Method.read_closure_input ctxt (Token.input_of tok);
val _ = Token.assign (SOME (Token.Source src)) tok;
in text end
val readText: Token.T Token.context_parser = Scan.lift (Parse.token Parse.text)
val text_and_texts_closure: (Method.text * Method.text list) Token.context_parser =
(Args.context -- readText -- (Scan.lift \<^keyword>‹and› |-- Scan.repeat readText)) >> (fn ((ctxt, tok), t) =>
(tok_to_method_text ctxt tok, map (tok_to_method_text ctxt) t));
›
method_setup split_with =
‹text_and_texts_closure >> (fn (m, ms) => fn ctxt => fn facts =>
let
fun tac m st' =
method_evaluate m ctxt facts
fun tac' m i st' =
Goal.restrict i 1 st'
|> method_evaluate m ctxt facts
|> Seq.map (Goal.unrestrict i)
handle THM _ => Seq.empty
val initialT: int -> tactic = tac' m
val nextTs: (int -> tactic) list = map tac' ms
in SIMPLE_METHOD (HEADGOAL (split_with_tac initialT nextTs)) facts end)
›
lemma
assumes r: "P ⟹ Q ⟹ R"
and p: "P"
and q: "Q"
shows "R"
by (split_with ‹rule r› and ‹rule p› ‹rule q›)
method conj_one_by_one uses pre = (
match pre in
p: "?P ∧ ?Q" ⇒ ‹split_with ‹rule conjI› and
‹conj_one_by_one pre: p[THEN conjunct1]›
‹conj_one_by_one pre: p[THEN conjunct2]››
| insert pre)
lemma example:
assumes c: "a 0 ∧ a 1 ∧ a 2 ∧ a 3"
and imp: "⋀i. a i ⟹ a' i"
shows "a' 0 ∧ a' 1 ∧ a' 2 ∧ a' 3"
proof (conj_one_by_one pre: c)
show "a 0 ⟹ a' 0" by (rule imp)
show "a 1 ⟹ a' 1" by (rule imp)
show "a 2 ⟹ a' 2" by (rule imp)
show "a 3 ⟹ a' 3" by (rule imp)
qed

How can I prove irreflexivity of an inductively defined relation in Isabelle?

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).

Proving theorem of the form ~pvq using Induction in Isabelle

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.

How to show that 2 formula sare semantically equivalent in Isabelle

I want to (miss-)use Isabelle to show that two given formulas are syntactically equivalent. For example A ∧ B = B ∧ A.
I don't want to go into any detail with regards towards the logic behind the formulas. I don't want to care that A ∧ B is true when both A and B are true. I just want to compare the two formulas on a structural level and say that they are equivalent because of the commutative property.
Basicly i want to be abled to write lemmas comparing 2 formulas with some equality function and use the given, however they are to be specified, axioms.
So far i thought that this could and should be done using axiomatization, but everyone here tells me axiomatzation is bad.
This leads me to my question how should this task be done. How can 2, say propositional, formulas be compared in Isabelle with regards towards their syntactical equivalence. To give a concrete example:
formula ∧ formula | formula ∨ formula
are given as operators, if possible as datatype
and distributive and commutative property are given as rules.
A ∧ B = B ∧ A, if stated in a theorem should be provable.
That's what i want to do, i hope the idea is clear and someone can explain to me how to pursue this properly in Isabelle. Thanks in advance.
From what you've written, I'm pretty sure you mean syntactic equivalence. Two formulas are semantically equivalent if they evaluate to the same result for all valuations of variables; two formulas are syntactically equivalent if you can rewrite one to another given a certain set of rewrite rules (or, more generally, prove their equivalence using a certain set of inference rules). Semantic equivalence looks only at the values of the expressions and not at their structure; Syntactic equivalence looks only at the structure of the expressions and not the values they produce.
Now, on to answer the question of how to do this in Isabelle.
Defining the relation
The standard way is to define a datatype for your formulas (I added some nicer infix syntax for it):
type_synonym vname = nat
datatype formula =
Atom vname
| FTrue
| Neg formula
| Conj formula formula (infixl "and" 60)
| Disj formula formula (infixl "or" 50)
definition "FFalse = Neg FTrue"
Then you can define the concept of ‘evaluating’ such a formula w.r.t. a given variable valuation:
primrec eval_formula :: "(vname ⇒ bool) ⇒ formula ⇒ bool" where
"eval_formula s (Atom x) ⟷ s x"
| "eval_formula _ FTrue ⟷ True"
| "eval_formula s (Neg a) ⟷ ¬eval_formula s a"
| "eval_formula s (a and b) ⟷ eval_formula s a ∧ eval_formula s b"
| "eval_formula s (a or b) ⟷ eval_formula s a ∨ eval_formula s b"
lemma eval_formula_False [simp]: "eval_formula s FFalse = False"
by (simp add: FFalse_def)
And, building on that, you can define the concept of semantic equivalence: two formulas are semantically equivalent if they evaluate to the same thing for all valuations:
definition formula_equiv_sem :: "formula ⇒ formula ⇒ bool" (infixl "≈" 40) where
"a ≈ b ⟷ (∀s. eval_formula s a = eval_formula s b)"
From your question, I gather that what you want to do is to define some kind of equivalence relation based on rewrite rules: two formulas are syntactically equivalent if you can transform one into another by applying some set of given rewrite rules.
That can be done e.g. with Isabelle's package for inductive predicates:
inductive formula_equiv :: "formula ⇒ formula ⇒ bool" (infixl "∼" 40) where
formula_refl [simp]: "a ∼ a"
| formula_sym: "a ∼ b ⟹ b ∼ a"
| formula_trans [trans]: "a ∼ b ⟹ b ∼ c ⟹ a ∼ c"
| neg_cong: "a ∼ b ⟹ Neg a ∼ Neg b"
| conj_cong: "a1 ∼ b1 ⟹ a2 ∼ b2 ⟹ a1 and a2 ∼ b1 and b2"
| disj_cong: "a1 ∼ b1 ⟹ a2 ∼ b2 ⟹ a1 or a2 ∼ b1 or b2"
| conj_commute: "a and b ∼ b and a"
| disj_commute: "a or b ∼ b or a"
| conj_assoc: "(a and b) and c ∼ a and (b and c)"
| disj_assoc: "(a or b) or c ∼ a or (b or c)"
| disj_conj: "a or (b and c) ∼ (a or b) and (a or c)"
| conj_disj: "a and (b or c) ∼ (a and b) or (a and c)"
| de_morgan1: "Neg (a and b) ∼ Neg a or Neg b"
| de_morgan2: "Neg (a or b) ∼ Neg a and Neg b"
| neg_neg: "Neg (Neg a) ∼ a"
| tnd: "a or Neg a ∼ FTrue"
| contr: "a and Neg a ∼ FFalse"
| disj_idem: "a or a ∼ a"
| conj_idem: "a and a ∼ a"
| conj_True: "a and FTrue ∼ a"
| disj_True: "a or FTrue ∼ FTrue"
The first six rules essentially set up rewriting (you can rewrite anything to itself, you can rewrite from left to right or from right to left, you can chain rewrite steps, if you can rewrite subterms, you can also rewrite the full term). The remaining rules are a few example rules that you may want to have in there (with no claim of completeness).
For more information on inductive predicates, refer to the manual of the predicate tool.
Proving stuff about it
So what can you do with this? Well, I expect that you will want to show that this is sound, i.e. that two formulas that are syntactically equivalent are also semantically equivalent:
lemma formula_equiv_syntactic_imp_semantic:
"a ∼ b ⟹ a ≈ b"
by (induction a b rule: formula_equiv.induct)
(auto simp: formula_equiv_sem_def)
You might also want to prove a few derived syntactical rules. For this, it is useful to have some convenient transitivity rules and setup the simplifier with the congruence rules:
lemmas formula_congs [simp] = neg_cong conj_cong disj_cong
lemma formula_trans_cong1 [trans]:
"a ∼ f b ⟹ b ∼ c ⟹ (⋀x y. x ∼ y ⟹ f x ∼ f y) ⟹ a ∼ f c"
by (rule formula_trans) simp_all
lemma formula_trans_cong2 [trans]:
"a ∼ b ⟹ f b ∼ f c ⟹ (⋀x y. x ∼ y ⟹ f x ∼ f y) ⟹ f a ∼ f c"
by (rule formula_trans) simp_all
Then we can do proofs like this:
lemma conj_False: "a and FFalse ∼ FFalse"
proof -
have "a and FFalse ∼ Neg (Neg (a and FFalse))"
by (rule formula_sym, rule neg_neg)
also have "Neg (a and FFalse) ∼ Neg a or Neg FFalse"
by (rule de_morgan1)
also have "Neg FFalse ∼ FTrue" unfolding FFalse_def by (rule neg_neg)
also have "Neg a or FTrue ∼ FTrue" by (rule disj_True)
also have "Neg FTrue = FFalse" unfolding FFalse_def ..
finally show ?thesis by - simp
qed
lemma disj_False: "a or FFalse ∼ a"
proof -
have "a or FFalse ∼ Neg (Neg (a or FFalse))" by (rule formula_sym, rule neg_neg)
also have "Neg (a or FFalse) ∼ Neg a and Neg FFalse" by (rule de_morgan2)
also have "Neg FFalse ∼ FTrue" unfolding FFalse_def by (rule neg_neg)
also have "Neg a and FTrue ∼ Neg a" by (rule conj_True)
also have "Neg (Neg a) ∼ a" by (rule neg_neg)
finally show ?thesis by - simp
qed
One would, of course, also like to prove completeness, i.e. that any two formulas that are semantically equivalent are also syntactically equivalent. For that, I think you will need a few more rules and then the proof is pretty complicated.
Why not axiomatization?
You mentioned axiomatization, and you might ask why you were advised not to use this for this purpose. Well, one reason is that axiomatization allows you to introduce inconsistencies into the system. You might ‘define’ two things to be equivalent and also define something else at another place that implies that they are not equivalent and then you have derive False and break everything. With the inductive predicate package, that cannot happen, because it proves automatically that your definitions are consistent. (by restricting them to be monotonic)
A more practical reason is that, as you can see above, you can do induction over an inductive predication, i.e. if you have two formulas that are synactically equivalent, you can do induction over the proof tree of their syntactic equivalence. In particular, you know that if two formulas are syntactically equivalent, that must be provable from the rules that you specified. If you just do axiomatization, you have no such guarantee – there may be many more syntactically equivalent formulas; with axiomatization, you could not even disprove something like Atom 0 ≈ Atom 1, unless you axiomatize something like that as well, which will be very ugly and very prone to accidental inconsistencies.
It is very rare for an Isabelle user to use axiomatization. I have been working with and on Isabelle for years and I have never used axiomatization. It is a very low-level feature designed to set up the basic underlying logic and much work has been invested in high-level definitional tools like typedef, datatype, fun, inductive, codatatype etc. to offer a definitional interface to the user that (hopefully) guarantee consistency to the user.
Appendix: Deciding semantic equivalence
If you're interested in using code generation to do interesting things on formulas: We can even decide semantic equivalence: we can simply test if the two expressions evaluate to the same results on the set of variables they containt. (syntactic equivalence is possible as well, but more difficult because you will have to get the inductive predicate package to compile usable code for it, and I did not manage to do that)
primrec vars :: "formula ⇒ vname list" where
"vars (Atom x) = [x]"
| "vars FTrue = []"
| "vars (Neg a) = vars a"
| "vars (Conj a b) = vars a # vars b"
| "vars (Disj a b) = vars a # vars b"
lemma eval_formula_cong:
"(⋀x. x ∈ set (vars a) ⟹ s x = s' x) ⟹ eval_formula s a = eval_formula s' a"
by (induction a) simp_all
primrec valuations :: "vname list ⇒ (vname ⇒ bool) list" where
"valuations [] = [λ_. False]"
| "valuations (x#xs) = [f' . f ← valuations xs, f' ← [f, fun_upd f x True]]"
lemma set_valuations: "set (valuations xs) = {f. ∀x. x∉set xs ⟶ f x = False}"
proof safe
case (goal2 f)
thus ?case
proof (induction xs arbitrary: f)
case (Cons x xs)
def f' ≡ "fun_upd f x False"
from Cons.prems have f': "f' ∈ set (valuations xs)"
by (intro Cons) (auto simp: f'_def)
show ?case
proof (cases "f x")
case False
hence "f' = f" by (intro ext) (simp add: f'_def)
with f' show ?thesis by auto
next
case True
hence "fun_upd f' x True = f" by (intro ext) (simp add: f'_def)
with f' show ?thesis by auto
qed
qed auto
qed (induction xs, auto)
lemma formula_equiv_sem_code [code]:
"a ≈ b ⟷ (∀s∈set (valuations (remdups (vars a # vars b))).
eval_formula s a = eval_formula s b)"
unfolding formula_equiv_sem_def
proof (rule iffI; rule ballI allI)
case (goal2 s)
def s' ≡ "λx. if x ∈ set (vars a # vars b) then s x else False"
have "s' ∈ set (valuations (remdups (vars a # vars b)))"
by (subst set_valuations) (auto simp: s'_def)
with goal2 have "eval_formula s' a = eval_formula s' b" by blast
also have "eval_formula s' a = eval_formula s a"
by (intro eval_formula_cong) (auto simp: s'_def)
also have "eval_formula s' b = eval_formula s b"
by (intro eval_formula_cong) (auto simp: s'_def)
finally show ?case .
qed auto
We can now simply ask Isabelle to compute whether two formulas are semantically equivalent:
value "Atom 0 and Atom 1 ≈ Atom 1 and Atom 0" (* True *)
value "Atom 0 and Atom 1 ≈ Atom 1 or Atom 0" (* False *)
You could even go further and write an automated proof method that decides a ≈ b for any formulas a and b by substituting all free variables with fresh atoms and then deciding the equivalence of those formulas (e.g. decide a and FFalse ≈ FFalse by deciding the equivalent Atom 0 and FFalse ≈ FFalse).

Induction on Recursive Function with a twist

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

Resources