How to define a termination order for the function with fmmap_keys? - isabelle

I'm trying to define a supremum operation for a datatype based on fmap:
datatype t = A | B | C "(nat, t) fmap"
abbreviation
"supc f xs ys ≡
fmmap_keys
(λk x. f x (the (fmlookup ys k)))
(fmfilter (λk. k |∈| fmdom ys) xs)"
fun sup_t (infixl "⊔" 65) where
"A ⊔ _ = A"
| "B ⊔ B = B"
| "B ⊔ _ = A"
| "C xs ⊔ C ys = C (supc (⊔) xs ys)"
| "C xs ⊔ _ = A"
And get the error:
Unfinished subgoals:
(a, 1, <):
1. ⋀ys x. size (the (fmlookup ys x)) < Suc (∑x∈fset (fset_of_fmap ys). Suc (case x of (a, x) ⇒ size x))
(a, 1, <=):
1. ⋀ys x. size (the (fmlookup ys x)) ≤ Suc (∑x∈fset (fset_of_fmap ys). Suc (case x of (a, x) ⇒ size x))
(a, 2, <):
1. ⋀xs xa. size xa < Suc (∑x∈fset (fset_of_fmap xs). Suc (case x of (a, x) ⇒ size x))
(a, 2, <=):
1. ⋀xs xa. size xa ≤ Suc (∑x∈fset (fset_of_fmap xs). Suc (case x of (a, x) ⇒ size x))
(a, 3, <):
1. False
Calls:
a) (C xs, C ys) ~> (xa, the (fmlookup ys x))
Measures:
1) λp. size (snd p)
2) λp. size (fst p)
3) size
Result matrix:
1 2 3
a: ? ? <=
Could not find lexicographic termination order.
If I simplify the function passed as the first argument to fmmap_keys, then the error disappears:
abbreviation
"supc f xs ys ≡
fmmap_keys
(λk x. x)
(fmfilter (λk. k |∈| fmdom ys) xs)"
So I guess, that the error is caused by a complex recursive call of sup_t. The only possible source of non-termination is structures of the form C («[x ↦ C (...)]»). But an external C is removed on each recursive call so the function should terminate.
Could you suggest how to fix this error or redefine supc?
UPDATE
Here is an alternative definition:
abbreviation
"supc f xs ys ≡
fmap_of_list (map
(λ(k, x). (k, f x (the (fmlookup ys k))))
(sorted_list_of_fmap (fmfilter (λk. k |∈| fmdom ys) xs)))"
function sup_t (infixl "⊔" 65) where
"A ⊔ _ = A"
| "B ⊔ x = (if x = B then B else A)"
| "C xs ⊔ x = (case x of C ys ⇒ C (supc sup_t xs ys) | _ ⇒ A)"
by pat_completeness auto
termination
apply auto
I have to prove the following subgoal:
⋀a b. sup_t_dom (a, b)
How to unfold sup_t_dom?

