# HG changeset patch # User wenzelm # Date 1256061124 -7200 # Node ID 9cf389429f6df980c84b68fb90238c58900fad0b # Parent 8f35633c492244048ecd6367e5e15abd0cd768cc modernized session Metis_Examples; diff -r 8f35633c4922 -r 9cf389429f6d Admin/isatest/isatest-stats --- a/Admin/isatest/isatest-stats Tue Oct 20 19:37:09 2009 +0200 +++ b/Admin/isatest/isatest-stats Tue Oct 20 19:52:04 2009 +0200 @@ -21,7 +21,7 @@ HOL-HoareParallel \ HOL-Lambda \ HOL-Library \ - HOL-MetisExamples \ + HOL-Metis_Examples \ HOL-MicroJava \ HOL-NSA \ HOL-Nominal-Examples \ diff -r 8f35633c4922 -r 9cf389429f6d src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Tue Oct 20 19:37:09 2009 +0200 +++ b/src/HOL/IsaMakefile Tue Oct 20 19:52:04 2009 +0200 @@ -29,7 +29,7 @@ HOL-Lambda \ HOL-Lattice \ HOL-Matrix \ - HOL-MetisExamples \ + HOL-Metis_Examples \ HOL-MicroJava \ HOL-Mirabelle \ HOL-Modelcheck \ @@ -556,16 +556,16 @@ @$(ISABELLE_TOOL) usedir -g true $(OUT)/HOL Hoare_Parallel -## HOL-MetisExamples +## HOL-Metis_Examples -HOL-MetisExamples: HOL $(LOG)/HOL-MetisExamples.gz +HOL-Metis_Examples: HOL $(LOG)/HOL-Metis_Examples.gz -$(LOG)/HOL-MetisExamples.gz: $(OUT)/HOL MetisExamples/ROOT.ML \ - MetisExamples/Abstraction.thy MetisExamples/BigO.thy \ - MetisExamples/BT.thy MetisExamples/Message.thy \ - MetisExamples/Tarski.thy MetisExamples/TransClosure.thy \ - MetisExamples/set.thy - @$(ISABELLE_TOOL) usedir $(OUT)/HOL MetisExamples +$(LOG)/HOL-Metis_Examples.gz: $(OUT)/HOL Metis_Examples/ROOT.ML \ + Metis_Examples/Abstraction.thy Metis_Examples/BigO.thy \ + Metis_Examples/BT.thy Metis_Examples/Message.thy \ + Metis_Examples/Tarski.thy Metis_Examples/TransClosure.thy \ + Metis_Examples/set.thy + @$(ISABELLE_TOOL) usedir $(OUT)/HOL Metis_Examples ## HOL-Algebra diff -r 8f35633c4922 -r 9cf389429f6d src/HOL/MetisExamples/Abstraction.thy --- a/src/HOL/MetisExamples/Abstraction.thy Tue Oct 20 19:37:09 2009 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,298 +0,0 @@ -(* Title: HOL/MetisExamples/Abstraction.thy - ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - -Testing the metis method -*) - -theory Abstraction -imports Main 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" - -declare [[ atp_problem_prefix = "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) - - -declare [[ atp_problem_prefix = "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*} - -declare [[ atp_problem_prefix = "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) - -declare [[ atp_problem_prefix = "Abstraction__Sigma_Collect" ]] -lemma "(a,b) \ (SIGMA x: A. {y. x = f y}) ==> a \ A & a = f b" -(*???metis says this is satisfiable! -by (metis CollectD SigmaD1 SigmaD2) -*) -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 Un_empty_right 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 Un_empty_right 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 Un_empty_right 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 - - -declare [[ atp_problem_prefix = "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 - -declare [[ atp_problem_prefix = "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 - - -declare [[ atp_problem_prefix = "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 - - -declare [[ atp_problem_prefix = "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 - -declare [[ atp_problem_prefix = "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*} -*) - -declare [[ atp_problem_prefix = "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) -*) - -declare [[ atp_problem_prefix = "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 fast -(*??no longer terminates, with combinators -by (metis Collect_mem_eq SigmaD2 subsetD) -*) - -declare [[ atp_problem_prefix = "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) -*) - -declare [[ atp_problem_prefix = "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 - -declare [[ atp_problem_prefix = "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 - -declare [[ atp_problem_prefix = "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 - -declare [[ atp_problem_prefix = "Abstraction__image_evenA" ]] -lemma "(%x. Suc(f x)) ` {x. even x} <= A ==> (\x. even x --> Suc(f x) \ A)"; -(*sledgehammer*) -by auto - -declare [[ atp_problem_prefix = "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 - -declare [[ atp_problem_prefix = "Abstraction__image_curry" ]] -lemma "f \ (%u v. b \ u \ v) ` A ==> \u v. P (b \ u \ v) ==> P(f y)" -(*sledgehammer*) -by auto - -declare [[ atp_problem_prefix = "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 -using [[ atp_problem_prefix = "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*) - -declare [[ atp_problem_prefix = "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 - -declare [[ atp_problem_prefix = "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 diff -r 8f35633c4922 -r 9cf389429f6d src/HOL/MetisExamples/BT.thy --- a/src/HOL/MetisExamples/BT.thy Tue Oct 20 19:37:09 2009 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,243 +0,0 @@ -(* Title: HOL/MetisTest/BT.thy - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - -Testing the metis method -*) - -header {* Binary trees *} - -theory BT -imports Main -begin - - -datatype 'a bt = - Lf - | Br 'a "'a bt" "'a bt" - -consts - n_nodes :: "'a bt => nat" - n_leaves :: "'a bt => nat" - depth :: "'a bt => nat" - reflect :: "'a bt => 'a bt" - bt_map :: "('a => 'b) => ('a bt => 'b bt)" - preorder :: "'a bt => 'a list" - inorder :: "'a bt => 'a list" - postorder :: "'a bt => 'a list" - appnd :: "'a bt => 'a bt => 'a bt" - -primrec - "n_nodes Lf = 0" - "n_nodes (Br a t1 t2) = Suc (n_nodes t1 + n_nodes t2)" - -primrec - "n_leaves Lf = Suc 0" - "n_leaves (Br a t1 t2) = n_leaves t1 + n_leaves t2" - -primrec - "depth Lf = 0" - "depth (Br a t1 t2) = Suc (max (depth t1) (depth t2))" - -primrec - "reflect Lf = Lf" - "reflect (Br a t1 t2) = Br a (reflect t2) (reflect t1)" - -primrec - "bt_map f Lf = Lf" - "bt_map f (Br a t1 t2) = Br (f a) (bt_map f t1) (bt_map f t2)" - -primrec - "preorder Lf = []" - "preorder (Br a t1 t2) = [a] @ (preorder t1) @ (preorder t2)" - -primrec - "inorder Lf = []" - "inorder (Br a t1 t2) = (inorder t1) @ [a] @ (inorder t2)" - -primrec - "postorder Lf = []" - "postorder (Br a t1 t2) = (postorder t1) @ (postorder t2) @ [a]" - -primrec - "appnd Lf t = t" - "appnd (Br a t1 t2) t = Br a (appnd t1 t) (appnd t2 t)" - - -text {* \medskip BT simplification *} - -declare [[ atp_problem_prefix = "BT__n_leaves_reflect" ]] -lemma n_leaves_reflect: "n_leaves (reflect t) = n_leaves t" - apply (induct t) - apply (metis add_right_cancel n_leaves.simps(1) reflect.simps(1)) - apply (metis add_commute n_leaves.simps(2) reflect.simps(2)) - done - -declare [[ atp_problem_prefix = "BT__n_nodes_reflect" ]] -lemma n_nodes_reflect: "n_nodes (reflect t) = n_nodes t" - apply (induct t) - apply (metis reflect.simps(1)) - apply (metis n_nodes.simps(2) nat_add_commute reflect.simps(2)) - done - -declare [[ atp_problem_prefix = "BT__depth_reflect" ]] -lemma depth_reflect: "depth (reflect t) = depth t" - apply (induct t) - apply (metis depth.simps(1) reflect.simps(1)) - apply (metis depth.simps(2) min_max.sup_commute reflect.simps(2)) - done - -text {* - The famous relationship between the numbers of leaves and nodes. -*} - -declare [[ atp_problem_prefix = "BT__n_leaves_nodes" ]] -lemma n_leaves_nodes: "n_leaves t = Suc (n_nodes t)" - apply (induct t) - apply (metis n_leaves.simps(1) n_nodes.simps(1)) - apply auto - done - -declare [[ atp_problem_prefix = "BT__reflect_reflect_ident" ]] -lemma reflect_reflect_ident: "reflect (reflect t) = t" - apply (induct t) - apply (metis add_right_cancel reflect.simps(1)); - apply (metis reflect.simps(2)) - done - -declare [[ atp_problem_prefix = "BT__bt_map_ident" ]] -lemma bt_map_ident: "bt_map (%x. x) = (%y. y)" -apply (rule ext) -apply (induct_tac y) - apply (metis bt_map.simps(1)) -txt{*BUG involving flex-flex pairs*} -(* apply (metis bt_map.simps(2)) *) -apply auto -done - - -declare [[ atp_problem_prefix = "BT__bt_map_appnd" ]] -lemma bt_map_appnd: "bt_map f (appnd t u) = appnd (bt_map f t) (bt_map f u)" -apply (induct t) - apply (metis appnd.simps(1) bt_map.simps(1)) - apply (metis appnd.simps(2) bt_map.simps(2)) (*slow!!*) -done - - -declare [[ atp_problem_prefix = "BT__bt_map_compose" ]] -lemma bt_map_compose: "bt_map (f o g) t = bt_map f (bt_map g t)" -apply (induct t) - apply (metis bt_map.simps(1)) -txt{*Metis runs forever*} -(* apply (metis bt_map.simps(2) o_apply)*) -apply auto -done - - -declare [[ atp_problem_prefix = "BT__bt_map_reflect" ]] -lemma bt_map_reflect: "bt_map f (reflect t) = reflect (bt_map f t)" - apply (induct t) - apply (metis add_right_cancel bt_map.simps(1) reflect.simps(1)) - apply (metis add_right_cancel bt_map.simps(2) reflect.simps(2)) - done - -declare [[ atp_problem_prefix = "BT__preorder_bt_map" ]] -lemma preorder_bt_map: "preorder (bt_map f t) = map f (preorder t)" - apply (induct t) - apply (metis bt_map.simps(1) map.simps(1) preorder.simps(1)) - apply simp - done - -declare [[ atp_problem_prefix = "BT__inorder_bt_map" ]] -lemma inorder_bt_map: "inorder (bt_map f t) = map f (inorder t)" - apply (induct t) - apply (metis bt_map.simps(1) inorder.simps(1) map.simps(1)) - apply simp - done - -declare [[ atp_problem_prefix = "BT__postorder_bt_map" ]] -lemma postorder_bt_map: "postorder (bt_map f t) = map f (postorder t)" - apply (induct t) - apply (metis bt_map.simps(1) map.simps(1) postorder.simps(1)) - apply simp - done - -declare [[ atp_problem_prefix = "BT__depth_bt_map" ]] -lemma depth_bt_map [simp]: "depth (bt_map f t) = depth t" - apply (induct t) - apply (metis bt_map.simps(1) depth.simps(1)) - apply simp - done - -declare [[ atp_problem_prefix = "BT__n_leaves_bt_map" ]] -lemma n_leaves_bt_map [simp]: "n_leaves (bt_map f t) = n_leaves t" - apply (induct t) - apply (metis One_nat_def Suc_eq_plus1 bt_map.simps(1) less_add_one less_antisym linorder_neq_iff n_leaves.simps(1)) - apply (metis bt_map.simps(2) n_leaves.simps(2)) - done - - -declare [[ atp_problem_prefix = "BT__preorder_reflect" ]] -lemma preorder_reflect: "preorder (reflect t) = rev (postorder t)" - apply (induct t) - apply (metis postorder.simps(1) preorder.simps(1) reflect.simps(1) rev_is_Nil_conv) - apply (metis append_Nil Cons_eq_append_conv postorder.simps(2) preorder.simps(2) reflect.simps(2) rev.simps(2) rev_append rev_rev_ident) - done - -declare [[ atp_problem_prefix = "BT__inorder_reflect" ]] -lemma inorder_reflect: "inorder (reflect t) = rev (inorder t)" - apply (induct t) - apply (metis inorder.simps(1) reflect.simps(1) rev.simps(1)) - apply simp - done - -declare [[ atp_problem_prefix = "BT__postorder_reflect" ]] -lemma postorder_reflect: "postorder (reflect t) = rev (preorder t)" - apply (induct t) - apply (metis postorder.simps(1) preorder.simps(1) reflect.simps(1) rev.simps(1)) - apply (metis Cons_eq_appendI postorder.simps(2) preorder.simps(2) reflect.simps(2) rev.simps(2) rev_append self_append_conv2) - done - -text {* - Analogues of the standard properties of the append function for lists. -*} - -declare [[ atp_problem_prefix = "BT__appnd_assoc" ]] -lemma appnd_assoc [simp]: - "appnd (appnd t1 t2) t3 = appnd t1 (appnd t2 t3)" - apply (induct t1) - apply (metis appnd.simps(1)) - apply (metis appnd.simps(2)) - done - -declare [[ atp_problem_prefix = "BT__appnd_Lf2" ]] -lemma appnd_Lf2 [simp]: "appnd t Lf = t" - apply (induct t) - apply (metis appnd.simps(1)) - apply (metis appnd.simps(2)) - done - -declare [[ atp_problem_prefix = "BT__depth_appnd" ]] - declare max_add_distrib_left [simp] -lemma depth_appnd [simp]: "depth (appnd t1 t2) = depth t1 + depth t2" - apply (induct t1) - apply (metis add_0 appnd.simps(1) depth.simps(1)) -apply (simp add: ); - done - -declare [[ atp_problem_prefix = "BT__n_leaves_appnd" ]] -lemma n_leaves_appnd [simp]: - "n_leaves (appnd t1 t2) = n_leaves t1 * n_leaves t2" - apply (induct t1) - apply (metis One_nat_def appnd.simps(1) less_irrefl less_linear n_leaves.simps(1) nat_mult_1) - apply (simp add: left_distrib) - done - -declare [[ atp_problem_prefix = "BT__bt_map_appnd" ]] -lemma (*bt_map_appnd:*) - "bt_map f (appnd t1 t2) = appnd (bt_map f t1) (bt_map f t2)" - apply (induct t1) - apply (metis appnd.simps(1) bt_map_appnd) - apply (metis bt_map_appnd) - done - -end diff -r 8f35633c4922 -r 9cf389429f6d src/HOL/MetisExamples/BigO.thy --- a/src/HOL/MetisExamples/BigO.thy Tue Oct 20 19:37:09 2009 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1231 +0,0 @@ -(* Title: HOL/MetisExamples/BigO.thy - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - -Testing the metis method -*) - -header {* Big O notation *} - -theory BigO -imports "~~/src/HOL/Decision_Procs/Dense_Linear_Order" Main SetsAndFunctions -begin - -subsection {* Definitions *} - -definition bigo :: "('a => 'b::ordered_idom) => ('a => 'b) set" ("(1O'(_'))") where - "O(f::('a => 'b)) == {h. EX c. ALL x. abs (h x) <= c * abs (f x)}" - -declare [[ atp_problem_prefix = "BigO__bigo_pos_const" ]] -lemma bigo_pos_const: "(EX (c::'a::ordered_idom). - ALL x. (abs (h x)) <= (c * (abs (f x)))) - = (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))" - apply auto - apply (case_tac "c = 0", simp) - apply (rule_tac x = "1" in exI, simp) - apply (rule_tac x = "abs c" in exI, auto) - apply (metis abs_ge_minus_self abs_ge_zero abs_minus_cancel abs_of_nonneg equation_minus_iff Orderings.xt1(6) abs_mult) - done - -(*** Now various verions with an increasing modulus ***) - -declare [[sledgehammer_modulus = 1]] - -lemma (*bigo_pos_const:*) "(EX (c::'a::ordered_idom). - ALL x. (abs (h x)) <= (c * (abs (f x)))) - = (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))" - apply auto - apply (case_tac "c = 0", simp) - apply (rule_tac x = "1" in exI, simp) - apply (rule_tac x = "abs c" in exI, auto) -proof (neg_clausify) -fix c x -have 0: "\(X1\'a\ordered_idom) X2\'a\ordered_idom. \X1 * X2\ = \X2 * X1\" - by (metis abs_mult mult_commute) -have 1: "\(X1\'a\ordered_idom) X2\'a\ordered_idom. - X1 \ (0\'a\ordered_idom) \ \X2\ * X1 = \X2 * X1\" - by (metis abs_mult_pos linorder_linear) -have 2: "\(X1\'a\ordered_idom) X2\'a\ordered_idom. - \ (0\'a\ordered_idom) < X1 * X2 \ - \ (0\'a\ordered_idom) \ X2 \ \ X1 \ (0\'a\ordered_idom)" - by (metis linorder_not_less mult_nonneg_nonpos2) -assume 3: "\x\'b\type. - \(h\'b\type \ 'a\ordered_idom) x\ - \ (c\'a\ordered_idom) * \(f\'b\type \ 'a\ordered_idom) x\" -assume 4: "\ \(h\'b\type \ 'a\ordered_idom) (x\'b\type)\ - \ \c\'a\ordered_idom\ * \(f\'b\type \ 'a\ordered_idom) x\" -have 5: "\ \(h\'b\type \ 'a\ordered_idom) (x\'b\type)\ - \ \(c\'a\ordered_idom) * (f\'b\type \ 'a\ordered_idom) x\" - by (metis 4 abs_mult) -have 6: "\(X1\'a\ordered_idom) X2\'a\ordered_idom. - \ X1 \ (0\'a\ordered_idom) \ X1 \ \X2\" - by (metis abs_ge_zero xt1(6)) -have 7: "\(X1\'a\ordered_idom) X2\'a\ordered_idom. - X1 \ \X2\ \ (0\'a\ordered_idom) < X1" - by (metis not_leE 6) -have 8: "(0\'a\ordered_idom) < \(h\'b\type \ 'a\ordered_idom) (x\'b\type)\" - by (metis 5 7) -have 9: "\X1\'a\ordered_idom. - \ \(h\'b\type \ 'a\ordered_idom) (x\'b\type)\ \ X1 \ - (0\'a\ordered_idom) < X1" - by (metis 8 order_less_le_trans) -have 10: "(0\'a\ordered_idom) -< (c\'a\ordered_idom) * \(f\'b\type \ 'a\ordered_idom) (x\'b\type)\" - by (metis 3 9) -have 11: "\ (c\'a\ordered_idom) \ (0\'a\ordered_idom)" - by (metis abs_ge_zero 2 10) -have 12: "\X1\'a\ordered_idom. (c\'a\ordered_idom) * \X1\ = \X1 * c\" - by (metis mult_commute 1 11) -have 13: "\X1\'b\type. - - (h\'b\type \ 'a\ordered_idom) X1 - \ (c\'a\ordered_idom) * \(f\'b\type \ 'a\ordered_idom) X1\" - by (metis 3 abs_le_D2) -have 14: "\X1\'b\type. - - (h\'b\type \ 'a\ordered_idom) X1 - \ \(c\'a\ordered_idom) * (f\'b\type \ 'a\ordered_idom) X1\" - by (metis 0 12 13) -have 15: "\(X1\'a\ordered_idom) X2\'a\ordered_idom. \X1 * \X2\\ = \X1 * X2\" - by (metis abs_mult abs_mult_pos abs_ge_zero) -have 16: "\(X1\'a\ordered_idom) X2\'a\ordered_idom. X1 \ \X2\ \ \ X1 \ X2" - by (metis xt1(6) abs_ge_self) -have 17: "\(X1\'a\ordered_idom) X2\'a\ordered_idom. \ \X1\ \ X2 \ X1 \ \X2\" - by (metis 16 abs_le_D1) -have 18: "\X1\'b\type. - (h\'b\type \ 'a\ordered_idom) X1 - \ \(c\'a\ordered_idom) * (f\'b\type \ 'a\ordered_idom) X1\" - by (metis 17 3 15) -show "False" - by (metis abs_le_iff 5 18 14) -qed - -declare [[sledgehammer_modulus = 2]] - -lemma (*bigo_pos_const:*) "(EX (c::'a::ordered_idom). - ALL x. (abs (h x)) <= (c * (abs (f x)))) - = (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))" - apply auto - apply (case_tac "c = 0", simp) - apply (rule_tac x = "1" in exI, simp) - apply (rule_tac x = "abs c" in exI, auto); -proof (neg_clausify) -fix c x -have 0: "\(X1\'a\ordered_idom) X2\'a\ordered_idom. \X1 * X2\ = \X2 * X1\" - by (metis abs_mult mult_commute) -assume 1: "\x\'b\type. - \(h\'b\type \ 'a\ordered_idom) x\ - \ (c\'a\ordered_idom) * \(f\'b\type \ 'a\ordered_idom) x\" -assume 2: "\ \(h\'b\type \ 'a\ordered_idom) (x\'b\type)\ - \ \c\'a\ordered_idom\ * \(f\'b\type \ 'a\ordered_idom) x\" -have 3: "\ \(h\'b\type \ 'a\ordered_idom) (x\'b\type)\ - \ \(c\'a\ordered_idom) * (f\'b\type \ 'a\ordered_idom) x\" - by (metis 2 abs_mult) -have 4: "\(X1\'a\ordered_idom) X2\'a\ordered_idom. - \ X1 \ (0\'a\ordered_idom) \ X1 \ \X2\" - by (metis abs_ge_zero xt1(6)) -have 5: "(0\'a\ordered_idom) < \(h\'b\type \ 'a\ordered_idom) (x\'b\type)\" - by (metis not_leE 4 3) -have 6: "(0\'a\ordered_idom) -< (c\'a\ordered_idom) * \(f\'b\type \ 'a\ordered_idom) (x\'b\type)\" - by (metis 1 order_less_le_trans 5) -have 7: "\X1\'a\ordered_idom. (c\'a\ordered_idom) * \X1\ = \X1 * c\" - by (metis abs_ge_zero linorder_not_less mult_nonneg_nonpos2 6 linorder_linear abs_mult_pos mult_commute) -have 8: "\X1\'b\type. - - (h\'b\type \ 'a\ordered_idom) X1 - \ \(c\'a\ordered_idom) * (f\'b\type \ 'a\ordered_idom) X1\" - by (metis 0 7 abs_le_D2 1) -have 9: "\(X1\'a\ordered_idom) X2\'a\ordered_idom. \ \X1\ \ X2 \ X1 \ \X2\" - by (metis abs_ge_self xt1(6) abs_le_D1) -show "False" - by (metis 8 abs_ge_zero abs_mult_pos abs_mult 1 9 3 abs_le_iff) -qed - -declare [[sledgehammer_modulus = 3]] - -lemma (*bigo_pos_const:*) "(EX (c::'a::ordered_idom). - ALL x. (abs (h x)) <= (c * (abs (f x)))) - = (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))" - apply auto - apply (case_tac "c = 0", simp) - apply (rule_tac x = "1" in exI, simp) - apply (rule_tac x = "abs c" in exI, auto); -proof (neg_clausify) -fix c x -assume 0: "\x\'b\type. - \(h\'b\type \ 'a\ordered_idom) x\ - \ (c\'a\ordered_idom) * \(f\'b\type \ 'a\ordered_idom) x\" -assume 1: "\ \(h\'b\type \ 'a\ordered_idom) (x\'b\type)\ - \ \c\'a\ordered_idom\ * \(f\'b\type \ 'a\ordered_idom) x\" -have 2: "\(X1\'a\ordered_idom) X2\'a\ordered_idom. - X1 \ \X2\ \ (0\'a\ordered_idom) < X1" - by (metis abs_ge_zero xt1(6) not_leE) -have 3: "\ (c\'a\ordered_idom) \ (0\'a\ordered_idom)" - by (metis abs_ge_zero mult_nonneg_nonpos2 linorder_not_less order_less_le_trans 1 abs_mult 2 0) -have 4: "\(X1\'a\ordered_idom) X2\'a\ordered_idom. \X1 * \X2\\ = \X1 * X2\" - by (metis abs_ge_zero abs_mult_pos abs_mult) -have 5: "\X1\'b\type. - (h\'b\type \ 'a\ordered_idom) X1 - \ \(c\'a\ordered_idom) * (f\'b\type \ 'a\ordered_idom) X1\" - by (metis 4 0 xt1(6) abs_ge_self abs_le_D1) -show "False" - by (metis abs_mult mult_commute 3 abs_mult_pos linorder_linear 0 abs_le_D2 5 1 abs_le_iff) -qed - - -declare [[sledgehammer_modulus = 1]] - -lemma (*bigo_pos_const:*) "(EX (c::'a::ordered_idom). - ALL x. (abs (h x)) <= (c * (abs (f x)))) - = (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))" - apply auto - apply (case_tac "c = 0", simp) - apply (rule_tac x = "1" in exI, simp) - apply (rule_tac x = "abs c" in exI, auto); -proof (neg_clausify) -fix c x (*sort/type constraint inserted by hand!*) -have 0: "\(X1\'a\ordered_idom) X2. \X1 * \X2\\ = \X1 * X2\" - by (metis abs_ge_zero abs_mult_pos abs_mult) -assume 1: "\A. \h A\ \ c * \f A\" -have 2: "\X1 X2. \ \X1\ \ X2 \ (0\'a) \ X2" - by (metis abs_ge_zero order_trans) -have 3: "\X1. (0\'a) \ c * \f X1\" - by (metis 1 2) -have 4: "\X1. c * \f X1\ = \c * f X1\" - by (metis 0 abs_of_nonneg 3) -have 5: "\X1. - h X1 \ c * \f X1\" - by (metis 1 abs_le_D2) -have 6: "\X1. - h X1 \ \c * f X1\" - by (metis 4 5) -have 7: "\X1. h X1 \ c * \f X1\" - by (metis 1 abs_le_D1) -have 8: "\X1. h X1 \ \c * f X1\" - by (metis 4 7) -assume 9: "\ \h x\ \ \c\ * \f x\" -have 10: "\ \h x\ \ \c * f x\" - by (metis abs_mult 9) -show "False" - by (metis 6 8 10 abs_leI) -qed - - -declare [[sledgehammer_sorts = true]] - -lemma bigo_alt_def: "O(f) = - {h. EX c. (0 < c & (ALL x. abs (h x) <= c * abs (f x)))}" -by (auto simp add: bigo_def bigo_pos_const) - -declare [[ atp_problem_prefix = "BigO__bigo_elt_subset" ]] -lemma bigo_elt_subset [intro]: "f : O(g) ==> O(f) <= O(g)" - apply (auto simp add: bigo_alt_def) - apply (rule_tac x = "ca * c" in exI) - apply (rule conjI) - apply (rule mult_pos_pos) - apply (assumption)+ -(*sledgehammer*); - apply (rule allI) - apply (drule_tac x = "xa" in spec)+ - apply (subgoal_tac "ca * abs(f xa) <= ca * (c * abs(g xa))"); - apply (erule order_trans) - apply (simp add: mult_ac) - apply (rule mult_left_mono, assumption) - apply (rule order_less_imp_le, assumption); -done - - -declare [[ atp_problem_prefix = "BigO__bigo_refl" ]] -lemma bigo_refl [intro]: "f : O(f)" - apply(auto simp add: bigo_def) -proof (neg_clausify) -fix x -assume 0: "\xa. \ \f (x xa)\ \ xa * \f (x xa)\" -have 1: "\X2. X2 \ (1\'b) * X2 \ \ (1\'b) \ (1\'b)" - by (metis mult_le_cancel_right1 order_eq_iff) -have 2: "\X2. X2 \ (1\'b) * X2" - by (metis order_eq_iff 1) -show "False" - by (metis 0 2) -qed - -declare [[ atp_problem_prefix = "BigO__bigo_zero" ]] -lemma bigo_zero: "0 : O(g)" - apply (auto simp add: bigo_def func_zero) -proof (neg_clausify) -fix x -assume 0: "\xa. \ (0\'b) \ xa * \g (x xa)\" -have 1: "\ (0\'b) \ (0\'b)" - by (metis 0 mult_eq_0_iff) -show "False" - by (metis 1 linorder_neq_iff linorder_antisym_conv1) -qed - -lemma bigo_zero2: "O(%x.0) = {%x.0}" - apply (auto simp add: bigo_def) - apply (rule ext) - apply auto -done - -lemma bigo_plus_self_subset [intro]: - "O(f) \ O(f) <= O(f)" - apply (auto simp add: bigo_alt_def set_plus_def) - apply (rule_tac x = "c + ca" in exI) - apply auto - apply (simp add: ring_distribs func_plus) - apply (blast intro:order_trans abs_triangle_ineq add_mono elim:) -done - -lemma bigo_plus_idemp [simp]: "O(f) \ O(f) = O(f)" - apply (rule equalityI) - apply (rule bigo_plus_self_subset) - apply (rule set_zero_plus2) - apply (rule bigo_zero) -done - -lemma bigo_plus_subset [intro]: "O(f + g) <= O(f) \ O(g)" - apply (rule subsetI) - apply (auto simp add: bigo_def bigo_pos_const func_plus set_plus_def) - apply (subst bigo_pos_const [symmetric])+ - apply (rule_tac x = - "%n. if abs (g n) <= (abs (f n)) then x n else 0" in exI) - apply (rule conjI) - apply (rule_tac x = "c + c" in exI) - apply (clarsimp) - apply (auto) - apply (subgoal_tac "c * abs (f xa + g xa) <= (c + c) * abs (f xa)") - apply (erule_tac x = xa in allE) - apply (erule order_trans) - apply (simp) - apply (subgoal_tac "c * abs (f xa + g xa) <= c * (abs (f xa) + abs (g xa))") - apply (erule order_trans) - apply (simp add: ring_distribs) - apply (rule mult_left_mono) - apply assumption - apply (simp add: order_less_le) - apply (rule mult_left_mono) - apply (simp add: abs_triangle_ineq) - apply (simp add: order_less_le) - apply (rule mult_nonneg_nonneg) - apply (rule add_nonneg_nonneg) - apply auto - apply (rule_tac x = "%n. if (abs (f n)) < abs (g n) then x n else 0" - in exI) - apply (rule conjI) - apply (rule_tac x = "c + c" in exI) - apply auto - apply (subgoal_tac "c * abs (f xa + g xa) <= (c + c) * abs (g xa)") - apply (erule_tac x = xa in allE) - apply (erule order_trans) - apply (simp) - apply (subgoal_tac "c * abs (f xa + g xa) <= c * (abs (f xa) + abs (g xa))") - apply (erule order_trans) - apply (simp add: ring_distribs) - apply (rule mult_left_mono) - apply (simp add: order_less_le) - apply (simp add: order_less_le) - apply (rule mult_left_mono) - apply (rule abs_triangle_ineq) - apply (simp add: order_less_le) -apply (metis abs_not_less_zero double_less_0_iff less_not_permute linorder_not_less mult_less_0_iff) - apply (rule ext) - apply (auto simp add: if_splits linorder_not_le) -done - -lemma bigo_plus_subset2 [intro]: "A <= O(f) ==> B <= O(f) ==> A \ B <= O(f)" - apply (subgoal_tac "A \ B <= O(f) \ O(f)") - apply (erule order_trans) - apply simp - apply (auto del: subsetI simp del: bigo_plus_idemp) -done - -declare [[ atp_problem_prefix = "BigO__bigo_plus_eq" ]] -lemma bigo_plus_eq: "ALL x. 0 <= f x ==> ALL x. 0 <= g x ==> - O(f + g) = O(f) \ O(g)" - apply (rule equalityI) - apply (rule bigo_plus_subset) - apply (simp add: bigo_alt_def set_plus_def func_plus) - apply clarify -(*sledgehammer*); - apply (rule_tac x = "max c ca" in exI) - apply (rule conjI) - apply (metis Orderings.less_max_iff_disj) - apply clarify - apply (drule_tac x = "xa" in spec)+ - apply (subgoal_tac "0 <= f xa + g xa") - apply (simp add: ring_distribs) - apply (subgoal_tac "abs(a xa + b xa) <= abs(a xa) + abs(b xa)") - apply (subgoal_tac "abs(a xa) + abs(b xa) <= - max c ca * f xa + max c ca * g xa") - apply (blast intro: order_trans) - defer 1 - apply (rule abs_triangle_ineq) - apply (metis add_nonneg_nonneg) - apply (rule add_mono) -using [[ atp_problem_prefix = "BigO__bigo_plus_eq_simpler" ]] -(*Found by SPASS; SLOW*) -apply (metis le_maxI2 linorder_linear linorder_not_le min_max.sup_absorb1 mult_le_cancel_right order_trans) -apply (metis le_maxI2 linorder_not_le mult_le_cancel_right order_trans) -done - -declare [[ atp_problem_prefix = "BigO__bigo_bounded_alt" ]] -lemma bigo_bounded_alt: "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> - f : O(g)" - apply (auto simp add: bigo_def) -(*Version 1: one-shot proof*) - apply (metis OrderedGroup.abs_le_D1 linorder_class.not_less order_less_le Orderings.xt1(12) Ring_and_Field.abs_mult) - done - -lemma (*bigo_bounded_alt:*) "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> - f : O(g)" - apply (auto simp add: bigo_def) -(*Version 2: single-step proof*) -proof (neg_clausify) -fix x -assume 0: "\x. f x \ c * g x" -assume 1: "\xa. \ f (x xa) \ xa * \g (x xa)\" -have 2: "\X3. c * g X3 = f X3 \ \ c * g X3 \ f X3" - by (metis 0 order_antisym_conv) -have 3: "\X3. \ f (x \X3\) \ \X3 * g (x \X3\)\" - by (metis 1 abs_mult) -have 4: "\X1 X3\'b\ordered_idom. X3 \ X1 \ X1 \ \X3\" - by (metis linorder_linear abs_le_D1) -have 5: "\X3::'b. \X3\ * \X3\ = X3 * X3" - by (metis abs_mult_self) -have 6: "\X3. \ X3 * X3 < (0\'b\ordered_idom)" - by (metis not_square_less_zero) -have 7: "\X1 X3::'b. \X1\ * \X3\ = \X3 * X1\" - by (metis abs_mult mult_commute) -have 8: "\X3::'b. X3 * X3 = \X3 * X3\" - by (metis abs_mult 5) -have 9: "\X3. X3 * g (x \X3\) \ f (x \X3\)" - by (metis 3 4) -have 10: "c * g (x \c\) = f (x \c\)" - by (metis 2 9) -have 11: "\X3::'b. \X3\ * \\X3\\ = \X3\ * \X3\" - by (metis abs_idempotent abs_mult 8) -have 12: "\X3::'b. \X3 * \X3\\ = \X3\ * \X3\" - by (metis mult_commute 7 11) -have 13: "\X3::'b. \X3 * \X3\\ = X3 * X3" - by (metis 8 7 12) -have 14: "\X3. X3 \ \X3\ \ X3 < (0\'b)" - by (metis abs_ge_self abs_le_D1 abs_if) -have 15: "\X3. X3 \ \X3\ \ \X3\ < (0\'b)" - by (metis abs_ge_self abs_le_D1 abs_if) -have 16: "\X3. X3 * X3 < (0\'b) \ X3 * \X3\ \ X3 * X3" - by (metis 15 13) -have 17: "\X3::'b. X3 * \X3\ \ X3 * X3" - by (metis 16 6) -have 18: "\X3. X3 \ \X3\ \ \ X3 < (0\'b)" - by (metis mult_le_cancel_left 17) -have 19: "\X3::'b. X3 \ \X3\" - by (metis 18 14) -have 20: "\ f (x \c\) \ \f (x \c\)\" - by (metis 3 10) -show "False" - by (metis 20 19) -qed - - -text{*So here is the easier (and more natural) problem using transitivity*} -declare [[ atp_problem_prefix = "BigO__bigo_bounded_alt_trans" ]] -lemma "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> f : O(g)" - apply (auto simp add: bigo_def) - (*Version 1: one-shot proof*) - apply (metis Orderings.leD Orderings.leI abs_ge_self abs_le_D1 abs_mult abs_of_nonneg order_le_less) - done - -text{*So here is the easier (and more natural) problem using transitivity*} -declare [[ atp_problem_prefix = "BigO__bigo_bounded_alt_trans" ]] -lemma "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> f : O(g)" - apply (auto simp add: bigo_def) -(*Version 2: single-step proof*) -proof (neg_clausify) -fix x -assume 0: "\A\'a\type. - (f\'a\type \ 'b\ordered_idom) A - \ (c\'b\ordered_idom) * (g\'a\type \ 'b\ordered_idom) A" -assume 1: "\A\'b\ordered_idom. - \ (f\'a\type \ 'b\ordered_idom) ((x\'b\ordered_idom \ 'a\type) A) - \ A * \(g\'a\type \ 'b\ordered_idom) (x A)\" -have 2: "\X2\'a\type. - \ (c\'b\ordered_idom) * (g\'a\type \ 'b\ordered_idom) X2 - < (f\'a\type \ 'b\ordered_idom) X2" - by (metis 0 linorder_not_le) -have 3: "\X2\'b\ordered_idom. - \ (f\'a\type \ 'b\ordered_idom) ((x\'b\ordered_idom \ 'a\type) \X2\) - \ \X2 * (g\'a\type \ 'b\ordered_idom) (x \X2\)\" - by (metis abs_mult 1) -have 4: "\X2\'b\ordered_idom. - \X2 * (g\'a\type \ 'b\ordered_idom) ((x\'b\ordered_idom \ 'a\type) \X2\)\ - < (f\'a\type \ 'b\ordered_idom) (x \X2\)" - by (metis 3 linorder_not_less) -have 5: "\X2\'b\ordered_idom. - X2 * (g\'a\type \ 'b\ordered_idom) ((x\'b\ordered_idom \ 'a\type) \X2\) - < (f\'a\type \ 'b\ordered_idom) (x \X2\)" - by (metis abs_less_iff 4) -show "False" - by (metis 2 5) -qed - - -lemma bigo_bounded: "ALL x. 0 <= f x ==> ALL x. f x <= g x ==> - f : O(g)" - apply (erule bigo_bounded_alt [of f 1 g]) - apply simp -done - -declare [[ atp_problem_prefix = "BigO__bigo_bounded2" ]] -lemma bigo_bounded2: "ALL x. lb x <= f x ==> ALL x. f x <= lb x + g x ==> - f : lb +o O(g)" - apply (rule set_minus_imp_plus) - apply (rule bigo_bounded) - apply (auto simp add: diff_minus fun_Compl_def func_plus) - prefer 2 - apply (drule_tac x = x in spec)+ - apply arith (*not clear that it's provable otherwise*) -proof (neg_clausify) -fix x -assume 0: "\y. lb y \ f y" -assume 1: "\ (0\'b) \ f x + - lb x" -have 2: "\X3. (0\'b) + X3 = X3" - by (metis diff_eq_eq right_minus_eq) -have 3: "\ (0\'b) \ f x - lb x" - by (metis 1 diff_minus) -have 4: "\ (0\'b) + lb x \ f x" - by (metis 3 le_diff_eq) -show "False" - by (metis 4 2 0) -qed - -declare [[ atp_problem_prefix = "BigO__bigo_abs" ]] -lemma bigo_abs: "(%x. abs(f x)) =o O(f)" - apply (unfold bigo_def) - apply auto -proof (neg_clausify) -fix x -assume 0: "\xa. \ \f (x xa)\ \ xa * \f (x xa)\" -have 1: "\X2. X2 \ (1\'b) * X2 \ \ (1\'b) \ (1\'b)" - by (metis mult_le_cancel_right1 order_eq_iff) -have 2: "\X2. X2 \ (1\'b) * X2" - by (metis order_eq_iff 1) -show "False" - by (metis 0 2) -qed - -declare [[ atp_problem_prefix = "BigO__bigo_abs2" ]] -lemma bigo_abs2: "f =o O(%x. abs(f x))" - apply (unfold bigo_def) - apply auto -proof (neg_clausify) -fix x -assume 0: "\xa. \ \f (x xa)\ \ xa * \f (x xa)\" -have 1: "\X2. X2 \ (1\'b) * X2 \ \ (1\'b) \ (1\'b)" - by (metis mult_le_cancel_right1 order_eq_iff) -have 2: "\X2. X2 \ (1\'b) * X2" - by (metis order_eq_iff 1) -show "False" - by (metis 0 2) -qed - -lemma bigo_abs3: "O(f) = O(%x. abs(f x))" - apply (rule equalityI) - apply (rule bigo_elt_subset) - apply (rule bigo_abs2) - apply (rule bigo_elt_subset) - apply (rule bigo_abs) -done - -lemma bigo_abs4: "f =o g +o O(h) ==> - (%x. abs (f x)) =o (%x. abs (g x)) +o O(h)" - apply (drule set_plus_imp_minus) - apply (rule set_minus_imp_plus) - apply (subst fun_diff_def) -proof - - assume a: "f - g : O(h)" - have "(%x. abs (f x) - abs (g x)) =o O(%x. abs(abs (f x) - abs (g x)))" - by (rule bigo_abs2) - also have "... <= O(%x. abs (f x - g x))" - apply (rule bigo_elt_subset) - apply (rule bigo_bounded) - apply force - apply (rule allI) - apply (rule abs_triangle_ineq3) - done - also have "... <= O(f - g)" - apply (rule bigo_elt_subset) - apply (subst fun_diff_def) - apply (rule bigo_abs) - done - also have "... <= O(h)" - using a by (rule bigo_elt_subset) - finally show "(%x. abs (f x) - abs (g x)) : O(h)". -qed - -lemma bigo_abs5: "f =o O(g) ==> (%x. abs(f x)) =o O(g)" -by (unfold bigo_def, auto) - -lemma bigo_elt_subset2 [intro]: "f : g +o O(h) ==> O(f) <= O(g) \ O(h)" -proof - - assume "f : g +o O(h)" - also have "... <= O(g) \ O(h)" - by (auto del: subsetI) - also have "... = O(%x. abs(g x)) \ O(%x. abs(h x))" - apply (subst bigo_abs3 [symmetric])+ - apply (rule refl) - done - also have "... = O((%x. abs(g x)) + (%x. abs(h x)))" - by (rule bigo_plus_eq [symmetric], auto) - finally have "f : ...". - then have "O(f) <= ..." - by (elim bigo_elt_subset) - also have "... = O(%x. abs(g x)) \ O(%x. abs(h x))" - by (rule bigo_plus_eq, auto) - finally show ?thesis - by (simp add: bigo_abs3 [symmetric]) -qed - -declare [[ atp_problem_prefix = "BigO__bigo_mult" ]] -lemma bigo_mult [intro]: "O(f)\O(g) <= O(f * g)" - apply (rule subsetI) - apply (subst bigo_def) - apply (auto simp del: abs_mult mult_ac - simp add: bigo_alt_def set_times_def func_times) -(*sledgehammer*); - apply (rule_tac x = "c * ca" in exI) - apply(rule allI) - apply(erule_tac x = x in allE)+ - apply(subgoal_tac "c * ca * abs(f x * g x) = - (c * abs(f x)) * (ca * abs(g x))") -using [[ atp_problem_prefix = "BigO__bigo_mult_simpler" ]] -prefer 2 -apply (metis mult_assoc mult_left_commute - OrderedGroup.abs_of_pos OrderedGroup.mult_left_commute - Ring_and_Field.abs_mult Ring_and_Field.mult_pos_pos) - apply (erule ssubst) - apply (subst abs_mult) -(*not qute BigO__bigo_mult_simpler_1 (a hard problem!) as abs_mult has - just been done*) -proof (neg_clausify) -fix a c b ca x -assume 0: "(0\'b\ordered_idom) < (c\'b\ordered_idom)" -assume 1: "\(a\'a \ 'b\ordered_idom) (x\'a)\ -\ (c\'b\ordered_idom) * \(f\'a \ 'b\ordered_idom) x\" -assume 2: "\(b\'a \ 'b\ordered_idom) (x\'a)\ -\ (ca\'b\ordered_idom) * \(g\'a \ 'b\ordered_idom) x\" -assume 3: "\ \(a\'a \ 'b\ordered_idom) (x\'a)\ * - \(b\'a \ 'b\ordered_idom) x\ - \ (c\'b\ordered_idom) * \(f\'a \ 'b\ordered_idom) x\ * - ((ca\'b\ordered_idom) * \(g\'a \ 'b\ordered_idom) x\)" -have 4: "\c\'b\ordered_idom\ = c" - by (metis OrderedGroup.abs_of_pos 0) -have 5: "\X1\'b\ordered_idom. (c\'b\ordered_idom) * \X1\ = \c * X1\" - by (metis Ring_and_Field.abs_mult 4) -have 6: "(0\'b\ordered_idom) = (1\'b\ordered_idom) \ -(0\'b\ordered_idom) < (1\'b\ordered_idom)" - by (metis OrderedGroup.abs_not_less_zero Ring_and_Field.abs_one Ring_and_Field.linorder_neqE_ordered_idom) -have 7: "(0\'b\ordered_idom) < (1\'b\ordered_idom)" - by (metis 6 Ring_and_Field.one_neq_zero) -have 8: "\1\'b\ordered_idom\ = (1\'b\ordered_idom)" - by (metis OrderedGroup.abs_of_pos 7) -have 9: "\X1\'b\ordered_idom. (0\'b\ordered_idom) \ (c\'b\ordered_idom) * \X1\" - by (metis OrderedGroup.abs_ge_zero 5) -have 10: "\X1\'b\ordered_idom. X1 * (1\'b\ordered_idom) = X1" - by (metis Ring_and_Field.mult_cancel_right2 mult_commute) -have 11: "\X1\'b\ordered_idom. \\X1\\ = \X1\ * \1\'b\ordered_idom\" - by (metis Ring_and_Field.abs_mult OrderedGroup.abs_idempotent 10) -have 12: "\X1\'b\ordered_idom. \\X1\\ = \X1\" - by (metis 11 8 10) -have 13: "\X1\'b\ordered_idom. (0\'b\ordered_idom) \ \X1\" - by (metis OrderedGroup.abs_ge_zero 12) -have 14: "\ (0\'b\ordered_idom) - \ (c\'b\ordered_idom) * \(f\'a \ 'b\ordered_idom) (x\'a)\ \ -\ (0\'b\ordered_idom) \ \(b\'a \ 'b\ordered_idom) x\ \ -\ \b x\ \ (ca\'b\ordered_idom) * \(g\'a \ 'b\ordered_idom) x\ \ -\ \(a\'a \ 'b\ordered_idom) x\ \ c * \f x\" - by (metis 3 Ring_and_Field.mult_mono) -have 15: "\ (0\'b\ordered_idom) \ \(b\'a \ 'b\ordered_idom) (x\'a)\ \ -\ \b x\ \ (ca\'b\ordered_idom) * \(g\'a \ 'b\ordered_idom) x\ \ -\ \(a\'a \ 'b\ordered_idom) x\ - \ (c\'b\ordered_idom) * \(f\'a \ 'b\ordered_idom) x\" - by (metis 14 9) -have 16: "\ \(b\'a \ 'b\ordered_idom) (x\'a)\ - \ (ca\'b\ordered_idom) * \(g\'a \ 'b\ordered_idom) x\ \ -\ \(a\'a \ 'b\ordered_idom) x\ - \ (c\'b\ordered_idom) * \(f\'a \ 'b\ordered_idom) x\" - by (metis 15 13) -have 17: "\ \(a\'a \ 'b\ordered_idom) (x\'a)\ - \ (c\'b\ordered_idom) * \(f\'a \ 'b\ordered_idom) x\" - by (metis 16 2) -show 18: "False" - by (metis 17 1) -qed - - -declare [[ atp_problem_prefix = "BigO__bigo_mult2" ]] -lemma bigo_mult2 [intro]: "f *o O(g) <= O(f * g)" - apply (auto simp add: bigo_def elt_set_times_def func_times abs_mult) -(*sledgehammer*); - apply (rule_tac x = c in exI) - apply clarify - apply (drule_tac x = x in spec) -using [[ atp_problem_prefix = "BigO__bigo_mult2_simpler" ]] -(*sledgehammer [no luck]*); - apply (subgoal_tac "abs(f x) * abs(b x) <= abs(f x) * (c * abs(g x))") - apply (simp add: mult_ac) - apply (rule mult_left_mono, assumption) - apply (rule abs_ge_zero) -done - -declare [[ atp_problem_prefix = "BigO__bigo_mult3" ]] -lemma bigo_mult3: "f : O(h) ==> g : O(j) ==> f * g : O(h * j)" -by (metis bigo_mult set_times_intro subset_iff) - -declare [[ atp_problem_prefix = "BigO__bigo_mult4" ]] -lemma bigo_mult4 [intro]:"f : k +o O(h) ==> g * f : (g * k) +o O(g * h)" -by (metis bigo_mult2 set_plus_mono_b set_times_intro2 set_times_plus_distrib) - - -lemma bigo_mult5: "ALL x. f x ~= 0 ==> - O(f * g) <= (f::'a => ('b::ordered_field)) *o O(g)" -proof - - assume "ALL x. f x ~= 0" - show "O(f * g) <= f *o O(g)" - proof - fix h - assume "h : O(f * g)" - then have "(%x. 1 / (f x)) * h : (%x. 1 / f x) *o O(f * g)" - by auto - also have "... <= O((%x. 1 / f x) * (f * g))" - by (rule bigo_mult2) - also have "(%x. 1 / f x) * (f * g) = g" - apply (simp add: func_times) - apply (rule ext) - apply (simp add: prems nonzero_divide_eq_eq mult_ac) - done - finally have "(%x. (1::'b) / f x) * h : O(g)". - then have "f * ((%x. (1::'b) / f x) * h) : f *o O(g)" - by auto - also have "f * ((%x. (1::'b) / f x) * h) = h" - apply (simp add: func_times) - apply (rule ext) - apply (simp add: prems nonzero_divide_eq_eq mult_ac) - done - finally show "h : f *o O(g)". - qed -qed - -declare [[ atp_problem_prefix = "BigO__bigo_mult6" ]] -lemma bigo_mult6: "ALL x. f x ~= 0 ==> - O(f * g) = (f::'a => ('b::ordered_field)) *o O(g)" -by (metis bigo_mult2 bigo_mult5 order_antisym) - -(*proof requires relaxing relevance: 2007-01-25*) -declare [[ atp_problem_prefix = "BigO__bigo_mult7" ]] - declare bigo_mult6 [simp] -lemma bigo_mult7: "ALL x. f x ~= 0 ==> - O(f * g) <= O(f::'a => ('b::ordered_field)) \ O(g)" -(*sledgehammer*) - apply (subst bigo_mult6) - apply assumption - apply (rule set_times_mono3) - apply (rule bigo_refl) -done - declare bigo_mult6 [simp del] - -declare [[ atp_problem_prefix = "BigO__bigo_mult8" ]] - declare bigo_mult7[intro!] -lemma bigo_mult8: "ALL x. f x ~= 0 ==> - O(f * g) = O(f::'a => ('b::ordered_field)) \ O(g)" -by (metis bigo_mult bigo_mult7 order_antisym_conv) - -lemma bigo_minus [intro]: "f : O(g) ==> - f : O(g)" - by (auto simp add: bigo_def fun_Compl_def) - -lemma bigo_minus2: "f : g +o O(h) ==> -f : -g +o O(h)" - apply (rule set_minus_imp_plus) - apply (drule set_plus_imp_minus) - apply (drule bigo_minus) - apply (simp add: diff_minus) -done - -lemma bigo_minus3: "O(-f) = O(f)" - by (auto simp add: bigo_def fun_Compl_def abs_minus_cancel) - -lemma bigo_plus_absorb_lemma1: "f : O(g) ==> f +o O(g) <= O(g)" -proof - - assume a: "f : O(g)" - show "f +o O(g) <= O(g)" - proof - - have "f : O(f)" by auto - then have "f +o O(g) <= O(f) \ O(g)" - by (auto del: subsetI) - also have "... <= O(g) \ O(g)" - proof - - from a have "O(f) <= O(g)" by (auto del: subsetI) - thus ?thesis by (auto del: subsetI) - qed - also have "... <= O(g)" by (simp add: bigo_plus_idemp) - finally show ?thesis . - qed -qed - -lemma bigo_plus_absorb_lemma2: "f : O(g) ==> O(g) <= f +o O(g)" -proof - - assume a: "f : O(g)" - show "O(g) <= f +o O(g)" - proof - - from a have "-f : O(g)" by auto - then have "-f +o O(g) <= O(g)" by (elim bigo_plus_absorb_lemma1) - then have "f +o (-f +o O(g)) <= f +o O(g)" by auto - also have "f +o (-f +o O(g)) = O(g)" - by (simp add: set_plus_rearranges) - finally show ?thesis . - qed -qed - -declare [[ atp_problem_prefix = "BigO__bigo_plus_absorb" ]] -lemma bigo_plus_absorb [simp]: "f : O(g) ==> f +o O(g) = O(g)" -by (metis bigo_plus_absorb_lemma1 bigo_plus_absorb_lemma2 order_eq_iff); - -lemma bigo_plus_absorb2 [intro]: "f : O(g) ==> A <= O(g) ==> f +o A <= O(g)" - apply (subgoal_tac "f +o A <= f +o O(g)") - apply force+ -done - -lemma bigo_add_commute_imp: "f : g +o O(h) ==> g : f +o O(h)" - apply (subst set_minus_plus [symmetric]) - apply (subgoal_tac "g - f = - (f - g)") - apply (erule ssubst) - apply (rule bigo_minus) - apply (subst set_minus_plus) - apply assumption - apply (simp add: diff_minus add_ac) -done - -lemma bigo_add_commute: "(f : g +o O(h)) = (g : f +o O(h))" - apply (rule iffI) - apply (erule bigo_add_commute_imp)+ -done - -lemma bigo_const1: "(%x. c) : O(%x. 1)" -by (auto simp add: bigo_def mult_ac) - -declare [[ atp_problem_prefix = "BigO__bigo_const2" ]] -lemma (*bigo_const2 [intro]:*) "O(%x. c) <= O(%x. 1)" -by (metis bigo_const1 bigo_elt_subset); - -lemma bigo_const2 [intro]: "O(%x. c::'b::ordered_idom) <= O(%x. 1)"; -(*??FAILS because the two occurrences of COMBK have different polymorphic types -proof (neg_clausify) -assume 0: "\ O(COMBK (c\'b\ordered_idom)) \ O(COMBK (1\'b\ordered_idom))" -have 1: "COMBK (c\'b\ordered_idom) \ O(COMBK (1\'b\ordered_idom))" -apply (rule notI) -apply (rule 0 [THEN notE]) -apply (rule bigo_elt_subset) -apply assumption; -sorry - by (metis 0 bigo_elt_subset) loops?? -show "False" - by (metis 1 bigo_const1) -qed -*) - apply (rule bigo_elt_subset) - apply (rule bigo_const1) -done - -declare [[ atp_problem_prefix = "BigO__bigo_const3" ]] -lemma bigo_const3: "(c::'a::ordered_field) ~= 0 ==> (%x. 1) : O(%x. c)" -apply (simp add: bigo_def) -proof (neg_clausify) -assume 0: "(c\'a\ordered_field) \ (0\'a\ordered_field)" -assume 1: "\A\'a\ordered_field. \ (1\'a\ordered_field) \ A * \c\'a\ordered_field\" -have 2: "(0\'a\ordered_field) = \c\'a\ordered_field\ \ -\ (1\'a\ordered_field) \ (1\'a\ordered_field)" - by (metis 1 field_inverse) -have 3: "\c\'a\ordered_field\ = (0\'a\ordered_field)" - by (metis linorder_neq_iff linorder_antisym_conv1 2) -have 4: "(0\'a\ordered_field) = (c\'a\ordered_field)" - by (metis 3 abs_eq_0) -show "False" - by (metis 0 4) -qed - -lemma bigo_const4: "(c::'a::ordered_field) ~= 0 ==> O(%x. 1) <= O(%x. c)" -by (rule bigo_elt_subset, rule bigo_const3, assumption) - -lemma bigo_const [simp]: "(c::'a::ordered_field) ~= 0 ==> - O(%x. c) = O(%x. 1)" -by (rule equalityI, rule bigo_const2, rule bigo_const4, assumption) - -declare [[ atp_problem_prefix = "BigO__bigo_const_mult1" ]] -lemma bigo_const_mult1: "(%x. c * f x) : O(f)" - apply (simp add: bigo_def abs_mult) -proof (neg_clausify) -fix x -assume 0: "\xa\'b\ordered_idom. - \ \c\'b\ordered_idom\ * - \(f\'a\type \ 'b\ordered_idom) ((x\'b\ordered_idom \ 'a\type) xa)\ - \ xa * \f (x xa)\" -show "False" - by (metis linorder_neq_iff linorder_antisym_conv1 0) -qed - -lemma bigo_const_mult2: "O(%x. c * f x) <= O(f)" -by (rule bigo_elt_subset, rule bigo_const_mult1) - -declare [[ atp_problem_prefix = "BigO__bigo_const_mult3" ]] -lemma bigo_const_mult3: "(c::'a::ordered_field) ~= 0 ==> f : O(%x. c * f x)" - apply (simp add: bigo_def) -(*sledgehammer [no luck]*); - apply (rule_tac x = "abs(inverse c)" in exI) - apply (simp only: abs_mult [symmetric] mult_assoc [symmetric]) -apply (subst left_inverse) -apply (auto ); -done - -lemma bigo_const_mult4: "(c::'a::ordered_field) ~= 0 ==> - O(f) <= O(%x. c * f x)" -by (rule bigo_elt_subset, rule bigo_const_mult3, assumption) - -lemma bigo_const_mult [simp]: "(c::'a::ordered_field) ~= 0 ==> - O(%x. c * f x) = O(f)" -by (rule equalityI, rule bigo_const_mult2, erule bigo_const_mult4) - -declare [[ atp_problem_prefix = "BigO__bigo_const_mult5" ]] -lemma bigo_const_mult5 [simp]: "(c::'a::ordered_field) ~= 0 ==> - (%x. c) *o O(f) = O(f)" - apply (auto del: subsetI) - apply (rule order_trans) - apply (rule bigo_mult2) - apply (simp add: func_times) - apply (auto intro!: subsetI simp add: bigo_def elt_set_times_def func_times) - apply (rule_tac x = "%y. inverse c * x y" in exI) - apply (rename_tac g d) - apply safe - apply (rule_tac [2] ext) - prefer 2 - apply simp - apply (simp add: mult_assoc [symmetric] abs_mult) - (*couldn't get this proof without the step above; SLOW*) - apply (metis mult_assoc abs_ge_zero mult_left_mono) -done - - -declare [[ atp_problem_prefix = "BigO__bigo_const_mult6" ]] -lemma bigo_const_mult6 [intro]: "(%x. c) *o O(f) <= O(f)" - apply (auto intro!: subsetI - simp add: bigo_def elt_set_times_def func_times - simp del: abs_mult mult_ac) -(*sledgehammer*); - apply (rule_tac x = "ca * (abs c)" in exI) - apply (rule allI) - apply (subgoal_tac "ca * abs(c) * abs(f x) = abs(c) * (ca * abs(f x))") - apply (erule ssubst) - apply (subst abs_mult) - apply (rule mult_left_mono) - apply (erule spec) - apply simp - apply(simp add: mult_ac) -done - -lemma bigo_const_mult7 [intro]: "f =o O(g) ==> (%x. c * f x) =o O(g)" -proof - - assume "f =o O(g)" - then have "(%x. c) * f =o (%x. c) *o O(g)" - by auto - also have "(%x. c) * f = (%x. c * f x)" - by (simp add: func_times) - also have "(%x. c) *o O(g) <= O(g)" - by (auto del: subsetI) - finally show ?thesis . -qed - -lemma bigo_compose1: "f =o O(g) ==> (%x. f(k x)) =o O(%x. g(k x))" -by (unfold bigo_def, auto) - -lemma bigo_compose2: "f =o g +o O(h) ==> (%x. f(k x)) =o (%x. g(k x)) +o - O(%x. h(k x))" - apply (simp only: set_minus_plus [symmetric] diff_minus fun_Compl_def - func_plus) - apply (erule bigo_compose1) -done - -subsection {* Setsum *} - -lemma bigo_setsum_main: "ALL x. ALL y : A x. 0 <= h x y ==> - EX c. ALL x. ALL y : A x. abs(f x y) <= c * (h x y) ==> - (%x. SUM y : A x. f x y) =o O(%x. SUM y : A x. h x y)" - apply (auto simp add: bigo_def) - apply (rule_tac x = "abs c" in exI) - apply (subst abs_of_nonneg) back back - apply (rule setsum_nonneg) - apply force - apply (subst setsum_right_distrib) - apply (rule allI) - apply (rule order_trans) - apply (rule setsum_abs) - apply (rule setsum_mono) -apply (blast intro: order_trans mult_right_mono abs_ge_self) -done - -declare [[ atp_problem_prefix = "BigO__bigo_setsum1" ]] -lemma bigo_setsum1: "ALL x y. 0 <= h x y ==> - EX c. ALL x y. abs(f x y) <= c * (h x y) ==> - (%x. SUM y : A x. f x y) =o O(%x. SUM y : A x. h x y)" - apply (rule bigo_setsum_main) -(*sledgehammer*); - apply force - apply clarsimp - apply (rule_tac x = c in exI) - apply force -done - -lemma bigo_setsum2: "ALL y. 0 <= h y ==> - EX c. ALL y. abs(f y) <= c * (h y) ==> - (%x. SUM y : A x. f y) =o O(%x. SUM y : A x. h y)" -by (rule bigo_setsum1, auto) - -declare [[ atp_problem_prefix = "BigO__bigo_setsum3" ]] -lemma bigo_setsum3: "f =o O(h) ==> - (%x. SUM y : A x. (l x y) * f(k x y)) =o - O(%x. SUM y : A x. abs(l x y * h(k x y)))" - apply (rule bigo_setsum1) - apply (rule allI)+ - apply (rule abs_ge_zero) - apply (unfold bigo_def) - apply (auto simp add: abs_mult); -(*sledgehammer*); - apply (rule_tac x = c in exI) - apply (rule allI)+ - apply (subst mult_left_commute) - apply (rule mult_left_mono) - apply (erule spec) - apply (rule abs_ge_zero) -done - -lemma bigo_setsum4: "f =o g +o O(h) ==> - (%x. SUM y : A x. l x y * f(k x y)) =o - (%x. SUM y : A x. l x y * g(k x y)) +o - O(%x. SUM y : A x. abs(l x y * h(k x y)))" - apply (rule set_minus_imp_plus) - apply (subst fun_diff_def) - apply (subst setsum_subtractf [symmetric]) - apply (subst right_diff_distrib [symmetric]) - apply (rule bigo_setsum3) - apply (subst fun_diff_def [symmetric]) - apply (erule set_plus_imp_minus) -done - -declare [[ atp_problem_prefix = "BigO__bigo_setsum5" ]] -lemma bigo_setsum5: "f =o O(h) ==> ALL x y. 0 <= l x y ==> - ALL x. 0 <= h x ==> - (%x. SUM y : A x. (l x y) * f(k x y)) =o - O(%x. SUM y : A x. (l x y) * h(k x y))" - apply (subgoal_tac "(%x. SUM y : A x. (l x y) * h(k x y)) = - (%x. SUM y : A x. abs((l x y) * h(k x y)))") - apply (erule ssubst) - apply (erule bigo_setsum3) - apply (rule ext) - apply (rule setsum_cong2) - apply (thin_tac "f \ O(h)") -apply (metis abs_of_nonneg zero_le_mult_iff) -done - -lemma bigo_setsum6: "f =o g +o O(h) ==> ALL x y. 0 <= l x y ==> - ALL x. 0 <= h x ==> - (%x. SUM y : A x. (l x y) * f(k x y)) =o - (%x. SUM y : A x. (l x y) * g(k x y)) +o - O(%x. SUM y : A x. (l x y) * h(k x y))" - apply (rule set_minus_imp_plus) - apply (subst fun_diff_def) - apply (subst setsum_subtractf [symmetric]) - apply (subst right_diff_distrib [symmetric]) - apply (rule bigo_setsum5) - apply (subst fun_diff_def [symmetric]) - apply (drule set_plus_imp_minus) - apply auto -done - -subsection {* Misc useful stuff *} - -lemma bigo_useful_intro: "A <= O(f) ==> B <= O(f) ==> - A \ B <= O(f)" - apply (subst bigo_plus_idemp [symmetric]) - apply (rule set_plus_mono2) - apply assumption+ -done - -lemma bigo_useful_add: "f =o O(h) ==> g =o O(h) ==> f + g =o O(h)" - apply (subst bigo_plus_idemp [symmetric]) - apply (rule set_plus_intro) - apply assumption+ -done - -lemma bigo_useful_const_mult: "(c::'a::ordered_field) ~= 0 ==> - (%x. c) * f =o O(h) ==> f =o O(h)" - apply (rule subsetD) - apply (subgoal_tac "(%x. 1 / c) *o O(h) <= O(h)") - apply assumption - apply (rule bigo_const_mult6) - apply (subgoal_tac "f = (%x. 1 / c) * ((%x. c) * f)") - apply (erule ssubst) - apply (erule set_times_intro2) - apply (simp add: func_times) -done - -declare [[ atp_problem_prefix = "BigO__bigo_fix" ]] -lemma bigo_fix: "(%x. f ((x::nat) + 1)) =o O(%x. h(x + 1)) ==> f 0 = 0 ==> - f =o O(h)" - apply (simp add: bigo_alt_def) -(*sledgehammer*); - apply clarify - apply (rule_tac x = c in exI) - apply safe - apply (case_tac "x = 0") -apply (metis OrderedGroup.abs_ge_zero OrderedGroup.abs_zero order_less_le Ring_and_Field.split_mult_pos_le) - apply (subgoal_tac "x = Suc (x - 1)") - apply metis - apply simp - done - - -lemma bigo_fix2: - "(%x. f ((x::nat) + 1)) =o (%x. g(x + 1)) +o O(%x. h(x + 1)) ==> - f 0 = g 0 ==> f =o g +o O(h)" - apply (rule set_minus_imp_plus) - apply (rule bigo_fix) - apply (subst fun_diff_def) - apply (subst fun_diff_def [symmetric]) - apply (rule set_plus_imp_minus) - apply simp - apply (simp add: fun_diff_def) -done - -subsection {* Less than or equal to *} - -constdefs - lesso :: "('a => 'b::ordered_idom) => ('a => 'b) => ('a => 'b)" - (infixl " ALL x. abs (g x) <= abs (f x) ==> - g =o O(h)" - apply (unfold bigo_def) - apply clarsimp -apply (blast intro: order_trans) -done - -lemma bigo_lesseq2: "f =o O(h) ==> ALL x. abs (g x) <= f x ==> - g =o O(h)" - apply (erule bigo_lesseq1) -apply (blast intro: abs_ge_self order_trans) -done - -lemma bigo_lesseq3: "f =o O(h) ==> ALL x. 0 <= g x ==> ALL x. g x <= f x ==> - g =o O(h)" - apply (erule bigo_lesseq2) - apply (rule allI) - apply (subst abs_of_nonneg) - apply (erule spec)+ -done - -lemma bigo_lesseq4: "f =o O(h) ==> - ALL x. 0 <= g x ==> ALL x. g x <= abs (f x) ==> - g =o O(h)" - apply (erule bigo_lesseq1) - apply (rule allI) - apply (subst abs_of_nonneg) - apply (erule spec)+ -done - -declare [[ atp_problem_prefix = "BigO__bigo_lesso1" ]] -lemma bigo_lesso1: "ALL x. f x <= g x ==> f - ALL x. 0 <= k x ==> ALL x. k x <= f x ==> - k A. k A \ f A" -have 1: "\(X1\'b\ordered_idom) X2. \ max X1 X2 < X1" - by (metis linorder_not_less le_maxI1) (*sort inserted by hand*) -assume 2: "(0\'b) \ k x - g x" -have 3: "\ k x - g x < (0\'b)" - by (metis 2 linorder_not_less) -have 4: "\X1 X2. min X1 (k X2) \ f X2" - by (metis min_max.inf_le2 min_max.le_inf_iff min_max.le_iff_inf 0) -have 5: "\g x - f x\ = f x - g x" - by (metis abs_minus_commute combine_common_factor mult_zero_right minus_add_cancel minus_zero abs_if diff_less_eq min_max.inf_commute 4 linorder_not_le min_max.le_iff_inf 3 diff_less_0_iff_less linorder_not_less) -have 6: "max (0\'b) (k x - g x) = k x - g x" - by (metis min_max.le_iff_sup 2) -assume 7: "\ max (k x - g x) (0\'b) \ \f x - g x\" -have 8: "\ k x - g x \ f x - g x" - by (metis 5 abs_minus_commute 7 min_max.sup_commute 6) -show "False" - by (metis min_max.sup_commute min_max.inf_commute min_max.sup_inf_absorb min_max.le_iff_inf 0 max_diff_distrib_left 1 linorder_not_le 8) -qed - -declare [[ atp_problem_prefix = "BigO__bigo_lesso3" ]] -lemma bigo_lesso3: "f =o g +o O(h) ==> - ALL x. 0 <= k x ==> ALL x. g x <= k x ==> - f 'b::ordered_field) ==> - g =o h +o O(k) ==> f - EX C. ALL x. f x <= g x + C * abs(h x)" - apply (simp only: lesso_def bigo_alt_def) - apply clarsimp - apply (metis abs_if abs_mult add_commute diff_le_eq less_not_permute) -done - -end diff -r 8f35633c4922 -r 9cf389429f6d src/HOL/MetisExamples/Message.thy --- a/src/HOL/MetisExamples/Message.thy Tue Oct 20 19:37:09 2009 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,812 +0,0 @@ -(* Title: HOL/MetisTest/Message.thy - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - -Testing the metis method. -*) - -theory Message imports Main begin - -(*Needed occasionally with spy_analz_tac, e.g. in analz_insert_Key_newK*) -lemma strange_Un_eq [simp]: "A \ (B \ A) = B \ A" -by blast - -types - key = nat - -consts - all_symmetric :: bool --{*true if all keys are symmetric*} - invKey :: "key=>key" --{*inverse of a symmetric key*} - -specification (invKey) - invKey [simp]: "invKey (invKey K) = K" - invKey_symmetric: "all_symmetric --> invKey = id" - by (rule exI [of _ id], auto) - - -text{*The inverse of a symmetric key is itself; that of a public key - is the private key and vice versa*} - -constdefs - symKeys :: "key set" - "symKeys == {K. invKey K = K}" - -datatype --{*We allow any number of friendly agents*} - agent = Server | Friend nat | Spy - -datatype - msg = Agent agent --{*Agent names*} - | Number nat --{*Ordinary integers, timestamps, ...*} - | Nonce nat --{*Unguessable nonces*} - | Key key --{*Crypto keys*} - | Hash msg --{*Hashing*} - | MPair msg msg --{*Compound messages*} - | Crypt key msg --{*Encryption, public- or shared-key*} - - -text{*Concrete syntax: messages appear as {|A,B,NA|}, etc...*} -syntax - "@MTuple" :: "['a, args] => 'a * 'b" ("(2{|_,/ _|})") - -syntax (xsymbols) - "@MTuple" :: "['a, args] => 'a * 'b" ("(2\_,/ _\)") - -translations - "{|x, y, z|}" == "{|x, {|y, z|}|}" - "{|x, y|}" == "MPair x y" - - -constdefs - HPair :: "[msg,msg] => msg" ("(4Hash[_] /_)" [0, 1000]) - --{*Message Y paired with a MAC computed with the help of X*} - "Hash[X] Y == {| Hash{|X,Y|}, Y|}" - - keysFor :: "msg set => key set" - --{*Keys useful to decrypt elements of a message set*} - "keysFor H == invKey ` {K. \X. Crypt K X \ H}" - - -subsubsection{*Inductive Definition of All Parts" of a Message*} - -inductive_set - parts :: "msg set => msg set" - for H :: "msg set" - where - Inj [intro]: "X \ H ==> X \ parts H" - | Fst: "{|X,Y|} \ parts H ==> X \ parts H" - | Snd: "{|X,Y|} \ parts H ==> Y \ parts H" - | Body: "Crypt K X \ parts H ==> X \ parts H" - - -declare [[ atp_problem_prefix = "Message__parts_mono" ]] -lemma parts_mono: "G \ H ==> parts(G) \ parts(H)" -apply auto -apply (erule parts.induct) -apply (metis Inj set_mp) -apply (metis Fst) -apply (metis Snd) -apply (metis Body) -done - - -text{*Equations hold because constructors are injective.*} -lemma Friend_image_eq [simp]: "(Friend x \ Friend`A) = (x:A)" -by auto - -lemma Key_image_eq [simp]: "(Key x \ Key`A) = (x\A)" -by auto - -lemma Nonce_Key_image_eq [simp]: "(Nonce x \ Key`A)" -by auto - - -subsubsection{*Inverse of keys *} - -declare [[ atp_problem_prefix = "Message__invKey_eq" ]] -lemma invKey_eq [simp]: "(invKey K = invKey K') = (K=K')" -by (metis invKey) - - -subsection{*keysFor operator*} - -lemma keysFor_empty [simp]: "keysFor {} = {}" -by (unfold keysFor_def, blast) - -lemma keysFor_Un [simp]: "keysFor (H \ H') = keysFor H \ keysFor H'" -by (unfold keysFor_def, blast) - -lemma keysFor_UN [simp]: "keysFor (\i\A. H i) = (\i\A. keysFor (H i))" -by (unfold keysFor_def, blast) - -text{*Monotonicity*} -lemma keysFor_mono: "G \ H ==> keysFor(G) \ keysFor(H)" -by (unfold keysFor_def, blast) - -lemma keysFor_insert_Agent [simp]: "keysFor (insert (Agent A) H) = keysFor H" -by (unfold keysFor_def, auto) - -lemma keysFor_insert_Nonce [simp]: "keysFor (insert (Nonce N) H) = keysFor H" -by (unfold keysFor_def, auto) - -lemma keysFor_insert_Number [simp]: "keysFor (insert (Number N) H) = keysFor H" -by (unfold keysFor_def, auto) - -lemma keysFor_insert_Key [simp]: "keysFor (insert (Key K) H) = keysFor H" -by (unfold keysFor_def, auto) - -lemma keysFor_insert_Hash [simp]: "keysFor (insert (Hash X) H) = keysFor H" -by (unfold keysFor_def, auto) - -lemma keysFor_insert_MPair [simp]: "keysFor (insert {|X,Y|} H) = keysFor H" -by (unfold keysFor_def, auto) - -lemma keysFor_insert_Crypt [simp]: - "keysFor (insert (Crypt K X) H) = insert (invKey K) (keysFor H)" -by (unfold keysFor_def, auto) - -lemma keysFor_image_Key [simp]: "keysFor (Key`E) = {}" -by (unfold keysFor_def, auto) - -lemma Crypt_imp_invKey_keysFor: "Crypt K X \ H ==> invKey K \ keysFor H" -by (unfold keysFor_def, blast) - - -subsection{*Inductive relation "parts"*} - -lemma MPair_parts: - "[| {|X,Y|} \ parts H; - [| X \ parts H; Y \ parts H |] ==> P |] ==> P" -by (blast dest: parts.Fst parts.Snd) - - declare MPair_parts [elim!] parts.Body [dest!] -text{*NB These two rules are UNSAFE in the formal sense, as they discard the - compound message. They work well on THIS FILE. - @{text MPair_parts} is left as SAFE because it speeds up proofs. - The Crypt rule is normally kept UNSAFE to avoid breaking up certificates.*} - -lemma parts_increasing: "H \ parts(H)" -by blast - -lemmas parts_insertI = subset_insertI [THEN parts_mono, THEN subsetD, standard] - -lemma parts_empty [simp]: "parts{} = {}" -apply safe -apply (erule parts.induct) -apply blast+ -done - -lemma parts_emptyE [elim!]: "X\ parts{} ==> P" -by simp - -text{*WARNING: loops if H = {Y}, therefore must not be repeated!*} -lemma parts_singleton: "X\ parts H ==> \Y\H. X\ parts {Y}" -apply (erule parts.induct) -apply fast+ -done - - -subsubsection{*Unions *} - -lemma parts_Un_subset1: "parts(G) \ parts(H) \ parts(G \ H)" -by (intro Un_least parts_mono Un_upper1 Un_upper2) - -lemma parts_Un_subset2: "parts(G \ H) \ parts(G) \ parts(H)" -apply (rule subsetI) -apply (erule parts.induct, blast+) -done - -lemma parts_Un [simp]: "parts(G \ H) = parts(G) \ parts(H)" -by (intro equalityI parts_Un_subset1 parts_Un_subset2) - -lemma parts_insert: "parts (insert X H) = parts {X} \ parts H" -apply (subst insert_is_Un [of _ H]) -apply (simp only: parts_Un) -done - -declare [[ atp_problem_prefix = "Message__parts_insert_two" ]] -lemma parts_insert2: - "parts (insert X (insert Y H)) = parts {X} \ parts {Y} \ parts H" -by (metis Un_commute Un_empty_left Un_empty_right Un_insert_left Un_insert_right parts_Un) - - -lemma parts_UN_subset1: "(\x\A. parts(H x)) \ parts(\x\A. H x)" -by (intro UN_least parts_mono UN_upper) - -lemma parts_UN_subset2: "parts(\x\A. H x) \ (\x\A. parts(H x))" -apply (rule subsetI) -apply (erule parts.induct, blast+) -done - -lemma parts_UN [simp]: "parts(\x\A. H x) = (\x\A. parts(H x))" -by (intro equalityI parts_UN_subset1 parts_UN_subset2) - -text{*Added to simplify arguments to parts, analz and synth. - NOTE: the UN versions are no longer used!*} - - -text{*This allows @{text blast} to simplify occurrences of - @{term "parts(G\H)"} in the assumption.*} -lemmas in_parts_UnE = parts_Un [THEN equalityD1, THEN subsetD, THEN UnE] -declare in_parts_UnE [elim!] - -lemma parts_insert_subset: "insert X (parts H) \ parts(insert X H)" -by (blast intro: parts_mono [THEN [2] rev_subsetD]) - -subsubsection{*Idempotence and transitivity *} - -lemma parts_partsD [dest!]: "X\ parts (parts H) ==> X\ parts H" -by (erule parts.induct, blast+) - -lemma parts_idem [simp]: "parts (parts H) = parts H" -by blast - -declare [[ atp_problem_prefix = "Message__parts_subset_iff" ]] -lemma parts_subset_iff [simp]: "(parts G \ parts H) = (G \ parts H)" -apply (rule iffI) -apply (metis Un_absorb1 Un_subset_iff parts_Un parts_increasing) -apply (metis parts_idem parts_mono) -done - -lemma parts_trans: "[| X\ parts G; G \ parts H |] ==> X\ parts H" -by (blast dest: parts_mono); - - -declare [[ atp_problem_prefix = "Message__parts_cut" ]] -lemma parts_cut: "[|Y\ parts(insert X G); X\ parts H|] ==> Y\ parts(G \ H)" -by (metis Un_subset_iff insert_subset parts_increasing parts_trans) - - - -subsubsection{*Rewrite rules for pulling out atomic messages *} - -lemmas parts_insert_eq_I = equalityI [OF subsetI parts_insert_subset] - - -lemma parts_insert_Agent [simp]: - "parts (insert (Agent agt) H) = insert (Agent agt) (parts H)" -apply (rule parts_insert_eq_I) -apply (erule parts.induct, auto) -done - -lemma parts_insert_Nonce [simp]: - "parts (insert (Nonce N) H) = insert (Nonce N) (parts H)" -apply (rule parts_insert_eq_I) -apply (erule parts.induct, auto) -done - -lemma parts_insert_Number [simp]: - "parts (insert (Number N) H) = insert (Number N) (parts H)" -apply (rule parts_insert_eq_I) -apply (erule parts.induct, auto) -done - -lemma parts_insert_Key [simp]: - "parts (insert (Key K) H) = insert (Key K) (parts H)" -apply (rule parts_insert_eq_I) -apply (erule parts.induct, auto) -done - -lemma parts_insert_Hash [simp]: - "parts (insert (Hash X) H) = insert (Hash X) (parts H)" -apply (rule parts_insert_eq_I) -apply (erule parts.induct, auto) -done - -lemma parts_insert_Crypt [simp]: - "parts (insert (Crypt K X) H) = - insert (Crypt K X) (parts (insert X H))" -apply (rule equalityI) -apply (rule subsetI) -apply (erule parts.induct, auto) -apply (blast intro: parts.Body) -done - -lemma parts_insert_MPair [simp]: - "parts (insert {|X,Y|} H) = - insert {|X,Y|} (parts (insert X (insert Y H)))" -apply (rule equalityI) -apply (rule subsetI) -apply (erule parts.induct, auto) -apply (blast intro: parts.Fst parts.Snd)+ -done - -lemma parts_image_Key [simp]: "parts (Key`N) = Key`N" -apply auto -apply (erule parts.induct, auto) -done - - -declare [[ atp_problem_prefix = "Message__msg_Nonce_supply" ]] -lemma msg_Nonce_supply: "\N. \n. N\n --> Nonce n \ parts {msg}" -apply (induct_tac "msg") -apply (simp_all add: parts_insert2) -apply (metis Suc_n_not_le_n) -apply (metis le_trans linorder_linear) -done - -subsection{*Inductive relation "analz"*} - -text{*Inductive definition of "analz" -- what can be broken down from a set of - messages, including keys. A form of downward closure. Pairs can - be taken apart; messages decrypted with known keys. *} - -inductive_set - analz :: "msg set => msg set" - for H :: "msg set" - where - Inj [intro,simp] : "X \ H ==> X \ analz H" - | Fst: "{|X,Y|} \ analz H ==> X \ analz H" - | Snd: "{|X,Y|} \ analz H ==> Y \ analz H" - | Decrypt [dest]: - "[|Crypt K X \ analz H; Key(invKey K): analz H|] ==> X \ analz H" - - -text{*Monotonicity; Lemma 1 of Lowe's paper*} -lemma analz_mono: "G\H ==> analz(G) \ analz(H)" -apply auto -apply (erule analz.induct) -apply (auto dest: analz.Fst analz.Snd) -done - -text{*Making it safe speeds up proofs*} -lemma MPair_analz [elim!]: - "[| {|X,Y|} \ analz H; - [| X \ analz H; Y \ analz H |] ==> P - |] ==> P" -by (blast dest: analz.Fst analz.Snd) - -lemma analz_increasing: "H \ analz(H)" -by blast - -lemma analz_subset_parts: "analz H \ parts H" -apply (rule subsetI) -apply (erule analz.induct, blast+) -done - -lemmas analz_into_parts = analz_subset_parts [THEN subsetD, standard] - -lemmas not_parts_not_analz = analz_subset_parts [THEN contra_subsetD, standard] - - -declare [[ atp_problem_prefix = "Message__parts_analz" ]] -lemma parts_analz [simp]: "parts (analz H) = parts H" -apply (rule equalityI) -apply (metis analz_subset_parts parts_subset_iff) -apply (metis analz_increasing parts_mono) -done - - -lemma analz_parts [simp]: "analz (parts H) = parts H" -apply auto -apply (erule analz.induct, auto) -done - -lemmas analz_insertI = subset_insertI [THEN analz_mono, THEN [2] rev_subsetD, standard] - -subsubsection{*General equational properties *} - -lemma analz_empty [simp]: "analz{} = {}" -apply safe -apply (erule analz.induct, blast+) -done - -text{*Converse fails: we can analz more from the union than from the - separate parts, as a key in one might decrypt a message in the other*} -lemma analz_Un: "analz(G) \ analz(H) \ analz(G \ H)" -by (intro Un_least analz_mono Un_upper1 Un_upper2) - -lemma analz_insert: "insert X (analz H) \ analz(insert X H)" -by (blast intro: analz_mono [THEN [2] rev_subsetD]) - -subsubsection{*Rewrite rules for pulling out atomic messages *} - -lemmas analz_insert_eq_I = equalityI [OF subsetI analz_insert] - -lemma analz_insert_Agent [simp]: - "analz (insert (Agent agt) H) = insert (Agent agt) (analz H)" -apply (rule analz_insert_eq_I) -apply (erule analz.induct, auto) -done - -lemma analz_insert_Nonce [simp]: - "analz (insert (Nonce N) H) = insert (Nonce N) (analz H)" -apply (rule analz_insert_eq_I) -apply (erule analz.induct, auto) -done - -lemma analz_insert_Number [simp]: - "analz (insert (Number N) H) = insert (Number N) (analz H)" -apply (rule analz_insert_eq_I) -apply (erule analz.induct, auto) -done - -lemma analz_insert_Hash [simp]: - "analz (insert (Hash X) H) = insert (Hash X) (analz H)" -apply (rule analz_insert_eq_I) -apply (erule analz.induct, auto) -done - -text{*Can only pull out Keys if they are not needed to decrypt the rest*} -lemma analz_insert_Key [simp]: - "K \ keysFor (analz H) ==> - analz (insert (Key K) H) = insert (Key K) (analz H)" -apply (unfold keysFor_def) -apply (rule analz_insert_eq_I) -apply (erule analz.induct, auto) -done - -lemma analz_insert_MPair [simp]: - "analz (insert {|X,Y|} H) = - insert {|X,Y|} (analz (insert X (insert Y H)))" -apply (rule equalityI) -apply (rule subsetI) -apply (erule analz.induct, auto) -apply (erule analz.induct) -apply (blast intro: analz.Fst analz.Snd)+ -done - -text{*Can pull out enCrypted message if the Key is not known*} -lemma analz_insert_Crypt: - "Key (invKey K) \ analz H - ==> analz (insert (Crypt K X) H) = insert (Crypt K X) (analz H)" -apply (rule analz_insert_eq_I) -apply (erule analz.induct, auto) - -done - -lemma lemma1: "Key (invKey K) \ analz H ==> - analz (insert (Crypt K X) H) \ - insert (Crypt K X) (analz (insert X H))" -apply (rule subsetI) -apply (erule_tac x = x in analz.induct, auto) -done - -lemma lemma2: "Key (invKey K) \ analz H ==> - insert (Crypt K X) (analz (insert X H)) \ - analz (insert (Crypt K X) H)" -apply auto -apply (erule_tac x = x in analz.induct, auto) -apply (blast intro: analz_insertI analz.Decrypt) -done - -lemma analz_insert_Decrypt: - "Key (invKey K) \ analz H ==> - analz (insert (Crypt K X) H) = - insert (Crypt K X) (analz (insert X H))" -by (intro equalityI lemma1 lemma2) - -text{*Case analysis: either the message is secure, or it is not! Effective, -but can cause subgoals to blow up! Use with @{text "split_if"}; apparently -@{text "split_tac"} does not cope with patterns such as @{term"analz (insert -(Crypt K X) H)"} *} -lemma analz_Crypt_if [simp]: - "analz (insert (Crypt K X) H) = - (if (Key (invKey K) \ analz H) - then insert (Crypt K X) (analz (insert X H)) - else insert (Crypt K X) (analz H))" -by (simp add: analz_insert_Crypt analz_insert_Decrypt) - - -text{*This rule supposes "for the sake of argument" that we have the key.*} -lemma analz_insert_Crypt_subset: - "analz (insert (Crypt K X) H) \ - insert (Crypt K X) (analz (insert X H))" -apply (rule subsetI) -apply (erule analz.induct, auto) -done - - -lemma analz_image_Key [simp]: "analz (Key`N) = Key`N" -apply auto -apply (erule analz.induct, auto) -done - - -subsubsection{*Idempotence and transitivity *} - -lemma analz_analzD [dest!]: "X\ analz (analz H) ==> X\ analz H" -by (erule analz.induct, blast+) - -lemma analz_idem [simp]: "analz (analz H) = analz H" -by blast - -lemma analz_subset_iff [simp]: "(analz G \ analz H) = (G \ analz H)" -apply (rule iffI) -apply (iprover intro: subset_trans analz_increasing) -apply (frule analz_mono, simp) -done - -lemma analz_trans: "[| X\ analz G; G \ analz H |] ==> X\ analz H" -by (drule analz_mono, blast) - - -declare [[ atp_problem_prefix = "Message__analz_cut" ]] - declare analz_trans[intro] -lemma analz_cut: "[| Y\ analz (insert X H); X\ analz H |] ==> Y\ analz H" -(*TOO SLOW -by (metis analz_idem analz_increasing analz_mono insert_absorb insert_mono insert_subset) --{*317s*} -??*) -by (erule analz_trans, blast) - - -text{*This rewrite rule helps in the simplification of messages that involve - the forwarding of unknown components (X). Without it, removing occurrences - of X can be very complicated. *} -lemma analz_insert_eq: "X\ analz H ==> analz (insert X H) = analz H" -by (blast intro: analz_cut analz_insertI) - - -text{*A congruence rule for "analz" *} - -declare [[ atp_problem_prefix = "Message__analz_subset_cong" ]] -lemma analz_subset_cong: - "[| analz G \ analz G'; analz H \ analz H' |] - ==> analz (G \ H) \ analz (G' \ H')" -apply simp -apply (metis Un_absorb2 Un_commute Un_subset_iff Un_upper1 Un_upper2 analz_mono) -done - - -lemma analz_cong: - "[| analz G = analz G'; analz H = analz H' - |] ==> analz (G \ H) = analz (G' \ H')" -by (intro equalityI analz_subset_cong, simp_all) - -lemma analz_insert_cong: - "analz H = analz H' ==> analz(insert X H) = analz(insert X H')" -by (force simp only: insert_def intro!: analz_cong) - -text{*If there are no pairs or encryptions then analz does nothing*} -lemma analz_trivial: - "[| \X Y. {|X,Y|} \ H; \X K. Crypt K X \ H |] ==> analz H = H" -apply safe -apply (erule analz.induct, blast+) -done - -text{*These two are obsolete (with a single Spy) but cost little to prove...*} -lemma analz_UN_analz_lemma: - "X\ analz (\i\A. analz (H i)) ==> X\ analz (\i\A. H i)" -apply (erule analz.induct) -apply (blast intro: analz_mono [THEN [2] rev_subsetD])+ -done - -lemma analz_UN_analz [simp]: "analz (\i\A. analz (H i)) = analz (\i\A. H i)" -by (blast intro: analz_UN_analz_lemma analz_mono [THEN [2] rev_subsetD]) - - -subsection{*Inductive relation "synth"*} - -text{*Inductive definition of "synth" -- what can be built up from a set of - messages. A form of upward closure. Pairs can be built, messages - encrypted with known keys. Agent names are public domain. - Numbers can be guessed, but Nonces cannot be. *} - -inductive_set - synth :: "msg set => msg set" - for H :: "msg set" - where - Inj [intro]: "X \ H ==> X \ synth H" - | Agent [intro]: "Agent agt \ synth H" - | Number [intro]: "Number n \ synth H" - | Hash [intro]: "X \ synth H ==> Hash X \ synth H" - | MPair [intro]: "[|X \ synth H; Y \ synth H|] ==> {|X,Y|} \ synth H" - | Crypt [intro]: "[|X \ synth H; Key(K) \ H|] ==> Crypt K X \ synth H" - -text{*Monotonicity*} -lemma synth_mono: "G\H ==> synth(G) \ synth(H)" - by (auto, erule synth.induct, auto) - -text{*NO @{text Agent_synth}, as any Agent name can be synthesized. - The same holds for @{term Number}*} -inductive_cases Nonce_synth [elim!]: "Nonce n \ synth H" -inductive_cases Key_synth [elim!]: "Key K \ synth H" -inductive_cases Hash_synth [elim!]: "Hash X \ synth H" -inductive_cases MPair_synth [elim!]: "{|X,Y|} \ synth H" -inductive_cases Crypt_synth [elim!]: "Crypt K X \ synth H" - - -lemma synth_increasing: "H \ synth(H)" -by blast - -subsubsection{*Unions *} - -text{*Converse fails: we can synth more from the union than from the - separate parts, building a compound message using elements of each.*} -lemma synth_Un: "synth(G) \ synth(H) \ synth(G \ H)" -by (intro Un_least synth_mono Un_upper1 Un_upper2) - - -declare [[ atp_problem_prefix = "Message__synth_insert" ]] - -lemma synth_insert: "insert X (synth H) \ synth(insert X H)" -by (metis insert_iff insert_subset subset_insertI synth.Inj synth_mono) - -subsubsection{*Idempotence and transitivity *} - -lemma synth_synthD [dest!]: "X\ synth (synth H) ==> X\ synth H" -by (erule synth.induct, blast+) - -lemma synth_idem: "synth (synth H) = synth H" -by blast - -lemma synth_subset_iff [simp]: "(synth G \ synth H) = (G \ synth H)" -apply (rule iffI) -apply (iprover intro: subset_trans synth_increasing) -apply (frule synth_mono, simp add: synth_idem) -done - -lemma synth_trans: "[| X\ synth G; G \ synth H |] ==> X\ synth H" -by (drule synth_mono, blast) - -declare [[ atp_problem_prefix = "Message__synth_cut" ]] -lemma synth_cut: "[| Y\ synth (insert X H); X\ synth H |] ==> Y\ synth H" -(*TOO SLOW -by (metis insert_absorb insert_mono insert_subset synth_idem synth_increasing synth_mono) -*) -by (erule synth_trans, blast) - - -lemma Agent_synth [simp]: "Agent A \ synth H" -by blast - -lemma Number_synth [simp]: "Number n \ synth H" -by blast - -lemma Nonce_synth_eq [simp]: "(Nonce N \ synth H) = (Nonce N \ H)" -by blast - -lemma Key_synth_eq [simp]: "(Key K \ synth H) = (Key K \ H)" -by blast - -lemma Crypt_synth_eq [simp]: - "Key K \ H ==> (Crypt K X \ synth H) = (Crypt K X \ H)" -by blast - - -lemma keysFor_synth [simp]: - "keysFor (synth H) = keysFor H \ invKey`{K. Key K \ H}" -by (unfold keysFor_def, blast) - - -subsubsection{*Combinations of parts, analz and synth *} - -declare [[ atp_problem_prefix = "Message__parts_synth" ]] -lemma parts_synth [simp]: "parts (synth H) = parts H \ synth H" -apply (rule equalityI) -apply (rule subsetI) -apply (erule parts.induct) -apply (metis UnCI) -apply (metis MPair_synth UnCI UnE insert_absorb insert_subset parts.Fst parts_increasing) -apply (metis MPair_synth UnCI UnE insert_absorb insert_subset parts.Snd parts_increasing) -apply (metis Body Crypt_synth UnCI UnE insert_absorb insert_subset parts_increasing) -apply (metis Un_subset_iff parts_increasing parts_mono synth_increasing) -done - - - - -declare [[ atp_problem_prefix = "Message__analz_analz_Un" ]] -lemma analz_analz_Un [simp]: "analz (analz G \ H) = analz (G \ H)" -apply (rule equalityI); -apply (metis analz_idem analz_subset_cong order_eq_refl) -apply (metis analz_increasing analz_subset_cong order_eq_refl) -done - -declare [[ atp_problem_prefix = "Message__analz_synth_Un" ]] - declare analz_mono [intro] analz.Fst [intro] analz.Snd [intro] Un_least [intro] -lemma analz_synth_Un [simp]: "analz (synth G \ H) = analz (G \ H) \ synth G" -apply (rule equalityI) -apply (rule subsetI) -apply (erule analz.induct) -apply (metis UnCI UnE Un_commute analz.Inj) -apply (metis MPair_synth UnCI UnE Un_commute Un_upper1 analz.Fst analz_increasing analz_mono insert_absorb insert_subset) -apply (metis MPair_synth UnCI UnE Un_commute Un_upper1 analz.Snd analz_increasing analz_mono insert_absorb insert_subset) -apply (blast intro: analz.Decrypt) -apply blast -done - - -declare [[ atp_problem_prefix = "Message__analz_synth" ]] -lemma analz_synth [simp]: "analz (synth H) = analz H \ synth H" -proof (neg_clausify) -assume 0: "analz (synth H) \ analz H \ synth H" -have 1: "\X1 X3. sup (analz (sup X3 X1)) (synth X3) = analz (sup (synth X3) X1)" - by (metis analz_synth_Un) -have 2: "sup (analz H) (synth H) \ analz (synth H)" - by (metis 0) -have 3: "\X1 X3. sup (synth X3) (analz (sup X3 X1)) = analz (sup (synth X3) X1)" - by (metis 1 Un_commute) -have 4: "\X3. sup (synth X3) (analz X3) = analz (sup (synth X3) {})" - by (metis 3 Un_empty_right) -have 5: "\X3. sup (synth X3) (analz X3) = analz (synth X3)" - by (metis 4 Un_empty_right) -have 6: "\X3. sup (analz X3) (synth X3) = analz (synth X3)" - by (metis 5 Un_commute) -show "False" - by (metis 2 6) -qed - - -subsubsection{*For reasoning about the Fake rule in traces *} - -declare [[ atp_problem_prefix = "Message__parts_insert_subset_Un" ]] -lemma parts_insert_subset_Un: "X\ G ==> parts(insert X H) \ parts G \ parts H" -proof (neg_clausify) -assume 0: "X \ G" -assume 1: "\ parts (insert X H) \ parts G \ parts H" -have 2: "\ parts (insert X H) \ parts (G \ H)" - by (metis 1 parts_Un) -have 3: "\ insert X H \ G \ H" - by (metis 2 parts_mono) -have 4: "X \ G \ H \ \ H \ G \ H" - by (metis 3 insert_subset) -have 5: "X \ G \ H" - by (metis 4 Un_upper2) -have 6: "X \ G" - by (metis 5 UnCI) -show "False" - by (metis 6 0) -qed - -declare [[ atp_problem_prefix = "Message__Fake_parts_insert" ]] -lemma Fake_parts_insert: - "X \ synth (analz H) ==> - parts (insert X H) \ synth (analz H) \ parts H" -proof (neg_clausify) -assume 0: "X \ synth (analz H)" -assume 1: "\ parts (insert X H) \ synth (analz H) \ parts H" -have 2: "\X3. parts X3 \ synth (analz X3) = parts (synth (analz X3))" - by (metis parts_synth parts_analz) -have 3: "\X3. analz X3 \ synth (analz X3) = analz (synth (analz X3))" - by (metis analz_synth analz_idem) -have 4: "\X3. analz X3 \ analz (synth X3)" - by (metis Un_upper1 analz_synth) -have 5: "\ parts (insert X H) \ parts H \ synth (analz H)" - by (metis 1 Un_commute) -have 6: "\ parts (insert X H) \ parts (synth (analz H))" - by (metis 5 2) -have 7: "\ insert X H \ synth (analz H)" - by (metis 6 parts_mono) -have 8: "X \ synth (analz H) \ \ H \ synth (analz H)" - by (metis 7 insert_subset) -have 9: "\ H \ synth (analz H)" - by (metis 8 0) -have 10: "\X3. X3 \ analz (synth X3)" - by (metis analz_subset_iff 4) -have 11: "\X3. X3 \ analz (synth (analz X3))" - by (metis analz_subset_iff 10) -have 12: "\X3. analz (synth (analz X3)) = synth (analz X3) \ - \ analz X3 \ synth (analz X3)" - by (metis Un_absorb1 3) -have 13: "\X3. analz (synth (analz X3)) = synth (analz X3)" - by (metis 12 synth_increasing) -have 14: "\X3. X3 \ synth (analz X3)" - by (metis 11 13) -show "False" - by (metis 9 14) -qed - -lemma Fake_parts_insert_in_Un: - "[|Z \ parts (insert X H); X: synth (analz H)|] - ==> Z \ synth (analz H) \ parts H"; -by (blast dest: Fake_parts_insert [THEN subsetD, dest]) - -declare [[ atp_problem_prefix = "Message__Fake_analz_insert" ]] - declare analz_mono [intro] synth_mono [intro] -lemma Fake_analz_insert: - "X\ synth (analz G) ==> - analz (insert X H) \ synth (analz G) \ analz (G \ H)" -by (metis Un_commute Un_insert_left Un_insert_right Un_upper1 analz_analz_Un analz_mono analz_synth_Un equalityE insert_absorb order_le_less xt1(12)) - -declare [[ atp_problem_prefix = "Message__Fake_analz_insert_simpler" ]] -(*simpler problems? BUT METIS CAN'T PROVE -lemma Fake_analz_insert_simpler: - "X\ synth (analz G) ==> - analz (insert X H) \ synth (analz G) \ analz (G \ H)" -apply (rule subsetI) -apply (subgoal_tac "x \ analz (synth (analz G) \ H) ") -apply (metis Un_commute analz_analz_Un analz_synth_Un) -apply (metis Un_commute Un_upper1 Un_upper2 analz_cut analz_increasing analz_mono insert_absorb insert_mono insert_subset) -done -*) - -end diff -r 8f35633c4922 -r 9cf389429f6d src/HOL/MetisExamples/ROOT.ML --- a/src/HOL/MetisExamples/ROOT.ML Tue Oct 20 19:37:09 2009 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,9 +0,0 @@ -(* Title: HOL/MetisExamples/ROOT.ML - ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - -Testing the metis method. -*) - -use_thys ["set", "BigO", "Abstraction", "BT", "Message", "Tarski", "TransClosure"]; - diff -r 8f35633c4922 -r 9cf389429f6d src/HOL/MetisExamples/Tarski.thy --- a/src/HOL/MetisExamples/Tarski.thy Tue Oct 20 19:37:09 2009 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1110 +0,0 @@ -(* Title: HOL/MetisTest/Tarski.thy - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - -Testing the metis method. -*) - -header {* The Full Theorem of Tarski *} - -theory Tarski -imports Main FuncSet -begin - -(*Many of these higher-order problems appear to be impossible using the -current linkup. They often seem to need either higher-order unification -or explicit reasoning about connectives such as conjunction. The numerous -set comprehensions are to blame.*) - - -record 'a potype = - pset :: "'a set" - order :: "('a * 'a) set" - -constdefs - monotone :: "['a => 'a, 'a set, ('a *'a)set] => bool" - "monotone f A r == \x\A. \y\A. (x, y): r --> ((f x), (f y)) : r" - - least :: "['a => bool, 'a potype] => 'a" - "least P po == @ x. x: pset po & P x & - (\y \ pset po. P y --> (x,y): order po)" - - greatest :: "['a => bool, 'a potype] => 'a" - "greatest P po == @ x. x: pset po & P x & - (\y \ pset po. P y --> (y,x): order po)" - - lub :: "['a set, 'a potype] => 'a" - "lub S po == least (%x. \y\S. (y,x): order po) po" - - glb :: "['a set, 'a potype] => 'a" - "glb S po == greatest (%x. \y\S. (x,y): order po) po" - - isLub :: "['a set, 'a potype, 'a] => bool" - "isLub S po == %L. (L: pset po & (\y\S. (y,L): order po) & - (\z\pset po. (\y\S. (y,z): order po) --> (L,z): order po))" - - isGlb :: "['a set, 'a potype, 'a] => bool" - "isGlb S po == %G. (G: pset po & (\y\S. (G,y): order po) & - (\z \ pset po. (\y\S. (z,y): order po) --> (z,G): order po))" - - "fix" :: "[('a => 'a), 'a set] => 'a set" - "fix f A == {x. x: A & f x = x}" - - interval :: "[('a*'a) set,'a, 'a ] => 'a set" - "interval r a b == {x. (a,x): r & (x,b): r}" - -constdefs - Bot :: "'a potype => 'a" - "Bot po == least (%x. True) po" - - Top :: "'a potype => 'a" - "Top po == greatest (%x. True) po" - - PartialOrder :: "('a potype) set" - "PartialOrder == {P. refl_on (pset P) (order P) & antisym (order P) & - trans (order P)}" - - CompleteLattice :: "('a potype) set" - "CompleteLattice == {cl. cl: PartialOrder & - (\S. S \ pset cl --> (\L. isLub S cl L)) & - (\S. S \ pset cl --> (\G. isGlb S cl G))}" - - induced :: "['a set, ('a * 'a) set] => ('a *'a)set" - "induced A r == {(a,b). a : A & b: A & (a,b): r}" - -constdefs - sublattice :: "('a potype * 'a set)set" - "sublattice == - SIGMA cl: CompleteLattice. - {S. S \ pset cl & - (| pset = S, order = induced S (order cl) |): CompleteLattice }" - -syntax - "@SL" :: "['a set, 'a potype] => bool" ("_ <<= _" [51,50]50) - -translations - "S <<= cl" == "S : sublattice `` {cl}" - -constdefs - dual :: "'a potype => 'a potype" - "dual po == (| pset = pset po, order = converse (order po) |)" - -locale PO = - fixes cl :: "'a potype" - and A :: "'a set" - and r :: "('a * 'a) set" - assumes cl_po: "cl : PartialOrder" - defines A_def: "A == pset cl" - and r_def: "r == order cl" - -locale CL = PO + - assumes cl_co: "cl : CompleteLattice" - -definition CLF_set :: "('a potype * ('a => 'a)) set" where - "CLF_set = (SIGMA cl: CompleteLattice. - {f. f: pset cl -> pset cl & monotone f (pset cl) (order cl)})" - -locale CLF = CL + - fixes f :: "'a => 'a" - and P :: "'a set" - assumes f_cl: "(cl,f) : CLF_set" (*was the equivalent "f : CLF``{cl}"*) - defines P_def: "P == fix f A" - - -locale Tarski = CLF + - fixes Y :: "'a set" - and intY1 :: "'a set" - and v :: "'a" - assumes - Y_ss: "Y \ P" - defines - intY1_def: "intY1 == interval r (lub Y cl) (Top cl)" - and v_def: "v == glb {x. ((%x: intY1. f x) x, x): induced intY1 r & - x: intY1} - (| pset=intY1, order=induced intY1 r|)" - - -subsection {* Partial Order *} - -lemma (in PO) PO_imp_refl_on: "refl_on A r" -apply (insert cl_po) -apply (simp add: PartialOrder_def A_def r_def) -done - -lemma (in PO) PO_imp_sym: "antisym r" -apply (insert cl_po) -apply (simp add: PartialOrder_def r_def) -done - -lemma (in PO) PO_imp_trans: "trans r" -apply (insert cl_po) -apply (simp add: PartialOrder_def r_def) -done - -lemma (in PO) reflE: "x \ A ==> (x, x) \ r" -apply (insert cl_po) -apply (simp add: PartialOrder_def refl_on_def A_def r_def) -done - -lemma (in PO) antisymE: "[| (a, b) \ r; (b, a) \ r |] ==> a = b" -apply (insert cl_po) -apply (simp add: PartialOrder_def antisym_def r_def) -done - -lemma (in PO) transE: "[| (a, b) \ r; (b, c) \ r|] ==> (a,c) \ r" -apply (insert cl_po) -apply (simp add: PartialOrder_def r_def) -apply (unfold trans_def, fast) -done - -lemma (in PO) monotoneE: - "[| monotone f A r; x \ A; y \ A; (x, y) \ r |] ==> (f x, f y) \ r" -by (simp add: monotone_def) - -lemma (in PO) po_subset_po: - "S \ A ==> (| pset = S, order = induced S r |) \ PartialOrder" -apply (simp (no_asm) add: PartialOrder_def) -apply auto --- {* refl *} -apply (simp add: refl_on_def induced_def) -apply (blast intro: reflE) --- {* antisym *} -apply (simp add: antisym_def induced_def) -apply (blast intro: antisymE) --- {* trans *} -apply (simp add: trans_def induced_def) -apply (blast intro: transE) -done - -lemma (in PO) indE: "[| (x, y) \ induced S r; S \ A |] ==> (x, y) \ r" -by (simp add: add: induced_def) - -lemma (in PO) indI: "[| (x, y) \ r; x \ S; y \ S |] ==> (x, y) \ induced S r" -by (simp add: add: induced_def) - -lemma (in CL) CL_imp_ex_isLub: "S \ A ==> \L. isLub S cl L" -apply (insert cl_co) -apply (simp add: CompleteLattice_def A_def) -done - -declare (in CL) cl_co [simp] - -lemma isLub_lub: "(\L. isLub S cl L) = isLub S cl (lub S cl)" -by (simp add: lub_def least_def isLub_def some_eq_ex [symmetric]) - -lemma isGlb_glb: "(\G. isGlb S cl G) = isGlb S cl (glb S cl)" -by (simp add: glb_def greatest_def isGlb_def some_eq_ex [symmetric]) - -lemma isGlb_dual_isLub: "isGlb S cl = isLub S (dual cl)" -by (simp add: isLub_def isGlb_def dual_def converse_def) - -lemma isLub_dual_isGlb: "isLub S cl = isGlb S (dual cl)" -by (simp add: isLub_def isGlb_def dual_def converse_def) - -lemma (in PO) dualPO: "dual cl \ PartialOrder" -apply (insert cl_po) -apply (simp add: PartialOrder_def dual_def refl_on_converse - trans_converse antisym_converse) -done - -lemma Rdual: - "\S. (S \ A -->( \L. isLub S (| pset = A, order = r|) L)) - ==> \S. (S \ A --> (\G. isGlb S (| pset = A, order = r|) G))" -apply safe -apply (rule_tac x = "lub {y. y \ A & (\k \ S. (y, k) \ r)} - (|pset = A, order = r|) " in exI) -apply (drule_tac x = "{y. y \ A & (\k \ S. (y,k) \ r) }" in spec) -apply (drule mp, fast) -apply (simp add: isLub_lub isGlb_def) -apply (simp add: isLub_def, blast) -done - -lemma lub_dual_glb: "lub S cl = glb S (dual cl)" -by (simp add: lub_def glb_def least_def greatest_def dual_def converse_def) - -lemma glb_dual_lub: "glb S cl = lub S (dual cl)" -by (simp add: lub_def glb_def least_def greatest_def dual_def converse_def) - -lemma CL_subset_PO: "CompleteLattice \ PartialOrder" -by (simp add: PartialOrder_def CompleteLattice_def, fast) - -lemmas CL_imp_PO = CL_subset_PO [THEN subsetD] - -declare PO.PO_imp_refl_on [OF PO.intro [OF CL_imp_PO], simp] -declare PO.PO_imp_sym [OF PO.intro [OF CL_imp_PO], simp] -declare PO.PO_imp_trans [OF PO.intro [OF CL_imp_PO], simp] - -lemma (in CL) CO_refl_on: "refl_on A r" -by (rule PO_imp_refl_on) - -lemma (in CL) CO_antisym: "antisym r" -by (rule PO_imp_sym) - -lemma (in CL) CO_trans: "trans r" -by (rule PO_imp_trans) - -lemma CompleteLatticeI: - "[| po \ PartialOrder; (\S. S \ pset po --> (\L. isLub S po L)); - (\S. S \ pset po --> (\G. isGlb S po G))|] - ==> po \ CompleteLattice" -apply (unfold CompleteLattice_def, blast) -done - -lemma (in CL) CL_dualCL: "dual cl \ CompleteLattice" -apply (insert cl_co) -apply (simp add: CompleteLattice_def dual_def) -apply (fold dual_def) -apply (simp add: isLub_dual_isGlb [symmetric] isGlb_dual_isLub [symmetric] - dualPO) -done - -lemma (in PO) dualA_iff: "pset (dual cl) = pset cl" -by (simp add: dual_def) - -lemma (in PO) dualr_iff: "((x, y) \ (order(dual cl))) = ((y, x) \ order cl)" -by (simp add: dual_def) - -lemma (in PO) monotone_dual: - "monotone f (pset cl) (order cl) - ==> monotone f (pset (dual cl)) (order(dual cl))" -by (simp add: monotone_def dualA_iff dualr_iff) - -lemma (in PO) interval_dual: - "[| x \ A; y \ A|] ==> interval r x y = interval (order(dual cl)) y x" -apply (simp add: interval_def dualr_iff) -apply (fold r_def, fast) -done - -lemma (in PO) interval_not_empty: - "[| trans r; interval r a b \ {} |] ==> (a, b) \ r" -apply (simp add: interval_def) -apply (unfold trans_def, blast) -done - -lemma (in PO) interval_imp_mem: "x \ interval r a b ==> (a, x) \ r" -by (simp add: interval_def) - -lemma (in PO) left_in_interval: - "[| a \ A; b \ A; interval r a b \ {} |] ==> a \ interval r a b" -apply (simp (no_asm_simp) add: interval_def) -apply (simp add: PO_imp_trans interval_not_empty) -apply (simp add: reflE) -done - -lemma (in PO) right_in_interval: - "[| a \ A; b \ A; interval r a b \ {} |] ==> b \ interval r a b" -apply (simp (no_asm_simp) add: interval_def) -apply (simp add: PO_imp_trans interval_not_empty) -apply (simp add: reflE) -done - - -subsection {* sublattice *} - -lemma (in PO) sublattice_imp_CL: - "S <<= cl ==> (| pset = S, order = induced S r |) \ CompleteLattice" -by (simp add: sublattice_def CompleteLattice_def A_def r_def) - -lemma (in CL) sublatticeI: - "[| S \ A; (| pset = S, order = induced S r |) \ CompleteLattice |] - ==> S <<= cl" -by (simp add: sublattice_def A_def r_def) - - -subsection {* lub *} - -lemma (in CL) lub_unique: "[| S \ A; isLub S cl x; isLub S cl L|] ==> x = L" -apply (rule antisymE) -apply (auto simp add: isLub_def r_def) -done - -lemma (in CL) lub_upper: "[|S \ A; x \ S|] ==> (x, lub S cl) \ r" -apply (rule CL_imp_ex_isLub [THEN exE], assumption) -apply (unfold lub_def least_def) -apply (rule some_equality [THEN ssubst]) - apply (simp add: isLub_def) - apply (simp add: lub_unique A_def isLub_def) -apply (simp add: isLub_def r_def) -done - -lemma (in CL) lub_least: - "[| S \ A; L \ A; \x \ S. (x,L) \ r |] ==> (lub S cl, L) \ r" -apply (rule CL_imp_ex_isLub [THEN exE], assumption) -apply (unfold lub_def least_def) -apply (rule_tac s=x in some_equality [THEN ssubst]) - apply (simp add: isLub_def) - apply (simp add: lub_unique A_def isLub_def) -apply (simp add: isLub_def r_def A_def) -done - -lemma (in CL) lub_in_lattice: "S \ A ==> lub S cl \ A" -apply (rule CL_imp_ex_isLub [THEN exE], assumption) -apply (unfold lub_def least_def) -apply (subst some_equality) -apply (simp add: isLub_def) -prefer 2 apply (simp add: isLub_def A_def) -apply (simp add: lub_unique A_def isLub_def) -done - -lemma (in CL) lubI: - "[| S \ A; L \ A; \x \ S. (x,L) \ r; - \z \ A. (\y \ S. (y,z) \ r) --> (L,z) \ r |] ==> L = lub S cl" -apply (rule lub_unique, assumption) -apply (simp add: isLub_def A_def r_def) -apply (unfold isLub_def) -apply (rule conjI) -apply (fold A_def r_def) -apply (rule lub_in_lattice, assumption) -apply (simp add: lub_upper lub_least) -done - -lemma (in CL) lubIa: "[| S \ A; isLub S cl L |] ==> L = lub S cl" -by (simp add: lubI isLub_def A_def r_def) - -lemma (in CL) isLub_in_lattice: "isLub S cl L ==> L \ A" -by (simp add: isLub_def A_def) - -lemma (in CL) isLub_upper: "[|isLub S cl L; y \ S|] ==> (y, L) \ r" -by (simp add: isLub_def r_def) - -lemma (in CL) isLub_least: - "[| isLub S cl L; z \ A; \y \ S. (y, z) \ r|] ==> (L, z) \ r" -by (simp add: isLub_def A_def r_def) - -lemma (in CL) isLubI: - "[| L \ A; \y \ S. (y, L) \ r; - (\z \ A. (\y \ S. (y, z):r) --> (L, z) \ r)|] ==> isLub S cl L" -by (simp add: isLub_def A_def r_def) - - - -subsection {* glb *} - -lemma (in CL) glb_in_lattice: "S \ A ==> glb S cl \ A" -apply (subst glb_dual_lub) -apply (simp add: A_def) -apply (rule dualA_iff [THEN subst]) -apply (rule CL.lub_in_lattice) -apply (rule CL.intro) -apply (rule PO.intro) -apply (rule dualPO) -apply (rule CL_axioms.intro) -apply (rule CL_dualCL) -apply (simp add: dualA_iff) -done - -lemma (in CL) glb_lower: "[|S \ A; x \ S|] ==> (glb S cl, x) \ r" -apply (subst glb_dual_lub) -apply (simp add: r_def) -apply (rule dualr_iff [THEN subst]) -apply (rule CL.lub_upper) -apply (rule CL.intro) -apply (rule PO.intro) -apply (rule dualPO) -apply (rule CL_axioms.intro) -apply (rule CL_dualCL) -apply (simp add: dualA_iff A_def, assumption) -done - -text {* - Reduce the sublattice property by using substructural properties; - abandoned see @{text "Tarski_4.ML"}. -*} - -declare (in CLF) f_cl [simp] - -(*never proved, 2007-01-22: Tarski__CLF_unnamed_lemma - NOT PROVABLE because of the conjunction used in the definition: we don't - allow reasoning with rules like conjE, which is essential here.*) -declare [[ atp_problem_prefix = "Tarski__CLF_unnamed_lemma" ]] -lemma (in CLF) [simp]: - "f: pset cl -> pset cl & monotone f (pset cl) (order cl)" -apply (insert f_cl) -apply (unfold CLF_set_def) -apply (erule SigmaE2) -apply (erule CollectE) -apply assumption -done - -lemma (in CLF) f_in_funcset: "f \ A -> A" -by (simp add: A_def) - -lemma (in CLF) monotone_f: "monotone f A r" -by (simp add: A_def r_def) - -(*never proved, 2007-01-22*) -declare [[ atp_problem_prefix = "Tarski__CLF_CLF_dual" ]] -declare (in CLF) CLF_set_def [simp] CL_dualCL [simp] monotone_dual [simp] dualA_iff [simp] - -lemma (in CLF) CLF_dual: "(dual cl, f) \ CLF_set" -apply (simp del: dualA_iff) -apply (simp) -done - -declare (in CLF) CLF_set_def[simp del] CL_dualCL[simp del] monotone_dual[simp del] - dualA_iff[simp del] - - -subsection {* fixed points *} - -lemma fix_subset: "fix f A \ A" -by (simp add: fix_def, fast) - -lemma fix_imp_eq: "x \ fix f A ==> f x = x" -by (simp add: fix_def) - -lemma fixf_subset: - "[| A \ B; x \ fix (%y: A. f y) A |] ==> x \ fix f B" -by (simp add: fix_def, auto) - - -subsection {* lemmas for Tarski, lub *} - -(*never proved, 2007-01-22*) -declare [[ atp_problem_prefix = "Tarski__CLF_lubH_le_flubH" ]] - declare CL.lub_least[intro] CLF.f_in_funcset[intro] funcset_mem[intro] CL.lub_in_lattice[intro] PO.transE[intro] PO.monotoneE[intro] CLF.monotone_f[intro] CL.lub_upper[intro] -lemma (in CLF) lubH_le_flubH: - "H = {x. (x, f x) \ r & x \ A} ==> (lub H cl, f (lub H cl)) \ r" -apply (rule lub_least, fast) -apply (rule f_in_funcset [THEN funcset_mem]) -apply (rule lub_in_lattice, fast) --- {* @{text "\x:H. (x, f (lub H r)) \ r"} *} -apply (rule ballI) -(*never proved, 2007-01-22*) -using [[ atp_problem_prefix = "Tarski__CLF_lubH_le_flubH_simpler" ]] -apply (rule transE) --- {* instantiates @{text "(x, ?z) \ order cl to (x, f x)"}, *} --- {* because of the def of @{text H} *} -apply fast --- {* so it remains to show @{text "(f x, f (lub H cl)) \ r"} *} -apply (rule_tac f = "f" in monotoneE) -apply (rule monotone_f, fast) -apply (rule lub_in_lattice, fast) -apply (rule lub_upper, fast) -apply assumption -done - declare CL.lub_least[rule del] CLF.f_in_funcset[rule del] - funcset_mem[rule del] CL.lub_in_lattice[rule del] - PO.transE[rule del] PO.monotoneE[rule del] - CLF.monotone_f[rule del] CL.lub_upper[rule del] - -(*never proved, 2007-01-22*) -declare [[ atp_problem_prefix = "Tarski__CLF_flubH_le_lubH" ]] - declare CLF.f_in_funcset[intro] funcset_mem[intro] CL.lub_in_lattice[intro] - PO.monotoneE[intro] CLF.monotone_f[intro] CL.lub_upper[intro] - CLF.lubH_le_flubH[simp] -lemma (in CLF) flubH_le_lubH: - "[| H = {x. (x, f x) \ r & x \ A} |] ==> (f (lub H cl), lub H cl) \ r" -apply (rule lub_upper, fast) -apply (rule_tac t = "H" in ssubst, assumption) -apply (rule CollectI) -apply (rule conjI) -using [[ atp_problem_prefix = "Tarski__CLF_flubH_le_lubH_simpler" ]] -(*??no longer terminates, with combinators -apply (metis CO_refl_on lubH_le_flubH monotone_def monotone_f reflD1 reflD2) -*) -apply (metis CO_refl_on lubH_le_flubH monotoneE [OF monotone_f] refl_onD1 refl_onD2) -apply (metis CO_refl_on lubH_le_flubH refl_onD2) -done - declare CLF.f_in_funcset[rule del] funcset_mem[rule del] - CL.lub_in_lattice[rule del] PO.monotoneE[rule del] - CLF.monotone_f[rule del] CL.lub_upper[rule del] - CLF.lubH_le_flubH[simp del] - - -(*never proved, 2007-01-22*) -declare [[ atp_problem_prefix = "Tarski__CLF_lubH_is_fixp" ]] -(*Single-step version fails. The conjecture clauses refer to local abstraction -functions (Frees), which prevents expand_defs_tac from removing those -"definitions" at the end of the proof. *) -lemma (in CLF) lubH_is_fixp: - "H = {x. (x, f x) \ r & x \ A} ==> lub H cl \ fix f A" -apply (simp add: fix_def) -apply (rule conjI) -proof (neg_clausify) -assume 0: "H = -Collect - (COMBS (COMBB op \ (COMBC (COMBB op \ (COMBS Pair f)) r)) (COMBC op \ A))" -assume 1: "lub (Collect - (COMBS (COMBB op \ (COMBC (COMBB op \ (COMBS Pair f)) r)) - (COMBC op \ A))) - cl -\ A" -have 2: "lub H cl \ A" - by (metis 1 0) -have 3: "(lub H cl, f (lub H cl)) \ r" - by (metis lubH_le_flubH 0) -have 4: "(f (lub H cl), lub H cl) \ r" - by (metis flubH_le_lubH 0) -have 5: "lub H cl = f (lub H cl) \ (lub H cl, f (lub H cl)) \ r" - by (metis antisymE 4) -have 6: "lub H cl = f (lub H cl)" - by (metis 5 3) -have 7: "(lub H cl, lub H cl) \ r" - by (metis 6 4) -have 8: "\X1. lub H cl \ X1 \ \ refl_on X1 r" - by (metis 7 refl_onD2) -have 9: "\ refl_on A r" - by (metis 8 2) -show "False" - by (metis CO_refl_on 9); -next --{*apparently the way to insert a second structured proof*} - show "H = {x. (x, f x) \ r \ x \ A} \ - f (lub {x. (x, f x) \ r \ x \ A} cl) = lub {x. (x, f x) \ r \ x \ A} cl" - proof (neg_clausify) - assume 0: "H = - Collect - (COMBS (COMBB op \ (COMBC (COMBB op \ (COMBS Pair f)) r)) (COMBC op \ A))" - assume 1: "f (lub (Collect - (COMBS (COMBB op \ (COMBC (COMBB op \ (COMBS Pair f)) r)) - (COMBC op \ A))) - cl) \ - lub (Collect - (COMBS (COMBB op \ (COMBC (COMBB op \ (COMBS Pair f)) r)) - (COMBC op \ A))) - cl" - have 2: "f (lub H cl) \ - lub (Collect - (COMBS (COMBB op \ (COMBC (COMBB op \ (COMBS Pair f)) r)) - (COMBC op \ A))) - cl" - by (metis 1 0) - have 3: "f (lub H cl) \ lub H cl" - by (metis 2 0) - have 4: "(lub H cl, f (lub H cl)) \ r" - by (metis lubH_le_flubH 0) - have 5: "(f (lub H cl), lub H cl) \ r" - by (metis flubH_le_lubH 0) - have 6: "lub H cl = f (lub H cl) \ (lub H cl, f (lub H cl)) \ r" - by (metis antisymE 5) - have 7: "lub H cl = f (lub H cl)" - by (metis 6 4) - show "False" - by (metis 3 7) - qed -qed - -lemma (in CLF) (*lubH_is_fixp:*) - "H = {x. (x, f x) \ r & x \ A} ==> lub H cl \ fix f A" -apply (simp add: fix_def) -apply (rule conjI) -using [[ atp_problem_prefix = "Tarski__CLF_lubH_is_fixp_simpler" ]] -apply (metis CO_refl_on lubH_le_flubH refl_onD1) -apply (metis antisymE flubH_le_lubH lubH_le_flubH) -done - -lemma (in CLF) fix_in_H: - "[| H = {x. (x, f x) \ r & x \ A}; x \ P |] ==> x \ H" -by (simp add: P_def fix_imp_eq [of _ f A] reflE CO_refl_on - fix_subset [of f A, THEN subsetD]) - - -lemma (in CLF) fixf_le_lubH: - "H = {x. (x, f x) \ r & x \ A} ==> \x \ fix f A. (x, lub H cl) \ r" -apply (rule ballI) -apply (rule lub_upper, fast) -apply (rule fix_in_H) -apply (simp_all add: P_def) -done - -declare [[ atp_problem_prefix = "Tarski__CLF_lubH_least_fixf" ]] -lemma (in CLF) lubH_least_fixf: - "H = {x. (x, f x) \ r & x \ A} - ==> \L. (\y \ fix f A. (y,L) \ r) --> (lub H cl, L) \ r" -apply (metis P_def lubH_is_fixp) -done - -subsection {* Tarski fixpoint theorem 1, first part *} -declare [[ atp_problem_prefix = "Tarski__CLF_T_thm_1_lub" ]] - declare CL.lubI[intro] fix_subset[intro] CL.lub_in_lattice[intro] - CLF.fixf_le_lubH[simp] CLF.lubH_least_fixf[simp] -lemma (in CLF) T_thm_1_lub: "lub P cl = lub {x. (x, f x) \ r & x \ A} cl" -(*sledgehammer;*) -apply (rule sym) -apply (simp add: P_def) -apply (rule lubI) -using [[ atp_problem_prefix = "Tarski__CLF_T_thm_1_lub_simpler" ]] -apply (metis P_def fix_subset) -apply (metis Collect_conj_eq Collect_mem_eq Int_commute Int_lower1 lub_in_lattice vimage_def) -(*??no longer terminates, with combinators -apply (metis P_def fix_def fixf_le_lubH) -apply (metis P_def fix_def lubH_least_fixf) -*) -apply (simp add: fixf_le_lubH) -apply (simp add: lubH_least_fixf) -done - declare CL.lubI[rule del] fix_subset[rule del] CL.lub_in_lattice[rule del] - CLF.fixf_le_lubH[simp del] CLF.lubH_least_fixf[simp del] - - -(*never proved, 2007-01-22*) -declare [[ atp_problem_prefix = "Tarski__CLF_glbH_is_fixp" ]] - declare glb_dual_lub[simp] PO.dualA_iff[intro] CLF.lubH_is_fixp[intro] - PO.dualPO[intro] CL.CL_dualCL[intro] PO.dualr_iff[simp] -lemma (in CLF) glbH_is_fixp: "H = {x. (f x, x) \ r & x \ A} ==> glb H cl \ P" - -- {* Tarski for glb *} -(*sledgehammer;*) -apply (simp add: glb_dual_lub P_def A_def r_def) -apply (rule dualA_iff [THEN subst]) -apply (rule CLF.lubH_is_fixp) -apply (rule CLF.intro) -apply (rule CL.intro) -apply (rule PO.intro) -apply (rule dualPO) -apply (rule CL_axioms.intro) -apply (rule CL_dualCL) -apply (rule CLF_axioms.intro) -apply (rule CLF_dual) -apply (simp add: dualr_iff dualA_iff) -done - declare glb_dual_lub[simp del] PO.dualA_iff[rule del] CLF.lubH_is_fixp[rule del] - PO.dualPO[rule del] CL.CL_dualCL[rule del] PO.dualr_iff[simp del] - - -(*never proved, 2007-01-22*) -declare [[ atp_problem_prefix = "Tarski__T_thm_1_glb" ]] (*ALL THEOREMS*) -lemma (in CLF) T_thm_1_glb: "glb P cl = glb {x. (f x, x) \ r & x \ A} cl" -(*sledgehammer;*) -apply (simp add: glb_dual_lub P_def A_def r_def) -apply (rule dualA_iff [THEN subst]) -(*never proved, 2007-01-22*) -using [[ atp_problem_prefix = "Tarski__T_thm_1_glb_simpler" ]] (*ALL THEOREMS*) -(*sledgehammer;*) -apply (simp add: CLF.T_thm_1_lub [of _ f, OF CLF.intro, OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro, - OF dualPO CL_dualCL] dualPO CL_dualCL CLF_dual dualr_iff) -done - -subsection {* interval *} - - -declare [[ atp_problem_prefix = "Tarski__rel_imp_elem" ]] - declare (in CLF) CO_refl_on[simp] refl_on_def [simp] -lemma (in CLF) rel_imp_elem: "(x, y) \ r ==> x \ A" -by (metis CO_refl_on refl_onD1) - declare (in CLF) CO_refl_on[simp del] refl_on_def [simp del] - -declare [[ atp_problem_prefix = "Tarski__interval_subset" ]] - declare (in CLF) rel_imp_elem[intro] - declare interval_def [simp] -lemma (in CLF) interval_subset: "[| a \ A; b \ A |] ==> interval r a b \ A" -by (metis CO_refl_on interval_imp_mem refl_onD refl_onD2 rel_imp_elem subset_eq) - declare (in CLF) rel_imp_elem[rule del] - declare interval_def [simp del] - - -lemma (in CLF) intervalI: - "[| (a, x) \ r; (x, b) \ r |] ==> x \ interval r a b" -by (simp add: interval_def) - -lemma (in CLF) interval_lemma1: - "[| S \ interval r a b; x \ S |] ==> (a, x) \ r" -by (unfold interval_def, fast) - -lemma (in CLF) interval_lemma2: - "[| S \ interval r a b; x \ S |] ==> (x, b) \ r" -by (unfold interval_def, fast) - -lemma (in CLF) a_less_lub: - "[| S \ A; S \ {}; - \x \ S. (a,x) \ r; \y \ S. (y, L) \ r |] ==> (a,L) \ r" -by (blast intro: transE) - -lemma (in CLF) glb_less_b: - "[| S \ A; S \ {}; - \x \ S. (x,b) \ r; \y \ S. (G, y) \ r |] ==> (G,b) \ r" -by (blast intro: transE) - -lemma (in CLF) S_intv_cl: - "[| a \ A; b \ A; S \ interval r a b |]==> S \ A" -by (simp add: subset_trans [OF _ interval_subset]) - -declare [[ atp_problem_prefix = "Tarski__L_in_interval" ]] (*ALL THEOREMS*) -lemma (in CLF) L_in_interval: - "[| a \ A; b \ A; S \ interval r a b; - S \ {}; isLub S cl L; interval r a b \ {} |] ==> L \ interval r a b" -(*WON'T TERMINATE -apply (metis CO_trans intervalI interval_lemma1 interval_lemma2 isLub_least isLub_upper subset_empty subset_iff trans_def) -*) -apply (rule intervalI) -apply (rule a_less_lub) -prefer 2 apply assumption -apply (simp add: S_intv_cl) -apply (rule ballI) -apply (simp add: interval_lemma1) -apply (simp add: isLub_upper) --- {* @{text "(L, b) \ r"} *} -apply (simp add: isLub_least interval_lemma2) -done - -(*never proved, 2007-01-22*) -declare [[ atp_problem_prefix = "Tarski__G_in_interval" ]] (*ALL THEOREMS*) -lemma (in CLF) G_in_interval: - "[| a \ A; b \ A; interval r a b \ {}; S \ interval r a b; isGlb S cl G; - S \ {} |] ==> G \ interval r a b" -apply (simp add: interval_dual) -apply (simp add: CLF.L_in_interval [of _ f, OF CLF.intro, OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro] - dualA_iff A_def dualPO CL_dualCL CLF_dual isGlb_dual_isLub) -done - -declare [[ atp_problem_prefix = "Tarski__intervalPO" ]] (*ALL THEOREMS*) -lemma (in CLF) intervalPO: - "[| a \ A; b \ A; interval r a b \ {} |] - ==> (| pset = interval r a b, order = induced (interval r a b) r |) - \ PartialOrder" -proof (neg_clausify) -assume 0: "a \ A" -assume 1: "b \ A" -assume 2: "\pset = interval r a b, order = induced (interval r a b) r\ \ PartialOrder" -have 3: "\ interval r a b \ A" - by (metis 2 po_subset_po) -have 4: "b \ A \ a \ A" - by (metis 3 interval_subset) -have 5: "a \ A" - by (metis 4 1) -show "False" - by (metis 5 0) -qed - -lemma (in CLF) intv_CL_lub: - "[| a \ A; b \ A; interval r a b \ {} |] - ==> \S. S \ interval r a b --> - (\L. isLub S (| pset = interval r a b, - order = induced (interval r a b) r |) L)" -apply (intro strip) -apply (frule S_intv_cl [THEN CL_imp_ex_isLub]) -prefer 2 apply assumption -apply assumption -apply (erule exE) --- {* define the lub for the interval as *} -apply (rule_tac x = "if S = {} then a else L" in exI) -apply (simp (no_asm_simp) add: isLub_def split del: split_if) -apply (intro impI conjI) --- {* @{text "(if S = {} then a else L) \ interval r a b"} *} -apply (simp add: CL_imp_PO L_in_interval) -apply (simp add: left_in_interval) --- {* lub prop 1 *} -apply (case_tac "S = {}") --- {* @{text "S = {}, y \ S = False => everything"} *} -apply fast --- {* @{text "S \ {}"} *} -apply simp --- {* @{text "\y:S. (y, L) \ induced (interval r a b) r"} *} -apply (rule ballI) -apply (simp add: induced_def L_in_interval) -apply (rule conjI) -apply (rule subsetD) -apply (simp add: S_intv_cl, assumption) -apply (simp add: isLub_upper) --- {* @{text "\z:interval r a b. (\y:S. (y, z) \ induced (interval r a b) r \ (if S = {} then a else L, z) \ induced (interval r a b) r"} *} -apply (rule ballI) -apply (rule impI) -apply (case_tac "S = {}") --- {* @{text "S = {}"} *} -apply simp -apply (simp add: induced_def interval_def) -apply (rule conjI) -apply (rule reflE, assumption) -apply (rule interval_not_empty) -apply (rule CO_trans) -apply (simp add: interval_def) --- {* @{text "S \ {}"} *} -apply simp -apply (simp add: induced_def L_in_interval) -apply (rule isLub_least, assumption) -apply (rule subsetD) -prefer 2 apply assumption -apply (simp add: S_intv_cl, fast) -done - -lemmas (in CLF) intv_CL_glb = intv_CL_lub [THEN Rdual] - -(*never proved, 2007-01-22*) -declare [[ atp_problem_prefix = "Tarski__interval_is_sublattice" ]] (*ALL THEOREMS*) -lemma (in CLF) interval_is_sublattice: - "[| a \ A; b \ A; interval r a b \ {} |] - ==> interval r a b <<= cl" -(*sledgehammer *) -apply (rule sublatticeI) -apply (simp add: interval_subset) -(*never proved, 2007-01-22*) -using [[ atp_problem_prefix = "Tarski__interval_is_sublattice_simpler" ]] -(*sledgehammer *) -apply (rule CompleteLatticeI) -apply (simp add: intervalPO) - apply (simp add: intv_CL_lub) -apply (simp add: intv_CL_glb) -done - -lemmas (in CLF) interv_is_compl_latt = - interval_is_sublattice [THEN sublattice_imp_CL] - - -subsection {* Top and Bottom *} -lemma (in CLF) Top_dual_Bot: "Top cl = Bot (dual cl)" -by (simp add: Top_def Bot_def least_def greatest_def dualA_iff dualr_iff) - -lemma (in CLF) Bot_dual_Top: "Bot cl = Top (dual cl)" -by (simp add: Top_def Bot_def least_def greatest_def dualA_iff dualr_iff) - -declare [[ atp_problem_prefix = "Tarski__Bot_in_lattice" ]] (*ALL THEOREMS*) -lemma (in CLF) Bot_in_lattice: "Bot cl \ A" -(*sledgehammer; *) -apply (simp add: Bot_def least_def) -apply (rule_tac a="glb A cl" in someI2) -apply (simp_all add: glb_in_lattice glb_lower - r_def [symmetric] A_def [symmetric]) -done - -(*first proved 2007-01-25 after relaxing relevance*) -declare [[ atp_problem_prefix = "Tarski__Top_in_lattice" ]] (*ALL THEOREMS*) -lemma (in CLF) Top_in_lattice: "Top cl \ A" -(*sledgehammer;*) -apply (simp add: Top_dual_Bot A_def) -(*first proved 2007-01-25 after relaxing relevance*) -using [[ atp_problem_prefix = "Tarski__Top_in_lattice_simpler" ]] (*ALL THEOREMS*) -(*sledgehammer*) -apply (rule dualA_iff [THEN subst]) -apply (blast intro!: CLF.Bot_in_lattice [OF CLF.intro, OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro] dualPO CL_dualCL CLF_dual) -done - -lemma (in CLF) Top_prop: "x \ A ==> (x, Top cl) \ r" -apply (simp add: Top_def greatest_def) -apply (rule_tac a="lub A cl" in someI2) -apply (rule someI2) -apply (simp_all add: lub_in_lattice lub_upper - r_def [symmetric] A_def [symmetric]) -done - -(*never proved, 2007-01-22*) -declare [[ atp_problem_prefix = "Tarski__Bot_prop" ]] (*ALL THEOREMS*) -lemma (in CLF) Bot_prop: "x \ A ==> (Bot cl, x) \ r" -(*sledgehammer*) -apply (simp add: Bot_dual_Top r_def) -apply (rule dualr_iff [THEN subst]) -apply (simp add: CLF.Top_prop [of _ f, OF CLF.intro, OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro] - dualA_iff A_def dualPO CL_dualCL CLF_dual) -done - -declare [[ atp_problem_prefix = "Tarski__Bot_in_lattice" ]] (*ALL THEOREMS*) -lemma (in CLF) Top_intv_not_empty: "x \ A ==> interval r x (Top cl) \ {}" -apply (metis Top_in_lattice Top_prop empty_iff intervalI reflE) -done - -declare [[ atp_problem_prefix = "Tarski__Bot_intv_not_empty" ]] (*ALL THEOREMS*) -lemma (in CLF) Bot_intv_not_empty: "x \ A ==> interval r (Bot cl) x \ {}" -apply (metis Bot_prop ex_in_conv intervalI reflE rel_imp_elem) -done - - -subsection {* fixed points form a partial order *} - -lemma (in CLF) fixf_po: "(| pset = P, order = induced P r|) \ PartialOrder" -by (simp add: P_def fix_subset po_subset_po) - -(*first proved 2007-01-25 after relaxing relevance*) -declare [[ atp_problem_prefix = "Tarski__Y_subset_A" ]] - declare (in Tarski) P_def[simp] Y_ss [simp] - declare fix_subset [intro] subset_trans [intro] -lemma (in Tarski) Y_subset_A: "Y \ A" -(*sledgehammer*) -apply (rule subset_trans [OF _ fix_subset]) -apply (rule Y_ss [simplified P_def]) -done - declare (in Tarski) P_def[simp del] Y_ss [simp del] - declare fix_subset [rule del] subset_trans [rule del] - - -lemma (in Tarski) lubY_in_A: "lub Y cl \ A" - by (rule Y_subset_A [THEN lub_in_lattice]) - -(*never proved, 2007-01-22*) -declare [[ atp_problem_prefix = "Tarski__lubY_le_flubY" ]] (*ALL THEOREMS*) -lemma (in Tarski) lubY_le_flubY: "(lub Y cl, f (lub Y cl)) \ r" -(*sledgehammer*) -apply (rule lub_least) -apply (rule Y_subset_A) -apply (rule f_in_funcset [THEN funcset_mem]) -apply (rule lubY_in_A) --- {* @{text "Y \ P ==> f x = x"} *} -apply (rule ballI) -using [[ atp_problem_prefix = "Tarski__lubY_le_flubY_simpler" ]] (*ALL THEOREMS*) -(*sledgehammer *) -apply (rule_tac t = "x" in fix_imp_eq [THEN subst]) -apply (erule Y_ss [simplified P_def, THEN subsetD]) --- {* @{text "reduce (f x, f (lub Y cl)) \ r to (x, lub Y cl) \ r"} by monotonicity *} -using [[ atp_problem_prefix = "Tarski__lubY_le_flubY_simplest" ]] (*ALL THEOREMS*) -(*sledgehammer*) -apply (rule_tac f = "f" in monotoneE) -apply (rule monotone_f) -apply (simp add: Y_subset_A [THEN subsetD]) -apply (rule lubY_in_A) -apply (simp add: lub_upper Y_subset_A) -done - -(*first proved 2007-01-25 after relaxing relevance*) -declare [[ atp_problem_prefix = "Tarski__intY1_subset" ]] (*ALL THEOREMS*) -lemma (in Tarski) intY1_subset: "intY1 \ A" -(*sledgehammer*) -apply (unfold intY1_def) -apply (rule interval_subset) -apply (rule lubY_in_A) -apply (rule Top_in_lattice) -done - -lemmas (in Tarski) intY1_elem = intY1_subset [THEN subsetD] - -(*never proved, 2007-01-22*) -declare [[ atp_problem_prefix = "Tarski__intY1_f_closed" ]] (*ALL THEOREMS*) -lemma (in Tarski) intY1_f_closed: "x \ intY1 \ f x \ intY1" -(*sledgehammer*) -apply (simp add: intY1_def interval_def) -apply (rule conjI) -apply (rule transE) -apply (rule lubY_le_flubY) --- {* @{text "(f (lub Y cl), f x) \ r"} *} -using [[ atp_problem_prefix = "Tarski__intY1_f_closed_simpler" ]] (*ALL THEOREMS*) -(*sledgehammer [has been proved before now...]*) -apply (rule_tac f=f in monotoneE) -apply (rule monotone_f) -apply (rule lubY_in_A) -apply (simp add: intY1_def interval_def intY1_elem) -apply (simp add: intY1_def interval_def) --- {* @{text "(f x, Top cl) \ r"} *} -apply (rule Top_prop) -apply (rule f_in_funcset [THEN funcset_mem]) -apply (simp add: intY1_def interval_def intY1_elem) -done - -declare [[ atp_problem_prefix = "Tarski__intY1_func" ]] (*ALL THEOREMS*) -lemma (in Tarski) intY1_func: "(%x: intY1. f x) \ intY1 -> intY1" -apply (rule restrict_in_funcset) -apply (metis intY1_f_closed restrict_in_funcset) -done - -declare [[ atp_problem_prefix = "Tarski__intY1_mono" ]] (*ALL THEOREMS*) -lemma (in Tarski) intY1_mono: - "monotone (%x: intY1. f x) intY1 (induced intY1 r)" -(*sledgehammer *) -apply (auto simp add: monotone_def induced_def intY1_f_closed) -apply (blast intro: intY1_elem monotone_f [THEN monotoneE]) -done - -(*proof requires relaxing relevance: 2007-01-25*) -declare [[ atp_problem_prefix = "Tarski__intY1_is_cl" ]] (*ALL THEOREMS*) -lemma (in Tarski) intY1_is_cl: - "(| pset = intY1, order = induced intY1 r |) \ CompleteLattice" -(*sledgehammer*) -apply (unfold intY1_def) -apply (rule interv_is_compl_latt) -apply (rule lubY_in_A) -apply (rule Top_in_lattice) -apply (rule Top_intv_not_empty) -apply (rule lubY_in_A) -done - -(*never proved, 2007-01-22*) -declare [[ atp_problem_prefix = "Tarski__v_in_P" ]] (*ALL THEOREMS*) -lemma (in Tarski) v_in_P: "v \ P" -(*sledgehammer*) -apply (unfold P_def) -apply (rule_tac A = "intY1" in fixf_subset) -apply (rule intY1_subset) -apply (simp add: CLF.glbH_is_fixp [OF CLF.intro, OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro, OF _ intY1_is_cl, simplified] - v_def CL_imp_PO intY1_is_cl CLF_set_def intY1_func intY1_mono) -done - -declare [[ atp_problem_prefix = "Tarski__z_in_interval" ]] (*ALL THEOREMS*) -lemma (in Tarski) z_in_interval: - "[| z \ P; \y\Y. (y, z) \ induced P r |] ==> z \ intY1" -(*sledgehammer *) -apply (unfold intY1_def P_def) -apply (rule intervalI) -prefer 2 - apply (erule fix_subset [THEN subsetD, THEN Top_prop]) -apply (rule lub_least) -apply (rule Y_subset_A) -apply (fast elim!: fix_subset [THEN subsetD]) -apply (simp add: induced_def) -done - -declare [[ atp_problem_prefix = "Tarski__fz_in_int_rel" ]] (*ALL THEOREMS*) -lemma (in Tarski) f'z_in_int_rel: "[| z \ P; \y\Y. (y, z) \ induced P r |] - ==> ((%x: intY1. f x) z, z) \ induced intY1 r" -apply (metis P_def acc_def fix_imp_eq fix_subset indI reflE restrict_apply subset_eq z_in_interval) -done - -(*never proved, 2007-01-22*) -declare [[ atp_problem_prefix = "Tarski__tarski_full_lemma" ]] (*ALL THEOREMS*) -lemma (in Tarski) tarski_full_lemma: - "\L. isLub Y (| pset = P, order = induced P r |) L" -apply (rule_tac x = "v" in exI) -apply (simp add: isLub_def) --- {* @{text "v \ P"} *} -apply (simp add: v_in_P) -apply (rule conjI) -(*sledgehammer*) --- {* @{text v} is lub *} --- {* @{text "1. \y:Y. (y, v) \ induced P r"} *} -apply (rule ballI) -apply (simp add: induced_def subsetD v_in_P) -apply (rule conjI) -apply (erule Y_ss [THEN subsetD]) -apply (rule_tac b = "lub Y cl" in transE) -apply (rule lub_upper) -apply (rule Y_subset_A, assumption) -apply (rule_tac b = "Top cl" in interval_imp_mem) -apply (simp add: v_def) -apply (fold intY1_def) -apply (rule CL.glb_in_lattice [OF CL.intro, OF PO.intro CL_axioms.intro, OF _ intY1_is_cl, simplified]) - apply (simp add: CL_imp_PO intY1_is_cl, force) --- {* @{text v} is LEAST ub *} -apply clarify -apply (rule indI) - prefer 3 apply assumption - prefer 2 apply (simp add: v_in_P) -apply (unfold v_def) -(*never proved, 2007-01-22*) -using [[ atp_problem_prefix = "Tarski__tarski_full_lemma_simpler" ]] -(*sledgehammer*) -apply (rule indE) -apply (rule_tac [2] intY1_subset) -(*never proved, 2007-01-22*) -using [[ atp_problem_prefix = "Tarski__tarski_full_lemma_simplest" ]] -(*sledgehammer*) -apply (rule CL.glb_lower [OF CL.intro, OF PO.intro CL_axioms.intro, OF _ intY1_is_cl, simplified]) - apply (simp add: CL_imp_PO intY1_is_cl) - apply force -apply (simp add: induced_def intY1_f_closed z_in_interval) -apply (simp add: P_def fix_imp_eq [of _ f A] reflE - fix_subset [of f A, THEN subsetD]) -done - -lemma CompleteLatticeI_simp: - "[| (| pset = A, order = r |) \ PartialOrder; - \S. S \ A --> (\L. isLub S (| pset = A, order = r |) L) |] - ==> (| pset = A, order = r |) \ CompleteLattice" -by (simp add: CompleteLatticeI Rdual) - - -(*never proved, 2007-01-22*) -declare [[ atp_problem_prefix = "Tarski__Tarski_full" ]] - declare (in CLF) fixf_po[intro] P_def [simp] A_def [simp] r_def [simp] - Tarski.tarski_full_lemma [intro] cl_po [intro] cl_co [intro] - CompleteLatticeI_simp [intro] -theorem (in CLF) Tarski_full: - "(| pset = P, order = induced P r|) \ CompleteLattice" -(*sledgehammer*) -apply (rule CompleteLatticeI_simp) -apply (rule fixf_po, clarify) -(*never proved, 2007-01-22*) -using [[ atp_problem_prefix = "Tarski__Tarski_full_simpler" ]] -(*sledgehammer*) -apply (simp add: P_def A_def r_def) -apply (blast intro!: Tarski.tarski_full_lemma [OF Tarski.intro, OF CLF.intro Tarski_axioms.intro, - OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro] cl_po cl_co f_cl) -done - declare (in CLF) fixf_po[rule del] P_def [simp del] A_def [simp del] r_def [simp del] - Tarski.tarski_full_lemma [rule del] cl_po [rule del] cl_co [rule del] - CompleteLatticeI_simp [rule del] - - -end diff -r 8f35633c4922 -r 9cf389429f6d src/HOL/MetisExamples/TransClosure.thy --- a/src/HOL/MetisExamples/TransClosure.thy Tue Oct 20 19:37:09 2009 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,62 +0,0 @@ -(* Title: HOL/MetisTest/TransClosure.thy - ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - -Testing the metis method -*) - -theory TransClosure -imports Main -begin - -types addr = nat - -datatype val - = Unit -- "dummy result value of void expressions" - | Null -- "null reference" - | Bool bool -- "Boolean value" - | Intg int -- "integer value" - | Addr addr -- "addresses of objects in the heap" - -consts R::"(addr \ addr) set" - -consts f:: "addr \ val" - -declare [[ atp_problem_prefix = "TransClosure__test" ]] -lemma "\ f c = Intg x; \ y. f b = Intg y \ y \ x; (a,b) \ R\<^sup>*; (b,c) \ R\<^sup>* \ - \ \ c. (b,c) \ R \ (a,c) \ R\<^sup>*" -by (metis Transitive_Closure.rtrancl_into_rtrancl converse_rtranclE trancl_reflcl) - -lemma "\ f c = Intg x; \ y. f b = Intg y \ y \ x; (a,b) \ R\<^sup>*; (b,c) \ R\<^sup>* \ - \ \ c. (b,c) \ R \ (a,c) \ R\<^sup>*" -proof (neg_clausify) -assume 0: "f c = Intg x" -assume 1: "(a, b) \ R\<^sup>*" -assume 2: "(b, c) \ R\<^sup>*" -assume 3: "f b \ Intg x" -assume 4: "\A. (b, A) \ R \ (a, A) \ R\<^sup>*" -have 5: "b = c \ b \ Domain R" - by (metis Not_Domain_rtrancl 2) -have 6: "\X1. (a, X1) \ R\<^sup>* \ (b, X1) \ R" - by (metis Transitive_Closure.rtrancl_into_rtrancl 1) -have 7: "\X1. (b, X1) \ R" - by (metis 6 4) -have 8: "b \ Domain R" - by (metis 7 DomainE) -have 9: "c = b" - by (metis 5 8) -have 10: "f b = Intg x" - by (metis 0 9) -show "False" - by (metis 10 3) -qed - -declare [[ atp_problem_prefix = "TransClosure__test_simpler" ]] -lemma "\ f c = Intg x; \ y. f b = Intg y \ y \ x; (a,b) \ R\<^sup>*; (b,c) \ R\<^sup>* \ - \ \ c. (b,c) \ R \ (a,c) \ R\<^sup>*" -apply (erule_tac x="b" in converse_rtranclE) -apply (metis rel_pow_0_E rel_pow_0_I) -apply (metis DomainE Domain_iff Transitive_Closure.rtrancl_into_rtrancl) -done - -end diff -r 8f35633c4922 -r 9cf389429f6d src/HOL/MetisExamples/set.thy --- a/src/HOL/MetisExamples/set.thy Tue Oct 20 19:37:09 2009 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,283 +0,0 @@ -(* Title: HOL/MetisExamples/set.thy - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - -Testing the metis method -*) - -theory set imports Main - -begin - -lemma "EX x X. ALL y. EX z Z. (~P(y,y) | P(x,x) | ~S(z,x)) & - (S(x,y) | ~S(y,z) | Q(Z,Z)) & - (Q(X,y) | ~Q(y,Z) | S(X,X))" -by metis -(*??But metis can't prove the single-step version...*) - - - -lemma "P(n::nat) ==> ~P(0) ==> n ~= 0" -by metis - -declare [[sledgehammer_modulus = 1]] - - -(*multiple versions of this example*) -lemma (*equal_union: *) - "(X = Y \ Z) = - (Y \ X \ Z \ X \ (\V. Y \ V \ Z \ V \ X \ V))" -proof (neg_clausify) -fix x -assume 0: "Y \ X \ X = Y \ Z" -assume 1: "Z \ X \ X = Y \ Z" -assume 2: "(\ Y \ X \ \ Z \ X \ Y \ x) \ X \ Y \ Z" -assume 3: "(\ Y \ X \ \ Z \ X \ Z \ x) \ X \ Y \ Z" -assume 4: "(\ Y \ X \ \ Z \ X \ \ X \ x) \ X \ Y \ Z" -assume 5: "\V. ((\ Y \ V \ \ Z \ V) \ X \ V) \ X = Y \ Z" -have 6: "sup Y Z = X \ Y \ X" - by (metis 0) -have 7: "sup Y Z = X \ Z \ X" - by (metis 1) -have 8: "\X3. sup Y Z = X \ X \ X3 \ \ Y \ X3 \ \ Z \ X3" - by (metis 5) -have 9: "Y \ x \ sup Y Z \ X \ \ Y \ X \ \ Z \ X" - by (metis 2) -have 10: "Z \ x \ sup Y Z \ X \ \ Y \ X \ \ Z \ X" - by (metis 3) -have 11: "sup Y Z \ X \ \ X \ x \ \ Y \ X \ \ Z \ X" - by (metis 4) -have 12: "Z \ X" - by (metis Un_upper2 7) -have 13: "\X3. sup Y Z = X \ X \ sup X3 Z \ \ Y \ sup X3 Z" - by (metis 8 Un_upper2) -have 14: "Y \ X" - by (metis Un_upper1 6) -have 15: "Z \ x \ sup Y Z \ X \ \ Y \ X" - by (metis 10 12) -have 16: "Y \ x \ sup Y Z \ X \ \ Y \ X" - by (metis 9 12) -have 17: "sup Y Z \ X \ \ X \ x \ \ Y \ X" - by (metis 11 12) -have 18: "sup Y Z \ X \ \ X \ x" - by (metis 17 14) -have 19: "Z \ x \ sup Y Z \ X" - by (metis 15 14) -have 20: "Y \ x \ sup Y Z \ X" - by (metis 16 14) -have 21: "sup Y Z = X \ X \ sup Y Z" - by (metis 13 Un_upper1) -have 22: "sup Y Z = X \ \ sup Y Z \ X" - by (metis equalityI 21) -have 23: "sup Y Z = X \ \ Z \ X \ \ Y \ X" - by (metis 22 Un_least) -have 24: "sup Y Z = X \ \ Y \ X" - by (metis 23 12) -have 25: "sup Y Z = X" - by (metis 24 14) -have 26: "\X3. X \ X3 \ \ Z \ X3 \ \ Y \ X3" - by (metis Un_least 25) -have 27: "Y \ x" - by (metis 20 25) -have 28: "Z \ x" - by (metis 19 25) -have 29: "\ X \ x" - by (metis 18 25) -have 30: "X \ x \ \ Y \ x" - by (metis 26 28) -have 31: "X \ x" - by (metis 30 27) -show "False" - by (metis 31 29) -qed - -declare [[sledgehammer_modulus = 2]] - -lemma (*equal_union: *) - "(X = Y \ Z) = - (Y \ X \ Z \ X \ (\V. Y \ V \ Z \ V \ X \ V))" -proof (neg_clausify) -fix x -assume 0: "Y \ X \ X = Y \ Z" -assume 1: "Z \ X \ X = Y \ Z" -assume 2: "(\ Y \ X \ \ Z \ X \ Y \ x) \ X \ Y \ Z" -assume 3: "(\ Y \ X \ \ Z \ X \ Z \ x) \ X \ Y \ Z" -assume 4: "(\ Y \ X \ \ Z \ X \ \ X \ x) \ X \ Y \ Z" -assume 5: "\V. ((\ Y \ V \ \ Z \ V) \ X \ V) \ X = Y \ Z" -have 6: "sup Y Z = X \ Y \ X" - by (metis 0) -have 7: "Y \ x \ sup Y Z \ X \ \ Y \ X \ \ Z \ X" - by (metis 2) -have 8: "sup Y Z \ X \ \ X \ x \ \ Y \ X \ \ Z \ X" - by (metis 4) -have 9: "\X3. sup Y Z = X \ X \ sup X3 Z \ \ Y \ sup X3 Z" - by (metis 5 Un_upper2) -have 10: "Z \ x \ sup Y Z \ X \ \ Y \ X" - by (metis 3 Un_upper2) -have 11: "sup Y Z \ X \ \ X \ x \ \ Y \ X" - by (metis 8 Un_upper2) -have 12: "Z \ x \ sup Y Z \ X" - by (metis 10 Un_upper1) -have 13: "sup Y Z = X \ X \ sup Y Z" - by (metis 9 Un_upper1) -have 14: "sup Y Z = X \ \ Z \ X \ \ Y \ X" - by (metis equalityI 13 Un_least) -have 15: "sup Y Z = X" - by (metis 14 1 6) -have 16: "Y \ x" - by (metis 7 Un_upper2 Un_upper1 15) -have 17: "\ X \ x" - by (metis 11 Un_upper1 15) -have 18: "X \ x" - by (metis Un_least 15 12 15 16) -show "False" - by (metis 18 17) -qed - -declare [[sledgehammer_modulus = 3]] - -lemma (*equal_union: *) - "(X = Y \ Z) = - (Y \ X \ Z \ X \ (\V. Y \ V \ Z \ V \ X \ V))" -proof (neg_clausify) -fix x -assume 0: "Y \ X \ X = Y \ Z" -assume 1: "Z \ X \ X = Y \ Z" -assume 2: "(\ Y \ X \ \ Z \ X \ Y \ x) \ X \ Y \ Z" -assume 3: "(\ Y \ X \ \ Z \ X \ Z \ x) \ X \ Y \ Z" -assume 4: "(\ Y \ X \ \ Z \ X \ \ X \ x) \ X \ Y \ Z" -assume 5: "\V. ((\ Y \ V \ \ Z \ V) \ X \ V) \ X = Y \ Z" -have 6: "Z \ x \ sup Y Z \ X \ \ Y \ X \ \ Z \ X" - by (metis 3) -have 7: "\X3. sup Y Z = X \ X \ sup X3 Z \ \ Y \ sup X3 Z" - by (metis 5 Un_upper2) -have 8: "Y \ x \ sup Y Z \ X \ \ Y \ X" - by (metis 2 Un_upper2) -have 9: "Z \ x \ sup Y Z \ X" - by (metis 6 Un_upper2 Un_upper1) -have 10: "sup Y Z = X \ \ sup Y Z \ X" - by (metis equalityI 7 Un_upper1) -have 11: "sup Y Z = X" - by (metis 10 Un_least 1 0) -have 12: "Z \ x" - by (metis 9 11) -have 13: "X \ x" - by (metis Un_least 11 12 8 Un_upper1 11) -show "False" - by (metis 13 4 Un_upper2 Un_upper1 11) -qed - -(*Example included in TPHOLs paper*) - -declare [[sledgehammer_modulus = 4]] - -lemma (*equal_union: *) - "(X = Y \ Z) = - (Y \ X \ Z \ X \ (\V. Y \ V \ Z \ V \ X \ V))" -proof (neg_clausify) -fix x -assume 0: "Y \ X \ X = Y \ Z" -assume 1: "Z \ X \ X = Y \ Z" -assume 2: "(\ Y \ X \ \ Z \ X \ Y \ x) \ X \ Y \ Z" -assume 3: "(\ Y \ X \ \ Z \ X \ Z \ x) \ X \ Y \ Z" -assume 4: "(\ Y \ X \ \ Z \ X \ \ X \ x) \ X \ Y \ Z" -assume 5: "\V. ((\ Y \ V \ \ Z \ V) \ X \ V) \ X = Y \ Z" -have 6: "sup Y Z \ X \ \ X \ x \ \ Y \ X \ \ Z \ X" - by (metis 4) -have 7: "Z \ x \ sup Y Z \ X \ \ Y \ X" - by (metis 3 Un_upper2) -have 8: "Z \ x \ sup Y Z \ X" - by (metis 7 Un_upper1) -have 9: "sup Y Z = X \ \ Z \ X \ \ Y \ X" - by (metis equalityI 5 Un_upper2 Un_upper1 Un_least) -have 10: "Y \ x" - by (metis 2 Un_upper2 1 Un_upper1 0 9 Un_upper2 1 Un_upper1 0) -have 11: "X \ x" - by (metis Un_least 9 Un_upper2 1 Un_upper1 0 8 9 Un_upper2 1 Un_upper1 0 10) -show "False" - by (metis 11 6 Un_upper2 1 Un_upper1 0 9 Un_upper2 1 Un_upper1 0) -qed - -declare [[ atp_problem_prefix = "set__equal_union" ]] -lemma (*equal_union: *) - "(X = Y \ Z) = - (Y \ X \ Z \ X \ (\V. Y \ V \ Z \ V \ X \ V))" -(*One shot proof: hand-reduced. Metis can't do the full proof any more.*) -by (metis Un_least Un_upper1 Un_upper2 set_eq_subset) - - -declare [[ atp_problem_prefix = "set__equal_inter" ]] -lemma "(X = Y \ Z) = - (X \ Y \ X \ Z \ (\V. V \ Y \ V \ Z \ V \ X))" -by (metis Int_greatest Int_lower1 Int_lower2 set_eq_subset) - -declare [[ atp_problem_prefix = "set__fixedpoint" ]] -lemma fixedpoint: - "\!x. f (g x) = x \ \!y. g (f y) = y" -by metis - -lemma (*fixedpoint:*) - "\!x. f (g x) = x \ \!y. g (f y) = y" -proof (neg_clausify) -fix x xa -assume 0: "f (g x) = x" -assume 1: "\y. y = x \ f (g y) \ y" -assume 2: "\x. g (f (xa x)) = xa x \ g (f x) \ x" -assume 3: "\x. g (f x) \ x \ xa x \ x" -have 4: "\X1. g (f X1) \ X1 \ g x \ X1" - by (metis 3 1 2) -show "False" - by (metis 4 0) -qed - -declare [[ atp_problem_prefix = "set__singleton_example" ]] -lemma (*singleton_example_2:*) - "\x \ S. \S \ x \ \z. S \ {z}" -by (metis Set.subsetI Union_upper insertCI set_eq_subset) - --{*found by SPASS*} - -lemma (*singleton_example_2:*) - "\x \ S. \S \ x \ \z. S \ {z}" -by (metis Set.subsetI Union_upper insert_iff set_eq_subset) - -lemma singleton_example_2: - "\x \ S. \S \ x \ \z. S \ {z}" -proof (neg_clausify) -assume 0: "\x. \ S \ {x}" -assume 1: "\x. x \ S \ \S \ x" -have 2: "\X3. X3 = \S \ \ X3 \ \S \ X3 \ S" - by (metis set_eq_subset 1) -have 3: "\X3. S \ insert (\S) X3" - by (metis insert_iff Set.subsetI Union_upper 2 Set.subsetI) -show "False" - by (metis 3 0) -qed - - - -text {* - From W. W. Bledsoe and Guohui Feng, SET-VAR. JAR 11 (3), 1993, pages - 293-314. -*} - -declare [[ atp_problem_prefix = "set__Bledsoe_Fung" ]] -(*Notes: 1, the numbering doesn't completely agree with the paper. -2, we must rename set variables to avoid type clashes.*) -lemma "\B. (\x \ B. x \ (0::int))" - "D \ F \ \G. \A \ G. \B \ F. A \ B" - "P a \ \A. (\x \ A. P x) \ (\y. y \ A)" - "a < b \ b < (c::int) \ \B. a \ B \ b \ B \ c \ B" - "P (f b) \ \s A. (\x \ A. P x) \ f s \ A" - "P (f b) \ \s A. (\x \ A. P x) \ f s \ A" - "\A. a \ A" - "(\C. (0, 0) \ C \ (\x y. (x, y) \ C \ (Suc x, Suc y) \ C) \ (n, m) \ C) \ Q n \ Q m" -apply (metis atMost_iff) -apply (metis emptyE) -apply (metis insert_iff singletonE) -apply (metis insertCI singletonE zless_le) -apply (metis Collect_def Collect_mem_eq) -apply (metis Collect_def Collect_mem_eq) -apply (metis DiffE) -apply (metis pair_in_Id_conv) -done - -end diff -r 8f35633c4922 -r 9cf389429f6d src/HOL/Metis_Examples/Abstraction.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Metis_Examples/Abstraction.thy Tue Oct 20 19:52:04 2009 +0200 @@ -0,0 +1,297 @@ +(* Title: HOL/Metis_Examples/Abstraction.thy + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + +Testing the metis method. +*) + +theory Abstraction +imports Main 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" + +declare [[ atp_problem_prefix = "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) + + +declare [[ atp_problem_prefix = "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*} + +declare [[ atp_problem_prefix = "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) + +declare [[ atp_problem_prefix = "Abstraction__Sigma_Collect" ]] +lemma "(a,b) \ (SIGMA x: A. {y. x = f y}) ==> a \ A & a = f b" +(*???metis says this is satisfiable! +by (metis CollectD SigmaD1 SigmaD2) +*) +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 Un_empty_right 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 Un_empty_right 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 Un_empty_right 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 + + +declare [[ atp_problem_prefix = "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 + +declare [[ atp_problem_prefix = "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 + + +declare [[ atp_problem_prefix = "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 + + +declare [[ atp_problem_prefix = "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 + +declare [[ atp_problem_prefix = "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*} +*) + +declare [[ atp_problem_prefix = "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) +*) + +declare [[ atp_problem_prefix = "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 fast +(*??no longer terminates, with combinators +by (metis Collect_mem_eq SigmaD2 subsetD) +*) + +declare [[ atp_problem_prefix = "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) +*) + +declare [[ atp_problem_prefix = "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 + +declare [[ atp_problem_prefix = "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 + +declare [[ atp_problem_prefix = "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 + +declare [[ atp_problem_prefix = "Abstraction__image_evenA" ]] +lemma "(%x. Suc(f x)) ` {x. even x} <= A ==> (\x. even x --> Suc(f x) \ A)"; +(*sledgehammer*) +by auto + +declare [[ atp_problem_prefix = "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 + +declare [[ atp_problem_prefix = "Abstraction__image_curry" ]] +lemma "f \ (%u v. b \ u \ v) ` A ==> \u v. P (b \ u \ v) ==> P(f y)" +(*sledgehammer*) +by auto + +declare [[ atp_problem_prefix = "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 +using [[ atp_problem_prefix = "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*) + +declare [[ atp_problem_prefix = "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 + +declare [[ atp_problem_prefix = "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 diff -r 8f35633c4922 -r 9cf389429f6d src/HOL/Metis_Examples/BT.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Metis_Examples/BT.thy Tue Oct 20 19:52:04 2009 +0200 @@ -0,0 +1,243 @@ +(* Title: HOL/MetisTest/BT.thy + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + +Testing the metis method +*) + +header {* Binary trees *} + +theory BT +imports Main +begin + + +datatype 'a bt = + Lf + | Br 'a "'a bt" "'a bt" + +consts + n_nodes :: "'a bt => nat" + n_leaves :: "'a bt => nat" + depth :: "'a bt => nat" + reflect :: "'a bt => 'a bt" + bt_map :: "('a => 'b) => ('a bt => 'b bt)" + preorder :: "'a bt => 'a list" + inorder :: "'a bt => 'a list" + postorder :: "'a bt => 'a list" + appnd :: "'a bt => 'a bt => 'a bt" + +primrec + "n_nodes Lf = 0" + "n_nodes (Br a t1 t2) = Suc (n_nodes t1 + n_nodes t2)" + +primrec + "n_leaves Lf = Suc 0" + "n_leaves (Br a t1 t2) = n_leaves t1 + n_leaves t2" + +primrec + "depth Lf = 0" + "depth (Br a t1 t2) = Suc (max (depth t1) (depth t2))" + +primrec + "reflect Lf = Lf" + "reflect (Br a t1 t2) = Br a (reflect t2) (reflect t1)" + +primrec + "bt_map f Lf = Lf" + "bt_map f (Br a t1 t2) = Br (f a) (bt_map f t1) (bt_map f t2)" + +primrec + "preorder Lf = []" + "preorder (Br a t1 t2) = [a] @ (preorder t1) @ (preorder t2)" + +primrec + "inorder Lf = []" + "inorder (Br a t1 t2) = (inorder t1) @ [a] @ (inorder t2)" + +primrec + "postorder Lf = []" + "postorder (Br a t1 t2) = (postorder t1) @ (postorder t2) @ [a]" + +primrec + "appnd Lf t = t" + "appnd (Br a t1 t2) t = Br a (appnd t1 t) (appnd t2 t)" + + +text {* \medskip BT simplification *} + +declare [[ atp_problem_prefix = "BT__n_leaves_reflect" ]] +lemma n_leaves_reflect: "n_leaves (reflect t) = n_leaves t" + apply (induct t) + apply (metis add_right_cancel n_leaves.simps(1) reflect.simps(1)) + apply (metis add_commute n_leaves.simps(2) reflect.simps(2)) + done + +declare [[ atp_problem_prefix = "BT__n_nodes_reflect" ]] +lemma n_nodes_reflect: "n_nodes (reflect t) = n_nodes t" + apply (induct t) + apply (metis reflect.simps(1)) + apply (metis n_nodes.simps(2) nat_add_commute reflect.simps(2)) + done + +declare [[ atp_problem_prefix = "BT__depth_reflect" ]] +lemma depth_reflect: "depth (reflect t) = depth t" + apply (induct t) + apply (metis depth.simps(1) reflect.simps(1)) + apply (metis depth.simps(2) min_max.sup_commute reflect.simps(2)) + done + +text {* + The famous relationship between the numbers of leaves and nodes. +*} + +declare [[ atp_problem_prefix = "BT__n_leaves_nodes" ]] +lemma n_leaves_nodes: "n_leaves t = Suc (n_nodes t)" + apply (induct t) + apply (metis n_leaves.simps(1) n_nodes.simps(1)) + apply auto + done + +declare [[ atp_problem_prefix = "BT__reflect_reflect_ident" ]] +lemma reflect_reflect_ident: "reflect (reflect t) = t" + apply (induct t) + apply (metis add_right_cancel reflect.simps(1)); + apply (metis reflect.simps(2)) + done + +declare [[ atp_problem_prefix = "BT__bt_map_ident" ]] +lemma bt_map_ident: "bt_map (%x. x) = (%y. y)" +apply (rule ext) +apply (induct_tac y) + apply (metis bt_map.simps(1)) +txt{*BUG involving flex-flex pairs*} +(* apply (metis bt_map.simps(2)) *) +apply auto +done + + +declare [[ atp_problem_prefix = "BT__bt_map_appnd" ]] +lemma bt_map_appnd: "bt_map f (appnd t u) = appnd (bt_map f t) (bt_map f u)" +apply (induct t) + apply (metis appnd.simps(1) bt_map.simps(1)) + apply (metis appnd.simps(2) bt_map.simps(2)) (*slow!!*) +done + + +declare [[ atp_problem_prefix = "BT__bt_map_compose" ]] +lemma bt_map_compose: "bt_map (f o g) t = bt_map f (bt_map g t)" +apply (induct t) + apply (metis bt_map.simps(1)) +txt{*Metis runs forever*} +(* apply (metis bt_map.simps(2) o_apply)*) +apply auto +done + + +declare [[ atp_problem_prefix = "BT__bt_map_reflect" ]] +lemma bt_map_reflect: "bt_map f (reflect t) = reflect (bt_map f t)" + apply (induct t) + apply (metis add_right_cancel bt_map.simps(1) reflect.simps(1)) + apply (metis add_right_cancel bt_map.simps(2) reflect.simps(2)) + done + +declare [[ atp_problem_prefix = "BT__preorder_bt_map" ]] +lemma preorder_bt_map: "preorder (bt_map f t) = map f (preorder t)" + apply (induct t) + apply (metis bt_map.simps(1) map.simps(1) preorder.simps(1)) + apply simp + done + +declare [[ atp_problem_prefix = "BT__inorder_bt_map" ]] +lemma inorder_bt_map: "inorder (bt_map f t) = map f (inorder t)" + apply (induct t) + apply (metis bt_map.simps(1) inorder.simps(1) map.simps(1)) + apply simp + done + +declare [[ atp_problem_prefix = "BT__postorder_bt_map" ]] +lemma postorder_bt_map: "postorder (bt_map f t) = map f (postorder t)" + apply (induct t) + apply (metis bt_map.simps(1) map.simps(1) postorder.simps(1)) + apply simp + done + +declare [[ atp_problem_prefix = "BT__depth_bt_map" ]] +lemma depth_bt_map [simp]: "depth (bt_map f t) = depth t" + apply (induct t) + apply (metis bt_map.simps(1) depth.simps(1)) + apply simp + done + +declare [[ atp_problem_prefix = "BT__n_leaves_bt_map" ]] +lemma n_leaves_bt_map [simp]: "n_leaves (bt_map f t) = n_leaves t" + apply (induct t) + apply (metis One_nat_def Suc_eq_plus1 bt_map.simps(1) less_add_one less_antisym linorder_neq_iff n_leaves.simps(1)) + apply (metis bt_map.simps(2) n_leaves.simps(2)) + done + + +declare [[ atp_problem_prefix = "BT__preorder_reflect" ]] +lemma preorder_reflect: "preorder (reflect t) = rev (postorder t)" + apply (induct t) + apply (metis postorder.simps(1) preorder.simps(1) reflect.simps(1) rev_is_Nil_conv) + apply (metis append_Nil Cons_eq_append_conv postorder.simps(2) preorder.simps(2) reflect.simps(2) rev.simps(2) rev_append rev_rev_ident) + done + +declare [[ atp_problem_prefix = "BT__inorder_reflect" ]] +lemma inorder_reflect: "inorder (reflect t) = rev (inorder t)" + apply (induct t) + apply (metis inorder.simps(1) reflect.simps(1) rev.simps(1)) + apply simp + done + +declare [[ atp_problem_prefix = "BT__postorder_reflect" ]] +lemma postorder_reflect: "postorder (reflect t) = rev (preorder t)" + apply (induct t) + apply (metis postorder.simps(1) preorder.simps(1) reflect.simps(1) rev.simps(1)) + apply (metis Cons_eq_appendI postorder.simps(2) preorder.simps(2) reflect.simps(2) rev.simps(2) rev_append self_append_conv2) + done + +text {* + Analogues of the standard properties of the append function for lists. +*} + +declare [[ atp_problem_prefix = "BT__appnd_assoc" ]] +lemma appnd_assoc [simp]: + "appnd (appnd t1 t2) t3 = appnd t1 (appnd t2 t3)" + apply (induct t1) + apply (metis appnd.simps(1)) + apply (metis appnd.simps(2)) + done + +declare [[ atp_problem_prefix = "BT__appnd_Lf2" ]] +lemma appnd_Lf2 [simp]: "appnd t Lf = t" + apply (induct t) + apply (metis appnd.simps(1)) + apply (metis appnd.simps(2)) + done + +declare [[ atp_problem_prefix = "BT__depth_appnd" ]] + declare max_add_distrib_left [simp] +lemma depth_appnd [simp]: "depth (appnd t1 t2) = depth t1 + depth t2" + apply (induct t1) + apply (metis add_0 appnd.simps(1) depth.simps(1)) +apply (simp add: ); + done + +declare [[ atp_problem_prefix = "BT__n_leaves_appnd" ]] +lemma n_leaves_appnd [simp]: + "n_leaves (appnd t1 t2) = n_leaves t1 * n_leaves t2" + apply (induct t1) + apply (metis One_nat_def appnd.simps(1) less_irrefl less_linear n_leaves.simps(1) nat_mult_1) + apply (simp add: left_distrib) + done + +declare [[ atp_problem_prefix = "BT__bt_map_appnd" ]] +lemma (*bt_map_appnd:*) + "bt_map f (appnd t1 t2) = appnd (bt_map f t1) (bt_map f t2)" + apply (induct t1) + apply (metis appnd.simps(1) bt_map_appnd) + apply (metis bt_map_appnd) + done + +end diff -r 8f35633c4922 -r 9cf389429f6d src/HOL/Metis_Examples/BigO.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Metis_Examples/BigO.thy Tue Oct 20 19:52:04 2009 +0200 @@ -0,0 +1,1231 @@ +(* Title: HOL/Metis_Examples/BigO.thy + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + +Testing the metis method. +*) + +header {* Big O notation *} + +theory BigO +imports "~~/src/HOL/Decision_Procs/Dense_Linear_Order" Main SetsAndFunctions +begin + +subsection {* Definitions *} + +definition bigo :: "('a => 'b::ordered_idom) => ('a => 'b) set" ("(1O'(_'))") where + "O(f::('a => 'b)) == {h. EX c. ALL x. abs (h x) <= c * abs (f x)}" + +declare [[ atp_problem_prefix = "BigO__bigo_pos_const" ]] +lemma bigo_pos_const: "(EX (c::'a::ordered_idom). + ALL x. (abs (h x)) <= (c * (abs (f x)))) + = (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))" + apply auto + apply (case_tac "c = 0", simp) + apply (rule_tac x = "1" in exI, simp) + apply (rule_tac x = "abs c" in exI, auto) + apply (metis abs_ge_minus_self abs_ge_zero abs_minus_cancel abs_of_nonneg equation_minus_iff Orderings.xt1(6) abs_mult) + done + +(*** Now various verions with an increasing modulus ***) + +declare [[sledgehammer_modulus = 1]] + +lemma (*bigo_pos_const:*) "(EX (c::'a::ordered_idom). + ALL x. (abs (h x)) <= (c * (abs (f x)))) + = (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))" + apply auto + apply (case_tac "c = 0", simp) + apply (rule_tac x = "1" in exI, simp) + apply (rule_tac x = "abs c" in exI, auto) +proof (neg_clausify) +fix c x +have 0: "\(X1\'a\ordered_idom) X2\'a\ordered_idom. \X1 * X2\ = \X2 * X1\" + by (metis abs_mult mult_commute) +have 1: "\(X1\'a\ordered_idom) X2\'a\ordered_idom. + X1 \ (0\'a\ordered_idom) \ \X2\ * X1 = \X2 * X1\" + by (metis abs_mult_pos linorder_linear) +have 2: "\(X1\'a\ordered_idom) X2\'a\ordered_idom. + \ (0\'a\ordered_idom) < X1 * X2 \ + \ (0\'a\ordered_idom) \ X2 \ \ X1 \ (0\'a\ordered_idom)" + by (metis linorder_not_less mult_nonneg_nonpos2) +assume 3: "\x\'b\type. + \(h\'b\type \ 'a\ordered_idom) x\ + \ (c\'a\ordered_idom) * \(f\'b\type \ 'a\ordered_idom) x\" +assume 4: "\ \(h\'b\type \ 'a\ordered_idom) (x\'b\type)\ + \ \c\'a\ordered_idom\ * \(f\'b\type \ 'a\ordered_idom) x\" +have 5: "\ \(h\'b\type \ 'a\ordered_idom) (x\'b\type)\ + \ \(c\'a\ordered_idom) * (f\'b\type \ 'a\ordered_idom) x\" + by (metis 4 abs_mult) +have 6: "\(X1\'a\ordered_idom) X2\'a\ordered_idom. + \ X1 \ (0\'a\ordered_idom) \ X1 \ \X2\" + by (metis abs_ge_zero xt1(6)) +have 7: "\(X1\'a\ordered_idom) X2\'a\ordered_idom. + X1 \ \X2\ \ (0\'a\ordered_idom) < X1" + by (metis not_leE 6) +have 8: "(0\'a\ordered_idom) < \(h\'b\type \ 'a\ordered_idom) (x\'b\type)\" + by (metis 5 7) +have 9: "\X1\'a\ordered_idom. + \ \(h\'b\type \ 'a\ordered_idom) (x\'b\type)\ \ X1 \ + (0\'a\ordered_idom) < X1" + by (metis 8 order_less_le_trans) +have 10: "(0\'a\ordered_idom) +< (c\'a\ordered_idom) * \(f\'b\type \ 'a\ordered_idom) (x\'b\type)\" + by (metis 3 9) +have 11: "\ (c\'a\ordered_idom) \ (0\'a\ordered_idom)" + by (metis abs_ge_zero 2 10) +have 12: "\X1\'a\ordered_idom. (c\'a\ordered_idom) * \X1\ = \X1 * c\" + by (metis mult_commute 1 11) +have 13: "\X1\'b\type. + - (h\'b\type \ 'a\ordered_idom) X1 + \ (c\'a\ordered_idom) * \(f\'b\type \ 'a\ordered_idom) X1\" + by (metis 3 abs_le_D2) +have 14: "\X1\'b\type. + - (h\'b\type \ 'a\ordered_idom) X1 + \ \(c\'a\ordered_idom) * (f\'b\type \ 'a\ordered_idom) X1\" + by (metis 0 12 13) +have 15: "\(X1\'a\ordered_idom) X2\'a\ordered_idom. \X1 * \X2\\ = \X1 * X2\" + by (metis abs_mult abs_mult_pos abs_ge_zero) +have 16: "\(X1\'a\ordered_idom) X2\'a\ordered_idom. X1 \ \X2\ \ \ X1 \ X2" + by (metis xt1(6) abs_ge_self) +have 17: "\(X1\'a\ordered_idom) X2\'a\ordered_idom. \ \X1\ \ X2 \ X1 \ \X2\" + by (metis 16 abs_le_D1) +have 18: "\X1\'b\type. + (h\'b\type \ 'a\ordered_idom) X1 + \ \(c\'a\ordered_idom) * (f\'b\type \ 'a\ordered_idom) X1\" + by (metis 17 3 15) +show "False" + by (metis abs_le_iff 5 18 14) +qed + +declare [[sledgehammer_modulus = 2]] + +lemma (*bigo_pos_const:*) "(EX (c::'a::ordered_idom). + ALL x. (abs (h x)) <= (c * (abs (f x)))) + = (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))" + apply auto + apply (case_tac "c = 0", simp) + apply (rule_tac x = "1" in exI, simp) + apply (rule_tac x = "abs c" in exI, auto); +proof (neg_clausify) +fix c x +have 0: "\(X1\'a\ordered_idom) X2\'a\ordered_idom. \X1 * X2\ = \X2 * X1\" + by (metis abs_mult mult_commute) +assume 1: "\x\'b\type. + \(h\'b\type \ 'a\ordered_idom) x\ + \ (c\'a\ordered_idom) * \(f\'b\type \ 'a\ordered_idom) x\" +assume 2: "\ \(h\'b\type \ 'a\ordered_idom) (x\'b\type)\ + \ \c\'a\ordered_idom\ * \(f\'b\type \ 'a\ordered_idom) x\" +have 3: "\ \(h\'b\type \ 'a\ordered_idom) (x\'b\type)\ + \ \(c\'a\ordered_idom) * (f\'b\type \ 'a\ordered_idom) x\" + by (metis 2 abs_mult) +have 4: "\(X1\'a\ordered_idom) X2\'a\ordered_idom. + \ X1 \ (0\'a\ordered_idom) \ X1 \ \X2\" + by (metis abs_ge_zero xt1(6)) +have 5: "(0\'a\ordered_idom) < \(h\'b\type \ 'a\ordered_idom) (x\'b\type)\" + by (metis not_leE 4 3) +have 6: "(0\'a\ordered_idom) +< (c\'a\ordered_idom) * \(f\'b\type \ 'a\ordered_idom) (x\'b\type)\" + by (metis 1 order_less_le_trans 5) +have 7: "\X1\'a\ordered_idom. (c\'a\ordered_idom) * \X1\ = \X1 * c\" + by (metis abs_ge_zero linorder_not_less mult_nonneg_nonpos2 6 linorder_linear abs_mult_pos mult_commute) +have 8: "\X1\'b\type. + - (h\'b\type \ 'a\ordered_idom) X1 + \ \(c\'a\ordered_idom) * (f\'b\type \ 'a\ordered_idom) X1\" + by (metis 0 7 abs_le_D2 1) +have 9: "\(X1\'a\ordered_idom) X2\'a\ordered_idom. \ \X1\ \ X2 \ X1 \ \X2\" + by (metis abs_ge_self xt1(6) abs_le_D1) +show "False" + by (metis 8 abs_ge_zero abs_mult_pos abs_mult 1 9 3 abs_le_iff) +qed + +declare [[sledgehammer_modulus = 3]] + +lemma (*bigo_pos_const:*) "(EX (c::'a::ordered_idom). + ALL x. (abs (h x)) <= (c * (abs (f x)))) + = (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))" + apply auto + apply (case_tac "c = 0", simp) + apply (rule_tac x = "1" in exI, simp) + apply (rule_tac x = "abs c" in exI, auto); +proof (neg_clausify) +fix c x +assume 0: "\x\'b\type. + \(h\'b\type \ 'a\ordered_idom) x\ + \ (c\'a\ordered_idom) * \(f\'b\type \ 'a\ordered_idom) x\" +assume 1: "\ \(h\'b\type \ 'a\ordered_idom) (x\'b\type)\ + \ \c\'a\ordered_idom\ * \(f\'b\type \ 'a\ordered_idom) x\" +have 2: "\(X1\'a\ordered_idom) X2\'a\ordered_idom. + X1 \ \X2\ \ (0\'a\ordered_idom) < X1" + by (metis abs_ge_zero xt1(6) not_leE) +have 3: "\ (c\'a\ordered_idom) \ (0\'a\ordered_idom)" + by (metis abs_ge_zero mult_nonneg_nonpos2 linorder_not_less order_less_le_trans 1 abs_mult 2 0) +have 4: "\(X1\'a\ordered_idom) X2\'a\ordered_idom. \X1 * \X2\\ = \X1 * X2\" + by (metis abs_ge_zero abs_mult_pos abs_mult) +have 5: "\X1\'b\type. + (h\'b\type \ 'a\ordered_idom) X1 + \ \(c\'a\ordered_idom) * (f\'b\type \ 'a\ordered_idom) X1\" + by (metis 4 0 xt1(6) abs_ge_self abs_le_D1) +show "False" + by (metis abs_mult mult_commute 3 abs_mult_pos linorder_linear 0 abs_le_D2 5 1 abs_le_iff) +qed + + +declare [[sledgehammer_modulus = 1]] + +lemma (*bigo_pos_const:*) "(EX (c::'a::ordered_idom). + ALL x. (abs (h x)) <= (c * (abs (f x)))) + = (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))" + apply auto + apply (case_tac "c = 0", simp) + apply (rule_tac x = "1" in exI, simp) + apply (rule_tac x = "abs c" in exI, auto); +proof (neg_clausify) +fix c x (*sort/type constraint inserted by hand!*) +have 0: "\(X1\'a\ordered_idom) X2. \X1 * \X2\\ = \X1 * X2\" + by (metis abs_ge_zero abs_mult_pos abs_mult) +assume 1: "\A. \h A\ \ c * \f A\" +have 2: "\X1 X2. \ \X1\ \ X2 \ (0\'a) \ X2" + by (metis abs_ge_zero order_trans) +have 3: "\X1. (0\'a) \ c * \f X1\" + by (metis 1 2) +have 4: "\X1. c * \f X1\ = \c * f X1\" + by (metis 0 abs_of_nonneg 3) +have 5: "\X1. - h X1 \ c * \f X1\" + by (metis 1 abs_le_D2) +have 6: "\X1. - h X1 \ \c * f X1\" + by (metis 4 5) +have 7: "\X1. h X1 \ c * \f X1\" + by (metis 1 abs_le_D1) +have 8: "\X1. h X1 \ \c * f X1\" + by (metis 4 7) +assume 9: "\ \h x\ \ \c\ * \f x\" +have 10: "\ \h x\ \ \c * f x\" + by (metis abs_mult 9) +show "False" + by (metis 6 8 10 abs_leI) +qed + + +declare [[sledgehammer_sorts = true]] + +lemma bigo_alt_def: "O(f) = + {h. EX c. (0 < c & (ALL x. abs (h x) <= c * abs (f x)))}" +by (auto simp add: bigo_def bigo_pos_const) + +declare [[ atp_problem_prefix = "BigO__bigo_elt_subset" ]] +lemma bigo_elt_subset [intro]: "f : O(g) ==> O(f) <= O(g)" + apply (auto simp add: bigo_alt_def) + apply (rule_tac x = "ca * c" in exI) + apply (rule conjI) + apply (rule mult_pos_pos) + apply (assumption)+ +(*sledgehammer*); + apply (rule allI) + apply (drule_tac x = "xa" in spec)+ + apply (subgoal_tac "ca * abs(f xa) <= ca * (c * abs(g xa))"); + apply (erule order_trans) + apply (simp add: mult_ac) + apply (rule mult_left_mono, assumption) + apply (rule order_less_imp_le, assumption); +done + + +declare [[ atp_problem_prefix = "BigO__bigo_refl" ]] +lemma bigo_refl [intro]: "f : O(f)" + apply(auto simp add: bigo_def) +proof (neg_clausify) +fix x +assume 0: "\xa. \ \f (x xa)\ \ xa * \f (x xa)\" +have 1: "\X2. X2 \ (1\'b) * X2 \ \ (1\'b) \ (1\'b)" + by (metis mult_le_cancel_right1 order_eq_iff) +have 2: "\X2. X2 \ (1\'b) * X2" + by (metis order_eq_iff 1) +show "False" + by (metis 0 2) +qed + +declare [[ atp_problem_prefix = "BigO__bigo_zero" ]] +lemma bigo_zero: "0 : O(g)" + apply (auto simp add: bigo_def func_zero) +proof (neg_clausify) +fix x +assume 0: "\xa. \ (0\'b) \ xa * \g (x xa)\" +have 1: "\ (0\'b) \ (0\'b)" + by (metis 0 mult_eq_0_iff) +show "False" + by (metis 1 linorder_neq_iff linorder_antisym_conv1) +qed + +lemma bigo_zero2: "O(%x.0) = {%x.0}" + apply (auto simp add: bigo_def) + apply (rule ext) + apply auto +done + +lemma bigo_plus_self_subset [intro]: + "O(f) \ O(f) <= O(f)" + apply (auto simp add: bigo_alt_def set_plus_def) + apply (rule_tac x = "c + ca" in exI) + apply auto + apply (simp add: ring_distribs func_plus) + apply (blast intro:order_trans abs_triangle_ineq add_mono elim:) +done + +lemma bigo_plus_idemp [simp]: "O(f) \ O(f) = O(f)" + apply (rule equalityI) + apply (rule bigo_plus_self_subset) + apply (rule set_zero_plus2) + apply (rule bigo_zero) +done + +lemma bigo_plus_subset [intro]: "O(f + g) <= O(f) \ O(g)" + apply (rule subsetI) + apply (auto simp add: bigo_def bigo_pos_const func_plus set_plus_def) + apply (subst bigo_pos_const [symmetric])+ + apply (rule_tac x = + "%n. if abs (g n) <= (abs (f n)) then x n else 0" in exI) + apply (rule conjI) + apply (rule_tac x = "c + c" in exI) + apply (clarsimp) + apply (auto) + apply (subgoal_tac "c * abs (f xa + g xa) <= (c + c) * abs (f xa)") + apply (erule_tac x = xa in allE) + apply (erule order_trans) + apply (simp) + apply (subgoal_tac "c * abs (f xa + g xa) <= c * (abs (f xa) + abs (g xa))") + apply (erule order_trans) + apply (simp add: ring_distribs) + apply (rule mult_left_mono) + apply assumption + apply (simp add: order_less_le) + apply (rule mult_left_mono) + apply (simp add: abs_triangle_ineq) + apply (simp add: order_less_le) + apply (rule mult_nonneg_nonneg) + apply (rule add_nonneg_nonneg) + apply auto + apply (rule_tac x = "%n. if (abs (f n)) < abs (g n) then x n else 0" + in exI) + apply (rule conjI) + apply (rule_tac x = "c + c" in exI) + apply auto + apply (subgoal_tac "c * abs (f xa + g xa) <= (c + c) * abs (g xa)") + apply (erule_tac x = xa in allE) + apply (erule order_trans) + apply (simp) + apply (subgoal_tac "c * abs (f xa + g xa) <= c * (abs (f xa) + abs (g xa))") + apply (erule order_trans) + apply (simp add: ring_distribs) + apply (rule mult_left_mono) + apply (simp add: order_less_le) + apply (simp add: order_less_le) + apply (rule mult_left_mono) + apply (rule abs_triangle_ineq) + apply (simp add: order_less_le) +apply (metis abs_not_less_zero double_less_0_iff less_not_permute linorder_not_less mult_less_0_iff) + apply (rule ext) + apply (auto simp add: if_splits linorder_not_le) +done + +lemma bigo_plus_subset2 [intro]: "A <= O(f) ==> B <= O(f) ==> A \ B <= O(f)" + apply (subgoal_tac "A \ B <= O(f) \ O(f)") + apply (erule order_trans) + apply simp + apply (auto del: subsetI simp del: bigo_plus_idemp) +done + +declare [[ atp_problem_prefix = "BigO__bigo_plus_eq" ]] +lemma bigo_plus_eq: "ALL x. 0 <= f x ==> ALL x. 0 <= g x ==> + O(f + g) = O(f) \ O(g)" + apply (rule equalityI) + apply (rule bigo_plus_subset) + apply (simp add: bigo_alt_def set_plus_def func_plus) + apply clarify +(*sledgehammer*); + apply (rule_tac x = "max c ca" in exI) + apply (rule conjI) + apply (metis Orderings.less_max_iff_disj) + apply clarify + apply (drule_tac x = "xa" in spec)+ + apply (subgoal_tac "0 <= f xa + g xa") + apply (simp add: ring_distribs) + apply (subgoal_tac "abs(a xa + b xa) <= abs(a xa) + abs(b xa)") + apply (subgoal_tac "abs(a xa) + abs(b xa) <= + max c ca * f xa + max c ca * g xa") + apply (blast intro: order_trans) + defer 1 + apply (rule abs_triangle_ineq) + apply (metis add_nonneg_nonneg) + apply (rule add_mono) +using [[ atp_problem_prefix = "BigO__bigo_plus_eq_simpler" ]] +(*Found by SPASS; SLOW*) +apply (metis le_maxI2 linorder_linear linorder_not_le min_max.sup_absorb1 mult_le_cancel_right order_trans) +apply (metis le_maxI2 linorder_not_le mult_le_cancel_right order_trans) +done + +declare [[ atp_problem_prefix = "BigO__bigo_bounded_alt" ]] +lemma bigo_bounded_alt: "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> + f : O(g)" + apply (auto simp add: bigo_def) +(*Version 1: one-shot proof*) + apply (metis OrderedGroup.abs_le_D1 linorder_class.not_less order_less_le Orderings.xt1(12) Ring_and_Field.abs_mult) + done + +lemma (*bigo_bounded_alt:*) "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> + f : O(g)" + apply (auto simp add: bigo_def) +(*Version 2: single-step proof*) +proof (neg_clausify) +fix x +assume 0: "\x. f x \ c * g x" +assume 1: "\xa. \ f (x xa) \ xa * \g (x xa)\" +have 2: "\X3. c * g X3 = f X3 \ \ c * g X3 \ f X3" + by (metis 0 order_antisym_conv) +have 3: "\X3. \ f (x \X3\) \ \X3 * g (x \X3\)\" + by (metis 1 abs_mult) +have 4: "\X1 X3\'b\ordered_idom. X3 \ X1 \ X1 \ \X3\" + by (metis linorder_linear abs_le_D1) +have 5: "\X3::'b. \X3\ * \X3\ = X3 * X3" + by (metis abs_mult_self) +have 6: "\X3. \ X3 * X3 < (0\'b\ordered_idom)" + by (metis not_square_less_zero) +have 7: "\X1 X3::'b. \X1\ * \X3\ = \X3 * X1\" + by (metis abs_mult mult_commute) +have 8: "\X3::'b. X3 * X3 = \X3 * X3\" + by (metis abs_mult 5) +have 9: "\X3. X3 * g (x \X3\) \ f (x \X3\)" + by (metis 3 4) +have 10: "c * g (x \c\) = f (x \c\)" + by (metis 2 9) +have 11: "\X3::'b. \X3\ * \\X3\\ = \X3\ * \X3\" + by (metis abs_idempotent abs_mult 8) +have 12: "\X3::'b. \X3 * \X3\\ = \X3\ * \X3\" + by (metis mult_commute 7 11) +have 13: "\X3::'b. \X3 * \X3\\ = X3 * X3" + by (metis 8 7 12) +have 14: "\X3. X3 \ \X3\ \ X3 < (0\'b)" + by (metis abs_ge_self abs_le_D1 abs_if) +have 15: "\X3. X3 \ \X3\ \ \X3\ < (0\'b)" + by (metis abs_ge_self abs_le_D1 abs_if) +have 16: "\X3. X3 * X3 < (0\'b) \ X3 * \X3\ \ X3 * X3" + by (metis 15 13) +have 17: "\X3::'b. X3 * \X3\ \ X3 * X3" + by (metis 16 6) +have 18: "\X3. X3 \ \X3\ \ \ X3 < (0\'b)" + by (metis mult_le_cancel_left 17) +have 19: "\X3::'b. X3 \ \X3\" + by (metis 18 14) +have 20: "\ f (x \c\) \ \f (x \c\)\" + by (metis 3 10) +show "False" + by (metis 20 19) +qed + + +text{*So here is the easier (and more natural) problem using transitivity*} +declare [[ atp_problem_prefix = "BigO__bigo_bounded_alt_trans" ]] +lemma "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> f : O(g)" + apply (auto simp add: bigo_def) + (*Version 1: one-shot proof*) + apply (metis Orderings.leD Orderings.leI abs_ge_self abs_le_D1 abs_mult abs_of_nonneg order_le_less) + done + +text{*So here is the easier (and more natural) problem using transitivity*} +declare [[ atp_problem_prefix = "BigO__bigo_bounded_alt_trans" ]] +lemma "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> f : O(g)" + apply (auto simp add: bigo_def) +(*Version 2: single-step proof*) +proof (neg_clausify) +fix x +assume 0: "\A\'a\type. + (f\'a\type \ 'b\ordered_idom) A + \ (c\'b\ordered_idom) * (g\'a\type \ 'b\ordered_idom) A" +assume 1: "\A\'b\ordered_idom. + \ (f\'a\type \ 'b\ordered_idom) ((x\'b\ordered_idom \ 'a\type) A) + \ A * \(g\'a\type \ 'b\ordered_idom) (x A)\" +have 2: "\X2\'a\type. + \ (c\'b\ordered_idom) * (g\'a\type \ 'b\ordered_idom) X2 + < (f\'a\type \ 'b\ordered_idom) X2" + by (metis 0 linorder_not_le) +have 3: "\X2\'b\ordered_idom. + \ (f\'a\type \ 'b\ordered_idom) ((x\'b\ordered_idom \ 'a\type) \X2\) + \ \X2 * (g\'a\type \ 'b\ordered_idom) (x \X2\)\" + by (metis abs_mult 1) +have 4: "\X2\'b\ordered_idom. + \X2 * (g\'a\type \ 'b\ordered_idom) ((x\'b\ordered_idom \ 'a\type) \X2\)\ + < (f\'a\type \ 'b\ordered_idom) (x \X2\)" + by (metis 3 linorder_not_less) +have 5: "\X2\'b\ordered_idom. + X2 * (g\'a\type \ 'b\ordered_idom) ((x\'b\ordered_idom \ 'a\type) \X2\) + < (f\'a\type \ 'b\ordered_idom) (x \X2\)" + by (metis abs_less_iff 4) +show "False" + by (metis 2 5) +qed + + +lemma bigo_bounded: "ALL x. 0 <= f x ==> ALL x. f x <= g x ==> + f : O(g)" + apply (erule bigo_bounded_alt [of f 1 g]) + apply simp +done + +declare [[ atp_problem_prefix = "BigO__bigo_bounded2" ]] +lemma bigo_bounded2: "ALL x. lb x <= f x ==> ALL x. f x <= lb x + g x ==> + f : lb +o O(g)" + apply (rule set_minus_imp_plus) + apply (rule bigo_bounded) + apply (auto simp add: diff_minus fun_Compl_def func_plus) + prefer 2 + apply (drule_tac x = x in spec)+ + apply arith (*not clear that it's provable otherwise*) +proof (neg_clausify) +fix x +assume 0: "\y. lb y \ f y" +assume 1: "\ (0\'b) \ f x + - lb x" +have 2: "\X3. (0\'b) + X3 = X3" + by (metis diff_eq_eq right_minus_eq) +have 3: "\ (0\'b) \ f x - lb x" + by (metis 1 diff_minus) +have 4: "\ (0\'b) + lb x \ f x" + by (metis 3 le_diff_eq) +show "False" + by (metis 4 2 0) +qed + +declare [[ atp_problem_prefix = "BigO__bigo_abs" ]] +lemma bigo_abs: "(%x. abs(f x)) =o O(f)" + apply (unfold bigo_def) + apply auto +proof (neg_clausify) +fix x +assume 0: "\xa. \ \f (x xa)\ \ xa * \f (x xa)\" +have 1: "\X2. X2 \ (1\'b) * X2 \ \ (1\'b) \ (1\'b)" + by (metis mult_le_cancel_right1 order_eq_iff) +have 2: "\X2. X2 \ (1\'b) * X2" + by (metis order_eq_iff 1) +show "False" + by (metis 0 2) +qed + +declare [[ atp_problem_prefix = "BigO__bigo_abs2" ]] +lemma bigo_abs2: "f =o O(%x. abs(f x))" + apply (unfold bigo_def) + apply auto +proof (neg_clausify) +fix x +assume 0: "\xa. \ \f (x xa)\ \ xa * \f (x xa)\" +have 1: "\X2. X2 \ (1\'b) * X2 \ \ (1\'b) \ (1\'b)" + by (metis mult_le_cancel_right1 order_eq_iff) +have 2: "\X2. X2 \ (1\'b) * X2" + by (metis order_eq_iff 1) +show "False" + by (metis 0 2) +qed + +lemma bigo_abs3: "O(f) = O(%x. abs(f x))" + apply (rule equalityI) + apply (rule bigo_elt_subset) + apply (rule bigo_abs2) + apply (rule bigo_elt_subset) + apply (rule bigo_abs) +done + +lemma bigo_abs4: "f =o g +o O(h) ==> + (%x. abs (f x)) =o (%x. abs (g x)) +o O(h)" + apply (drule set_plus_imp_minus) + apply (rule set_minus_imp_plus) + apply (subst fun_diff_def) +proof - + assume a: "f - g : O(h)" + have "(%x. abs (f x) - abs (g x)) =o O(%x. abs(abs (f x) - abs (g x)))" + by (rule bigo_abs2) + also have "... <= O(%x. abs (f x - g x))" + apply (rule bigo_elt_subset) + apply (rule bigo_bounded) + apply force + apply (rule allI) + apply (rule abs_triangle_ineq3) + done + also have "... <= O(f - g)" + apply (rule bigo_elt_subset) + apply (subst fun_diff_def) + apply (rule bigo_abs) + done + also have "... <= O(h)" + using a by (rule bigo_elt_subset) + finally show "(%x. abs (f x) - abs (g x)) : O(h)". +qed + +lemma bigo_abs5: "f =o O(g) ==> (%x. abs(f x)) =o O(g)" +by (unfold bigo_def, auto) + +lemma bigo_elt_subset2 [intro]: "f : g +o O(h) ==> O(f) <= O(g) \ O(h)" +proof - + assume "f : g +o O(h)" + also have "... <= O(g) \ O(h)" + by (auto del: subsetI) + also have "... = O(%x. abs(g x)) \ O(%x. abs(h x))" + apply (subst bigo_abs3 [symmetric])+ + apply (rule refl) + done + also have "... = O((%x. abs(g x)) + (%x. abs(h x)))" + by (rule bigo_plus_eq [symmetric], auto) + finally have "f : ...". + then have "O(f) <= ..." + by (elim bigo_elt_subset) + also have "... = O(%x. abs(g x)) \ O(%x. abs(h x))" + by (rule bigo_plus_eq, auto) + finally show ?thesis + by (simp add: bigo_abs3 [symmetric]) +qed + +declare [[ atp_problem_prefix = "BigO__bigo_mult" ]] +lemma bigo_mult [intro]: "O(f)\O(g) <= O(f * g)" + apply (rule subsetI) + apply (subst bigo_def) + apply (auto simp del: abs_mult mult_ac + simp add: bigo_alt_def set_times_def func_times) +(*sledgehammer*); + apply (rule_tac x = "c * ca" in exI) + apply(rule allI) + apply(erule_tac x = x in allE)+ + apply(subgoal_tac "c * ca * abs(f x * g x) = + (c * abs(f x)) * (ca * abs(g x))") +using [[ atp_problem_prefix = "BigO__bigo_mult_simpler" ]] +prefer 2 +apply (metis mult_assoc mult_left_commute + OrderedGroup.abs_of_pos OrderedGroup.mult_left_commute + Ring_and_Field.abs_mult Ring_and_Field.mult_pos_pos) + apply (erule ssubst) + apply (subst abs_mult) +(*not qute BigO__bigo_mult_simpler_1 (a hard problem!) as abs_mult has + just been done*) +proof (neg_clausify) +fix a c b ca x +assume 0: "(0\'b\ordered_idom) < (c\'b\ordered_idom)" +assume 1: "\(a\'a \ 'b\ordered_idom) (x\'a)\ +\ (c\'b\ordered_idom) * \(f\'a \ 'b\ordered_idom) x\" +assume 2: "\(b\'a \ 'b\ordered_idom) (x\'a)\ +\ (ca\'b\ordered_idom) * \(g\'a \ 'b\ordered_idom) x\" +assume 3: "\ \(a\'a \ 'b\ordered_idom) (x\'a)\ * + \(b\'a \ 'b\ordered_idom) x\ + \ (c\'b\ordered_idom) * \(f\'a \ 'b\ordered_idom) x\ * + ((ca\'b\ordered_idom) * \(g\'a \ 'b\ordered_idom) x\)" +have 4: "\c\'b\ordered_idom\ = c" + by (metis OrderedGroup.abs_of_pos 0) +have 5: "\X1\'b\ordered_idom. (c\'b\ordered_idom) * \X1\ = \c * X1\" + by (metis Ring_and_Field.abs_mult 4) +have 6: "(0\'b\ordered_idom) = (1\'b\ordered_idom) \ +(0\'b\ordered_idom) < (1\'b\ordered_idom)" + by (metis OrderedGroup.abs_not_less_zero Ring_and_Field.abs_one Ring_and_Field.linorder_neqE_ordered_idom) +have 7: "(0\'b\ordered_idom) < (1\'b\ordered_idom)" + by (metis 6 Ring_and_Field.one_neq_zero) +have 8: "\1\'b\ordered_idom\ = (1\'b\ordered_idom)" + by (metis OrderedGroup.abs_of_pos 7) +have 9: "\X1\'b\ordered_idom. (0\'b\ordered_idom) \ (c\'b\ordered_idom) * \X1\" + by (metis OrderedGroup.abs_ge_zero 5) +have 10: "\X1\'b\ordered_idom. X1 * (1\'b\ordered_idom) = X1" + by (metis Ring_and_Field.mult_cancel_right2 mult_commute) +have 11: "\X1\'b\ordered_idom. \\X1\\ = \X1\ * \1\'b\ordered_idom\" + by (metis Ring_and_Field.abs_mult OrderedGroup.abs_idempotent 10) +have 12: "\X1\'b\ordered_idom. \\X1\\ = \X1\" + by (metis 11 8 10) +have 13: "\X1\'b\ordered_idom. (0\'b\ordered_idom) \ \X1\" + by (metis OrderedGroup.abs_ge_zero 12) +have 14: "\ (0\'b\ordered_idom) + \ (c\'b\ordered_idom) * \(f\'a \ 'b\ordered_idom) (x\'a)\ \ +\ (0\'b\ordered_idom) \ \(b\'a \ 'b\ordered_idom) x\ \ +\ \b x\ \ (ca\'b\ordered_idom) * \(g\'a \ 'b\ordered_idom) x\ \ +\ \(a\'a \ 'b\ordered_idom) x\ \ c * \f x\" + by (metis 3 Ring_and_Field.mult_mono) +have 15: "\ (0\'b\ordered_idom) \ \(b\'a \ 'b\ordered_idom) (x\'a)\ \ +\ \b x\ \ (ca\'b\ordered_idom) * \(g\'a \ 'b\ordered_idom) x\ \ +\ \(a\'a \ 'b\ordered_idom) x\ + \ (c\'b\ordered_idom) * \(f\'a \ 'b\ordered_idom) x\" + by (metis 14 9) +have 16: "\ \(b\'a \ 'b\ordered_idom) (x\'a)\ + \ (ca\'b\ordered_idom) * \(g\'a \ 'b\ordered_idom) x\ \ +\ \(a\'a \ 'b\ordered_idom) x\ + \ (c\'b\ordered_idom) * \(f\'a \ 'b\ordered_idom) x\" + by (metis 15 13) +have 17: "\ \(a\'a \ 'b\ordered_idom) (x\'a)\ + \ (c\'b\ordered_idom) * \(f\'a \ 'b\ordered_idom) x\" + by (metis 16 2) +show 18: "False" + by (metis 17 1) +qed + + +declare [[ atp_problem_prefix = "BigO__bigo_mult2" ]] +lemma bigo_mult2 [intro]: "f *o O(g) <= O(f * g)" + apply (auto simp add: bigo_def elt_set_times_def func_times abs_mult) +(*sledgehammer*); + apply (rule_tac x = c in exI) + apply clarify + apply (drule_tac x = x in spec) +using [[ atp_problem_prefix = "BigO__bigo_mult2_simpler" ]] +(*sledgehammer [no luck]*); + apply (subgoal_tac "abs(f x) * abs(b x) <= abs(f x) * (c * abs(g x))") + apply (simp add: mult_ac) + apply (rule mult_left_mono, assumption) + apply (rule abs_ge_zero) +done + +declare [[ atp_problem_prefix = "BigO__bigo_mult3" ]] +lemma bigo_mult3: "f : O(h) ==> g : O(j) ==> f * g : O(h * j)" +by (metis bigo_mult set_times_intro subset_iff) + +declare [[ atp_problem_prefix = "BigO__bigo_mult4" ]] +lemma bigo_mult4 [intro]:"f : k +o O(h) ==> g * f : (g * k) +o O(g * h)" +by (metis bigo_mult2 set_plus_mono_b set_times_intro2 set_times_plus_distrib) + + +lemma bigo_mult5: "ALL x. f x ~= 0 ==> + O(f * g) <= (f::'a => ('b::ordered_field)) *o O(g)" +proof - + assume "ALL x. f x ~= 0" + show "O(f * g) <= f *o O(g)" + proof + fix h + assume "h : O(f * g)" + then have "(%x. 1 / (f x)) * h : (%x. 1 / f x) *o O(f * g)" + by auto + also have "... <= O((%x. 1 / f x) * (f * g))" + by (rule bigo_mult2) + also have "(%x. 1 / f x) * (f * g) = g" + apply (simp add: func_times) + apply (rule ext) + apply (simp add: prems nonzero_divide_eq_eq mult_ac) + done + finally have "(%x. (1::'b) / f x) * h : O(g)". + then have "f * ((%x. (1::'b) / f x) * h) : f *o O(g)" + by auto + also have "f * ((%x. (1::'b) / f x) * h) = h" + apply (simp add: func_times) + apply (rule ext) + apply (simp add: prems nonzero_divide_eq_eq mult_ac) + done + finally show "h : f *o O(g)". + qed +qed + +declare [[ atp_problem_prefix = "BigO__bigo_mult6" ]] +lemma bigo_mult6: "ALL x. f x ~= 0 ==> + O(f * g) = (f::'a => ('b::ordered_field)) *o O(g)" +by (metis bigo_mult2 bigo_mult5 order_antisym) + +(*proof requires relaxing relevance: 2007-01-25*) +declare [[ atp_problem_prefix = "BigO__bigo_mult7" ]] + declare bigo_mult6 [simp] +lemma bigo_mult7: "ALL x. f x ~= 0 ==> + O(f * g) <= O(f::'a => ('b::ordered_field)) \ O(g)" +(*sledgehammer*) + apply (subst bigo_mult6) + apply assumption + apply (rule set_times_mono3) + apply (rule bigo_refl) +done + declare bigo_mult6 [simp del] + +declare [[ atp_problem_prefix = "BigO__bigo_mult8" ]] + declare bigo_mult7[intro!] +lemma bigo_mult8: "ALL x. f x ~= 0 ==> + O(f * g) = O(f::'a => ('b::ordered_field)) \ O(g)" +by (metis bigo_mult bigo_mult7 order_antisym_conv) + +lemma bigo_minus [intro]: "f : O(g) ==> - f : O(g)" + by (auto simp add: bigo_def fun_Compl_def) + +lemma bigo_minus2: "f : g +o O(h) ==> -f : -g +o O(h)" + apply (rule set_minus_imp_plus) + apply (drule set_plus_imp_minus) + apply (drule bigo_minus) + apply (simp add: diff_minus) +done + +lemma bigo_minus3: "O(-f) = O(f)" + by (auto simp add: bigo_def fun_Compl_def abs_minus_cancel) + +lemma bigo_plus_absorb_lemma1: "f : O(g) ==> f +o O(g) <= O(g)" +proof - + assume a: "f : O(g)" + show "f +o O(g) <= O(g)" + proof - + have "f : O(f)" by auto + then have "f +o O(g) <= O(f) \ O(g)" + by (auto del: subsetI) + also have "... <= O(g) \ O(g)" + proof - + from a have "O(f) <= O(g)" by (auto del: subsetI) + thus ?thesis by (auto del: subsetI) + qed + also have "... <= O(g)" by (simp add: bigo_plus_idemp) + finally show ?thesis . + qed +qed + +lemma bigo_plus_absorb_lemma2: "f : O(g) ==> O(g) <= f +o O(g)" +proof - + assume a: "f : O(g)" + show "O(g) <= f +o O(g)" + proof - + from a have "-f : O(g)" by auto + then have "-f +o O(g) <= O(g)" by (elim bigo_plus_absorb_lemma1) + then have "f +o (-f +o O(g)) <= f +o O(g)" by auto + also have "f +o (-f +o O(g)) = O(g)" + by (simp add: set_plus_rearranges) + finally show ?thesis . + qed +qed + +declare [[ atp_problem_prefix = "BigO__bigo_plus_absorb" ]] +lemma bigo_plus_absorb [simp]: "f : O(g) ==> f +o O(g) = O(g)" +by (metis bigo_plus_absorb_lemma1 bigo_plus_absorb_lemma2 order_eq_iff); + +lemma bigo_plus_absorb2 [intro]: "f : O(g) ==> A <= O(g) ==> f +o A <= O(g)" + apply (subgoal_tac "f +o A <= f +o O(g)") + apply force+ +done + +lemma bigo_add_commute_imp: "f : g +o O(h) ==> g : f +o O(h)" + apply (subst set_minus_plus [symmetric]) + apply (subgoal_tac "g - f = - (f - g)") + apply (erule ssubst) + apply (rule bigo_minus) + apply (subst set_minus_plus) + apply assumption + apply (simp add: diff_minus add_ac) +done + +lemma bigo_add_commute: "(f : g +o O(h)) = (g : f +o O(h))" + apply (rule iffI) + apply (erule bigo_add_commute_imp)+ +done + +lemma bigo_const1: "(%x. c) : O(%x. 1)" +by (auto simp add: bigo_def mult_ac) + +declare [[ atp_problem_prefix = "BigO__bigo_const2" ]] +lemma (*bigo_const2 [intro]:*) "O(%x. c) <= O(%x. 1)" +by (metis bigo_const1 bigo_elt_subset); + +lemma bigo_const2 [intro]: "O(%x. c::'b::ordered_idom) <= O(%x. 1)"; +(*??FAILS because the two occurrences of COMBK have different polymorphic types +proof (neg_clausify) +assume 0: "\ O(COMBK (c\'b\ordered_idom)) \ O(COMBK (1\'b\ordered_idom))" +have 1: "COMBK (c\'b\ordered_idom) \ O(COMBK (1\'b\ordered_idom))" +apply (rule notI) +apply (rule 0 [THEN notE]) +apply (rule bigo_elt_subset) +apply assumption; +sorry + by (metis 0 bigo_elt_subset) loops?? +show "False" + by (metis 1 bigo_const1) +qed +*) + apply (rule bigo_elt_subset) + apply (rule bigo_const1) +done + +declare [[ atp_problem_prefix = "BigO__bigo_const3" ]] +lemma bigo_const3: "(c::'a::ordered_field) ~= 0 ==> (%x. 1) : O(%x. c)" +apply (simp add: bigo_def) +proof (neg_clausify) +assume 0: "(c\'a\ordered_field) \ (0\'a\ordered_field)" +assume 1: "\A\'a\ordered_field. \ (1\'a\ordered_field) \ A * \c\'a\ordered_field\" +have 2: "(0\'a\ordered_field) = \c\'a\ordered_field\ \ +\ (1\'a\ordered_field) \ (1\'a\ordered_field)" + by (metis 1 field_inverse) +have 3: "\c\'a\ordered_field\ = (0\'a\ordered_field)" + by (metis linorder_neq_iff linorder_antisym_conv1 2) +have 4: "(0\'a\ordered_field) = (c\'a\ordered_field)" + by (metis 3 abs_eq_0) +show "False" + by (metis 0 4) +qed + +lemma bigo_const4: "(c::'a::ordered_field) ~= 0 ==> O(%x. 1) <= O(%x. c)" +by (rule bigo_elt_subset, rule bigo_const3, assumption) + +lemma bigo_const [simp]: "(c::'a::ordered_field) ~= 0 ==> + O(%x. c) = O(%x. 1)" +by (rule equalityI, rule bigo_const2, rule bigo_const4, assumption) + +declare [[ atp_problem_prefix = "BigO__bigo_const_mult1" ]] +lemma bigo_const_mult1: "(%x. c * f x) : O(f)" + apply (simp add: bigo_def abs_mult) +proof (neg_clausify) +fix x +assume 0: "\xa\'b\ordered_idom. + \ \c\'b\ordered_idom\ * + \(f\'a\type \ 'b\ordered_idom) ((x\'b\ordered_idom \ 'a\type) xa)\ + \ xa * \f (x xa)\" +show "False" + by (metis linorder_neq_iff linorder_antisym_conv1 0) +qed + +lemma bigo_const_mult2: "O(%x. c * f x) <= O(f)" +by (rule bigo_elt_subset, rule bigo_const_mult1) + +declare [[ atp_problem_prefix = "BigO__bigo_const_mult3" ]] +lemma bigo_const_mult3: "(c::'a::ordered_field) ~= 0 ==> f : O(%x. c * f x)" + apply (simp add: bigo_def) +(*sledgehammer [no luck]*); + apply (rule_tac x = "abs(inverse c)" in exI) + apply (simp only: abs_mult [symmetric] mult_assoc [symmetric]) +apply (subst left_inverse) +apply (auto ); +done + +lemma bigo_const_mult4: "(c::'a::ordered_field) ~= 0 ==> + O(f) <= O(%x. c * f x)" +by (rule bigo_elt_subset, rule bigo_const_mult3, assumption) + +lemma bigo_const_mult [simp]: "(c::'a::ordered_field) ~= 0 ==> + O(%x. c * f x) = O(f)" +by (rule equalityI, rule bigo_const_mult2, erule bigo_const_mult4) + +declare [[ atp_problem_prefix = "BigO__bigo_const_mult5" ]] +lemma bigo_const_mult5 [simp]: "(c::'a::ordered_field) ~= 0 ==> + (%x. c) *o O(f) = O(f)" + apply (auto del: subsetI) + apply (rule order_trans) + apply (rule bigo_mult2) + apply (simp add: func_times) + apply (auto intro!: subsetI simp add: bigo_def elt_set_times_def func_times) + apply (rule_tac x = "%y. inverse c * x y" in exI) + apply (rename_tac g d) + apply safe + apply (rule_tac [2] ext) + prefer 2 + apply simp + apply (simp add: mult_assoc [symmetric] abs_mult) + (*couldn't get this proof without the step above; SLOW*) + apply (metis mult_assoc abs_ge_zero mult_left_mono) +done + + +declare [[ atp_problem_prefix = "BigO__bigo_const_mult6" ]] +lemma bigo_const_mult6 [intro]: "(%x. c) *o O(f) <= O(f)" + apply (auto intro!: subsetI + simp add: bigo_def elt_set_times_def func_times + simp del: abs_mult mult_ac) +(*sledgehammer*); + apply (rule_tac x = "ca * (abs c)" in exI) + apply (rule allI) + apply (subgoal_tac "ca * abs(c) * abs(f x) = abs(c) * (ca * abs(f x))") + apply (erule ssubst) + apply (subst abs_mult) + apply (rule mult_left_mono) + apply (erule spec) + apply simp + apply(simp add: mult_ac) +done + +lemma bigo_const_mult7 [intro]: "f =o O(g) ==> (%x. c * f x) =o O(g)" +proof - + assume "f =o O(g)" + then have "(%x. c) * f =o (%x. c) *o O(g)" + by auto + also have "(%x. c) * f = (%x. c * f x)" + by (simp add: func_times) + also have "(%x. c) *o O(g) <= O(g)" + by (auto del: subsetI) + finally show ?thesis . +qed + +lemma bigo_compose1: "f =o O(g) ==> (%x. f(k x)) =o O(%x. g(k x))" +by (unfold bigo_def, auto) + +lemma bigo_compose2: "f =o g +o O(h) ==> (%x. f(k x)) =o (%x. g(k x)) +o + O(%x. h(k x))" + apply (simp only: set_minus_plus [symmetric] diff_minus fun_Compl_def + func_plus) + apply (erule bigo_compose1) +done + +subsection {* Setsum *} + +lemma bigo_setsum_main: "ALL x. ALL y : A x. 0 <= h x y ==> + EX c. ALL x. ALL y : A x. abs(f x y) <= c * (h x y) ==> + (%x. SUM y : A x. f x y) =o O(%x. SUM y : A x. h x y)" + apply (auto simp add: bigo_def) + apply (rule_tac x = "abs c" in exI) + apply (subst abs_of_nonneg) back back + apply (rule setsum_nonneg) + apply force + apply (subst setsum_right_distrib) + apply (rule allI) + apply (rule order_trans) + apply (rule setsum_abs) + apply (rule setsum_mono) +apply (blast intro: order_trans mult_right_mono abs_ge_self) +done + +declare [[ atp_problem_prefix = "BigO__bigo_setsum1" ]] +lemma bigo_setsum1: "ALL x y. 0 <= h x y ==> + EX c. ALL x y. abs(f x y) <= c * (h x y) ==> + (%x. SUM y : A x. f x y) =o O(%x. SUM y : A x. h x y)" + apply (rule bigo_setsum_main) +(*sledgehammer*); + apply force + apply clarsimp + apply (rule_tac x = c in exI) + apply force +done + +lemma bigo_setsum2: "ALL y. 0 <= h y ==> + EX c. ALL y. abs(f y) <= c * (h y) ==> + (%x. SUM y : A x. f y) =o O(%x. SUM y : A x. h y)" +by (rule bigo_setsum1, auto) + +declare [[ atp_problem_prefix = "BigO__bigo_setsum3" ]] +lemma bigo_setsum3: "f =o O(h) ==> + (%x. SUM y : A x. (l x y) * f(k x y)) =o + O(%x. SUM y : A x. abs(l x y * h(k x y)))" + apply (rule bigo_setsum1) + apply (rule allI)+ + apply (rule abs_ge_zero) + apply (unfold bigo_def) + apply (auto simp add: abs_mult); +(*sledgehammer*); + apply (rule_tac x = c in exI) + apply (rule allI)+ + apply (subst mult_left_commute) + apply (rule mult_left_mono) + apply (erule spec) + apply (rule abs_ge_zero) +done + +lemma bigo_setsum4: "f =o g +o O(h) ==> + (%x. SUM y : A x. l x y * f(k x y)) =o + (%x. SUM y : A x. l x y * g(k x y)) +o + O(%x. SUM y : A x. abs(l x y * h(k x y)))" + apply (rule set_minus_imp_plus) + apply (subst fun_diff_def) + apply (subst setsum_subtractf [symmetric]) + apply (subst right_diff_distrib [symmetric]) + apply (rule bigo_setsum3) + apply (subst fun_diff_def [symmetric]) + apply (erule set_plus_imp_minus) +done + +declare [[ atp_problem_prefix = "BigO__bigo_setsum5" ]] +lemma bigo_setsum5: "f =o O(h) ==> ALL x y. 0 <= l x y ==> + ALL x. 0 <= h x ==> + (%x. SUM y : A x. (l x y) * f(k x y)) =o + O(%x. SUM y : A x. (l x y) * h(k x y))" + apply (subgoal_tac "(%x. SUM y : A x. (l x y) * h(k x y)) = + (%x. SUM y : A x. abs((l x y) * h(k x y)))") + apply (erule ssubst) + apply (erule bigo_setsum3) + apply (rule ext) + apply (rule setsum_cong2) + apply (thin_tac "f \ O(h)") +apply (metis abs_of_nonneg zero_le_mult_iff) +done + +lemma bigo_setsum6: "f =o g +o O(h) ==> ALL x y. 0 <= l x y ==> + ALL x. 0 <= h x ==> + (%x. SUM y : A x. (l x y) * f(k x y)) =o + (%x. SUM y : A x. (l x y) * g(k x y)) +o + O(%x. SUM y : A x. (l x y) * h(k x y))" + apply (rule set_minus_imp_plus) + apply (subst fun_diff_def) + apply (subst setsum_subtractf [symmetric]) + apply (subst right_diff_distrib [symmetric]) + apply (rule bigo_setsum5) + apply (subst fun_diff_def [symmetric]) + apply (drule set_plus_imp_minus) + apply auto +done + +subsection {* Misc useful stuff *} + +lemma bigo_useful_intro: "A <= O(f) ==> B <= O(f) ==> + A \ B <= O(f)" + apply (subst bigo_plus_idemp [symmetric]) + apply (rule set_plus_mono2) + apply assumption+ +done + +lemma bigo_useful_add: "f =o O(h) ==> g =o O(h) ==> f + g =o O(h)" + apply (subst bigo_plus_idemp [symmetric]) + apply (rule set_plus_intro) + apply assumption+ +done + +lemma bigo_useful_const_mult: "(c::'a::ordered_field) ~= 0 ==> + (%x. c) * f =o O(h) ==> f =o O(h)" + apply (rule subsetD) + apply (subgoal_tac "(%x. 1 / c) *o O(h) <= O(h)") + apply assumption + apply (rule bigo_const_mult6) + apply (subgoal_tac "f = (%x. 1 / c) * ((%x. c) * f)") + apply (erule ssubst) + apply (erule set_times_intro2) + apply (simp add: func_times) +done + +declare [[ atp_problem_prefix = "BigO__bigo_fix" ]] +lemma bigo_fix: "(%x. f ((x::nat) + 1)) =o O(%x. h(x + 1)) ==> f 0 = 0 ==> + f =o O(h)" + apply (simp add: bigo_alt_def) +(*sledgehammer*); + apply clarify + apply (rule_tac x = c in exI) + apply safe + apply (case_tac "x = 0") +apply (metis OrderedGroup.abs_ge_zero OrderedGroup.abs_zero order_less_le Ring_and_Field.split_mult_pos_le) + apply (subgoal_tac "x = Suc (x - 1)") + apply metis + apply simp + done + + +lemma bigo_fix2: + "(%x. f ((x::nat) + 1)) =o (%x. g(x + 1)) +o O(%x. h(x + 1)) ==> + f 0 = g 0 ==> f =o g +o O(h)" + apply (rule set_minus_imp_plus) + apply (rule bigo_fix) + apply (subst fun_diff_def) + apply (subst fun_diff_def [symmetric]) + apply (rule set_plus_imp_minus) + apply simp + apply (simp add: fun_diff_def) +done + +subsection {* Less than or equal to *} + +constdefs + lesso :: "('a => 'b::ordered_idom) => ('a => 'b) => ('a => 'b)" + (infixl " ALL x. abs (g x) <= abs (f x) ==> + g =o O(h)" + apply (unfold bigo_def) + apply clarsimp +apply (blast intro: order_trans) +done + +lemma bigo_lesseq2: "f =o O(h) ==> ALL x. abs (g x) <= f x ==> + g =o O(h)" + apply (erule bigo_lesseq1) +apply (blast intro: abs_ge_self order_trans) +done + +lemma bigo_lesseq3: "f =o O(h) ==> ALL x. 0 <= g x ==> ALL x. g x <= f x ==> + g =o O(h)" + apply (erule bigo_lesseq2) + apply (rule allI) + apply (subst abs_of_nonneg) + apply (erule spec)+ +done + +lemma bigo_lesseq4: "f =o O(h) ==> + ALL x. 0 <= g x ==> ALL x. g x <= abs (f x) ==> + g =o O(h)" + apply (erule bigo_lesseq1) + apply (rule allI) + apply (subst abs_of_nonneg) + apply (erule spec)+ +done + +declare [[ atp_problem_prefix = "BigO__bigo_lesso1" ]] +lemma bigo_lesso1: "ALL x. f x <= g x ==> f + ALL x. 0 <= k x ==> ALL x. k x <= f x ==> + k A. k A \ f A" +have 1: "\(X1\'b\ordered_idom) X2. \ max X1 X2 < X1" + by (metis linorder_not_less le_maxI1) (*sort inserted by hand*) +assume 2: "(0\'b) \ k x - g x" +have 3: "\ k x - g x < (0\'b)" + by (metis 2 linorder_not_less) +have 4: "\X1 X2. min X1 (k X2) \ f X2" + by (metis min_max.inf_le2 min_max.le_inf_iff min_max.le_iff_inf 0) +have 5: "\g x - f x\ = f x - g x" + by (metis abs_minus_commute combine_common_factor mult_zero_right minus_add_cancel minus_zero abs_if diff_less_eq min_max.inf_commute 4 linorder_not_le min_max.le_iff_inf 3 diff_less_0_iff_less linorder_not_less) +have 6: "max (0\'b) (k x - g x) = k x - g x" + by (metis min_max.le_iff_sup 2) +assume 7: "\ max (k x - g x) (0\'b) \ \f x - g x\" +have 8: "\ k x - g x \ f x - g x" + by (metis 5 abs_minus_commute 7 min_max.sup_commute 6) +show "False" + by (metis min_max.sup_commute min_max.inf_commute min_max.sup_inf_absorb min_max.le_iff_inf 0 max_diff_distrib_left 1 linorder_not_le 8) +qed + +declare [[ atp_problem_prefix = "BigO__bigo_lesso3" ]] +lemma bigo_lesso3: "f =o g +o O(h) ==> + ALL x. 0 <= k x ==> ALL x. g x <= k x ==> + f 'b::ordered_field) ==> + g =o h +o O(k) ==> f + EX C. ALL x. f x <= g x + C * abs(h x)" + apply (simp only: lesso_def bigo_alt_def) + apply clarsimp + apply (metis abs_if abs_mult add_commute diff_le_eq less_not_permute) +done + +end diff -r 8f35633c4922 -r 9cf389429f6d src/HOL/Metis_Examples/Message.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Metis_Examples/Message.thy Tue Oct 20 19:52:04 2009 +0200 @@ -0,0 +1,812 @@ +(* Title: HOL/MetisTest/Message.thy + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + +Testing the metis method. +*) + +theory Message imports Main begin + +(*Needed occasionally with spy_analz_tac, e.g. in analz_insert_Key_newK*) +lemma strange_Un_eq [simp]: "A \ (B \ A) = B \ A" +by blast + +types + key = nat + +consts + all_symmetric :: bool --{*true if all keys are symmetric*} + invKey :: "key=>key" --{*inverse of a symmetric key*} + +specification (invKey) + invKey [simp]: "invKey (invKey K) = K" + invKey_symmetric: "all_symmetric --> invKey = id" + by (rule exI [of _ id], auto) + + +text{*The inverse of a symmetric key is itself; that of a public key + is the private key and vice versa*} + +constdefs + symKeys :: "key set" + "symKeys == {K. invKey K = K}" + +datatype --{*We allow any number of friendly agents*} + agent = Server | Friend nat | Spy + +datatype + msg = Agent agent --{*Agent names*} + | Number nat --{*Ordinary integers, timestamps, ...*} + | Nonce nat --{*Unguessable nonces*} + | Key key --{*Crypto keys*} + | Hash msg --{*Hashing*} + | MPair msg msg --{*Compound messages*} + | Crypt key msg --{*Encryption, public- or shared-key*} + + +text{*Concrete syntax: messages appear as {|A,B,NA|}, etc...*} +syntax + "@MTuple" :: "['a, args] => 'a * 'b" ("(2{|_,/ _|})") + +syntax (xsymbols) + "@MTuple" :: "['a, args] => 'a * 'b" ("(2\_,/ _\)") + +translations + "{|x, y, z|}" == "{|x, {|y, z|}|}" + "{|x, y|}" == "MPair x y" + + +constdefs + HPair :: "[msg,msg] => msg" ("(4Hash[_] /_)" [0, 1000]) + --{*Message Y paired with a MAC computed with the help of X*} + "Hash[X] Y == {| Hash{|X,Y|}, Y|}" + + keysFor :: "msg set => key set" + --{*Keys useful to decrypt elements of a message set*} + "keysFor H == invKey ` {K. \X. Crypt K X \ H}" + + +subsubsection{*Inductive Definition of All Parts" of a Message*} + +inductive_set + parts :: "msg set => msg set" + for H :: "msg set" + where + Inj [intro]: "X \ H ==> X \ parts H" + | Fst: "{|X,Y|} \ parts H ==> X \ parts H" + | Snd: "{|X,Y|} \ parts H ==> Y \ parts H" + | Body: "Crypt K X \ parts H ==> X \ parts H" + + +declare [[ atp_problem_prefix = "Message__parts_mono" ]] +lemma parts_mono: "G \ H ==> parts(G) \ parts(H)" +apply auto +apply (erule parts.induct) +apply (metis Inj set_mp) +apply (metis Fst) +apply (metis Snd) +apply (metis Body) +done + + +text{*Equations hold because constructors are injective.*} +lemma Friend_image_eq [simp]: "(Friend x \ Friend`A) = (x:A)" +by auto + +lemma Key_image_eq [simp]: "(Key x \ Key`A) = (x\A)" +by auto + +lemma Nonce_Key_image_eq [simp]: "(Nonce x \ Key`A)" +by auto + + +subsubsection{*Inverse of keys *} + +declare [[ atp_problem_prefix = "Message__invKey_eq" ]] +lemma invKey_eq [simp]: "(invKey K = invKey K') = (K=K')" +by (metis invKey) + + +subsection{*keysFor operator*} + +lemma keysFor_empty [simp]: "keysFor {} = {}" +by (unfold keysFor_def, blast) + +lemma keysFor_Un [simp]: "keysFor (H \ H') = keysFor H \ keysFor H'" +by (unfold keysFor_def, blast) + +lemma keysFor_UN [simp]: "keysFor (\i\A. H i) = (\i\A. keysFor (H i))" +by (unfold keysFor_def, blast) + +text{*Monotonicity*} +lemma keysFor_mono: "G \ H ==> keysFor(G) \ keysFor(H)" +by (unfold keysFor_def, blast) + +lemma keysFor_insert_Agent [simp]: "keysFor (insert (Agent A) H) = keysFor H" +by (unfold keysFor_def, auto) + +lemma keysFor_insert_Nonce [simp]: "keysFor (insert (Nonce N) H) = keysFor H" +by (unfold keysFor_def, auto) + +lemma keysFor_insert_Number [simp]: "keysFor (insert (Number N) H) = keysFor H" +by (unfold keysFor_def, auto) + +lemma keysFor_insert_Key [simp]: "keysFor (insert (Key K) H) = keysFor H" +by (unfold keysFor_def, auto) + +lemma keysFor_insert_Hash [simp]: "keysFor (insert (Hash X) H) = keysFor H" +by (unfold keysFor_def, auto) + +lemma keysFor_insert_MPair [simp]: "keysFor (insert {|X,Y|} H) = keysFor H" +by (unfold keysFor_def, auto) + +lemma keysFor_insert_Crypt [simp]: + "keysFor (insert (Crypt K X) H) = insert (invKey K) (keysFor H)" +by (unfold keysFor_def, auto) + +lemma keysFor_image_Key [simp]: "keysFor (Key`E) = {}" +by (unfold keysFor_def, auto) + +lemma Crypt_imp_invKey_keysFor: "Crypt K X \ H ==> invKey K \ keysFor H" +by (unfold keysFor_def, blast) + + +subsection{*Inductive relation "parts"*} + +lemma MPair_parts: + "[| {|X,Y|} \ parts H; + [| X \ parts H; Y \ parts H |] ==> P |] ==> P" +by (blast dest: parts.Fst parts.Snd) + + declare MPair_parts [elim!] parts.Body [dest!] +text{*NB These two rules are UNSAFE in the formal sense, as they discard the + compound message. They work well on THIS FILE. + @{text MPair_parts} is left as SAFE because it speeds up proofs. + The Crypt rule is normally kept UNSAFE to avoid breaking up certificates.*} + +lemma parts_increasing: "H \ parts(H)" +by blast + +lemmas parts_insertI = subset_insertI [THEN parts_mono, THEN subsetD, standard] + +lemma parts_empty [simp]: "parts{} = {}" +apply safe +apply (erule parts.induct) +apply blast+ +done + +lemma parts_emptyE [elim!]: "X\ parts{} ==> P" +by simp + +text{*WARNING: loops if H = {Y}, therefore must not be repeated!*} +lemma parts_singleton: "X\ parts H ==> \Y\H. X\ parts {Y}" +apply (erule parts.induct) +apply fast+ +done + + +subsubsection{*Unions *} + +lemma parts_Un_subset1: "parts(G) \ parts(H) \ parts(G \ H)" +by (intro Un_least parts_mono Un_upper1 Un_upper2) + +lemma parts_Un_subset2: "parts(G \ H) \ parts(G) \ parts(H)" +apply (rule subsetI) +apply (erule parts.induct, blast+) +done + +lemma parts_Un [simp]: "parts(G \ H) = parts(G) \ parts(H)" +by (intro equalityI parts_Un_subset1 parts_Un_subset2) + +lemma parts_insert: "parts (insert X H) = parts {X} \ parts H" +apply (subst insert_is_Un [of _ H]) +apply (simp only: parts_Un) +done + +declare [[ atp_problem_prefix = "Message__parts_insert_two" ]] +lemma parts_insert2: + "parts (insert X (insert Y H)) = parts {X} \ parts {Y} \ parts H" +by (metis Un_commute Un_empty_left Un_empty_right Un_insert_left Un_insert_right parts_Un) + + +lemma parts_UN_subset1: "(\x\A. parts(H x)) \ parts(\x\A. H x)" +by (intro UN_least parts_mono UN_upper) + +lemma parts_UN_subset2: "parts(\x\A. H x) \ (\x\A. parts(H x))" +apply (rule subsetI) +apply (erule parts.induct, blast+) +done + +lemma parts_UN [simp]: "parts(\x\A. H x) = (\x\A. parts(H x))" +by (intro equalityI parts_UN_subset1 parts_UN_subset2) + +text{*Added to simplify arguments to parts, analz and synth. + NOTE: the UN versions are no longer used!*} + + +text{*This allows @{text blast} to simplify occurrences of + @{term "parts(G\H)"} in the assumption.*} +lemmas in_parts_UnE = parts_Un [THEN equalityD1, THEN subsetD, THEN UnE] +declare in_parts_UnE [elim!] + +lemma parts_insert_subset: "insert X (parts H) \ parts(insert X H)" +by (blast intro: parts_mono [THEN [2] rev_subsetD]) + +subsubsection{*Idempotence and transitivity *} + +lemma parts_partsD [dest!]: "X\ parts (parts H) ==> X\ parts H" +by (erule parts.induct, blast+) + +lemma parts_idem [simp]: "parts (parts H) = parts H" +by blast + +declare [[ atp_problem_prefix = "Message__parts_subset_iff" ]] +lemma parts_subset_iff [simp]: "(parts G \ parts H) = (G \ parts H)" +apply (rule iffI) +apply (metis Un_absorb1 Un_subset_iff parts_Un parts_increasing) +apply (metis parts_idem parts_mono) +done + +lemma parts_trans: "[| X\ parts G; G \ parts H |] ==> X\ parts H" +by (blast dest: parts_mono); + + +declare [[ atp_problem_prefix = "Message__parts_cut" ]] +lemma parts_cut: "[|Y\ parts(insert X G); X\ parts H|] ==> Y\ parts(G \ H)" +by (metis Un_subset_iff insert_subset parts_increasing parts_trans) + + + +subsubsection{*Rewrite rules for pulling out atomic messages *} + +lemmas parts_insert_eq_I = equalityI [OF subsetI parts_insert_subset] + + +lemma parts_insert_Agent [simp]: + "parts (insert (Agent agt) H) = insert (Agent agt) (parts H)" +apply (rule parts_insert_eq_I) +apply (erule parts.induct, auto) +done + +lemma parts_insert_Nonce [simp]: + "parts (insert (Nonce N) H) = insert (Nonce N) (parts H)" +apply (rule parts_insert_eq_I) +apply (erule parts.induct, auto) +done + +lemma parts_insert_Number [simp]: + "parts (insert (Number N) H) = insert (Number N) (parts H)" +apply (rule parts_insert_eq_I) +apply (erule parts.induct, auto) +done + +lemma parts_insert_Key [simp]: + "parts (insert (Key K) H) = insert (Key K) (parts H)" +apply (rule parts_insert_eq_I) +apply (erule parts.induct, auto) +done + +lemma parts_insert_Hash [simp]: + "parts (insert (Hash X) H) = insert (Hash X) (parts H)" +apply (rule parts_insert_eq_I) +apply (erule parts.induct, auto) +done + +lemma parts_insert_Crypt [simp]: + "parts (insert (Crypt K X) H) = + insert (Crypt K X) (parts (insert X H))" +apply (rule equalityI) +apply (rule subsetI) +apply (erule parts.induct, auto) +apply (blast intro: parts.Body) +done + +lemma parts_insert_MPair [simp]: + "parts (insert {|X,Y|} H) = + insert {|X,Y|} (parts (insert X (insert Y H)))" +apply (rule equalityI) +apply (rule subsetI) +apply (erule parts.induct, auto) +apply (blast intro: parts.Fst parts.Snd)+ +done + +lemma parts_image_Key [simp]: "parts (Key`N) = Key`N" +apply auto +apply (erule parts.induct, auto) +done + + +declare [[ atp_problem_prefix = "Message__msg_Nonce_supply" ]] +lemma msg_Nonce_supply: "\N. \n. N\n --> Nonce n \ parts {msg}" +apply (induct_tac "msg") +apply (simp_all add: parts_insert2) +apply (metis Suc_n_not_le_n) +apply (metis le_trans linorder_linear) +done + +subsection{*Inductive relation "analz"*} + +text{*Inductive definition of "analz" -- what can be broken down from a set of + messages, including keys. A form of downward closure. Pairs can + be taken apart; messages decrypted with known keys. *} + +inductive_set + analz :: "msg set => msg set" + for H :: "msg set" + where + Inj [intro,simp] : "X \ H ==> X \ analz H" + | Fst: "{|X,Y|} \ analz H ==> X \ analz H" + | Snd: "{|X,Y|} \ analz H ==> Y \ analz H" + | Decrypt [dest]: + "[|Crypt K X \ analz H; Key(invKey K): analz H|] ==> X \ analz H" + + +text{*Monotonicity; Lemma 1 of Lowe's paper*} +lemma analz_mono: "G\H ==> analz(G) \ analz(H)" +apply auto +apply (erule analz.induct) +apply (auto dest: analz.Fst analz.Snd) +done + +text{*Making it safe speeds up proofs*} +lemma MPair_analz [elim!]: + "[| {|X,Y|} \ analz H; + [| X \ analz H; Y \ analz H |] ==> P + |] ==> P" +by (blast dest: analz.Fst analz.Snd) + +lemma analz_increasing: "H \ analz(H)" +by blast + +lemma analz_subset_parts: "analz H \ parts H" +apply (rule subsetI) +apply (erule analz.induct, blast+) +done + +lemmas analz_into_parts = analz_subset_parts [THEN subsetD, standard] + +lemmas not_parts_not_analz = analz_subset_parts [THEN contra_subsetD, standard] + + +declare [[ atp_problem_prefix = "Message__parts_analz" ]] +lemma parts_analz [simp]: "parts (analz H) = parts H" +apply (rule equalityI) +apply (metis analz_subset_parts parts_subset_iff) +apply (metis analz_increasing parts_mono) +done + + +lemma analz_parts [simp]: "analz (parts H) = parts H" +apply auto +apply (erule analz.induct, auto) +done + +lemmas analz_insertI = subset_insertI [THEN analz_mono, THEN [2] rev_subsetD, standard] + +subsubsection{*General equational properties *} + +lemma analz_empty [simp]: "analz{} = {}" +apply safe +apply (erule analz.induct, blast+) +done + +text{*Converse fails: we can analz more from the union than from the + separate parts, as a key in one might decrypt a message in the other*} +lemma analz_Un: "analz(G) \ analz(H) \ analz(G \ H)" +by (intro Un_least analz_mono Un_upper1 Un_upper2) + +lemma analz_insert: "insert X (analz H) \ analz(insert X H)" +by (blast intro: analz_mono [THEN [2] rev_subsetD]) + +subsubsection{*Rewrite rules for pulling out atomic messages *} + +lemmas analz_insert_eq_I = equalityI [OF subsetI analz_insert] + +lemma analz_insert_Agent [simp]: + "analz (insert (Agent agt) H) = insert (Agent agt) (analz H)" +apply (rule analz_insert_eq_I) +apply (erule analz.induct, auto) +done + +lemma analz_insert_Nonce [simp]: + "analz (insert (Nonce N) H) = insert (Nonce N) (analz H)" +apply (rule analz_insert_eq_I) +apply (erule analz.induct, auto) +done + +lemma analz_insert_Number [simp]: + "analz (insert (Number N) H) = insert (Number N) (analz H)" +apply (rule analz_insert_eq_I) +apply (erule analz.induct, auto) +done + +lemma analz_insert_Hash [simp]: + "analz (insert (Hash X) H) = insert (Hash X) (analz H)" +apply (rule analz_insert_eq_I) +apply (erule analz.induct, auto) +done + +text{*Can only pull out Keys if they are not needed to decrypt the rest*} +lemma analz_insert_Key [simp]: + "K \ keysFor (analz H) ==> + analz (insert (Key K) H) = insert (Key K) (analz H)" +apply (unfold keysFor_def) +apply (rule analz_insert_eq_I) +apply (erule analz.induct, auto) +done + +lemma analz_insert_MPair [simp]: + "analz (insert {|X,Y|} H) = + insert {|X,Y|} (analz (insert X (insert Y H)))" +apply (rule equalityI) +apply (rule subsetI) +apply (erule analz.induct, auto) +apply (erule analz.induct) +apply (blast intro: analz.Fst analz.Snd)+ +done + +text{*Can pull out enCrypted message if the Key is not known*} +lemma analz_insert_Crypt: + "Key (invKey K) \ analz H + ==> analz (insert (Crypt K X) H) = insert (Crypt K X) (analz H)" +apply (rule analz_insert_eq_I) +apply (erule analz.induct, auto) + +done + +lemma lemma1: "Key (invKey K) \ analz H ==> + analz (insert (Crypt K X) H) \ + insert (Crypt K X) (analz (insert X H))" +apply (rule subsetI) +apply (erule_tac x = x in analz.induct, auto) +done + +lemma lemma2: "Key (invKey K) \ analz H ==> + insert (Crypt K X) (analz (insert X H)) \ + analz (insert (Crypt K X) H)" +apply auto +apply (erule_tac x = x in analz.induct, auto) +apply (blast intro: analz_insertI analz.Decrypt) +done + +lemma analz_insert_Decrypt: + "Key (invKey K) \ analz H ==> + analz (insert (Crypt K X) H) = + insert (Crypt K X) (analz (insert X H))" +by (intro equalityI lemma1 lemma2) + +text{*Case analysis: either the message is secure, or it is not! Effective, +but can cause subgoals to blow up! Use with @{text "split_if"}; apparently +@{text "split_tac"} does not cope with patterns such as @{term"analz (insert +(Crypt K X) H)"} *} +lemma analz_Crypt_if [simp]: + "analz (insert (Crypt K X) H) = + (if (Key (invKey K) \ analz H) + then insert (Crypt K X) (analz (insert X H)) + else insert (Crypt K X) (analz H))" +by (simp add: analz_insert_Crypt analz_insert_Decrypt) + + +text{*This rule supposes "for the sake of argument" that we have the key.*} +lemma analz_insert_Crypt_subset: + "analz (insert (Crypt K X) H) \ + insert (Crypt K X) (analz (insert X H))" +apply (rule subsetI) +apply (erule analz.induct, auto) +done + + +lemma analz_image_Key [simp]: "analz (Key`N) = Key`N" +apply auto +apply (erule analz.induct, auto) +done + + +subsubsection{*Idempotence and transitivity *} + +lemma analz_analzD [dest!]: "X\ analz (analz H) ==> X\ analz H" +by (erule analz.induct, blast+) + +lemma analz_idem [simp]: "analz (analz H) = analz H" +by blast + +lemma analz_subset_iff [simp]: "(analz G \ analz H) = (G \ analz H)" +apply (rule iffI) +apply (iprover intro: subset_trans analz_increasing) +apply (frule analz_mono, simp) +done + +lemma analz_trans: "[| X\ analz G; G \ analz H |] ==> X\ analz H" +by (drule analz_mono, blast) + + +declare [[ atp_problem_prefix = "Message__analz_cut" ]] + declare analz_trans[intro] +lemma analz_cut: "[| Y\ analz (insert X H); X\ analz H |] ==> Y\ analz H" +(*TOO SLOW +by (metis analz_idem analz_increasing analz_mono insert_absorb insert_mono insert_subset) --{*317s*} +??*) +by (erule analz_trans, blast) + + +text{*This rewrite rule helps in the simplification of messages that involve + the forwarding of unknown components (X). Without it, removing occurrences + of X can be very complicated. *} +lemma analz_insert_eq: "X\ analz H ==> analz (insert X H) = analz H" +by (blast intro: analz_cut analz_insertI) + + +text{*A congruence rule for "analz" *} + +declare [[ atp_problem_prefix = "Message__analz_subset_cong" ]] +lemma analz_subset_cong: + "[| analz G \ analz G'; analz H \ analz H' |] + ==> analz (G \ H) \ analz (G' \ H')" +apply simp +apply (metis Un_absorb2 Un_commute Un_subset_iff Un_upper1 Un_upper2 analz_mono) +done + + +lemma analz_cong: + "[| analz G = analz G'; analz H = analz H' + |] ==> analz (G \ H) = analz (G' \ H')" +by (intro equalityI analz_subset_cong, simp_all) + +lemma analz_insert_cong: + "analz H = analz H' ==> analz(insert X H) = analz(insert X H')" +by (force simp only: insert_def intro!: analz_cong) + +text{*If there are no pairs or encryptions then analz does nothing*} +lemma analz_trivial: + "[| \X Y. {|X,Y|} \ H; \X K. Crypt K X \ H |] ==> analz H = H" +apply safe +apply (erule analz.induct, blast+) +done + +text{*These two are obsolete (with a single Spy) but cost little to prove...*} +lemma analz_UN_analz_lemma: + "X\ analz (\i\A. analz (H i)) ==> X\ analz (\i\A. H i)" +apply (erule analz.induct) +apply (blast intro: analz_mono [THEN [2] rev_subsetD])+ +done + +lemma analz_UN_analz [simp]: "analz (\i\A. analz (H i)) = analz (\i\A. H i)" +by (blast intro: analz_UN_analz_lemma analz_mono [THEN [2] rev_subsetD]) + + +subsection{*Inductive relation "synth"*} + +text{*Inductive definition of "synth" -- what can be built up from a set of + messages. A form of upward closure. Pairs can be built, messages + encrypted with known keys. Agent names are public domain. + Numbers can be guessed, but Nonces cannot be. *} + +inductive_set + synth :: "msg set => msg set" + for H :: "msg set" + where + Inj [intro]: "X \ H ==> X \ synth H" + | Agent [intro]: "Agent agt \ synth H" + | Number [intro]: "Number n \ synth H" + | Hash [intro]: "X \ synth H ==> Hash X \ synth H" + | MPair [intro]: "[|X \ synth H; Y \ synth H|] ==> {|X,Y|} \ synth H" + | Crypt [intro]: "[|X \ synth H; Key(K) \ H|] ==> Crypt K X \ synth H" + +text{*Monotonicity*} +lemma synth_mono: "G\H ==> synth(G) \ synth(H)" + by (auto, erule synth.induct, auto) + +text{*NO @{text Agent_synth}, as any Agent name can be synthesized. + The same holds for @{term Number}*} +inductive_cases Nonce_synth [elim!]: "Nonce n \ synth H" +inductive_cases Key_synth [elim!]: "Key K \ synth H" +inductive_cases Hash_synth [elim!]: "Hash X \ synth H" +inductive_cases MPair_synth [elim!]: "{|X,Y|} \ synth H" +inductive_cases Crypt_synth [elim!]: "Crypt K X \ synth H" + + +lemma synth_increasing: "H \ synth(H)" +by blast + +subsubsection{*Unions *} + +text{*Converse fails: we can synth more from the union than from the + separate parts, building a compound message using elements of each.*} +lemma synth_Un: "synth(G) \ synth(H) \ synth(G \ H)" +by (intro Un_least synth_mono Un_upper1 Un_upper2) + + +declare [[ atp_problem_prefix = "Message__synth_insert" ]] + +lemma synth_insert: "insert X (synth H) \ synth(insert X H)" +by (metis insert_iff insert_subset subset_insertI synth.Inj synth_mono) + +subsubsection{*Idempotence and transitivity *} + +lemma synth_synthD [dest!]: "X\ synth (synth H) ==> X\ synth H" +by (erule synth.induct, blast+) + +lemma synth_idem: "synth (synth H) = synth H" +by blast + +lemma synth_subset_iff [simp]: "(synth G \ synth H) = (G \ synth H)" +apply (rule iffI) +apply (iprover intro: subset_trans synth_increasing) +apply (frule synth_mono, simp add: synth_idem) +done + +lemma synth_trans: "[| X\ synth G; G \ synth H |] ==> X\ synth H" +by (drule synth_mono, blast) + +declare [[ atp_problem_prefix = "Message__synth_cut" ]] +lemma synth_cut: "[| Y\ synth (insert X H); X\ synth H |] ==> Y\ synth H" +(*TOO SLOW +by (metis insert_absorb insert_mono insert_subset synth_idem synth_increasing synth_mono) +*) +by (erule synth_trans, blast) + + +lemma Agent_synth [simp]: "Agent A \ synth H" +by blast + +lemma Number_synth [simp]: "Number n \ synth H" +by blast + +lemma Nonce_synth_eq [simp]: "(Nonce N \ synth H) = (Nonce N \ H)" +by blast + +lemma Key_synth_eq [simp]: "(Key K \ synth H) = (Key K \ H)" +by blast + +lemma Crypt_synth_eq [simp]: + "Key K \ H ==> (Crypt K X \ synth H) = (Crypt K X \ H)" +by blast + + +lemma keysFor_synth [simp]: + "keysFor (synth H) = keysFor H \ invKey`{K. Key K \ H}" +by (unfold keysFor_def, blast) + + +subsubsection{*Combinations of parts, analz and synth *} + +declare [[ atp_problem_prefix = "Message__parts_synth" ]] +lemma parts_synth [simp]: "parts (synth H) = parts H \ synth H" +apply (rule equalityI) +apply (rule subsetI) +apply (erule parts.induct) +apply (metis UnCI) +apply (metis MPair_synth UnCI UnE insert_absorb insert_subset parts.Fst parts_increasing) +apply (metis MPair_synth UnCI UnE insert_absorb insert_subset parts.Snd parts_increasing) +apply (metis Body Crypt_synth UnCI UnE insert_absorb insert_subset parts_increasing) +apply (metis Un_subset_iff parts_increasing parts_mono synth_increasing) +done + + + + +declare [[ atp_problem_prefix = "Message__analz_analz_Un" ]] +lemma analz_analz_Un [simp]: "analz (analz G \ H) = analz (G \ H)" +apply (rule equalityI); +apply (metis analz_idem analz_subset_cong order_eq_refl) +apply (metis analz_increasing analz_subset_cong order_eq_refl) +done + +declare [[ atp_problem_prefix = "Message__analz_synth_Un" ]] + declare analz_mono [intro] analz.Fst [intro] analz.Snd [intro] Un_least [intro] +lemma analz_synth_Un [simp]: "analz (synth G \ H) = analz (G \ H) \ synth G" +apply (rule equalityI) +apply (rule subsetI) +apply (erule analz.induct) +apply (metis UnCI UnE Un_commute analz.Inj) +apply (metis MPair_synth UnCI UnE Un_commute Un_upper1 analz.Fst analz_increasing analz_mono insert_absorb insert_subset) +apply (metis MPair_synth UnCI UnE Un_commute Un_upper1 analz.Snd analz_increasing analz_mono insert_absorb insert_subset) +apply (blast intro: analz.Decrypt) +apply blast +done + + +declare [[ atp_problem_prefix = "Message__analz_synth" ]] +lemma analz_synth [simp]: "analz (synth H) = analz H \ synth H" +proof (neg_clausify) +assume 0: "analz (synth H) \ analz H \ synth H" +have 1: "\X1 X3. sup (analz (sup X3 X1)) (synth X3) = analz (sup (synth X3) X1)" + by (metis analz_synth_Un) +have 2: "sup (analz H) (synth H) \ analz (synth H)" + by (metis 0) +have 3: "\X1 X3. sup (synth X3) (analz (sup X3 X1)) = analz (sup (synth X3) X1)" + by (metis 1 Un_commute) +have 4: "\X3. sup (synth X3) (analz X3) = analz (sup (synth X3) {})" + by (metis 3 Un_empty_right) +have 5: "\X3. sup (synth X3) (analz X3) = analz (synth X3)" + by (metis 4 Un_empty_right) +have 6: "\X3. sup (analz X3) (synth X3) = analz (synth X3)" + by (metis 5 Un_commute) +show "False" + by (metis 2 6) +qed + + +subsubsection{*For reasoning about the Fake rule in traces *} + +declare [[ atp_problem_prefix = "Message__parts_insert_subset_Un" ]] +lemma parts_insert_subset_Un: "X\ G ==> parts(insert X H) \ parts G \ parts H" +proof (neg_clausify) +assume 0: "X \ G" +assume 1: "\ parts (insert X H) \ parts G \ parts H" +have 2: "\ parts (insert X H) \ parts (G \ H)" + by (metis 1 parts_Un) +have 3: "\ insert X H \ G \ H" + by (metis 2 parts_mono) +have 4: "X \ G \ H \ \ H \ G \ H" + by (metis 3 insert_subset) +have 5: "X \ G \ H" + by (metis 4 Un_upper2) +have 6: "X \ G" + by (metis 5 UnCI) +show "False" + by (metis 6 0) +qed + +declare [[ atp_problem_prefix = "Message__Fake_parts_insert" ]] +lemma Fake_parts_insert: + "X \ synth (analz H) ==> + parts (insert X H) \ synth (analz H) \ parts H" +proof (neg_clausify) +assume 0: "X \ synth (analz H)" +assume 1: "\ parts (insert X H) \ synth (analz H) \ parts H" +have 2: "\X3. parts X3 \ synth (analz X3) = parts (synth (analz X3))" + by (metis parts_synth parts_analz) +have 3: "\X3. analz X3 \ synth (analz X3) = analz (synth (analz X3))" + by (metis analz_synth analz_idem) +have 4: "\X3. analz X3 \ analz (synth X3)" + by (metis Un_upper1 analz_synth) +have 5: "\ parts (insert X H) \ parts H \ synth (analz H)" + by (metis 1 Un_commute) +have 6: "\ parts (insert X H) \ parts (synth (analz H))" + by (metis 5 2) +have 7: "\ insert X H \ synth (analz H)" + by (metis 6 parts_mono) +have 8: "X \ synth (analz H) \ \ H \ synth (analz H)" + by (metis 7 insert_subset) +have 9: "\ H \ synth (analz H)" + by (metis 8 0) +have 10: "\X3. X3 \ analz (synth X3)" + by (metis analz_subset_iff 4) +have 11: "\X3. X3 \ analz (synth (analz X3))" + by (metis analz_subset_iff 10) +have 12: "\X3. analz (synth (analz X3)) = synth (analz X3) \ + \ analz X3 \ synth (analz X3)" + by (metis Un_absorb1 3) +have 13: "\X3. analz (synth (analz X3)) = synth (analz X3)" + by (metis 12 synth_increasing) +have 14: "\X3. X3 \ synth (analz X3)" + by (metis 11 13) +show "False" + by (metis 9 14) +qed + +lemma Fake_parts_insert_in_Un: + "[|Z \ parts (insert X H); X: synth (analz H)|] + ==> Z \ synth (analz H) \ parts H"; +by (blast dest: Fake_parts_insert [THEN subsetD, dest]) + +declare [[ atp_problem_prefix = "Message__Fake_analz_insert" ]] + declare analz_mono [intro] synth_mono [intro] +lemma Fake_analz_insert: + "X\ synth (analz G) ==> + analz (insert X H) \ synth (analz G) \ analz (G \ H)" +by (metis Un_commute Un_insert_left Un_insert_right Un_upper1 analz_analz_Un analz_mono analz_synth_Un equalityE insert_absorb order_le_less xt1(12)) + +declare [[ atp_problem_prefix = "Message__Fake_analz_insert_simpler" ]] +(*simpler problems? BUT METIS CAN'T PROVE +lemma Fake_analz_insert_simpler: + "X\ synth (analz G) ==> + analz (insert X H) \ synth (analz G) \ analz (G \ H)" +apply (rule subsetI) +apply (subgoal_tac "x \ analz (synth (analz G) \ H) ") +apply (metis Un_commute analz_analz_Un analz_synth_Un) +apply (metis Un_commute Un_upper1 Un_upper2 analz_cut analz_increasing analz_mono insert_absorb insert_mono insert_subset) +done +*) + +end diff -r 8f35633c4922 -r 9cf389429f6d src/HOL/Metis_Examples/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Metis_Examples/ROOT.ML Tue Oct 20 19:52:04 2009 +0200 @@ -0,0 +1,8 @@ +(* Title: HOL/Metis_Examples/ROOT.ML + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + +Testing the metis method. +*) + +use_thys ["set", "BigO", "Abstraction", "BT", "Message", "Tarski", "TransClosure"]; + diff -r 8f35633c4922 -r 9cf389429f6d src/HOL/Metis_Examples/Tarski.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Metis_Examples/Tarski.thy Tue Oct 20 19:52:04 2009 +0200 @@ -0,0 +1,1110 @@ +(* Title: HOL/MetisTest/Tarski.thy + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + +Testing the metis method. +*) + +header {* The Full Theorem of Tarski *} + +theory Tarski +imports Main FuncSet +begin + +(*Many of these higher-order problems appear to be impossible using the +current linkup. They often seem to need either higher-order unification +or explicit reasoning about connectives such as conjunction. The numerous +set comprehensions are to blame.*) + + +record 'a potype = + pset :: "'a set" + order :: "('a * 'a) set" + +constdefs + monotone :: "['a => 'a, 'a set, ('a *'a)set] => bool" + "monotone f A r == \x\A. \y\A. (x, y): r --> ((f x), (f y)) : r" + + least :: "['a => bool, 'a potype] => 'a" + "least P po == @ x. x: pset po & P x & + (\y \ pset po. P y --> (x,y): order po)" + + greatest :: "['a => bool, 'a potype] => 'a" + "greatest P po == @ x. x: pset po & P x & + (\y \ pset po. P y --> (y,x): order po)" + + lub :: "['a set, 'a potype] => 'a" + "lub S po == least (%x. \y\S. (y,x): order po) po" + + glb :: "['a set, 'a potype] => 'a" + "glb S po == greatest (%x. \y\S. (x,y): order po) po" + + isLub :: "['a set, 'a potype, 'a] => bool" + "isLub S po == %L. (L: pset po & (\y\S. (y,L): order po) & + (\z\pset po. (\y\S. (y,z): order po) --> (L,z): order po))" + + isGlb :: "['a set, 'a potype, 'a] => bool" + "isGlb S po == %G. (G: pset po & (\y\S. (G,y): order po) & + (\z \ pset po. (\y\S. (z,y): order po) --> (z,G): order po))" + + "fix" :: "[('a => 'a), 'a set] => 'a set" + "fix f A == {x. x: A & f x = x}" + + interval :: "[('a*'a) set,'a, 'a ] => 'a set" + "interval r a b == {x. (a,x): r & (x,b): r}" + +constdefs + Bot :: "'a potype => 'a" + "Bot po == least (%x. True) po" + + Top :: "'a potype => 'a" + "Top po == greatest (%x. True) po" + + PartialOrder :: "('a potype) set" + "PartialOrder == {P. refl_on (pset P) (order P) & antisym (order P) & + trans (order P)}" + + CompleteLattice :: "('a potype) set" + "CompleteLattice == {cl. cl: PartialOrder & + (\S. S \ pset cl --> (\L. isLub S cl L)) & + (\S. S \ pset cl --> (\G. isGlb S cl G))}" + + induced :: "['a set, ('a * 'a) set] => ('a *'a)set" + "induced A r == {(a,b). a : A & b: A & (a,b): r}" + +constdefs + sublattice :: "('a potype * 'a set)set" + "sublattice == + SIGMA cl: CompleteLattice. + {S. S \ pset cl & + (| pset = S, order = induced S (order cl) |): CompleteLattice }" + +syntax + "@SL" :: "['a set, 'a potype] => bool" ("_ <<= _" [51,50]50) + +translations + "S <<= cl" == "S : sublattice `` {cl}" + +constdefs + dual :: "'a potype => 'a potype" + "dual po == (| pset = pset po, order = converse (order po) |)" + +locale PO = + fixes cl :: "'a potype" + and A :: "'a set" + and r :: "('a * 'a) set" + assumes cl_po: "cl : PartialOrder" + defines A_def: "A == pset cl" + and r_def: "r == order cl" + +locale CL = PO + + assumes cl_co: "cl : CompleteLattice" + +definition CLF_set :: "('a potype * ('a => 'a)) set" where + "CLF_set = (SIGMA cl: CompleteLattice. + {f. f: pset cl -> pset cl & monotone f (pset cl) (order cl)})" + +locale CLF = CL + + fixes f :: "'a => 'a" + and P :: "'a set" + assumes f_cl: "(cl,f) : CLF_set" (*was the equivalent "f : CLF``{cl}"*) + defines P_def: "P == fix f A" + + +locale Tarski = CLF + + fixes Y :: "'a set" + and intY1 :: "'a set" + and v :: "'a" + assumes + Y_ss: "Y \ P" + defines + intY1_def: "intY1 == interval r (lub Y cl) (Top cl)" + and v_def: "v == glb {x. ((%x: intY1. f x) x, x): induced intY1 r & + x: intY1} + (| pset=intY1, order=induced intY1 r|)" + + +subsection {* Partial Order *} + +lemma (in PO) PO_imp_refl_on: "refl_on A r" +apply (insert cl_po) +apply (simp add: PartialOrder_def A_def r_def) +done + +lemma (in PO) PO_imp_sym: "antisym r" +apply (insert cl_po) +apply (simp add: PartialOrder_def r_def) +done + +lemma (in PO) PO_imp_trans: "trans r" +apply (insert cl_po) +apply (simp add: PartialOrder_def r_def) +done + +lemma (in PO) reflE: "x \ A ==> (x, x) \ r" +apply (insert cl_po) +apply (simp add: PartialOrder_def refl_on_def A_def r_def) +done + +lemma (in PO) antisymE: "[| (a, b) \ r; (b, a) \ r |] ==> a = b" +apply (insert cl_po) +apply (simp add: PartialOrder_def antisym_def r_def) +done + +lemma (in PO) transE: "[| (a, b) \ r; (b, c) \ r|] ==> (a,c) \ r" +apply (insert cl_po) +apply (simp add: PartialOrder_def r_def) +apply (unfold trans_def, fast) +done + +lemma (in PO) monotoneE: + "[| monotone f A r; x \ A; y \ A; (x, y) \ r |] ==> (f x, f y) \ r" +by (simp add: monotone_def) + +lemma (in PO) po_subset_po: + "S \ A ==> (| pset = S, order = induced S r |) \ PartialOrder" +apply (simp (no_asm) add: PartialOrder_def) +apply auto +-- {* refl *} +apply (simp add: refl_on_def induced_def) +apply (blast intro: reflE) +-- {* antisym *} +apply (simp add: antisym_def induced_def) +apply (blast intro: antisymE) +-- {* trans *} +apply (simp add: trans_def induced_def) +apply (blast intro: transE) +done + +lemma (in PO) indE: "[| (x, y) \ induced S r; S \ A |] ==> (x, y) \ r" +by (simp add: add: induced_def) + +lemma (in PO) indI: "[| (x, y) \ r; x \ S; y \ S |] ==> (x, y) \ induced S r" +by (simp add: add: induced_def) + +lemma (in CL) CL_imp_ex_isLub: "S \ A ==> \L. isLub S cl L" +apply (insert cl_co) +apply (simp add: CompleteLattice_def A_def) +done + +declare (in CL) cl_co [simp] + +lemma isLub_lub: "(\L. isLub S cl L) = isLub S cl (lub S cl)" +by (simp add: lub_def least_def isLub_def some_eq_ex [symmetric]) + +lemma isGlb_glb: "(\G. isGlb S cl G) = isGlb S cl (glb S cl)" +by (simp add: glb_def greatest_def isGlb_def some_eq_ex [symmetric]) + +lemma isGlb_dual_isLub: "isGlb S cl = isLub S (dual cl)" +by (simp add: isLub_def isGlb_def dual_def converse_def) + +lemma isLub_dual_isGlb: "isLub S cl = isGlb S (dual cl)" +by (simp add: isLub_def isGlb_def dual_def converse_def) + +lemma (in PO) dualPO: "dual cl \ PartialOrder" +apply (insert cl_po) +apply (simp add: PartialOrder_def dual_def refl_on_converse + trans_converse antisym_converse) +done + +lemma Rdual: + "\S. (S \ A -->( \L. isLub S (| pset = A, order = r|) L)) + ==> \S. (S \ A --> (\G. isGlb S (| pset = A, order = r|) G))" +apply safe +apply (rule_tac x = "lub {y. y \ A & (\k \ S. (y, k) \ r)} + (|pset = A, order = r|) " in exI) +apply (drule_tac x = "{y. y \ A & (\k \ S. (y,k) \ r) }" in spec) +apply (drule mp, fast) +apply (simp add: isLub_lub isGlb_def) +apply (simp add: isLub_def, blast) +done + +lemma lub_dual_glb: "lub S cl = glb S (dual cl)" +by (simp add: lub_def glb_def least_def greatest_def dual_def converse_def) + +lemma glb_dual_lub: "glb S cl = lub S (dual cl)" +by (simp add: lub_def glb_def least_def greatest_def dual_def converse_def) + +lemma CL_subset_PO: "CompleteLattice \ PartialOrder" +by (simp add: PartialOrder_def CompleteLattice_def, fast) + +lemmas CL_imp_PO = CL_subset_PO [THEN subsetD] + +declare PO.PO_imp_refl_on [OF PO.intro [OF CL_imp_PO], simp] +declare PO.PO_imp_sym [OF PO.intro [OF CL_imp_PO], simp] +declare PO.PO_imp_trans [OF PO.intro [OF CL_imp_PO], simp] + +lemma (in CL) CO_refl_on: "refl_on A r" +by (rule PO_imp_refl_on) + +lemma (in CL) CO_antisym: "antisym r" +by (rule PO_imp_sym) + +lemma (in CL) CO_trans: "trans r" +by (rule PO_imp_trans) + +lemma CompleteLatticeI: + "[| po \ PartialOrder; (\S. S \ pset po --> (\L. isLub S po L)); + (\S. S \ pset po --> (\G. isGlb S po G))|] + ==> po \ CompleteLattice" +apply (unfold CompleteLattice_def, blast) +done + +lemma (in CL) CL_dualCL: "dual cl \ CompleteLattice" +apply (insert cl_co) +apply (simp add: CompleteLattice_def dual_def) +apply (fold dual_def) +apply (simp add: isLub_dual_isGlb [symmetric] isGlb_dual_isLub [symmetric] + dualPO) +done + +lemma (in PO) dualA_iff: "pset (dual cl) = pset cl" +by (simp add: dual_def) + +lemma (in PO) dualr_iff: "((x, y) \ (order(dual cl))) = ((y, x) \ order cl)" +by (simp add: dual_def) + +lemma (in PO) monotone_dual: + "monotone f (pset cl) (order cl) + ==> monotone f (pset (dual cl)) (order(dual cl))" +by (simp add: monotone_def dualA_iff dualr_iff) + +lemma (in PO) interval_dual: + "[| x \ A; y \ A|] ==> interval r x y = interval (order(dual cl)) y x" +apply (simp add: interval_def dualr_iff) +apply (fold r_def, fast) +done + +lemma (in PO) interval_not_empty: + "[| trans r; interval r a b \ {} |] ==> (a, b) \ r" +apply (simp add: interval_def) +apply (unfold trans_def, blast) +done + +lemma (in PO) interval_imp_mem: "x \ interval r a b ==> (a, x) \ r" +by (simp add: interval_def) + +lemma (in PO) left_in_interval: + "[| a \ A; b \ A; interval r a b \ {} |] ==> a \ interval r a b" +apply (simp (no_asm_simp) add: interval_def) +apply (simp add: PO_imp_trans interval_not_empty) +apply (simp add: reflE) +done + +lemma (in PO) right_in_interval: + "[| a \ A; b \ A; interval r a b \ {} |] ==> b \ interval r a b" +apply (simp (no_asm_simp) add: interval_def) +apply (simp add: PO_imp_trans interval_not_empty) +apply (simp add: reflE) +done + + +subsection {* sublattice *} + +lemma (in PO) sublattice_imp_CL: + "S <<= cl ==> (| pset = S, order = induced S r |) \ CompleteLattice" +by (simp add: sublattice_def CompleteLattice_def A_def r_def) + +lemma (in CL) sublatticeI: + "[| S \ A; (| pset = S, order = induced S r |) \ CompleteLattice |] + ==> S <<= cl" +by (simp add: sublattice_def A_def r_def) + + +subsection {* lub *} + +lemma (in CL) lub_unique: "[| S \ A; isLub S cl x; isLub S cl L|] ==> x = L" +apply (rule antisymE) +apply (auto simp add: isLub_def r_def) +done + +lemma (in CL) lub_upper: "[|S \ A; x \ S|] ==> (x, lub S cl) \ r" +apply (rule CL_imp_ex_isLub [THEN exE], assumption) +apply (unfold lub_def least_def) +apply (rule some_equality [THEN ssubst]) + apply (simp add: isLub_def) + apply (simp add: lub_unique A_def isLub_def) +apply (simp add: isLub_def r_def) +done + +lemma (in CL) lub_least: + "[| S \ A; L \ A; \x \ S. (x,L) \ r |] ==> (lub S cl, L) \ r" +apply (rule CL_imp_ex_isLub [THEN exE], assumption) +apply (unfold lub_def least_def) +apply (rule_tac s=x in some_equality [THEN ssubst]) + apply (simp add: isLub_def) + apply (simp add: lub_unique A_def isLub_def) +apply (simp add: isLub_def r_def A_def) +done + +lemma (in CL) lub_in_lattice: "S \ A ==> lub S cl \ A" +apply (rule CL_imp_ex_isLub [THEN exE], assumption) +apply (unfold lub_def least_def) +apply (subst some_equality) +apply (simp add: isLub_def) +prefer 2 apply (simp add: isLub_def A_def) +apply (simp add: lub_unique A_def isLub_def) +done + +lemma (in CL) lubI: + "[| S \ A; L \ A; \x \ S. (x,L) \ r; + \z \ A. (\y \ S. (y,z) \ r) --> (L,z) \ r |] ==> L = lub S cl" +apply (rule lub_unique, assumption) +apply (simp add: isLub_def A_def r_def) +apply (unfold isLub_def) +apply (rule conjI) +apply (fold A_def r_def) +apply (rule lub_in_lattice, assumption) +apply (simp add: lub_upper lub_least) +done + +lemma (in CL) lubIa: "[| S \ A; isLub S cl L |] ==> L = lub S cl" +by (simp add: lubI isLub_def A_def r_def) + +lemma (in CL) isLub_in_lattice: "isLub S cl L ==> L \ A" +by (simp add: isLub_def A_def) + +lemma (in CL) isLub_upper: "[|isLub S cl L; y \ S|] ==> (y, L) \ r" +by (simp add: isLub_def r_def) + +lemma (in CL) isLub_least: + "[| isLub S cl L; z \ A; \y \ S. (y, z) \ r|] ==> (L, z) \ r" +by (simp add: isLub_def A_def r_def) + +lemma (in CL) isLubI: + "[| L \ A; \y \ S. (y, L) \ r; + (\z \ A. (\y \ S. (y, z):r) --> (L, z) \ r)|] ==> isLub S cl L" +by (simp add: isLub_def A_def r_def) + + + +subsection {* glb *} + +lemma (in CL) glb_in_lattice: "S \ A ==> glb S cl \ A" +apply (subst glb_dual_lub) +apply (simp add: A_def) +apply (rule dualA_iff [THEN subst]) +apply (rule CL.lub_in_lattice) +apply (rule CL.intro) +apply (rule PO.intro) +apply (rule dualPO) +apply (rule CL_axioms.intro) +apply (rule CL_dualCL) +apply (simp add: dualA_iff) +done + +lemma (in CL) glb_lower: "[|S \ A; x \ S|] ==> (glb S cl, x) \ r" +apply (subst glb_dual_lub) +apply (simp add: r_def) +apply (rule dualr_iff [THEN subst]) +apply (rule CL.lub_upper) +apply (rule CL.intro) +apply (rule PO.intro) +apply (rule dualPO) +apply (rule CL_axioms.intro) +apply (rule CL_dualCL) +apply (simp add: dualA_iff A_def, assumption) +done + +text {* + Reduce the sublattice property by using substructural properties; + abandoned see @{text "Tarski_4.ML"}. +*} + +declare (in CLF) f_cl [simp] + +(*never proved, 2007-01-22: Tarski__CLF_unnamed_lemma + NOT PROVABLE because of the conjunction used in the definition: we don't + allow reasoning with rules like conjE, which is essential here.*) +declare [[ atp_problem_prefix = "Tarski__CLF_unnamed_lemma" ]] +lemma (in CLF) [simp]: + "f: pset cl -> pset cl & monotone f (pset cl) (order cl)" +apply (insert f_cl) +apply (unfold CLF_set_def) +apply (erule SigmaE2) +apply (erule CollectE) +apply assumption +done + +lemma (in CLF) f_in_funcset: "f \ A -> A" +by (simp add: A_def) + +lemma (in CLF) monotone_f: "monotone f A r" +by (simp add: A_def r_def) + +(*never proved, 2007-01-22*) +declare [[ atp_problem_prefix = "Tarski__CLF_CLF_dual" ]] +declare (in CLF) CLF_set_def [simp] CL_dualCL [simp] monotone_dual [simp] dualA_iff [simp] + +lemma (in CLF) CLF_dual: "(dual cl, f) \ CLF_set" +apply (simp del: dualA_iff) +apply (simp) +done + +declare (in CLF) CLF_set_def[simp del] CL_dualCL[simp del] monotone_dual[simp del] + dualA_iff[simp del] + + +subsection {* fixed points *} + +lemma fix_subset: "fix f A \ A" +by (simp add: fix_def, fast) + +lemma fix_imp_eq: "x \ fix f A ==> f x = x" +by (simp add: fix_def) + +lemma fixf_subset: + "[| A \ B; x \ fix (%y: A. f y) A |] ==> x \ fix f B" +by (simp add: fix_def, auto) + + +subsection {* lemmas for Tarski, lub *} + +(*never proved, 2007-01-22*) +declare [[ atp_problem_prefix = "Tarski__CLF_lubH_le_flubH" ]] + declare CL.lub_least[intro] CLF.f_in_funcset[intro] funcset_mem[intro] CL.lub_in_lattice[intro] PO.transE[intro] PO.monotoneE[intro] CLF.monotone_f[intro] CL.lub_upper[intro] +lemma (in CLF) lubH_le_flubH: + "H = {x. (x, f x) \ r & x \ A} ==> (lub H cl, f (lub H cl)) \ r" +apply (rule lub_least, fast) +apply (rule f_in_funcset [THEN funcset_mem]) +apply (rule lub_in_lattice, fast) +-- {* @{text "\x:H. (x, f (lub H r)) \ r"} *} +apply (rule ballI) +(*never proved, 2007-01-22*) +using [[ atp_problem_prefix = "Tarski__CLF_lubH_le_flubH_simpler" ]] +apply (rule transE) +-- {* instantiates @{text "(x, ?z) \ order cl to (x, f x)"}, *} +-- {* because of the def of @{text H} *} +apply fast +-- {* so it remains to show @{text "(f x, f (lub H cl)) \ r"} *} +apply (rule_tac f = "f" in monotoneE) +apply (rule monotone_f, fast) +apply (rule lub_in_lattice, fast) +apply (rule lub_upper, fast) +apply assumption +done + declare CL.lub_least[rule del] CLF.f_in_funcset[rule del] + funcset_mem[rule del] CL.lub_in_lattice[rule del] + PO.transE[rule del] PO.monotoneE[rule del] + CLF.monotone_f[rule del] CL.lub_upper[rule del] + +(*never proved, 2007-01-22*) +declare [[ atp_problem_prefix = "Tarski__CLF_flubH_le_lubH" ]] + declare CLF.f_in_funcset[intro] funcset_mem[intro] CL.lub_in_lattice[intro] + PO.monotoneE[intro] CLF.monotone_f[intro] CL.lub_upper[intro] + CLF.lubH_le_flubH[simp] +lemma (in CLF) flubH_le_lubH: + "[| H = {x. (x, f x) \ r & x \ A} |] ==> (f (lub H cl), lub H cl) \ r" +apply (rule lub_upper, fast) +apply (rule_tac t = "H" in ssubst, assumption) +apply (rule CollectI) +apply (rule conjI) +using [[ atp_problem_prefix = "Tarski__CLF_flubH_le_lubH_simpler" ]] +(*??no longer terminates, with combinators +apply (metis CO_refl_on lubH_le_flubH monotone_def monotone_f reflD1 reflD2) +*) +apply (metis CO_refl_on lubH_le_flubH monotoneE [OF monotone_f] refl_onD1 refl_onD2) +apply (metis CO_refl_on lubH_le_flubH refl_onD2) +done + declare CLF.f_in_funcset[rule del] funcset_mem[rule del] + CL.lub_in_lattice[rule del] PO.monotoneE[rule del] + CLF.monotone_f[rule del] CL.lub_upper[rule del] + CLF.lubH_le_flubH[simp del] + + +(*never proved, 2007-01-22*) +declare [[ atp_problem_prefix = "Tarski__CLF_lubH_is_fixp" ]] +(*Single-step version fails. The conjecture clauses refer to local abstraction +functions (Frees), which prevents expand_defs_tac from removing those +"definitions" at the end of the proof. *) +lemma (in CLF) lubH_is_fixp: + "H = {x. (x, f x) \ r & x \ A} ==> lub H cl \ fix f A" +apply (simp add: fix_def) +apply (rule conjI) +proof (neg_clausify) +assume 0: "H = +Collect + (COMBS (COMBB op \ (COMBC (COMBB op \ (COMBS Pair f)) r)) (COMBC op \ A))" +assume 1: "lub (Collect + (COMBS (COMBB op \ (COMBC (COMBB op \ (COMBS Pair f)) r)) + (COMBC op \ A))) + cl +\ A" +have 2: "lub H cl \ A" + by (metis 1 0) +have 3: "(lub H cl, f (lub H cl)) \ r" + by (metis lubH_le_flubH 0) +have 4: "(f (lub H cl), lub H cl) \ r" + by (metis flubH_le_lubH 0) +have 5: "lub H cl = f (lub H cl) \ (lub H cl, f (lub H cl)) \ r" + by (metis antisymE 4) +have 6: "lub H cl = f (lub H cl)" + by (metis 5 3) +have 7: "(lub H cl, lub H cl) \ r" + by (metis 6 4) +have 8: "\X1. lub H cl \ X1 \ \ refl_on X1 r" + by (metis 7 refl_onD2) +have 9: "\ refl_on A r" + by (metis 8 2) +show "False" + by (metis CO_refl_on 9); +next --{*apparently the way to insert a second structured proof*} + show "H = {x. (x, f x) \ r \ x \ A} \ + f (lub {x. (x, f x) \ r \ x \ A} cl) = lub {x. (x, f x) \ r \ x \ A} cl" + proof (neg_clausify) + assume 0: "H = + Collect + (COMBS (COMBB op \ (COMBC (COMBB op \ (COMBS Pair f)) r)) (COMBC op \ A))" + assume 1: "f (lub (Collect + (COMBS (COMBB op \ (COMBC (COMBB op \ (COMBS Pair f)) r)) + (COMBC op \ A))) + cl) \ + lub (Collect + (COMBS (COMBB op \ (COMBC (COMBB op \ (COMBS Pair f)) r)) + (COMBC op \ A))) + cl" + have 2: "f (lub H cl) \ + lub (Collect + (COMBS (COMBB op \ (COMBC (COMBB op \ (COMBS Pair f)) r)) + (COMBC op \ A))) + cl" + by (metis 1 0) + have 3: "f (lub H cl) \ lub H cl" + by (metis 2 0) + have 4: "(lub H cl, f (lub H cl)) \ r" + by (metis lubH_le_flubH 0) + have 5: "(f (lub H cl), lub H cl) \ r" + by (metis flubH_le_lubH 0) + have 6: "lub H cl = f (lub H cl) \ (lub H cl, f (lub H cl)) \ r" + by (metis antisymE 5) + have 7: "lub H cl = f (lub H cl)" + by (metis 6 4) + show "False" + by (metis 3 7) + qed +qed + +lemma (in CLF) (*lubH_is_fixp:*) + "H = {x. (x, f x) \ r & x \ A} ==> lub H cl \ fix f A" +apply (simp add: fix_def) +apply (rule conjI) +using [[ atp_problem_prefix = "Tarski__CLF_lubH_is_fixp_simpler" ]] +apply (metis CO_refl_on lubH_le_flubH refl_onD1) +apply (metis antisymE flubH_le_lubH lubH_le_flubH) +done + +lemma (in CLF) fix_in_H: + "[| H = {x. (x, f x) \ r & x \ A}; x \ P |] ==> x \ H" +by (simp add: P_def fix_imp_eq [of _ f A] reflE CO_refl_on + fix_subset [of f A, THEN subsetD]) + + +lemma (in CLF) fixf_le_lubH: + "H = {x. (x, f x) \ r & x \ A} ==> \x \ fix f A. (x, lub H cl) \ r" +apply (rule ballI) +apply (rule lub_upper, fast) +apply (rule fix_in_H) +apply (simp_all add: P_def) +done + +declare [[ atp_problem_prefix = "Tarski__CLF_lubH_least_fixf" ]] +lemma (in CLF) lubH_least_fixf: + "H = {x. (x, f x) \ r & x \ A} + ==> \L. (\y \ fix f A. (y,L) \ r) --> (lub H cl, L) \ r" +apply (metis P_def lubH_is_fixp) +done + +subsection {* Tarski fixpoint theorem 1, first part *} +declare [[ atp_problem_prefix = "Tarski__CLF_T_thm_1_lub" ]] + declare CL.lubI[intro] fix_subset[intro] CL.lub_in_lattice[intro] + CLF.fixf_le_lubH[simp] CLF.lubH_least_fixf[simp] +lemma (in CLF) T_thm_1_lub: "lub P cl = lub {x. (x, f x) \ r & x \ A} cl" +(*sledgehammer;*) +apply (rule sym) +apply (simp add: P_def) +apply (rule lubI) +using [[ atp_problem_prefix = "Tarski__CLF_T_thm_1_lub_simpler" ]] +apply (metis P_def fix_subset) +apply (metis Collect_conj_eq Collect_mem_eq Int_commute Int_lower1 lub_in_lattice vimage_def) +(*??no longer terminates, with combinators +apply (metis P_def fix_def fixf_le_lubH) +apply (metis P_def fix_def lubH_least_fixf) +*) +apply (simp add: fixf_le_lubH) +apply (simp add: lubH_least_fixf) +done + declare CL.lubI[rule del] fix_subset[rule del] CL.lub_in_lattice[rule del] + CLF.fixf_le_lubH[simp del] CLF.lubH_least_fixf[simp del] + + +(*never proved, 2007-01-22*) +declare [[ atp_problem_prefix = "Tarski__CLF_glbH_is_fixp" ]] + declare glb_dual_lub[simp] PO.dualA_iff[intro] CLF.lubH_is_fixp[intro] + PO.dualPO[intro] CL.CL_dualCL[intro] PO.dualr_iff[simp] +lemma (in CLF) glbH_is_fixp: "H = {x. (f x, x) \ r & x \ A} ==> glb H cl \ P" + -- {* Tarski for glb *} +(*sledgehammer;*) +apply (simp add: glb_dual_lub P_def A_def r_def) +apply (rule dualA_iff [THEN subst]) +apply (rule CLF.lubH_is_fixp) +apply (rule CLF.intro) +apply (rule CL.intro) +apply (rule PO.intro) +apply (rule dualPO) +apply (rule CL_axioms.intro) +apply (rule CL_dualCL) +apply (rule CLF_axioms.intro) +apply (rule CLF_dual) +apply (simp add: dualr_iff dualA_iff) +done + declare glb_dual_lub[simp del] PO.dualA_iff[rule del] CLF.lubH_is_fixp[rule del] + PO.dualPO[rule del] CL.CL_dualCL[rule del] PO.dualr_iff[simp del] + + +(*never proved, 2007-01-22*) +declare [[ atp_problem_prefix = "Tarski__T_thm_1_glb" ]] (*ALL THEOREMS*) +lemma (in CLF) T_thm_1_glb: "glb P cl = glb {x. (f x, x) \ r & x \ A} cl" +(*sledgehammer;*) +apply (simp add: glb_dual_lub P_def A_def r_def) +apply (rule dualA_iff [THEN subst]) +(*never proved, 2007-01-22*) +using [[ atp_problem_prefix = "Tarski__T_thm_1_glb_simpler" ]] (*ALL THEOREMS*) +(*sledgehammer;*) +apply (simp add: CLF.T_thm_1_lub [of _ f, OF CLF.intro, OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro, + OF dualPO CL_dualCL] dualPO CL_dualCL CLF_dual dualr_iff) +done + +subsection {* interval *} + + +declare [[ atp_problem_prefix = "Tarski__rel_imp_elem" ]] + declare (in CLF) CO_refl_on[simp] refl_on_def [simp] +lemma (in CLF) rel_imp_elem: "(x, y) \ r ==> x \ A" +by (metis CO_refl_on refl_onD1) + declare (in CLF) CO_refl_on[simp del] refl_on_def [simp del] + +declare [[ atp_problem_prefix = "Tarski__interval_subset" ]] + declare (in CLF) rel_imp_elem[intro] + declare interval_def [simp] +lemma (in CLF) interval_subset: "[| a \ A; b \ A |] ==> interval r a b \ A" +by (metis CO_refl_on interval_imp_mem refl_onD refl_onD2 rel_imp_elem subset_eq) + declare (in CLF) rel_imp_elem[rule del] + declare interval_def [simp del] + + +lemma (in CLF) intervalI: + "[| (a, x) \ r; (x, b) \ r |] ==> x \ interval r a b" +by (simp add: interval_def) + +lemma (in CLF) interval_lemma1: + "[| S \ interval r a b; x \ S |] ==> (a, x) \ r" +by (unfold interval_def, fast) + +lemma (in CLF) interval_lemma2: + "[| S \ interval r a b; x \ S |] ==> (x, b) \ r" +by (unfold interval_def, fast) + +lemma (in CLF) a_less_lub: + "[| S \ A; S \ {}; + \x \ S. (a,x) \ r; \y \ S. (y, L) \ r |] ==> (a,L) \ r" +by (blast intro: transE) + +lemma (in CLF) glb_less_b: + "[| S \ A; S \ {}; + \x \ S. (x,b) \ r; \y \ S. (G, y) \ r |] ==> (G,b) \ r" +by (blast intro: transE) + +lemma (in CLF) S_intv_cl: + "[| a \ A; b \ A; S \ interval r a b |]==> S \ A" +by (simp add: subset_trans [OF _ interval_subset]) + +declare [[ atp_problem_prefix = "Tarski__L_in_interval" ]] (*ALL THEOREMS*) +lemma (in CLF) L_in_interval: + "[| a \ A; b \ A; S \ interval r a b; + S \ {}; isLub S cl L; interval r a b \ {} |] ==> L \ interval r a b" +(*WON'T TERMINATE +apply (metis CO_trans intervalI interval_lemma1 interval_lemma2 isLub_least isLub_upper subset_empty subset_iff trans_def) +*) +apply (rule intervalI) +apply (rule a_less_lub) +prefer 2 apply assumption +apply (simp add: S_intv_cl) +apply (rule ballI) +apply (simp add: interval_lemma1) +apply (simp add: isLub_upper) +-- {* @{text "(L, b) \ r"} *} +apply (simp add: isLub_least interval_lemma2) +done + +(*never proved, 2007-01-22*) +declare [[ atp_problem_prefix = "Tarski__G_in_interval" ]] (*ALL THEOREMS*) +lemma (in CLF) G_in_interval: + "[| a \ A; b \ A; interval r a b \ {}; S \ interval r a b; isGlb S cl G; + S \ {} |] ==> G \ interval r a b" +apply (simp add: interval_dual) +apply (simp add: CLF.L_in_interval [of _ f, OF CLF.intro, OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro] + dualA_iff A_def dualPO CL_dualCL CLF_dual isGlb_dual_isLub) +done + +declare [[ atp_problem_prefix = "Tarski__intervalPO" ]] (*ALL THEOREMS*) +lemma (in CLF) intervalPO: + "[| a \ A; b \ A; interval r a b \ {} |] + ==> (| pset = interval r a b, order = induced (interval r a b) r |) + \ PartialOrder" +proof (neg_clausify) +assume 0: "a \ A" +assume 1: "b \ A" +assume 2: "\pset = interval r a b, order = induced (interval r a b) r\ \ PartialOrder" +have 3: "\ interval r a b \ A" + by (metis 2 po_subset_po) +have 4: "b \ A \ a \ A" + by (metis 3 interval_subset) +have 5: "a \ A" + by (metis 4 1) +show "False" + by (metis 5 0) +qed + +lemma (in CLF) intv_CL_lub: + "[| a \ A; b \ A; interval r a b \ {} |] + ==> \S. S \ interval r a b --> + (\L. isLub S (| pset = interval r a b, + order = induced (interval r a b) r |) L)" +apply (intro strip) +apply (frule S_intv_cl [THEN CL_imp_ex_isLub]) +prefer 2 apply assumption +apply assumption +apply (erule exE) +-- {* define the lub for the interval as *} +apply (rule_tac x = "if S = {} then a else L" in exI) +apply (simp (no_asm_simp) add: isLub_def split del: split_if) +apply (intro impI conjI) +-- {* @{text "(if S = {} then a else L) \ interval r a b"} *} +apply (simp add: CL_imp_PO L_in_interval) +apply (simp add: left_in_interval) +-- {* lub prop 1 *} +apply (case_tac "S = {}") +-- {* @{text "S = {}, y \ S = False => everything"} *} +apply fast +-- {* @{text "S \ {}"} *} +apply simp +-- {* @{text "\y:S. (y, L) \ induced (interval r a b) r"} *} +apply (rule ballI) +apply (simp add: induced_def L_in_interval) +apply (rule conjI) +apply (rule subsetD) +apply (simp add: S_intv_cl, assumption) +apply (simp add: isLub_upper) +-- {* @{text "\z:interval r a b. (\y:S. (y, z) \ induced (interval r a b) r \ (if S = {} then a else L, z) \ induced (interval r a b) r"} *} +apply (rule ballI) +apply (rule impI) +apply (case_tac "S = {}") +-- {* @{text "S = {}"} *} +apply simp +apply (simp add: induced_def interval_def) +apply (rule conjI) +apply (rule reflE, assumption) +apply (rule interval_not_empty) +apply (rule CO_trans) +apply (simp add: interval_def) +-- {* @{text "S \ {}"} *} +apply simp +apply (simp add: induced_def L_in_interval) +apply (rule isLub_least, assumption) +apply (rule subsetD) +prefer 2 apply assumption +apply (simp add: S_intv_cl, fast) +done + +lemmas (in CLF) intv_CL_glb = intv_CL_lub [THEN Rdual] + +(*never proved, 2007-01-22*) +declare [[ atp_problem_prefix = "Tarski__interval_is_sublattice" ]] (*ALL THEOREMS*) +lemma (in CLF) interval_is_sublattice: + "[| a \ A; b \ A; interval r a b \ {} |] + ==> interval r a b <<= cl" +(*sledgehammer *) +apply (rule sublatticeI) +apply (simp add: interval_subset) +(*never proved, 2007-01-22*) +using [[ atp_problem_prefix = "Tarski__interval_is_sublattice_simpler" ]] +(*sledgehammer *) +apply (rule CompleteLatticeI) +apply (simp add: intervalPO) + apply (simp add: intv_CL_lub) +apply (simp add: intv_CL_glb) +done + +lemmas (in CLF) interv_is_compl_latt = + interval_is_sublattice [THEN sublattice_imp_CL] + + +subsection {* Top and Bottom *} +lemma (in CLF) Top_dual_Bot: "Top cl = Bot (dual cl)" +by (simp add: Top_def Bot_def least_def greatest_def dualA_iff dualr_iff) + +lemma (in CLF) Bot_dual_Top: "Bot cl = Top (dual cl)" +by (simp add: Top_def Bot_def least_def greatest_def dualA_iff dualr_iff) + +declare [[ atp_problem_prefix = "Tarski__Bot_in_lattice" ]] (*ALL THEOREMS*) +lemma (in CLF) Bot_in_lattice: "Bot cl \ A" +(*sledgehammer; *) +apply (simp add: Bot_def least_def) +apply (rule_tac a="glb A cl" in someI2) +apply (simp_all add: glb_in_lattice glb_lower + r_def [symmetric] A_def [symmetric]) +done + +(*first proved 2007-01-25 after relaxing relevance*) +declare [[ atp_problem_prefix = "Tarski__Top_in_lattice" ]] (*ALL THEOREMS*) +lemma (in CLF) Top_in_lattice: "Top cl \ A" +(*sledgehammer;*) +apply (simp add: Top_dual_Bot A_def) +(*first proved 2007-01-25 after relaxing relevance*) +using [[ atp_problem_prefix = "Tarski__Top_in_lattice_simpler" ]] (*ALL THEOREMS*) +(*sledgehammer*) +apply (rule dualA_iff [THEN subst]) +apply (blast intro!: CLF.Bot_in_lattice [OF CLF.intro, OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro] dualPO CL_dualCL CLF_dual) +done + +lemma (in CLF) Top_prop: "x \ A ==> (x, Top cl) \ r" +apply (simp add: Top_def greatest_def) +apply (rule_tac a="lub A cl" in someI2) +apply (rule someI2) +apply (simp_all add: lub_in_lattice lub_upper + r_def [symmetric] A_def [symmetric]) +done + +(*never proved, 2007-01-22*) +declare [[ atp_problem_prefix = "Tarski__Bot_prop" ]] (*ALL THEOREMS*) +lemma (in CLF) Bot_prop: "x \ A ==> (Bot cl, x) \ r" +(*sledgehammer*) +apply (simp add: Bot_dual_Top r_def) +apply (rule dualr_iff [THEN subst]) +apply (simp add: CLF.Top_prop [of _ f, OF CLF.intro, OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro] + dualA_iff A_def dualPO CL_dualCL CLF_dual) +done + +declare [[ atp_problem_prefix = "Tarski__Bot_in_lattice" ]] (*ALL THEOREMS*) +lemma (in CLF) Top_intv_not_empty: "x \ A ==> interval r x (Top cl) \ {}" +apply (metis Top_in_lattice Top_prop empty_iff intervalI reflE) +done + +declare [[ atp_problem_prefix = "Tarski__Bot_intv_not_empty" ]] (*ALL THEOREMS*) +lemma (in CLF) Bot_intv_not_empty: "x \ A ==> interval r (Bot cl) x \ {}" +apply (metis Bot_prop ex_in_conv intervalI reflE rel_imp_elem) +done + + +subsection {* fixed points form a partial order *} + +lemma (in CLF) fixf_po: "(| pset = P, order = induced P r|) \ PartialOrder" +by (simp add: P_def fix_subset po_subset_po) + +(*first proved 2007-01-25 after relaxing relevance*) +declare [[ atp_problem_prefix = "Tarski__Y_subset_A" ]] + declare (in Tarski) P_def[simp] Y_ss [simp] + declare fix_subset [intro] subset_trans [intro] +lemma (in Tarski) Y_subset_A: "Y \ A" +(*sledgehammer*) +apply (rule subset_trans [OF _ fix_subset]) +apply (rule Y_ss [simplified P_def]) +done + declare (in Tarski) P_def[simp del] Y_ss [simp del] + declare fix_subset [rule del] subset_trans [rule del] + + +lemma (in Tarski) lubY_in_A: "lub Y cl \ A" + by (rule Y_subset_A [THEN lub_in_lattice]) + +(*never proved, 2007-01-22*) +declare [[ atp_problem_prefix = "Tarski__lubY_le_flubY" ]] (*ALL THEOREMS*) +lemma (in Tarski) lubY_le_flubY: "(lub Y cl, f (lub Y cl)) \ r" +(*sledgehammer*) +apply (rule lub_least) +apply (rule Y_subset_A) +apply (rule f_in_funcset [THEN funcset_mem]) +apply (rule lubY_in_A) +-- {* @{text "Y \ P ==> f x = x"} *} +apply (rule ballI) +using [[ atp_problem_prefix = "Tarski__lubY_le_flubY_simpler" ]] (*ALL THEOREMS*) +(*sledgehammer *) +apply (rule_tac t = "x" in fix_imp_eq [THEN subst]) +apply (erule Y_ss [simplified P_def, THEN subsetD]) +-- {* @{text "reduce (f x, f (lub Y cl)) \ r to (x, lub Y cl) \ r"} by monotonicity *} +using [[ atp_problem_prefix = "Tarski__lubY_le_flubY_simplest" ]] (*ALL THEOREMS*) +(*sledgehammer*) +apply (rule_tac f = "f" in monotoneE) +apply (rule monotone_f) +apply (simp add: Y_subset_A [THEN subsetD]) +apply (rule lubY_in_A) +apply (simp add: lub_upper Y_subset_A) +done + +(*first proved 2007-01-25 after relaxing relevance*) +declare [[ atp_problem_prefix = "Tarski__intY1_subset" ]] (*ALL THEOREMS*) +lemma (in Tarski) intY1_subset: "intY1 \ A" +(*sledgehammer*) +apply (unfold intY1_def) +apply (rule interval_subset) +apply (rule lubY_in_A) +apply (rule Top_in_lattice) +done + +lemmas (in Tarski) intY1_elem = intY1_subset [THEN subsetD] + +(*never proved, 2007-01-22*) +declare [[ atp_problem_prefix = "Tarski__intY1_f_closed" ]] (*ALL THEOREMS*) +lemma (in Tarski) intY1_f_closed: "x \ intY1 \ f x \ intY1" +(*sledgehammer*) +apply (simp add: intY1_def interval_def) +apply (rule conjI) +apply (rule transE) +apply (rule lubY_le_flubY) +-- {* @{text "(f (lub Y cl), f x) \ r"} *} +using [[ atp_problem_prefix = "Tarski__intY1_f_closed_simpler" ]] (*ALL THEOREMS*) +(*sledgehammer [has been proved before now...]*) +apply (rule_tac f=f in monotoneE) +apply (rule monotone_f) +apply (rule lubY_in_A) +apply (simp add: intY1_def interval_def intY1_elem) +apply (simp add: intY1_def interval_def) +-- {* @{text "(f x, Top cl) \ r"} *} +apply (rule Top_prop) +apply (rule f_in_funcset [THEN funcset_mem]) +apply (simp add: intY1_def interval_def intY1_elem) +done + +declare [[ atp_problem_prefix = "Tarski__intY1_func" ]] (*ALL THEOREMS*) +lemma (in Tarski) intY1_func: "(%x: intY1. f x) \ intY1 -> intY1" +apply (rule restrict_in_funcset) +apply (metis intY1_f_closed restrict_in_funcset) +done + +declare [[ atp_problem_prefix = "Tarski__intY1_mono" ]] (*ALL THEOREMS*) +lemma (in Tarski) intY1_mono: + "monotone (%x: intY1. f x) intY1 (induced intY1 r)" +(*sledgehammer *) +apply (auto simp add: monotone_def induced_def intY1_f_closed) +apply (blast intro: intY1_elem monotone_f [THEN monotoneE]) +done + +(*proof requires relaxing relevance: 2007-01-25*) +declare [[ atp_problem_prefix = "Tarski__intY1_is_cl" ]] (*ALL THEOREMS*) +lemma (in Tarski) intY1_is_cl: + "(| pset = intY1, order = induced intY1 r |) \ CompleteLattice" +(*sledgehammer*) +apply (unfold intY1_def) +apply (rule interv_is_compl_latt) +apply (rule lubY_in_A) +apply (rule Top_in_lattice) +apply (rule Top_intv_not_empty) +apply (rule lubY_in_A) +done + +(*never proved, 2007-01-22*) +declare [[ atp_problem_prefix = "Tarski__v_in_P" ]] (*ALL THEOREMS*) +lemma (in Tarski) v_in_P: "v \ P" +(*sledgehammer*) +apply (unfold P_def) +apply (rule_tac A = "intY1" in fixf_subset) +apply (rule intY1_subset) +apply (simp add: CLF.glbH_is_fixp [OF CLF.intro, OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro, OF _ intY1_is_cl, simplified] + v_def CL_imp_PO intY1_is_cl CLF_set_def intY1_func intY1_mono) +done + +declare [[ atp_problem_prefix = "Tarski__z_in_interval" ]] (*ALL THEOREMS*) +lemma (in Tarski) z_in_interval: + "[| z \ P; \y\Y. (y, z) \ induced P r |] ==> z \ intY1" +(*sledgehammer *) +apply (unfold intY1_def P_def) +apply (rule intervalI) +prefer 2 + apply (erule fix_subset [THEN subsetD, THEN Top_prop]) +apply (rule lub_least) +apply (rule Y_subset_A) +apply (fast elim!: fix_subset [THEN subsetD]) +apply (simp add: induced_def) +done + +declare [[ atp_problem_prefix = "Tarski__fz_in_int_rel" ]] (*ALL THEOREMS*) +lemma (in Tarski) f'z_in_int_rel: "[| z \ P; \y\Y. (y, z) \ induced P r |] + ==> ((%x: intY1. f x) z, z) \ induced intY1 r" +apply (metis P_def acc_def fix_imp_eq fix_subset indI reflE restrict_apply subset_eq z_in_interval) +done + +(*never proved, 2007-01-22*) +declare [[ atp_problem_prefix = "Tarski__tarski_full_lemma" ]] (*ALL THEOREMS*) +lemma (in Tarski) tarski_full_lemma: + "\L. isLub Y (| pset = P, order = induced P r |) L" +apply (rule_tac x = "v" in exI) +apply (simp add: isLub_def) +-- {* @{text "v \ P"} *} +apply (simp add: v_in_P) +apply (rule conjI) +(*sledgehammer*) +-- {* @{text v} is lub *} +-- {* @{text "1. \y:Y. (y, v) \ induced P r"} *} +apply (rule ballI) +apply (simp add: induced_def subsetD v_in_P) +apply (rule conjI) +apply (erule Y_ss [THEN subsetD]) +apply (rule_tac b = "lub Y cl" in transE) +apply (rule lub_upper) +apply (rule Y_subset_A, assumption) +apply (rule_tac b = "Top cl" in interval_imp_mem) +apply (simp add: v_def) +apply (fold intY1_def) +apply (rule CL.glb_in_lattice [OF CL.intro, OF PO.intro CL_axioms.intro, OF _ intY1_is_cl, simplified]) + apply (simp add: CL_imp_PO intY1_is_cl, force) +-- {* @{text v} is LEAST ub *} +apply clarify +apply (rule indI) + prefer 3 apply assumption + prefer 2 apply (simp add: v_in_P) +apply (unfold v_def) +(*never proved, 2007-01-22*) +using [[ atp_problem_prefix = "Tarski__tarski_full_lemma_simpler" ]] +(*sledgehammer*) +apply (rule indE) +apply (rule_tac [2] intY1_subset) +(*never proved, 2007-01-22*) +using [[ atp_problem_prefix = "Tarski__tarski_full_lemma_simplest" ]] +(*sledgehammer*) +apply (rule CL.glb_lower [OF CL.intro, OF PO.intro CL_axioms.intro, OF _ intY1_is_cl, simplified]) + apply (simp add: CL_imp_PO intY1_is_cl) + apply force +apply (simp add: induced_def intY1_f_closed z_in_interval) +apply (simp add: P_def fix_imp_eq [of _ f A] reflE + fix_subset [of f A, THEN subsetD]) +done + +lemma CompleteLatticeI_simp: + "[| (| pset = A, order = r |) \ PartialOrder; + \S. S \ A --> (\L. isLub S (| pset = A, order = r |) L) |] + ==> (| pset = A, order = r |) \ CompleteLattice" +by (simp add: CompleteLatticeI Rdual) + + +(*never proved, 2007-01-22*) +declare [[ atp_problem_prefix = "Tarski__Tarski_full" ]] + declare (in CLF) fixf_po[intro] P_def [simp] A_def [simp] r_def [simp] + Tarski.tarski_full_lemma [intro] cl_po [intro] cl_co [intro] + CompleteLatticeI_simp [intro] +theorem (in CLF) Tarski_full: + "(| pset = P, order = induced P r|) \ CompleteLattice" +(*sledgehammer*) +apply (rule CompleteLatticeI_simp) +apply (rule fixf_po, clarify) +(*never proved, 2007-01-22*) +using [[ atp_problem_prefix = "Tarski__Tarski_full_simpler" ]] +(*sledgehammer*) +apply (simp add: P_def A_def r_def) +apply (blast intro!: Tarski.tarski_full_lemma [OF Tarski.intro, OF CLF.intro Tarski_axioms.intro, + OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro] cl_po cl_co f_cl) +done + declare (in CLF) fixf_po[rule del] P_def [simp del] A_def [simp del] r_def [simp del] + Tarski.tarski_full_lemma [rule del] cl_po [rule del] cl_co [rule del] + CompleteLatticeI_simp [rule del] + + +end diff -r 8f35633c4922 -r 9cf389429f6d src/HOL/Metis_Examples/TransClosure.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Metis_Examples/TransClosure.thy Tue Oct 20 19:52:04 2009 +0200 @@ -0,0 +1,62 @@ +(* Title: HOL/MetisTest/TransClosure.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + +Testing the metis method +*) + +theory TransClosure +imports Main +begin + +types addr = nat + +datatype val + = Unit -- "dummy result value of void expressions" + | Null -- "null reference" + | Bool bool -- "Boolean value" + | Intg int -- "integer value" + | Addr addr -- "addresses of objects in the heap" + +consts R::"(addr \ addr) set" + +consts f:: "addr \ val" + +declare [[ atp_problem_prefix = "TransClosure__test" ]] +lemma "\ f c = Intg x; \ y. f b = Intg y \ y \ x; (a,b) \ R\<^sup>*; (b,c) \ R\<^sup>* \ + \ \ c. (b,c) \ R \ (a,c) \ R\<^sup>*" +by (metis Transitive_Closure.rtrancl_into_rtrancl converse_rtranclE trancl_reflcl) + +lemma "\ f c = Intg x; \ y. f b = Intg y \ y \ x; (a,b) \ R\<^sup>*; (b,c) \ R\<^sup>* \ + \ \ c. (b,c) \ R \ (a,c) \ R\<^sup>*" +proof (neg_clausify) +assume 0: "f c = Intg x" +assume 1: "(a, b) \ R\<^sup>*" +assume 2: "(b, c) \ R\<^sup>*" +assume 3: "f b \ Intg x" +assume 4: "\A. (b, A) \ R \ (a, A) \ R\<^sup>*" +have 5: "b = c \ b \ Domain R" + by (metis Not_Domain_rtrancl 2) +have 6: "\X1. (a, X1) \ R\<^sup>* \ (b, X1) \ R" + by (metis Transitive_Closure.rtrancl_into_rtrancl 1) +have 7: "\X1. (b, X1) \ R" + by (metis 6 4) +have 8: "b \ Domain R" + by (metis 7 DomainE) +have 9: "c = b" + by (metis 5 8) +have 10: "f b = Intg x" + by (metis 0 9) +show "False" + by (metis 10 3) +qed + +declare [[ atp_problem_prefix = "TransClosure__test_simpler" ]] +lemma "\ f c = Intg x; \ y. f b = Intg y \ y \ x; (a,b) \ R\<^sup>*; (b,c) \ R\<^sup>* \ + \ \ c. (b,c) \ R \ (a,c) \ R\<^sup>*" +apply (erule_tac x="b" in converse_rtranclE) +apply (metis rel_pow_0_E rel_pow_0_I) +apply (metis DomainE Domain_iff Transitive_Closure.rtrancl_into_rtrancl) +done + +end diff -r 8f35633c4922 -r 9cf389429f6d src/HOL/Metis_Examples/set.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Metis_Examples/set.thy Tue Oct 20 19:52:04 2009 +0200 @@ -0,0 +1,283 @@ +(* Title: HOL/Metis_Examples/set.thy + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + +Testing the metis method. +*) + +theory set +imports Main +begin + +lemma "EX x X. ALL y. EX z Z. (~P(y,y) | P(x,x) | ~S(z,x)) & + (S(x,y) | ~S(y,z) | Q(Z,Z)) & + (Q(X,y) | ~Q(y,Z) | S(X,X))" +by metis +(*??But metis can't prove the single-step version...*) + + + +lemma "P(n::nat) ==> ~P(0) ==> n ~= 0" +by metis + +declare [[sledgehammer_modulus = 1]] + + +(*multiple versions of this example*) +lemma (*equal_union: *) + "(X = Y \ Z) = + (Y \ X \ Z \ X \ (\V. Y \ V \ Z \ V \ X \ V))" +proof (neg_clausify) +fix x +assume 0: "Y \ X \ X = Y \ Z" +assume 1: "Z \ X \ X = Y \ Z" +assume 2: "(\ Y \ X \ \ Z \ X \ Y \ x) \ X \ Y \ Z" +assume 3: "(\ Y \ X \ \ Z \ X \ Z \ x) \ X \ Y \ Z" +assume 4: "(\ Y \ X \ \ Z \ X \ \ X \ x) \ X \ Y \ Z" +assume 5: "\V. ((\ Y \ V \ \ Z \ V) \ X \ V) \ X = Y \ Z" +have 6: "sup Y Z = X \ Y \ X" + by (metis 0) +have 7: "sup Y Z = X \ Z \ X" + by (metis 1) +have 8: "\X3. sup Y Z = X \ X \ X3 \ \ Y \ X3 \ \ Z \ X3" + by (metis 5) +have 9: "Y \ x \ sup Y Z \ X \ \ Y \ X \ \ Z \ X" + by (metis 2) +have 10: "Z \ x \ sup Y Z \ X \ \ Y \ X \ \ Z \ X" + by (metis 3) +have 11: "sup Y Z \ X \ \ X \ x \ \ Y \ X \ \ Z \ X" + by (metis 4) +have 12: "Z \ X" + by (metis Un_upper2 7) +have 13: "\X3. sup Y Z = X \ X \ sup X3 Z \ \ Y \ sup X3 Z" + by (metis 8 Un_upper2) +have 14: "Y \ X" + by (metis Un_upper1 6) +have 15: "Z \ x \ sup Y Z \ X \ \ Y \ X" + by (metis 10 12) +have 16: "Y \ x \ sup Y Z \ X \ \ Y \ X" + by (metis 9 12) +have 17: "sup Y Z \ X \ \ X \ x \ \ Y \ X" + by (metis 11 12) +have 18: "sup Y Z \ X \ \ X \ x" + by (metis 17 14) +have 19: "Z \ x \ sup Y Z \ X" + by (metis 15 14) +have 20: "Y \ x \ sup Y Z \ X" + by (metis 16 14) +have 21: "sup Y Z = X \ X \ sup Y Z" + by (metis 13 Un_upper1) +have 22: "sup Y Z = X \ \ sup Y Z \ X" + by (metis equalityI 21) +have 23: "sup Y Z = X \ \ Z \ X \ \ Y \ X" + by (metis 22 Un_least) +have 24: "sup Y Z = X \ \ Y \ X" + by (metis 23 12) +have 25: "sup Y Z = X" + by (metis 24 14) +have 26: "\X3. X \ X3 \ \ Z \ X3 \ \ Y \ X3" + by (metis Un_least 25) +have 27: "Y \ x" + by (metis 20 25) +have 28: "Z \ x" + by (metis 19 25) +have 29: "\ X \ x" + by (metis 18 25) +have 30: "X \ x \ \ Y \ x" + by (metis 26 28) +have 31: "X \ x" + by (metis 30 27) +show "False" + by (metis 31 29) +qed + +declare [[sledgehammer_modulus = 2]] + +lemma (*equal_union: *) + "(X = Y \ Z) = + (Y \ X \ Z \ X \ (\V. Y \ V \ Z \ V \ X \ V))" +proof (neg_clausify) +fix x +assume 0: "Y \ X \ X = Y \ Z" +assume 1: "Z \ X \ X = Y \ Z" +assume 2: "(\ Y \ X \ \ Z \ X \ Y \ x) \ X \ Y \ Z" +assume 3: "(\ Y \ X \ \ Z \ X \ Z \ x) \ X \ Y \ Z" +assume 4: "(\ Y \ X \ \ Z \ X \ \ X \ x) \ X \ Y \ Z" +assume 5: "\V. ((\ Y \ V \ \ Z \ V) \ X \ V) \ X = Y \ Z" +have 6: "sup Y Z = X \ Y \ X" + by (metis 0) +have 7: "Y \ x \ sup Y Z \ X \ \ Y \ X \ \ Z \ X" + by (metis 2) +have 8: "sup Y Z \ X \ \ X \ x \ \ Y \ X \ \ Z \ X" + by (metis 4) +have 9: "\X3. sup Y Z = X \ X \ sup X3 Z \ \ Y \ sup X3 Z" + by (metis 5 Un_upper2) +have 10: "Z \ x \ sup Y Z \ X \ \ Y \ X" + by (metis 3 Un_upper2) +have 11: "sup Y Z \ X \ \ X \ x \ \ Y \ X" + by (metis 8 Un_upper2) +have 12: "Z \ x \ sup Y Z \ X" + by (metis 10 Un_upper1) +have 13: "sup Y Z = X \ X \ sup Y Z" + by (metis 9 Un_upper1) +have 14: "sup Y Z = X \ \ Z \ X \ \ Y \ X" + by (metis equalityI 13 Un_least) +have 15: "sup Y Z = X" + by (metis 14 1 6) +have 16: "Y \ x" + by (metis 7 Un_upper2 Un_upper1 15) +have 17: "\ X \ x" + by (metis 11 Un_upper1 15) +have 18: "X \ x" + by (metis Un_least 15 12 15 16) +show "False" + by (metis 18 17) +qed + +declare [[sledgehammer_modulus = 3]] + +lemma (*equal_union: *) + "(X = Y \ Z) = + (Y \ X \ Z \ X \ (\V. Y \ V \ Z \ V \ X \ V))" +proof (neg_clausify) +fix x +assume 0: "Y \ X \ X = Y \ Z" +assume 1: "Z \ X \ X = Y \ Z" +assume 2: "(\ Y \ X \ \ Z \ X \ Y \ x) \ X \ Y \ Z" +assume 3: "(\ Y \ X \ \ Z \ X \ Z \ x) \ X \ Y \ Z" +assume 4: "(\ Y \ X \ \ Z \ X \ \ X \ x) \ X \ Y \ Z" +assume 5: "\V. ((\ Y \ V \ \ Z \ V) \ X \ V) \ X = Y \ Z" +have 6: "Z \ x \ sup Y Z \ X \ \ Y \ X \ \ Z \ X" + by (metis 3) +have 7: "\X3. sup Y Z = X \ X \ sup X3 Z \ \ Y \ sup X3 Z" + by (metis 5 Un_upper2) +have 8: "Y \ x \ sup Y Z \ X \ \ Y \ X" + by (metis 2 Un_upper2) +have 9: "Z \ x \ sup Y Z \ X" + by (metis 6 Un_upper2 Un_upper1) +have 10: "sup Y Z = X \ \ sup Y Z \ X" + by (metis equalityI 7 Un_upper1) +have 11: "sup Y Z = X" + by (metis 10 Un_least 1 0) +have 12: "Z \ x" + by (metis 9 11) +have 13: "X \ x" + by (metis Un_least 11 12 8 Un_upper1 11) +show "False" + by (metis 13 4 Un_upper2 Un_upper1 11) +qed + +(*Example included in TPHOLs paper*) + +declare [[sledgehammer_modulus = 4]] + +lemma (*equal_union: *) + "(X = Y \ Z) = + (Y \ X \ Z \ X \ (\V. Y \ V \ Z \ V \ X \ V))" +proof (neg_clausify) +fix x +assume 0: "Y \ X \ X = Y \ Z" +assume 1: "Z \ X \ X = Y \ Z" +assume 2: "(\ Y \ X \ \ Z \ X \ Y \ x) \ X \ Y \ Z" +assume 3: "(\ Y \ X \ \ Z \ X \ Z \ x) \ X \ Y \ Z" +assume 4: "(\ Y \ X \ \ Z \ X \ \ X \ x) \ X \ Y \ Z" +assume 5: "\V. ((\ Y \ V \ \ Z \ V) \ X \ V) \ X = Y \ Z" +have 6: "sup Y Z \ X \ \ X \ x \ \ Y \ X \ \ Z \ X" + by (metis 4) +have 7: "Z \ x \ sup Y Z \ X \ \ Y \ X" + by (metis 3 Un_upper2) +have 8: "Z \ x \ sup Y Z \ X" + by (metis 7 Un_upper1) +have 9: "sup Y Z = X \ \ Z \ X \ \ Y \ X" + by (metis equalityI 5 Un_upper2 Un_upper1 Un_least) +have 10: "Y \ x" + by (metis 2 Un_upper2 1 Un_upper1 0 9 Un_upper2 1 Un_upper1 0) +have 11: "X \ x" + by (metis Un_least 9 Un_upper2 1 Un_upper1 0 8 9 Un_upper2 1 Un_upper1 0 10) +show "False" + by (metis 11 6 Un_upper2 1 Un_upper1 0 9 Un_upper2 1 Un_upper1 0) +qed + +declare [[ atp_problem_prefix = "set__equal_union" ]] +lemma (*equal_union: *) + "(X = Y \ Z) = + (Y \ X \ Z \ X \ (\V. Y \ V \ Z \ V \ X \ V))" +(*One shot proof: hand-reduced. Metis can't do the full proof any more.*) +by (metis Un_least Un_upper1 Un_upper2 set_eq_subset) + + +declare [[ atp_problem_prefix = "set__equal_inter" ]] +lemma "(X = Y \ Z) = + (X \ Y \ X \ Z \ (\V. V \ Y \ V \ Z \ V \ X))" +by (metis Int_greatest Int_lower1 Int_lower2 set_eq_subset) + +declare [[ atp_problem_prefix = "set__fixedpoint" ]] +lemma fixedpoint: + "\!x. f (g x) = x \ \!y. g (f y) = y" +by metis + +lemma (*fixedpoint:*) + "\!x. f (g x) = x \ \!y. g (f y) = y" +proof (neg_clausify) +fix x xa +assume 0: "f (g x) = x" +assume 1: "\y. y = x \ f (g y) \ y" +assume 2: "\x. g (f (xa x)) = xa x \ g (f x) \ x" +assume 3: "\x. g (f x) \ x \ xa x \ x" +have 4: "\X1. g (f X1) \ X1 \ g x \ X1" + by (metis 3 1 2) +show "False" + by (metis 4 0) +qed + +declare [[ atp_problem_prefix = "set__singleton_example" ]] +lemma (*singleton_example_2:*) + "\x \ S. \S \ x \ \z. S \ {z}" +by (metis Set.subsetI Union_upper insertCI set_eq_subset) + --{*found by SPASS*} + +lemma (*singleton_example_2:*) + "\x \ S. \S \ x \ \z. S \ {z}" +by (metis Set.subsetI Union_upper insert_iff set_eq_subset) + +lemma singleton_example_2: + "\x \ S. \S \ x \ \z. S \ {z}" +proof (neg_clausify) +assume 0: "\x. \ S \ {x}" +assume 1: "\x. x \ S \ \S \ x" +have 2: "\X3. X3 = \S \ \ X3 \ \S \ X3 \ S" + by (metis set_eq_subset 1) +have 3: "\X3. S \ insert (\S) X3" + by (metis insert_iff Set.subsetI Union_upper 2 Set.subsetI) +show "False" + by (metis 3 0) +qed + + + +text {* + From W. W. Bledsoe and Guohui Feng, SET-VAR. JAR 11 (3), 1993, pages + 293-314. +*} + +declare [[ atp_problem_prefix = "set__Bledsoe_Fung" ]] +(*Notes: 1, the numbering doesn't completely agree with the paper. +2, we must rename set variables to avoid type clashes.*) +lemma "\B. (\x \ B. x \ (0::int))" + "D \ F \ \G. \A \ G. \B \ F. A \ B" + "P a \ \A. (\x \ A. P x) \ (\y. y \ A)" + "a < b \ b < (c::int) \ \B. a \ B \ b \ B \ c \ B" + "P (f b) \ \s A. (\x \ A. P x) \ f s \ A" + "P (f b) \ \s A. (\x \ A. P x) \ f s \ A" + "\A. a \ A" + "(\C. (0, 0) \ C \ (\x y. (x, y) \ C \ (Suc x, Suc y) \ C) \ (n, m) \ C) \ Q n \ Q m" +apply (metis atMost_iff) +apply (metis emptyE) +apply (metis insert_iff singletonE) +apply (metis insertCI singletonE zless_le) +apply (metis Collect_def Collect_mem_eq) +apply (metis Collect_def Collect_mem_eq) +apply (metis DiffE) +apply (metis pair_in_Id_conv) +done + +end diff -r 8f35633c4922 -r 9cf389429f6d src/HOL/Tools/res_axioms.ML --- a/src/HOL/Tools/res_axioms.ML Tue Oct 20 19:37:09 2009 +0200 +++ b/src/HOL/Tools/res_axioms.ML Tue Oct 20 19:52:04 2009 +0200 @@ -301,7 +301,7 @@ then exists (excessive_lambdas_fm Ts) (#2 (strip_comb t)) else excessive_lambdas (t, max_lambda_nesting); -(*The max apply_depth of any metis call in MetisExamples (on 31-10-2007) was 11.*) +(*The max apply_depth of any metis call in Metis_Examples (on 31-10-2007) was 11.*) val max_apply_depth = 15; fun apply_depth (f$t) = Int.max (apply_depth f, apply_depth t + 1)