Up to index of Isabelle/HOL/MetisExamples
theory Abstraction(* Title: HOL/MetisExamples/Abstraction.thy ID: $Id: Abstraction.thy,v 1.7 2007/10/04 10:32:59 paulson Exp $ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Testing the metis method *) theory Abstraction imports FuncSet begin (*For Christoph Benzmueller*) lemma "x<1 & ((op=) = (op=)) ==> ((op=) = (op=)) & (x<(2::nat))"; by (metis One_nat_def less_Suc0 not_less0 not_less_eq numeral_2_eq_2) (*this is a theorem, but we can't prove it unless ext is applied explicitly lemma "(op=) = (%x y. y=x)" *) consts monotone :: "['a => 'a, 'a set, ('a *'a)set] => bool" pset :: "'a set => 'a set" order :: "'a set => ('a * 'a) set" ML{*ResAtp.problem_name := "Abstraction__Collect_triv"*} lemma (*Collect_triv:*) "a ∈ {x. P x} ==> P a" proof (neg_clausify) assume 0: "(a::'a::type) ∈ Collect (P::'a::type => bool)" assume 1: "¬ (P::'a::type => bool) (a::'a::type)" have 2: "(P::'a::type => bool) (a::'a::type)" by (metis CollectD 0) show "False" by (metis 2 1) qed lemma Collect_triv: "a ∈ {x. P x} ==> P a" by (metis mem_Collect_eq) ML{*ResAtp.problem_name := "Abstraction__Collect_mp"*} lemma "a ∈ {x. P x --> Q x} ==> a ∈ {x. P x} ==> a ∈ {x. Q x}" by (metis CollectI Collect_imp_eq ComplD UnE mem_Collect_eq); --{*34 secs*} ML{*ResAtp.problem_name := "Abstraction__Sigma_triv"*} lemma "(a,b) ∈ Sigma A B ==> a ∈ A & b ∈ B a" proof (neg_clausify) assume 0: "(a::'a::type, b::'b::type) ∈ Sigma (A::'a::type set) (B::'a::type => 'b::type set)" assume 1: "(a::'a::type) ∉ (A::'a::type set) ∨ (b::'b::type) ∉ (B::'a::type => 'b::type set) a" have 2: "(a::'a::type) ∈ (A::'a::type set)" by (metis SigmaD1 0) have 3: "(b::'b::type) ∈ (B::'a::type => 'b::type set) (a::'a::type)" by (metis SigmaD2 0) have 4: "(b::'b::type) ∉ (B::'a::type => 'b::type set) (a::'a::type)" by (metis 1 2) show "False" by (metis 3 4) qed lemma Sigma_triv: "(a,b) ∈ Sigma A B ==> a ∈ A & b ∈ B a" by (metis SigmaD1 SigmaD2) ML{*ResAtp.problem_name := "Abstraction__Sigma_Collect"*} lemma "(a,b) ∈ (SIGMA x: A. {y. x = f y}) ==> a ∈ A & a = f b" (*???metis cannot prove this by (metis CollectD SigmaD1 SigmaD2 UN_eq) Also, UN_eq is unnecessary*) by (meson CollectD SigmaD1 SigmaD2) (*single-step*) lemma "(a,b) ∈ (SIGMA x: A. {y. x = f y}) ==> a ∈ A & a = f b" by (metis SigmaD1 SigmaD2 insert_def singleton_conv2 union_empty2 vimage_Collect_eq vimage_def vimage_singleton_eq) lemma "(a,b) ∈ (SIGMA x: A. {y. x = f y}) ==> a ∈ A & a = f b" proof (neg_clausify) assume 0: "(a::'a::type, b::'b::type) ∈ Sigma (A::'a::type set) (COMBB Collect (COMBC (COMBB COMBB op =) (f::'b::type => 'a::type)))" assume 1: "(a::'a::type) ∉ (A::'a::type set) ∨ a ≠ (f::'b::type => 'a::type) (b::'b::type)" have 2: "(a::'a::type) ∈ (A::'a::type set)" by (metis 0 SigmaD1) have 3: "(b::'b::type) ∈ COMBB Collect (COMBC (COMBB COMBB op =) (f::'b::type => 'a::type)) (a::'a::type)" by (metis 0 SigmaD2) have 4: "(b::'b::type) ∈ Collect (COMBB (op = (a::'a::type)) (f::'b::type => 'a::type))" by (metis 3) have 5: "(f::'b::type => 'a::type) (b::'b::type) ≠ (a::'a::type)" by (metis 1 2) have 6: "(f::'b::type => 'a::type) (b::'b::type) = (a::'a::type)" by (metis 4 vimage_singleton_eq insert_def singleton_conv2 union_empty2 vimage_Collect_eq vimage_def) show "False" by (metis 5 6) qed (*Alternative structured proof, untyped*) lemma "(a,b) ∈ (SIGMA x: A. {y. x = f y}) ==> a ∈ A & a = f b" proof (neg_clausify) assume 0: "(a, b) ∈ Sigma A (COMBB Collect (COMBC (COMBB COMBB op =) f))" have 1: "b ∈ Collect (COMBB (op = a) f)" by (metis 0 SigmaD2) have 2: "f b = a" by (metis 1 vimage_Collect_eq singleton_conv2 insert_def union_empty2 vimage_singleton_eq vimage_def) assume 3: "a ∉ A ∨ a ≠ f b" have 4: "a ∈ A" by (metis 0 SigmaD1) have 5: "f b ≠ a" by (metis 4 3) show "False" by (metis 5 2) qed ML{*ResAtp.problem_name := "Abstraction__CLF_eq_in_pp"*} lemma "(cl,f) ∈ CLF ==> CLF = (SIGMA cl: CL.{f. f ∈ pset cl}) ==> f ∈ pset cl" by (metis Collect_mem_eq SigmaD2) lemma "(cl,f) ∈ CLF ==> CLF = (SIGMA cl: CL.{f. f ∈ pset cl}) ==> f ∈ pset cl" proof (neg_clausify) assume 0: "(cl, f) ∈ CLF" assume 1: "CLF = Sigma CL (COMBB Collect (COMBB (COMBC op ∈) pset))" assume 2: "f ∉ pset cl" have 3: "!!X1 X2. X2 ∈ COMBB Collect (COMBB (COMBC op ∈) pset) X1 ∨ (X1, X2) ∉ CLF" by (metis SigmaD2 1) have 4: "!!X1 X2. X2 ∈ pset X1 ∨ (X1, X2) ∉ CLF" by (metis 3 Collect_mem_eq) have 5: "(cl, f) ∉ CLF" by (metis 2 4) show "False" by (metis 5 0) qed ML{*ResAtp.problem_name := "Abstraction__Sigma_Collect_Pi"*} lemma "(cl,f) ∈ (SIGMA cl: CL. {f. f ∈ pset cl -> pset cl}) ==> f ∈ pset cl -> pset cl" proof (neg_clausify) assume 0: "f ∉ Pi (pset cl) (COMBK (pset cl))" assume 1: "(cl, f) ∈ Sigma CL (COMBB Collect (COMBB (COMBC op ∈) (COMBS (COMBB Pi pset) (COMBB COMBK pset))))" show "False" (* by (metis 0 Collect_mem_eq SigmaD2 1) ??doesn't terminate*) by (insert 0 1, simp add: COMBB_def COMBS_def COMBC_def) qed ML{*ResAtp.problem_name := "Abstraction__Sigma_Collect_Int"*} lemma "(cl,f) ∈ (SIGMA cl: CL. {f. f ∈ pset cl ∩ cl}) ==> f ∈ pset cl ∩ cl" proof (neg_clausify) assume 0: "(cl, f) ∈ Sigma CL (COMBB Collect (COMBB (COMBC op ∈) (COMBS (COMBB op ∩ pset) COMBI)))" assume 1: "f ∉ pset cl ∩ cl" have 2: "f ∈ COMBB Collect (COMBB (COMBC op ∈) (COMBS (COMBB op ∩ pset) COMBI)) cl" by (insert 0, simp add: COMBB_def) (* by (metis SigmaD2 0) ??doesn't terminate*) have 3: "f ∈ COMBS (COMBB op ∩ pset) COMBI cl" by (metis 2 Collect_mem_eq) have 4: "f ∉ cl ∩ pset cl" by (metis 1 Int_commute) have 5: "f ∈ cl ∩ pset cl" by (metis 3 Int_commute) show "False" by (metis 5 4) qed ML{*ResAtp.problem_name := "Abstraction__Sigma_Collect_Pi_mono"*} lemma "(cl,f) ∈ (SIGMA cl: CL. {f. f ∈ pset cl -> pset cl & monotone f (pset cl) (order cl)}) ==> (f ∈ pset cl -> pset cl) & (monotone f (pset cl) (order cl))" by auto ML{*ResAtp.problem_name := "Abstraction__CLF_subset_Collect_Int"*} lemma "(cl,f) ∈ CLF ==> CLF ⊆ (SIGMA cl: CL. {f. f ∈ pset cl ∩ cl}) ==> f ∈ pset cl ∩ cl" by auto (*??no longer terminates, with combinators by (metis Collect_mem_eq Int_def SigmaD2 UnCI Un_absorb1) --{*@{text Int_def} is redundant} *) ML{*ResAtp.problem_name := "Abstraction__CLF_eq_Collect_Int"*} lemma "(cl,f) ∈ CLF ==> CLF = (SIGMA cl: CL. {f. f ∈ pset cl ∩ cl}) ==> f ∈ pset cl ∩ cl" by auto (*??no longer terminates, with combinators by (metis Collect_mem_eq Int_commute SigmaD2) *) ML{*ResAtp.problem_name := "Abstraction__CLF_subset_Collect_Pi"*} lemma "(cl,f) ∈ CLF ==> CLF ⊆ (SIGMA cl': CL. {f. f ∈ pset cl' -> pset cl'}) ==> f ∈ pset cl -> pset cl" by auto (*??no longer terminates, with combinators by (metis Collect_mem_eq SigmaD2 subsetD) *) ML{*ResAtp.problem_name := "Abstraction__CLF_eq_Collect_Pi"*} lemma "(cl,f) ∈ CLF ==> CLF = (SIGMA cl: CL. {f. f ∈ pset cl -> pset cl}) ==> f ∈ pset cl -> pset cl" by auto (*??no longer terminates, with combinators by (metis Collect_mem_eq SigmaD2 contra_subsetD equalityE) *) ML{*ResAtp.problem_name := "Abstraction__CLF_eq_Collect_Pi_mono"*} lemma "(cl,f) ∈ CLF ==> CLF = (SIGMA cl: CL. {f. f ∈ pset cl -> pset cl & monotone f (pset cl) (order cl)}) ==> (f ∈ pset cl -> pset cl) & (monotone f (pset cl) (order cl))" by auto ML{*ResAtp.problem_name := "Abstraction__map_eq_zipA"*} lemma "map (%x. (f x, g x)) xs = zip (map f xs) (map g xs)" apply (induct xs) (*sledgehammer*) apply auto done ML{*ResAtp.problem_name := "Abstraction__map_eq_zipB"*} lemma "map (%w. (w -> w, w × w)) xs = zip (map (%w. w -> w) xs) (map (%w. w × w) xs)" apply (induct xs) (*sledgehammer*) apply auto done ML{*ResAtp.problem_name := "Abstraction__image_evenA"*} lemma "(%x. Suc(f x)) ` {x. even x} <= A ==> (∀x. even x --> Suc(f x) ∈ A)"; (*sledgehammer*) by auto ML{*ResAtp.problem_name := "Abstraction__image_evenB"*} lemma "(%x. f (f x)) ` ((%x. Suc(f x)) ` {x. even x}) <= A ==> (∀x. even x --> f (f (Suc(f x))) ∈ A)"; (*sledgehammer*) by auto ML{*ResAtp.problem_name := "Abstraction__image_curry"*} lemma "f ∈ (%u v. b × u × v) ` A ==> ∀u v. P (b × u × v) ==> P(f y)" (*sledgehammer*) by auto ML{*ResAtp.problem_name := "Abstraction__image_TimesA"*} lemma image_TimesA: "(%(x,y). (f x, g y)) ` (A × B) = (f`A) × (g`B)" (*sledgehammer*) apply (rule equalityI) (***Even the two inclusions are far too difficult ML{*ResAtp.problem_name := "Abstraction__image_TimesA_simpler"*} ***) apply (rule subsetI) apply (erule imageE) (*V manages from here with help: Abstraction__image_TimesA_simpler_1_b.p*) apply (erule ssubst) apply (erule SigmaE) (*V manages from here: Abstraction__image_TimesA_simpler_1_a.p*) apply (erule ssubst) apply (subst split_conv) apply (rule SigmaI) apply (erule imageI) + txt{*subgoal 2*} apply (clarify ); apply (simp add: ); apply (rule rev_image_eqI) apply (blast intro: elim:); apply (simp add: ); done (*Given the difficulty of the previous problem, these two are probably impossible*) ML{*ResAtp.problem_name := "Abstraction__image_TimesB"*} lemma image_TimesB: "(%(x,y,z). (f x, g y, h z)) ` (A × B × C) = (f`A) × (g`B) × (h`C)" (*sledgehammer*) by force ML{*ResAtp.problem_name := "Abstraction__image_TimesC"*} lemma image_TimesC: "(%(x,y). (x -> x, y × y)) ` (A × B) = ((%x. x -> x) ` A) × ((%y. y × y) ` B)" (*sledgehammer*) by auto end
lemma
x < 1 ∧ op = = op = ==> op = = op = ∧ x < 2
lemma
a ∈ {x. P x} ==> P a
lemma Collect_triv:
a ∈ {x. P x} ==> P a
lemma
[| a ∈ {x. P x --> Q x}; a ∈ {x. P x} |] ==> a ∈ {x. Q x}
lemma
(a, b) ∈ Sigma A B ==> a ∈ A ∧ b ∈ B a
lemma Sigma_triv:
(a, b) ∈ Sigma A B ==> a ∈ A ∧ b ∈ B a
lemma
(a, b) ∈ (SIGMA x:A. {y. x = f y}) ==> a ∈ A ∧ a = f b
lemma
(a, b) ∈ (SIGMA x:A. {y. x = f y}) ==> a ∈ A ∧ a = f b
lemma
(a, b) ∈ (SIGMA x:A. {y. x = f y}) ==> a ∈ A ∧ a = f b
lemma
(a, b) ∈ (SIGMA x:A. {y. x = f y}) ==> a ∈ A ∧ a = f b
lemma
[| (cl, f) ∈ CLF; CLF = (SIGMA cl:CL. {f. f ∈ pset cl}) |] ==> f ∈ pset cl
lemma
[| (cl, f) ∈ CLF; CLF = (SIGMA cl:CL. {f. f ∈ pset cl}) |] ==> f ∈ pset cl
lemma
(cl, f) ∈ (SIGMA cl:CL. {f. f ∈ pset cl -> pset cl}) ==> f ∈ pset cl -> pset cl
lemma
(cl, f) ∈ (SIGMA cl:CL. {f. f ∈ pset cl ∩ cl}) ==> f ∈ pset cl ∩ cl
lemma
(cl, f)
∈ (SIGMA cl:CL.
{f : pset cl -> pset cl. monotone f (pset cl) (Abstraction.order cl)})
==> f ∈ pset cl -> pset cl ∧ monotone f (pset cl) (Abstraction.order cl)
lemma
[| (cl, f) ∈ CLF; CLF ⊆ (SIGMA cl:CL. {f. f ∈ pset cl ∩ cl}) |]
==> f ∈ pset cl ∩ cl
lemma
[| (cl, f) ∈ CLF; CLF = (SIGMA cl:CL. {f. f ∈ pset cl ∩ cl}) |]
==> f ∈ pset cl ∩ cl
lemma
[| (cl, f) ∈ CLF; CLF ⊆ (SIGMA cl':CL. {f. f ∈ pset cl' -> pset cl'}) |]
==> f ∈ pset cl -> pset cl
lemma
[| (cl, f) ∈ CLF; CLF = (SIGMA cl:CL. {f. f ∈ pset cl -> pset cl}) |]
==> f ∈ pset cl -> pset cl
lemma
[| (cl, f) ∈ CLF;
CLF =
(SIGMA cl:CL.
{f : pset cl -> pset cl. monotone f (pset cl) (Abstraction.order cl)}) |]
==> f ∈ pset cl -> pset cl ∧ monotone f (pset cl) (Abstraction.order cl)
lemma
map (λx. (f x, g x)) xs = zip (map f xs) (map g xs)
lemma
map (λw. (w -> w, w × w)) xs = zip (map (λw. w -> w) xs) (map (λw. w × w) xs)
lemma
(λx. Suc (f x)) ` {x. even x} ⊆ A ==> ∀x. even x --> Suc (f x) ∈ A
lemma
(λx. f (f x)) ` (λx. Suc (f x)) ` {x. even x} ⊆ A
==> ∀x. even x --> f (f (Suc (f x))) ∈ A
lemma
[| f ∈ (λu v. b × u × v) ` A; ∀u v. P (b × u × v) |] ==> P (f y)
lemma image_TimesA:
(λ(x, y). (f x, g y)) ` (A × B) = f ` A × g ` B
lemma image_TimesB:
(λ(x, y, z). (f x, g y, h z)) ` (A × B × C) = f ` A × g ` B × h ` C
lemma image_TimesC:
(λ(x, y). (x -> x, y × y)) ` (A × B) = (λx. x -> x) ` A × (λy. y × y) ` B