Please find a potentially viable solution in the code listing below.
Background
The issue that you have encountered is described partially in the document "Defining Recursive Functions in Isabelle/HOL" written by Alexander Krauss (also known as "Tutorial on Function Definitions" in Isabelle documentation) and more comprehensively in the PhD thesis "Automating Recursive Definitions and Termination Proofs in Higher-Order Logic" that was also written by Alexander Krauss. In particular, see Chapter 4 in the tutorial and section 3.3 in the thesis.
Size of t
From the aforementioned references, it is possible to infer that one way to prove the termination of sup_t is to provide a suitable measure function. In this case, it is apparent that a measure function that is associated with the size of the datatype might be suitable for the application. Unfortunately, t is a nested type and (in this particular case) the default function size does not seem to capture the recursive nature of the datatype - this is not always the case (see section 3.3.2 in the thesis).
As the first step, I provided a new size function for t. The definition is based on the total number of Cs contained within x::t (the definition should be easy to modify to suit your needs for other applications).
Measure and Termination
I found that the measure function (λ(xs, ys). size ys) is suitable to prove the termination of sup_t. Also, this measure function is used in Isabelle to prove the termination of sup_t if it is declared with the command fun. However, in this case, it was not possible to prove that the arguments of recursive calls indeed decrease with respect to the relation that was established by the measure automatically. However, it would be sufficient to show that "size (the (fmlookup x k)) < size (C x)".
Unfortunately, the function supc (as stated in your question) is not self-certifying with respect to the property that the first argument that is passed to (λk x. f x (the (fmlookup ys k))) is in the domain of ys. Therefore, the (fmlookup ys k) can take the value the None. Given that this issue seems to be nearly orthogonal to the main topic of the question, I decided not to investigate it further and made an amendment to the function supc to ensure that it is guaranteed to return a concrete term of t (you may want to prove explicitly that the function specified below is identical in its behaviour to the one that you provided in your question or, otherwise, provide a better alternative that would be self-certifying):
abbreviation
"supc f xs ys ≡
fmmap_keys
(λk x. if (k |∈| fmdom ys) then (f x (the (fmlookup ys k))) else A)
(fmfilter (λk. k |∈| fmdom ys) xs)"
After this modification, the previous goal "size (the (fmlookup x k)) < size (C x)" was changed to "(k |∈| fmdom ys) ⟹ size (the (fmlookup x k)) < size (C x)", which could be easily proven (see lemma measure_cond). If this lemma is declared as an introduction rule then the termination of sup_t can be proven automatically if it is declared with the command fun.
Remarks
The main reason why I decided to investigate this issue and provide an answer is that I knew very little about some of the main topics of the question and wanted to learn them. As a result, my answer may be suboptimal due to the lack of experience/knowledge in these areas. Of course, if you also have doubts about whether the solution that I proposed here is optimal for the application, it may be worth trying to ask the question on the mailing list.
theory termination_problem
imports
Complex_Main
"HOL-Library.Finite_Map"
begin
datatype (plugins del: "size") t = A | B | C "(nat, t) fmap"
abbreviation "tcf ≡ (λ v::(nat × nat). (λ r::nat. snd v + r))"
interpretation tcf: comp_fun_commute tcf
proof
fix x y
show "tcf y ∘ tcf x = tcf x ∘ tcf y"
proof -
fix z
have "(tcf y ∘ tcf x) z = snd y + snd x + z" by auto
also have "(tcf x ∘ tcf y) z = snd y + snd x + z" by auto
ultimately have "(tcf y ∘ tcf x) z = (tcf x ∘ tcf y) z" by auto
then show "(tcf y ∘ tcf x) = (tcf x ∘ tcf y)" by auto
qed
qed
instantiation t :: size
begin
primrec t_size :: "t ⇒ nat" where
AR: "t_size A = 0" |
BR: "t_size B = 0" |
CR: "t_size (C x) =
(Suc 0) + ffold tcf 0 (fset_of_fmap (fmmap t_size x))"
definition size_t where
size_t_def: "size_t = t_size"
instance ..
end
lemma ffold_rec_exp:
assumes "k |∈| fmdom x"
and "ky = (k, the (fmlookup (fmmap t_size x) k))"
shows "ffold tcf 0 (fset_of_fmap (fmmap t_size x)) =
tcf ky (ffold tcf 0 ((fset_of_fmap (fmmap t_size x)) |-| {|ky|}))"
using assms tcf.ffold_rec by auto
lemma elem_le_ffold:
assumes "k |∈| fmdom x"
shows "t_size (the (fmlookup x k)) <
(Suc 0) + ffold tcf 0 (fset_of_fmap (fmmap t_size x))"
using ffold_rec_exp assms by auto
lemma measure_cond [intro]:
assumes "k |∈| fmdom x"
shows "size (the (fmlookup x k)) < size (C x)"
using assms elem_le_ffold size_t_def by auto
abbreviation
"supc f xs ys ≡
fmmap_keys
(λk x. if (k |∈| fmdom ys) then (f x (the (fmlookup ys k))) else A)
(fmfilter (λk. k |∈| fmdom ys) xs)"
fun sup_t (infixl "⊔" 65) where
"A ⊔ _ = A"
| "B ⊔ x = (if x = B then B else A)"
| "C xs ⊔ x = (case x of C ys ⇒ C (supc sup_t xs ys) | _ ⇒ A)"
(*Examples*)
abbreviation "list_1 ≡ fmap_of_list [(1::nat, B)]"
abbreviation "list_2 ≡ fmap_of_list [(1::nat, A), (2::nat, A)]"
value "(C list_1) ⊔ (C list_2)"
abbreviation "list_3 ≡ fmap_of_list [(1::nat, B), (3::nat, A)]"
abbreviation "list_4 ≡ fmap_of_list [(2::nat, A), (4::nat, B)]"
value "(C list_3) ⊔ (C list_4)"
abbreviation "list_5 ≡ fmap_of_list [(1::nat, B), (2::nat, B)]"
abbreviation "list_6 ≡ fmap_of_list [(2::nat, B), (4::nat, B)]"
value "(C list_5) ⊔ (C list_6)"
abbreviation "list_7 ≡
fmap_of_list [(1::nat, B), (2::nat, C list_5), (3::nat, A)]"
abbreviation "list_8 ≡ fmap_of_list [(2::nat, B), (4::nat, B)]"
value "(C list_7) ⊔ (C list_8)"
abbreviation "list_9 ≡
fmap_of_list [(1::nat, B), (2::nat, C list_5), (3::nat, A)]"
abbreviation "list_10 ≡ fmap_of_list [(2::nat, C list_6), (3::nat, B)]"
value "(C list_9) ⊔ (C list_10)"
end

Related

Isabelle Failed to refine any pending goal during instantiation

datatype 'a list = Cons 'a "'a list" | Nil
instantiation list :: (order) order
begin
fun less_eq_list :: "'a list ⇒ 'a list ⇒ bool" where
"less_eq_list Nil Nil = True" |
"less_eq_list (Cons _ _) Nil = True" |
"less_eq_list Nil (Cons _ _) = False" |
"less_eq_list (Cons _ a) (Cons _ b) = less_eq_list a b"
instance
proof
fix x y:: "'a list"
show "x ≤ x"
apply(induct_tac x)
apply(auto)
done
(* at this point the state is
show x ≤ x
Successful attempt to solve goal by exported rule:
?x2 ≤ ?x2
proof (state)
this:
x ≤ x
goal (3 subgoals):
1. ⋀x y. (x < y) = (x ≤ y ∧ ¬ y ≤ x)
2. ⋀x y z. x ≤ y ⟹ y ≤ z ⟹ x ≤ z
3. ⋀x y. x ≤ y ⟹ y ≤ x ⟹ x = y
*)
show "(x < y) = (x ≤ y ∧ ¬ y ≤ x)"
(* I get an error here
Failed to refine any pending goal
Local statement fails to refine any pending goal
Failed attempt to solve goal by exported rule:
(?x2 < ?y2) = (?x2 ≤ ?y2 ∧ ¬ ?y2 ≤ ?x2)
*)
qed
end
What is wrong with this? The proof of "x ≤ x" worked like a charm. Somehow "(x < y) = (x ≤ y ∧ ¬ y ≤ x)" doesn't match any subgoal.
Class order is a subclass of preorder, which in turn is a subclass of ord. Class ord requires you to define both less_eq (≤) and less (<). In your code, you have correctly defined less_eq_list but forgot to define less_list, and that's why you got an error when trying to prove (x < y) = (x ≤ y ∧ ¬ y ≤ x).

Isabelle structure proof

There is a set of some structures. I'm trying to prove that the cardinality of the set equals some number. Full theory is too long to post here. So here is a simplified one just to show the idea.
Let the objects (which I need to count) are sets containing natural numbers from 1 to n. The idea of the proof is as follows. I define a function which transforms sets to lists of 0 and 1. Here is the function and its inverse:
fun set_to_bitmap :: "nat set ⇒ nat ⇒ nat ⇒ nat list" where
"set_to_bitmap xs x 0 = []"
| "set_to_bitmap xs x (Suc n) =
(if x ∈ xs then Suc 0 else 0) # set_to_bitmap xs (Suc x) n"
fun bitmap_to_set :: "nat list ⇒ nat ⇒ nat set" where
"bitmap_to_set [] n = {}"
| "bitmap_to_set (x#xs) n =
(if x = Suc 0 then {n} else {}) ∪ bitmap_to_set xs (Suc n)"
value "set_to_bitmap {1,3,7,8} 1 8"
value "bitmap_to_set (set_to_bitmap {1,3,7,8} 1 8) 1"
Then I plan to prove that 1) a number of 0/1 lists with length n equals 2^^n,
2) the functions are bijections,
3) so the cardinality of the original set is 2^^n too.
Here are some auxiliary definitions and lemmas, which seems useful:
definition "valid_set xs n ≡ (∀a. a ∈ xs ⟶ 0 < a ∧ a ≤ n)"
definition "valid_bitmap ps n ≡ length ps = n ∧ set ps ⊆ {0, Suc 0}"
lemma length_set_to_bitmap:
"valid_set xs n ⟹
x = Suc 0 ⟹
length (set_to_bitmap xs x n) = n"
apply (induct xs x n rule: set_to_bitmap.induct)
apply simp
sorry
lemma bitmap_members:
"valid_set xs n ⟹
x = Suc 0 ⟹
set_to_bitmap xs x n = ps ⟹
set ps ⊆ {0, Suc 0}"
apply (induct xs x n arbitrary: ps rule: set_to_bitmap.induct)
apply simp
sorry
lemma valid_set_to_valid_bitmap:
"valid_set xs n ⟹
x = Suc 0 ⟹
set_to_bitmap xs x n = ps ⟹
valid_bitmap ps n"
unfolding valid_bitmap_def
using bitmap_members length_set_to_bitmap by auto
lemma valid_bitmap_to_valid_set:
"valid_bitmap ps n ⟹
x = Suc 0 ⟹
bitmap_to_set ps x = xs ⟹
valid_set xs n"
sorry
lemma set_to_bitmap_inj:
"valid_set xs n ⟹
valid_set xy n ⟹
x = Suc 0 ⟹
set_to_bitmap xs x n = ps ⟹
set_to_bitmap ys x n = qs ⟹
ps = qs ⟹
xs = ys"
sorry
lemma set_to_bitmap_surj:
"valid_bitmap ps n ⟹
x = Suc 0 ⟹
∃xs. set_to_bitmap xs x n = ps"
sorry
lemma bitmap_to_set_to_bitmap_id:
"valid_set xs n ⟹
x = Suc 0 ⟹
bitmap_to_set (set_to_bitmap xs x n) x = xs"
sorry
lemma set_to_bitmap_to_set_id:
"valid_bitmap ps n ⟹
x = Suc 0 ⟹
set_to_bitmap (bitmap_to_set ps x) x n = ps"
sorry
Here is a final lemma:
lemma valid_set_size:
"card {xs. valid_set xs n} = 2 ^^ n"
Does this approach seem valid? Are there any examples of such a proof? Could you suggest an idea on how to prove the lemmas? I'm stuck because the induction with set_to_bitmap.induct seems to be not applicable here.
In principle, that kind of approach does work: if you have a function f from a set A to a set B and an inverse function to it, you can prove bij_betw f A B (read: f is a bijection from A to B), and that then implies card A = card B.
However, there are a few comments that I have:
You should use bool lists instead of nat lists if you can only have 0 or 1 in them anyway.
It is usually better to use existing library functions than to define new ones yourself. Your two functions could be defined using library functions like this:
set_to_bitmap :: nat ⇒ nat ⇒ nat set ⇒ bool list
set_to_bitmap x n A = map (λi. i ∈ A) [x..<x+n]
bitmap_to_set :: nat ⇒ bool list ⇒ nat set
bitmap_to_set n xs = (λi. i + n) ` {i. i < length xs ∧ xs ! i}```
Side note: I would use upper-case letters for sets, not something like xs (which is usually used for lists).
Perhaps this is because you simplified your problem, but in its present form, valid_set A n is simply the same as A ⊆ {1..n} and the {A. valid_set A n} is simply Pow {1..n}. The cardinality of that is easy to show with results from the library:
lemma "card (Pow {1..(n::nat)}) = 2 ^ n"
by (simp add: card_Pow)`
As for your original questions: Your first few lemmas are provable, but for the induction to go through, you have to get rid of some of the unneeded assumptions first. The x = Suc 0 is the worst one – there is no way you can use induction if you have that as an assumption, because as soon as you do one induction step, you increase x by 1 and so you won't be able to apply your induction hypothesis. The following versions of your first three lemmas go through easily:
lemma length_set_to_bitmap:
"length (set_to_bitmap xs x n) = n"
by (induct xs x n rule: set_to_bitmap.induct) auto
lemma bitmap_members:
"set (set_to_bitmap xs x n) ⊆ {0, Suc 0}"
by (induct xs x n rule: set_to_bitmap.induct) auto
lemma valid_set_to_valid_bitmap: "valid_bitmap (set_to_bitmap xs x n) n"
unfolding valid_bitmap_def
using bitmap_members length_set_to_bitmap by auto
I also recommend not adding "abbreviations" like ps = set_to_bitmap xs x n as an assumption. It doesn't break anything, but it tends to complicate things needlessly.
The next lemma is a bit trickier. Due to your recursive definitions, you have to generalise the lemma first (valid_bitmap requires the set to be in the range from 1 to n, but once you make one induction step it has to be from 2 to n). The following works:
lemma valid_bitmap_to_valid_set_aux:
"bitmap_to_set ps x ⊆ {x..<x + length ps}"
by (induction ps x rule: bitmap_to_set.induct)
(auto simp: valid_bitmap_def valid_set_def)
lemma valid_bitmap_to_valid_set:
"valid_bitmap ps n ⟹ valid_set (bitmap_to_set ps 1) n"
using valid_bitmap_to_valid_set_aux unfolding valid_bitmap_def valid_set_def
by force
Injectivity and surjectivity (which is your ultimate goal) should follow from the fact that the two are inverse functions. Proving that will probably be doable with induction, but will require a few generalisations and auxiliary lemmas. It should be easier if you stick to the non-recursive definition using library functions that I sketched above.

Proving two bindings equal in Nominal Isabelle

Consider the following datatypes with bindings in Nominal Isabelle:
theory Example
imports "Nominal2.Nominal2"
begin
atom_decl vrs
nominal_datatype ty =
Tvar "vrs"
| Arrow x::vrs T::"ty" binds x in T
nominal_datatype trm =
Var "vrs"
| Abs x::"vrs" t::"trm" binds x in t
inductive
typing :: "trm ⇒ ty ⇒ bool" ("_ , _" [60,60] 60)
where
T_Abs[intro]: "(Abs x t) , (Arrow x T)"
equivariance typing
nominal_inductive typing done
lemma
assumes "(Abs x t), (Arrow y T)"
shows "x = y"
using assms
I want to prove that the two bindings appearing in the relation are equal. I see two ways an Isabelle user could help:
If you know Nominal Isabelle is it possible to do this?
Otherwise, are the two occurrences of x in the rule T_Abs equal for the assistant or are they sort of bound variable with different identity?
If you know Nominal Isabelle is it possible to do this?
Unfortunately, it is not possible to prove the theorem that you are trying to prove. Here is a counterexample (the proofs were Sledgehammered):
theory Scratch
imports "Nominal2.Nominal2"
begin
atom_decl vrs
nominal_datatype ty =
Tvar "vrs"
| Arrow x::vrs T::"ty" binds x in T
nominal_datatype trm =
Var "vrs"
| Abs x::"vrs" t::"trm" binds x in t
inductive
typing :: "trm ⇒ ty ⇒ bool" ("_ , _" [60,60] 60)
where
T_Abs[intro]: "(Abs x t) , (Arrow x T)"
equivariance typing
nominal_inductive typing .
abbreviation s where "s ≡ Sort ''Scratch.vrs'' []"
abbreviation v where "v n ≡ Abs_vrs (Atom s n)"
lemma neq: "Abs (v 1) (Var (v 0)), Arrow (v (Suc (Suc 0))) (Tvar (v 0))"
(is "?a, ?b")
proof-
have a_def: "Abs (v 1) (Var (v 0)) = Abs (v (Suc (Suc 0))) (Var (v 0))"
(*Sledgehammered*)
by simp (smt Abs_vrs_inverse atom.inject flip_at_base_simps(3) fresh_PairD(2)
fresh_at_base(2) mem_Collect_eq nat.distinct(1) sort_of.simps trm.fresh(1))
from typing.simps[of ?a ?b, unfolded this, THEN iffD2] have
"Abs (v (Suc (Suc 0))) (Var (v 0)) , Arrow (v (Suc (Suc 0))) (Tvar (v 0))"
by auto
then show ?thesis unfolding a_def by clarsimp
qed
lemma "∃x y t T. x ≠ y ∧ (Abs x t), (Arrow y T)"
proof(intro exI conjI)
show "v 1 ≠ v (Suc (Suc 0))"
(*Sledgehammered*)
by (smt Abs_vrs_inverse One_nat_def atom.inject mem_Collect_eq n_not_Suc_n
sort_of.simps)
show "Abs (v 1) (Var (v 0)) , Arrow (v (Suc (Suc 0))) (Tvar (v 0))"
by (rule neq)
qed
end
Otherwise, are the two occurrences of x in the rule T_Abs equal for
the assistant or are they sort of bound variable with different
identity?
I believe that you are thinking along the right lines and, hopefully, the example above will clarify any confusion that you might have. Generally, you could interpret the meaning of Abs x t1 = Abs y t2 as the alpha-equivalence of (λx. t1) and (λy. t2). Of course, (λx. t1) and (λy. t2) may be alpha equivalent without x and y being equal.

Basic Isabelle/Isar style (exercise 4.6)

I'm interested in using Isabelle/Isar for writing proofs which are both human-readable and machine checked, and I am looking to improve my style and streamline my proofs.
prog-prove has the following exercise:
Exercise 4.6. Define a recursive function elems :: 'a list ⇒ 'a set and prove x ∈ elems xs ⟹ ∃ ys zs. xs = ys # x # zs ∧ x ∉ elems ys.
Mimicking something similar to what I would write with pen and paper, my solution is
fun elems :: "'a list ⇒ 'a set" where
"elems [] = {}" |
"elems (x # xs) = {x} ∪ elems xs"
fun takeUntil :: "('a ⇒ bool) ⇒ 'a list ⇒ 'a list" where
"takeUntil f [] = []" |
"takeUntil f (x # xs) = (case (f x) of False ⇒ x # takeUntil f xs | True ⇒ [])"
theorem "x ∈ elems xs ⟹ ∃ ys zs. xs = ys # x # zs ∧ x ∉ elems ys"
proof -
assume 1: "x ∈ elems xs"
let ?ys = "takeUntil (λ z. z = x) xs"
let ?zs = "drop (length ?ys + 1) xs"
have "xs = ?ys # x # ?zs ∧ x ∉ elems ?ys"
proof
have 2: "x ∉ elems ?ys"
proof (induction xs)
case Nil
thus ?case by simp
next
case (Cons a xs)
thus ?case
proof -
{
assume "a = x"
hence "takeUntil (λz. z = x) (a # xs) = []" by simp
hence A: ?thesis by simp
}
note eq = this
{
assume "a ≠ x"
hence "takeUntil (λz. z = x) (a # xs) = a # takeUntil (λz. z = x) xs" by simp
hence ?thesis using Cons.IH by auto
}
note noteq = this
have "a = x ∨ a ≠ x" by simp
thus ?thesis using eq noteq by blast
qed
qed
from 1 have "xs = ?ys # x # ?zs"
proof (induction xs)
case Nil
hence False by simp
thus ?case by simp
next
case (Cons a xs)
{
assume 1: "a = x"
hence 2: "takeUntil (λz. z = x) (a # xs) = []" by simp
hence "length (takeUntil (λz. z = x) (a # xs)) + 1 = 1" by simp
hence 3: "drop (length (takeUntil (λz. z = x) (a # xs)) + 1) (a # xs) = xs" by simp
from 1 2 3 have ?case by simp
}
note eq = this
{
assume 1: "a ≠ x"
with Cons.prems have "x ∈ elems xs" by simp
with Cons.IH
have IH: "xs = takeUntil (λz. z = x) xs # x # drop (length (takeUntil (λz. z = x) xs) + 1) xs" by simp
from 1 have 2: "takeUntil (λz. z = x) (a # xs) = a # takeUntil (λz. z = x) (xs)" by simp
from 1 have "drop (length (takeUntil (λz. z = x) (a # xs)) + 1) (a # xs) = drop (length (takeUntil (λz. z = x) xs) + 1) xs" by simp
hence ?case using IH 2 by simp
}
note noteq = this
have "a = x ∨ a ≠ x" by simp
thus ?case using eq noteq by blast
qed
with 2 have 3: ?thesis by blast
thus "xs = takeUntil (λz. z = x) xs # x # drop (length (takeUntil (λz. z = x) xs) + 1) xs" by simp
from 3 show "x ∉ elems (takeUntil (λz. z = x) xs)" by simp
qed
thus ?thesis by blast
qed
but it seems rather long. In particular, I think invoking law of excluded middle here is cumbersome, and I feel like there ought to be some convenient schematic variable like ?goal which can refer to the current goal or something.
How can I make this proof shorter without sacrificing clarity?
Not really an answer to your specific question, but I would nonetheless like to point out, that a more concise prove can still be comprehensible.
lemma "x ∈ elems xs ⟹ ∃ ys zs. xs = ys # x # zs ∧ x ∉ elems ys"
proof (induction)
case (Cons l ls)
thus ?case
proof (cases "x ≠ l")
case True
hence "∃ys zs. ls = ys # x # zs ∧ x ∉ elems ys" using Cons by simp
thus ?thesis using ‹x ≠ l› Cons_eq_appendI by fastforce
qed (fastforce)
qed (simp)
Here's another shorter proof than your own:
fun elems :: ‹'a list ⇒ 'a set› where
‹elems [] = {}› |
‹elems (x#xs) = {x} ∪ elems xs›
lemma elems_prefix_suffix:
assumes ‹x ∈ elems xs›
shows ‹∃pre suf. xs = pre # [x] # suf ∧ x ∉ elems pre›
using assms proof(induction xs)
fix y ys
assume *: ‹x ∈ elems (y#ys)›
and IH: ‹x ∈ elems ys ⟹ ∃pre suf. ys = pre # [x] # suf ∧ x ∉ elems pre›
{
assume ‹x = y›
from this have ‹∃pre suf. y#ys = pre # [x] # suf ∧ x ∉ elems pre›
using * by fastforce
}
note L = this
{
assume ‹x ≠ y› and ‹x ∈ elems ys›
moreover from this obtain pre and suf where ‹ys = pre # [x] # suf› and ‹x ∉ elems pre›
using IH by auto
moreover have ‹y#ys = y#pre # [x] # suf› and ‹x ∉ elems (y#pre)›
by(simp add: calculation)+
ultimately have ‹∃pre suf. y#ys = pre # [x] # suf ∧ x ∉ elems pre›
by(metis append_Cons)
}
from this and L show ‹∃pre suf. y#ys = pre # [x] # suf ∧ x ∉ elems pre›
using * by auto
qed auto ― ‹Base case trivial›
I've used a few features of Isar to compress the proof:
Blocks within the braces {...} allow you to perform hypothetical reasoning.
Facts can be explicitly named using note.
The moreover keyword starts a calculation that implicitly "carries along" facts as they are established. The calculation "comes to a head" with the ultimately keyword. This style can significantly reduce the number of explicitly named facts that you need to introduce over the course of a proof.
The qed auto completes the proof by applying auto to all remaining subgoals. A comment notes that the subgoal remaining is the base case of the induction, which is trivial.

What happens during function proofs

I am trying to proof a property of the icmp6 checksum function (sum 16bit integers, add carry, invert 16bit integer).
I defined the functions in isabelle. (I know my proofs are terrible)
But for some reason, isabelle can't proof something about the icmp_csum function, it wants to have.
When I replace the oops in the paste with done it produces thousands of lines that just says:
"linarith_split_limit exceeded (current value is 9)"
theory Scratch
imports Main Int List
begin
fun norm_helper :: "nat ⇒ nat" where
"norm_helper x = (let y = divide x 65536 in (y + x - y * 65536))"
lemma "x ≥ 65536 ⟹ norm_helper x < x" by simp
lemma h: "norm_helper x ≤ x" by simp
fun normalize :: "nat ⇒ nat" where
"normalize x = (if x ≥ 65536
then normalize (norm_helper x)
else x)"
inductive norm_to :: "nat ⇒ nat ⇒ bool" where
"(x < 65536) ⟹ norm_to x x"
| "norm_to y z ⟹ y = norm_helper x ⟹ norm_to x z"
lemma ne: "norm_to x y ⟹ y = normalize x"
apply (induct x y rule: norm_to.induct) by simp+
lemma i: "norm_to x y ⟹ x ≥ y"
apply (induct x y rule: norm_to.induct) by simp+
lemma l: "norm_to x y ⟹ y < 65536"
apply (induct x y rule: norm_to.induct) by simp+
lemma en: "y = normalize x ⟹ norm_to x y"
apply (induct x rule: normalize.induct)
proof -
fix x :: nat
assume 1: "(x ≥ 65536 ⟹ y = Scratch.normalize (norm_helper x) ⟹ norm_to (norm_helper x) y)"
assume 2: "y = Scratch.normalize x"
show "norm_to x y"
proof (cases "x ≥ 65536")
show "¬ 65536 ≤ x ⟹ norm_to x y"
using norm_to.intros(1)[of x] 2 by simp
{
assume s: "65536 ≤ x"
have d: "y = normalize (norm_helper x)" using 2 s by simp
show "65536 ≤ x ⟹ norm_to x y"
using 1 d norm_to.intros(2)[of "norm_helper x" y x]
by blast
}
qed
qed
lemma "normalize x ≤ x" using en i by simp
lemma n[simp]: "normalize x < 65536" using en l by blast
fun sum :: "nat list ⇒ nat" where
"sum [] = 0"
| "sum (x#xs) = x + sum xs"
fun csum :: "nat list ⇒ nat" where
"csum xs = normalize (sum xs)"
fun invert :: "nat ⇒ nat" where
"invert x = 65535 - x"
lemma c: "csum xs ≤ 65535" using n[of "sum xs"] by simp
lemma ic: "invert (csum xs) ≥ 0" using c[of xs] by blast
lemma asdf:
assumes "xs = ys"
shows "invert (csum xs) = invert (csum ys)"
using HOL.arg_cong[of "csum xs" "csum ys" invert,
OF HOL.arg_cong[of xs ys csum]] assms(1)
by blast
function icmp_csum :: "nat list ⇒ nat" where
"icmp_csum xs = invert (csum xs)"
apply simp
apply (rule asdf)
apply simp
oops
end
I have no idea why there is tracing output from linarith there, but given that your definition is neither recursive nor performs pattern matching, you can write it as a definition:
definition icmp_csum :: "nat list ⇒ nat" where
"icmp_csum xs = invert (csum xs)"
Another possibility is to change invert to a definition instead of a fun. (In general, if it's neither recursive nor performs pattern matching, definition is preferable because it has much less overhead than fun.)
NB, just import Main, not Main Int List.
Edit: An explanation from Tobias Nipkow on the mailing list:
This is a known issue. In the outdated LNCS 2283 you can find a discussion what to do about it in Section 3.5.3 Simplification and Recursive Functions. The gist: don't use "if", use pattern matching or "case". Or disable if_split.

Resources