diff --git a/thys/Adaptive_State_Counting/ASC/ASC_LB.thy b/thys/Adaptive_State_Counting/ASC/ASC_LB.thy --- a/thys/Adaptive_State_Counting/ASC/ASC_LB.thy +++ b/thys/Adaptive_State_Counting/ASC/ASC_LB.thy @@ -1,2892 +1,2892 @@ theory ASC_LB imports "../ATC/ATC" "../FSM/FSM_Product" begin section \ The lower bound function \ text \ This theory defines the lower bound function @{verbatim LB} and its properties. Function @{verbatim LB} calculates a lower bound on the number of states of some FSM in order for some sequence to not contain certain repetitions. \ subsection \ Permutation function Perm \ text \ Function @{verbatim Perm} calculates all possible reactions of an FSM to a set of inputs sequences such that every set in the calculated set of reactions contains exactly one reaction for each input sequence. \ fun Perm :: "'in list set \ ('in, 'out, 'state) FSM \ ('in \ 'out) list set set" where "Perm V M = {image f V | f . \ v \ V . f v \ language_state_for_input M (initial M) v }" lemma perm_empty : assumes "is_det_state_cover M2 V" and "V'' \ Perm V M1" shows "[] \ V''" proof - have init_seq : "[] \ V" using det_state_cover_empty assms by simp obtain f where f_def : "V'' = image f V \ (\ v \ V . f v \ language_state_for_input M1 (initial M1) v)" using assms by auto then have "f [] = []" using init_seq by (metis language_state_for_input_empty singleton_iff) then show ?thesis using init_seq f_def by (metis image_eqI) qed lemma perm_elem_finite : assumes "is_det_state_cover M2 V" and "well_formed M2" and "V'' \ Perm V M1" shows "finite V''" proof - obtain f where "is_det_state_cover_ass M2 f \ V = f ` d_reachable M2 (initial M2)" using assms by auto moreover have "finite (d_reachable M2 (initial M2))" proof - have "finite (nodes M2)" using assms by auto moreover have "nodes M2 = reachable M2 (initial M2)" by auto ultimately have "finite (reachable M2 (initial M2))" by simp moreover have "d_reachable M2 (initial M2) \ reachable M2 (initial M2)" by auto ultimately show ?thesis using infinite_super by blast qed ultimately have "finite V" by auto moreover obtain f'' where "V'' = image f'' V \ (\ v \ V . f'' v \ language_state_for_input M1 (initial M1) v)" using assms(3) by auto ultimately show ?thesis by simp qed lemma perm_inputs : assumes "V'' \ Perm V M" and "vs \ V''" shows "map fst vs \ V" proof - obtain f where f_def : "V'' = image f V \ (\ v \ V . f v \ language_state_for_input M (initial M) v)" using assms by auto then obtain v where v_def : "v \ V \ f v = vs" using assms by auto then have "vs \ language_state_for_input M (initial M) v" using f_def by auto then show ?thesis using v_def unfolding language_state_for_input.simps by auto qed lemma perm_inputs_diff : assumes "V'' \ Perm V M" and "vs1 \ V''" and "vs2 \ V''" and "vs1 \ vs2" shows "map fst vs1 \ map fst vs2" proof - obtain f where f_def : "V'' = image f V \ (\ v \ V . f v \ language_state_for_input M (initial M) v)" using assms by auto then obtain v1 v2 where v_def : "v1 \ V \ f v1 = vs1 \ v2 \ V \ f v2 = vs2" using assms by auto then have "vs1 \ language_state_for_input M (initial M) v1" "vs2 \ language_state_for_input M (initial M) v2" using f_def by auto moreover have "v1 \ v2" using v_def assms(4) by blast ultimately show ?thesis by auto qed lemma perm_language : assumes "V'' \ Perm V M" and "vs \ V''" shows "vs \ L M" proof - obtain f where f_def : "image f V = V'' \ (\ v \ V . f v \ language_state_for_input M (initial M) v)" using assms(1) by auto then have "\ v . f v = vs \ f v \ language_state_for_input M (initial M) v" using assms(2) by blast then show ?thesis by auto qed subsection \ Helper predicates \ text \ The following predicates are used to combine often repeated assumption. \ abbreviation "asc_fault_domain M2 M1 m \ (inputs M2 = inputs M1 \ card (nodes M1) \ m )" lemma asc_fault_domain_props[elim!] : assumes "asc_fault_domain M2 M1 m" shows "inputs M2 = inputs M1" "card (nodes M1) \ m"using assms by auto abbreviation "test_tools M2 M1 FAIL PM V \ \ ( productF M2 M1 FAIL PM \ is_det_state_cover M2 V \ applicable_set M2 \ )" lemma test_tools_props[elim] : assumes "test_tools M2 M1 FAIL PM V \" and "asc_fault_domain M2 M1 m" shows "productF M2 M1 FAIL PM" "is_det_state_cover M2 V" "applicable_set M2 \" "applicable_set M1 \" proof - show "productF M2 M1 FAIL PM" using assms(1) by blast show "is_det_state_cover M2 V" using assms(1) by blast show "applicable_set M2 \" using assms(1) by blast then show "applicable_set M1 \" unfolding applicable_set.simps applicable.simps using asc_fault_domain_props(1)[OF assms(2)] by simp qed lemma perm_nonempty : assumes "is_det_state_cover M2 V" and "OFSM M1" and "OFSM M2" and "inputs M1 = inputs M2" shows "Perm V M1 \ {}" proof - have "finite (nodes M2)" using assms(3) by auto moreover have "d_reachable M2 (initial M2) \ nodes M2" by auto ultimately have "finite V" using det_state_cover_card[OF assms(1)] by (metis assms(1) finite_imageI infinite_super is_det_state_cover.elims(2)) have "[] \ V" using assms(1) det_state_cover_empty by blast have "\ VS . VS \ V \ VS \ {} \ Perm VS M1 \ {}" proof - fix VS assume "VS \ V \ VS \ {}" then have "finite VS" using \finite V\ using infinite_subset by auto then show "Perm VS M1 \ {}" using \VS \ V \ VS \ {}\ \finite VS\ proof (induction VS) case empty then show ?case by auto next case (insert vs F) then have "vs \ V" by blast obtain q2 where "d_reaches M2 (initial M2) vs q2" using det_state_cover_d_reachable[OF assms(1) \vs \ V\] by blast then obtain vs' vsP where io_path : "length vs = length vs' \ length vs = length vsP \ (path M2 ((vs || vs') || vsP) (initial M2)) \ target ((vs || vs') || vsP) (initial M2) = q2" by auto have "well_formed M2" using assms by auto have "map fst (map fst ((vs || vs') || vsP)) = vs" proof - have "length (vs || vs') = length vsP" using io_path by simp then show ?thesis using io_path by auto qed moreover have "set (map fst (map fst ((vs || vs') || vsP))) \ inputs M2" using path_input_containment[OF \well_formed M2\, of "(vs || vs') || vsP" "initial M2"] io_path by linarith ultimately have "set vs \ inputs M2" by presburger then have "set vs \ inputs M1" using assms by auto then have "L\<^sub>i\<^sub>n M1 {vs} \ {}" using assms(2) language_state_for_inputs_nonempty by (metis FSM.nodes.initial) then have "language_state_for_input M1 (initial M1) vs \ {}" by auto then obtain vs' where "vs' \ language_state_for_input M1 (initial M1) vs" by blast show ?case proof (cases "F = {}") case True moreover obtain f where "f vs = vs'" - by moura + by force ultimately have "image f (insert vs F) \ Perm (insert vs F) M1" using Perm.simps \vs' \ language_state_for_input M1 (initial M1) vs\ by blast then show ?thesis by blast next case False then obtain F'' where "F'' \ Perm F M1" using insert.IH insert.hyps(1) insert.prems(1) by blast then obtain f where "F'' = image f F" "(\ v \ F . f v \ language_state_for_input M1 (initial M1) v)" by auto let ?f = "f(vs := vs')" have "\ v \ (insert vs F) . ?f v \ language_state_for_input M1 (initial M1) v" proof fix v assume "v \ insert vs F" then show "?f v \ language_state_for_input M1 (initial M1) v" proof (cases "v = vs") case True then show ?thesis using \vs' \ language_state_for_input M1 (initial M1) vs\ by auto next case False then have "v \ F" using \v \ insert vs F\ by blast then show ?thesis using False \\v\F. f v \ language_state_for_input M1 (initial M1) v\ by auto qed qed then have "image ?f (insert vs F) \ Perm (insert vs F) M1" using Perm.simps by blast then show ?thesis by blast qed qed qed then show ?thesis using \[] \ V\ by blast qed lemma perm_elem : assumes "is_det_state_cover M2 V" and "OFSM M1" and "OFSM M2" and "inputs M1 = inputs M2" and "vs \ V" and "vs' \ language_state_for_input M1 (initial M1) vs" obtains V'' where "V'' \ Perm V M1" "vs' \ V''" proof - obtain V'' where "V'' \ Perm V M1" using perm_nonempty[OF assms(1-4)] by blast then obtain f where "V'' = image f V" "(\ v \ V . f v \ language_state_for_input M1 (initial M1) v)" by auto let ?f = "f(vs := vs')" have "\ v \ V . (?f v) \ (language_state_for_input M1 (initial M1) v)" using \\v\V. (f v) \ (language_state_for_input M1 (initial M1) v)\ assms(6) by fastforce then have "(image ?f V) \ Perm V M1" unfolding Perm.simps by blast moreover have "vs' \ image ?f V" by (metis assms(5) fun_upd_same imageI) ultimately show ?thesis using that by blast qed subsection \ Function R \ text \ Function @{verbatim R} calculates the set of suffixes of a sequence that reach a given state if applied after a given other sequence. \ fun R :: "('in, 'out, 'state) FSM \ 'state \ ('in \ 'out) list \ ('in \ 'out) list \ ('in \ 'out) list set" where "R M s vs xs = { vs@xs' | xs' . xs' \ [] \ prefix xs' xs \ s \ io_targets M (initial M) (vs@xs') }" lemma finite_R : "finite (R M s vs xs)" proof - have "R M s vs xs \ { vs @ xs' | xs' .prefix xs' xs }" by auto then have "R M s vs xs \ image (\ xs' . vs @ xs') {xs' . prefix xs' xs}" by auto moreover have "{xs' . prefix xs' xs} = {take n xs | n . n \ length xs}" proof show "{xs'. prefix xs' xs} \ {take n xs |n. n \ length xs}" proof fix xs' assume "xs' \ {xs'. prefix xs' xs}" then obtain zs' where "xs' @ zs' = xs" by (metis (full_types) mem_Collect_eq prefixE) then obtain i where "xs' = take i xs \ i \ length xs" by (metis (full_types) append_eq_conv_conj le_cases take_all) then show "xs' \ {take n xs |n. n \ length xs}" by auto qed show "{take n xs |n. n \ length xs} \ {xs'. prefix xs' xs}" using take_is_prefix by force qed moreover have "finite {take n xs | n . n \ length xs}" by auto ultimately show ?thesis by auto qed lemma card_union_of_singletons : assumes "\ S \ SS . (\ t . S = {t})" shows "card (\ SS) = card SS" proof - let ?f = "\ x . {x}" have "bij_betw ?f (\ SS) SS" unfolding bij_betw_def inj_on_def using assms by fastforce then show ?thesis using bij_betw_same_card by blast qed lemma card_union_of_distinct : assumes "\ S1 \ SS . \ S2 \ SS . S1 = S2 \ f S1 \ f S2 = {}" and "finite SS" and "\ S \ SS . f S \ {}" shows "card (image f SS) = card SS" proof - from assms(2) have "\ S1 \ SS . \ S2 \ SS . S1 = S2 \ f S1 \ f S2 = {} \ \ S \ SS . f S \ {} \ ?thesis" proof (induction SS) case empty then show ?case by auto next case (insert x F) then have "\ (\ y \ F . f y = f x)" by auto then have "f x \ image f F" by auto then have "card (image f (insert x F)) = Suc (card (image f F))" using insert by auto moreover have "card (f ` F) = card F" using insert by auto moreover have "card (insert x F) = Suc (card F)" using insert by auto ultimately show ?case by simp qed then show ?thesis using assms by simp qed lemma R_count : assumes "(vs @ xs) \ L M1 \ L M2" and "observable M1" and "observable M2" and "well_formed M1" and "well_formed M2" and "s \ nodes M2" and "productF M2 M1 FAIL PM" and "io_targets PM (initial PM) vs = {(q2,q1)}" and "path PM (xs || tr) (q2,q1)" and "length xs = length tr" and "distinct (states (xs || tr) (q2,q1))" shows "card (\ (image (io_targets M1 (initial M1)) (R M2 s vs xs))) = card (R M2 s vs xs)" \ \each sequence in the set calculated by R reaches a different state in M1\ proof - \ \Proof sketch: - states of PM reached by the sequences calculated by R can differ only in their second value - the sequences in the set calculated by R reach different states in PM due to distinctness\ have obs_PM : "observable PM" using observable_productF assms(2) assms(3) assms(7) by blast have state_component_2 : "\ io \ (R M2 s vs xs) . io_targets M2 (initial M2) io = {s}" proof fix io assume "io \ R M2 s vs xs" then have "s \ io_targets M2 (initial M2) io" by auto moreover have "io \ language_state M2 (initial M2)" using calculation by auto ultimately show "io_targets M2 (initial M2) io = {s}" using assms(3) io_targets_observable_singleton_ex by (metis singletonD) qed moreover have "\ io \ R M2 s vs xs . io_targets PM (initial PM) io = io_targets M2 (initial M2) io \ io_targets M1 (initial M1) io" proof fix io assume io_assm : "io \ R M2 s vs xs" then have io_prefix : "prefix io (vs @ xs)" by auto then have io_lang_subs : "io \ L M1 \ io \ L M2" using assms(1) unfolding prefix_def by (metis IntE language_state language_state_split) then have io_lang_inter : "io \ L M1 \ L M2" by simp then have io_lang_pm : "io \ L PM" using productF_language assms by blast moreover obtain p2 p1 where "(p2,p1) \ io_targets PM (initial PM) io" by (metis assms(2) assms(3) assms(7) calculation insert_absorb insert_ident insert_not_empty io_targets_observable_singleton_ob observable_productF singleton_insert_inj_eq subrelI) ultimately have targets_pm : "io_targets PM (initial PM) io = {(p2,p1)}" using assms io_targets_observable_singleton_ex singletonD by (metis observable_productF) then obtain trP where trP_def : "target (io || trP) (initial PM) = (p2,p1) \ path PM (io || trP) (initial PM) \ length io = length trP" proof - assume a1: "\trP. target (io || trP) (initial PM) = (p2, p1) \ path PM (io || trP) (initial PM) \ length io = length trP \ thesis" have "\ps. target (io || ps) (initial PM) = (p2, p1) \ path PM (io || ps) (initial PM) \ length io = length ps" using \(p2, p1) \ io_targets PM (initial PM) io\ by auto then show ?thesis using a1 by blast qed then have trP_unique : "{ tr . path PM (io || tr) (initial PM) \ length io = length tr } = { trP }" using observable_productF observable_path_unique_ex[of PM io "initial PM"] io_lang_pm assms(2) assms(3) assms(7) proof - obtain pps :: "('d \ 'c) list" where f1: "{ps. path PM (io || ps) (initial PM) \ length io = length ps} = {pps} \ \ observable PM" by (metis (no_types) \\thesis. \observable PM; io \ L PM; \tr. {t. path PM (io || t) (initial PM) \ length io = length t} = {tr} \ thesis\ \ thesis\ io_lang_pm) have f2: "observable PM" by (meson \observable M1\ \observable M2\ \productF M2 M1 FAIL PM\ observable_productF) then have "trP \ {pps}" using f1 trP_def by blast then show ?thesis using f2 f1 by force qed obtain trIO2 where trIO2_def : "{tr . path M2 (io||tr) (initial M2) \ length io = length tr} = { trIO2 }" using observable_path_unique_ex[of M2 io "initial M2"] io_lang_subs assms(3) by blast obtain trIO1 where trIO1_def : "{tr . path M1 (io||tr) (initial M1) \ length io = length tr} = { trIO1 }" using observable_path_unique_ex[of M1 io "initial M1"] io_lang_subs assms(2) by blast have "path PM (io || trIO2 || trIO1) (initial M2, initial M1) \ length io = length trIO2 \ length trIO2 = length trIO1" proof - have f1: "path M2 (io || trIO2) (initial M2) \ length io = length trIO2" using trIO2_def by auto have f2: "path M1 (io || trIO1) (initial M1) \ length io = length trIO1" using trIO1_def by auto then have "length trIO2 = length trIO1" using f1 by presburger then show ?thesis using f2 f1 assms(4) assms(5) assms(7) by blast qed then have trP_split : "path PM (io || trIO2 || trIO1) (initial PM) \ length io = length trIO2 \ length trIO2 = length trIO1" using assms(7) by auto then have trP_zip : "trIO2 || trIO1 = trP" using trP_def trP_unique using length_zip by fastforce have "target (io || trIO2) (initial M2) = p2 \ path M2 (io || trIO2) (initial M2) \ length io = length trIO2" using trP_zip trP_split assms(7) trP_def trIO2_def by auto then have "p2 \ io_targets M2 (initial M2) io" by auto then have targets_2 : "io_targets M2 (initial M2) io = {p2}" by (metis state_component_2 io_assm singletonD) have "target (io || trIO1) (initial M1) = p1 \ path M1 (io || trIO1) (initial M1) \ length io = length trIO1" using trP_zip trP_split assms(7) trP_def trIO1_def by auto then have "p1 \ io_targets M1 (initial M1) io" by auto then have targets_1 : "io_targets M1 (initial M1) io = {p1}" by (metis io_lang_subs assms(2) io_targets_observable_singleton_ex singletonD) have "io_targets M2 (initial M2) io \ io_targets M1 (initial M1) io = {(p2,p1)}" using targets_2 targets_1 by simp then show "io_targets PM (initial PM) io = io_targets M2 (initial M2) io \ io_targets M1 (initial M1) io" using targets_pm by simp qed ultimately have state_components : "\ io \ R M2 s vs xs . io_targets PM (initial PM) io = {s} \ io_targets M1 (initial M1) io" by auto then have "\ (image (io_targets PM (initial PM)) (R M2 s vs xs)) = \ (image (\ io . {s} \ io_targets M1 (initial M1) io) (R M2 s vs xs))" by auto then have "\ (image (io_targets PM (initial PM)) (R M2 s vs xs)) = {s} \ \ (image (io_targets M1 (initial M1)) (R M2 s vs xs))" by auto then have "card (\ (image (io_targets PM (initial PM)) (R M2 s vs xs))) = card (\ (image (io_targets M1 (initial M1)) (R M2 s vs xs)))" by (metis (no_types) card_cartesian_product_singleton) moreover have "card (\ (image (io_targets PM (initial PM)) (R M2 s vs xs))) = card (R M2 s vs xs)" proof (rule ccontr) assume assm : "card (\ (io_targets PM (initial PM) ` R M2 s vs xs) ) \ card (R M2 s vs xs)" have "\ io \ R M2 s vs xs . io \ L PM" proof fix io assume io_assm : "io \ R M2 s vs xs" then have "prefix io (vs @ xs)" by auto then have "io \ L M1 \ io \ L M2" using assms(1) unfolding prefix_def by (metis IntE language_state language_state_split) then show "io \ L PM" using productF_language assms by blast qed then have singletons : "\ io \ R M2 s vs xs . (\ t . io_targets PM (initial PM) io = {t})" using io_targets_observable_singleton_ex observable_productF assms by metis then have card_targets : "card (\(io_targets PM (initial PM) ` R M2 s vs xs)) = card (image (io_targets PM (initial PM)) (R M2 s vs xs))" using finite_R card_union_of_singletons [of "image (io_targets PM (initial PM)) (R M2 s vs xs)"] by simp moreover have "card (image (io_targets PM (initial PM)) (R M2 s vs xs)) \ card (R M2 s vs xs)" using finite_R by (metis card_image_le) ultimately have card_le : "card (\(io_targets PM (initial PM) ` R M2 s vs xs)) < card (R M2 s vs xs)" using assm by linarith have "\ io1 \ (R M2 s vs xs) . \ io2 \ (R M2 s vs xs) . io1 \ io2 \ io_targets PM (initial PM) io1 \ io_targets PM (initial PM) io2 \ {}" proof (rule ccontr) assume "\ (\io1\R M2 s vs xs. \io2\R M2 s vs xs. io1 \ io2 \ io_targets PM (initial PM) io1 \ io_targets PM (initial PM) io2 \ {})" then have "\io1\R M2 s vs xs. \io2\R M2 s vs xs. io1 = io2 \ io_targets PM (initial PM) io1 \ io_targets PM (initial PM) io2 = {}" by blast moreover have "\io\R M2 s vs xs. io_targets PM (initial PM) io \ {}" by (metis insert_not_empty singletons) ultimately have "card (image (io_targets PM (initial PM)) (R M2 s vs xs)) = card (R M2 s vs xs)" using finite_R[of M2 s vs xs] card_union_of_distinct [of "R M2 s vs xs" "(io_targets PM (initial PM))"] by blast then show "False" using card_le card_targets by linarith qed then have "\ io1 io2 . io1 \ (R M2 s vs xs) \ io2 \ (R M2 s vs xs) \ io1 \ io2 \ io_targets PM (initial PM) io1 \ io_targets PM (initial PM) io2 \ {}" by blast moreover have "\ io1 io2 . (io1 \ (R M2 s vs xs) \ io2 \ (R M2 s vs xs) \ io1 \ io2) \ length io1 \ length io2" proof (rule ccontr) assume " \ (\io1 io2. io1 \ R M2 s vs xs \ io2 \ R M2 s vs xs \ io1 \ io2 \ length io1 \ length io2)" then obtain io1 io2 where io_def : "io1 \ R M2 s vs xs \ io2 \ R M2 s vs xs \ io1 \ io2 \ length io1 = length io2" by auto then have "prefix io1 (vs @ xs) \ prefix io2 (vs @ xs)" by auto then have "io1 = take (length io1) (vs @ xs) \ io2 = take (length io2) (vs @ xs)" by (metis append_eq_conv_conj prefixE) then show "False" using io_def by auto qed ultimately obtain io1 io2 where rep_ios_def : "io1 \ (R M2 s vs xs) \ io2 \ (R M2 s vs xs) \ length io1 < length io2 \ io_targets PM (initial PM) io1 \ io_targets PM (initial PM) io2 \ {}" by (metis inf_sup_aci(1) linorder_neqE_nat) obtain rep where "(s,rep) \ io_targets PM (initial PM) io1 \ io_targets PM (initial PM) io2" proof - assume a1: "\rep. (s, rep) \ io_targets PM (initial PM) io1 \ io_targets PM (initial PM) io2 \ thesis" have "\f. Sigma {s} f \ (io_targets PM (initial PM) io1 \ io_targets PM (initial PM) io2) \ {}" by (metis (no_types) inf.left_idem rep_ios_def state_components) then show ?thesis using a1 by blast qed then have rep_state : "io_targets PM (initial PM) io1 = {(s,rep)} \ io_targets PM (initial PM) io2 = {(s,rep)}" by (metis Int_iff rep_ios_def singletonD singletons) obtain io1X io2X where rep_ios_split : "io1 = vs @ io1X \ prefix io1X xs \ io2 = vs @ io2X \ prefix io2X xs" using rep_ios_def by auto then have "length io1 > length vs" using rep_ios_def by auto \ \get a path from (initial PM) to (q2,q1)\ have "vs@xs \ L PM" by (metis (no_types) assms(1) assms(4) assms(5) assms(7) inf_commute productF_language) then have "vs \ L PM" by (meson language_state_prefix) then obtain trV where trV_def : "{tr . path PM (vs || tr) (initial PM) \ length vs = length tr} = { trV }" using observable_path_unique_ex[of PM vs "initial PM"] assms(2) assms(3) assms(7) observable_productF by blast let ?qv = "target (vs || trV) (initial PM)" have "?qv \ io_targets PM (initial PM) vs" using trV_def by auto then have qv_simp[simp] : "?qv = (q2,q1)" using singletons assms by blast then have "?qv \ nodes PM" using trV_def assms by blast \ \get a path using io1X from the state reached by vs in PM\ obtain tr1X_all where tr1X_all_def : "path PM (vs @ io1X || tr1X_all) (initial PM) \ length (vs @ io1X) = length tr1X_all" using rep_ios_def rep_ios_split by auto let ?tr1X = "drop (length vs) tr1X_all" have "take (length vs) tr1X_all = trV" proof - have "path PM (vs || take (length vs) tr1X_all) (initial PM) \ length vs = length (take (length vs) tr1X_all)" using tr1X_all_def trV_def by (metis (no_types, lifting) FSM.path_append_elim append_eq_conv_conj length_take zip_append1) then show "take (length vs) tr1X_all = trV" using trV_def by blast qed then have tr1X_def : "path PM (io1X || ?tr1X) ?qv \ length io1X = length ?tr1X" proof - have "length tr1X_all = length vs + length io1X" using tr1X_all_def by auto then have "length io1X = length tr1X_all - length vs" by presburger then show ?thesis by (metis (no_types) FSM.path_append_elim \take (length vs) tr1X_all = trV\ length_drop tr1X_all_def zip_append1) qed then have io1X_lang : "io1X \ language_state PM ?qv" by auto then obtain tr1X' where tr1X'_def : "{tr . path PM (io1X || tr) ?qv \ length io1X = length tr} = { tr1X' }" using observable_path_unique_ex[of PM io1X ?qv] assms(2) assms(3) assms(7) observable_productF by blast moreover have "?tr1X \ { tr . path PM (io1X || tr) ?qv \ length io1X = length tr }" using tr1X_def by auto ultimately have tr1x_unique : "tr1X' = ?tr1X" by simp \ \get a path using io2X from the state reached by vs in PM\ obtain tr2X_all where tr2X_all_def : "path PM (vs @ io2X || tr2X_all) (initial PM) \ length (vs @ io2X) = length tr2X_all" using rep_ios_def rep_ios_split by auto let ?tr2X = "drop (length vs) tr2X_all" have "take (length vs) tr2X_all = trV" proof - have "path PM (vs || take (length vs) tr2X_all) (initial PM) \ length vs = length (take (length vs) tr2X_all)" using tr2X_all_def trV_def by (metis (no_types, lifting) FSM.path_append_elim append_eq_conv_conj length_take zip_append1) then show "take (length vs) tr2X_all = trV" using trV_def by blast qed then have tr2X_def : "path PM (io2X || ?tr2X) ?qv \ length io2X = length ?tr2X" proof - have "length tr2X_all = length vs + length io2X" using tr2X_all_def by auto then have "length io2X = length tr2X_all - length vs" by presburger then show ?thesis by (metis (no_types) FSM.path_append_elim \take (length vs) tr2X_all = trV\ length_drop tr2X_all_def zip_append1) qed then have io2X_lang : "io2X \ language_state PM ?qv" by auto then obtain tr2X' where tr2X'_def : "{tr . path PM (io2X || tr) ?qv \ length io2X = length tr} = { tr2X' }" using observable_path_unique_ex[of PM io2X ?qv] assms(2) assms(3) assms(7) observable_productF by blast moreover have "?tr2X \ { tr . path PM (io2X || tr) ?qv \ length io2X = length tr }" using tr2X_def by auto ultimately have tr2x_unique : "tr2X' = ?tr2X" by simp \ \both paths reach the same state\ have "io_targets PM (initial PM) (vs @ io1X) = {(s,rep)}" using rep_state rep_ios_split by auto moreover have "io_targets PM (initial PM) vs = {?qv}" using assms(8) by auto ultimately have rep_via_1 : "io_targets PM ?qv io1X = {(s,rep)}" by (meson obs_PM observable_io_targets_split) then have rep_tgt_1 : "target (io1X || tr1X') ?qv = (s,rep)" using obs_PM observable_io_target_unique_target[of PM ?qv io1X "(s,rep)"] tr1X'_def by blast have length_1 : "length (io1X || tr1X') > 0" using \length vs < length io1\ rep_ios_split tr1X_def tr1x_unique by auto have tr1X_alt_def : "tr1X' = take (length io1X) tr" by (metis (no_types) assms(10) assms(9) obs_PM observable_path_prefix qv_simp rep_ios_split tr1X_def tr1x_unique) moreover have "io1X = take (length io1X) xs" using rep_ios_split by (metis append_eq_conv_conj prefixE) ultimately have "(io1X || tr1X') = take (length io1X) (xs || tr)" by (metis take_zip) moreover have "length (xs || tr) \ length (io1X || tr1X')" by (metis (no_types) \io1X = take (length io1X) xs\ assms(10) length_take length_zip nat_le_linear take_all tr1X_def tr1x_unique) ultimately have rep_idx_1 : "(states (xs || tr) ?qv) ! ((length io1X) - 1) = (s,rep)" by (metis (no_types, lifting) One_nat_def Suc_less_eq Suc_pred rep_tgt_1 length_1 less_Suc_eq_le map_snd_zip scan_length scan_nth states_alt_def tr1X_def tr1x_unique) have "io_targets PM (initial PM) (vs @ io2X) = {(s,rep)}" using rep_state rep_ios_split by auto moreover have "io_targets PM (initial PM) vs = {?qv}" using assms(8) by auto ultimately have rep_via_2 : "io_targets PM ?qv io2X = {(s,rep)}" by (meson obs_PM observable_io_targets_split) then have rep_tgt_2 : "target (io2X || tr2X') ?qv = (s,rep)" using obs_PM observable_io_target_unique_target[of PM ?qv io2X "(s,rep)"] tr2X'_def by blast moreover have length_2 : "length (io2X || tr2X') > 0" by (metis \length vs < length io1\ append.right_neutral length_0_conv length_zip less_asym min.idem neq0_conv rep_ios_def rep_ios_split tr2X_def tr2x_unique) have tr2X_alt_def : "tr2X' = take (length io2X) tr" by (metis (no_types) assms(10) assms(9) obs_PM observable_path_prefix qv_simp rep_ios_split tr2X_def tr2x_unique) moreover have "io2X = take (length io2X) xs" using rep_ios_split by (metis append_eq_conv_conj prefixE) ultimately have "(io2X || tr2X') = take (length io2X) (xs || tr)" by (metis take_zip) moreover have "length (xs || tr) \ length (io2X || tr2X')" using calculation by auto ultimately have rep_idx_2 : "(states (xs || tr) ?qv) ! ((length io2X) - 1) = (s,rep)" by (metis (no_types, lifting) One_nat_def Suc_less_eq Suc_pred rep_tgt_2 length_2 less_Suc_eq_le map_snd_zip scan_length scan_nth states_alt_def tr2X_def tr2x_unique) \ \thus the distinctness assumption is violated\ have "length io1X \ length io2X" by (metis \io1X = take (length io1X) xs\ \io2X = take (length io2X) xs\ less_irrefl rep_ios_def rep_ios_split) moreover have "(states (xs || tr) ?qv) ! ((length io1X) - 1) = (states (xs || tr) ?qv) ! ((length io2X) - 1)" using rep_idx_1 rep_idx_2 by simp ultimately have "\ (distinct (states (xs || tr) ?qv))" by (metis Suc_less_eq \io1X = take (length io1X) xs\ \io1X || tr1X' = take (length io1X) (xs || tr)\ \io2X = take (length io2X) xs\ \io2X || tr2X' = take (length io2X) (xs || tr)\ \length (io1X || tr1X') \ length (xs || tr)\ \length (io2X || tr2X') \ length (xs || tr)\ assms(10) diff_Suc_1 distinct_conv_nth gr0_conv_Suc le_imp_less_Suc length_1 length_2 length_take map_snd_zip scan_length states_alt_def) then show "False" by (metis assms(11) states_alt_def) qed ultimately show ?thesis by linarith qed lemma R_state_component_2 : assumes "io \ (R M2 s vs xs)" and "observable M2" shows "io_targets M2 (initial M2) io = {s}" proof - have "s \ io_targets M2 (initial M2) io" using assms(1) by auto moreover have "io \ language_state M2 (initial M2)" using calculation by auto ultimately show "io_targets M2 (initial M2) io = {s}" using assms(2) io_targets_observable_singleton_ex by (metis singletonD) qed lemma R_union_card_is_suffix_length : assumes "OFSM M2" and "io@xs \ L M2" shows "sum (\ q . card (R M2 q io xs)) (nodes M2) = length xs" using assms proof (induction xs rule: rev_induct) case Nil show ?case by (simp add: sum.neutral) next case (snoc x xs) have "finite (nodes M2)" using assms by auto have R_update : "\ q . R M2 q io (xs@[x]) = (if (q \ io_targets M2 (initial M2) (io @ xs @ [x])) then insert (io@xs@[x]) (R M2 q io xs) else R M2 q io xs)" by auto obtain q where "io_targets M2 (initial M2) (io @ xs @ [x]) = {q}" by (meson assms(1) io_targets_observable_singleton_ex snoc.prems(2)) then have "R M2 q io (xs@[x]) = insert (io@xs@[x]) (R M2 q io xs)" using R_update by auto moreover have "(io@xs@[x]) \ (R M2 q io xs)" by auto ultimately have "card (R M2 q io (xs@[x])) = Suc (card (R M2 q io xs))" by (metis card_insert_disjoint finite_R) have "q \ nodes M2" by (metis (full_types) FSM.nodes.initial \io_targets M2 (initial M2) (io@xs @ [x]) = {q}\ insertI1 io_targets_nodes) have "\ q' . q' \ q \ R M2 q' io (xs@[x]) = R M2 q' io xs" using \io_targets M2 (initial M2) (io@xs @ [x]) = {q}\ R_update by auto then have "\ q' . q' \ q \ card (R M2 q' io (xs@[x])) = card (R M2 q' io xs)" by auto then have "(\q\(nodes M2 - {q}). card (R M2 q io (xs@[x]))) = (\q\(nodes M2 - {q}). card (R M2 q io xs))" by auto moreover have "(\q\nodes M2. card (R M2 q io (xs@[x]))) = (\q\(nodes M2 - {q}). card (R M2 q io (xs@[x]))) + (card (R M2 q io (xs@[x])))" "(\q\nodes M2. card (R M2 q io xs)) = (\q\(nodes M2 - {q}). card (R M2 q io xs)) + (card (R M2 q io xs))" proof - have "\C c f. (infinite C \ (c::'c) \ C) \ sum f C = (f c::nat) + sum f (C - {c})" by (meson sum.remove) then show "(\q\nodes M2. card (R M2 q io (xs@[x]))) = (\q\(nodes M2 - {q}). card (R M2 q io (xs@[x]))) + (card (R M2 q io (xs@[x])))" "(\q\nodes M2. card (R M2 q io xs)) = (\q\(nodes M2 - {q}). card (R M2 q io xs)) + (card (R M2 q io xs))" using \finite (nodes M2)\ \q \ nodes M2\ by presburger+ qed ultimately have "(\q\nodes M2. card (R M2 q io (xs@[x]))) = Suc (\q\nodes M2. card (R M2 q io xs))" using \card (R M2 q io (xs@[x])) = Suc (card (R M2 q io xs))\ by presburger have "(\q\nodes M2. card (R M2 q io xs)) = length xs" using snoc.IH snoc.prems language_state_prefix[of "io@xs" "[x]" M2 "initial M2"] proof - show ?thesis by (metis (no_types) \(io @ xs) @ [x] \ L M2 \ io @ xs \ L M2\ \OFSM M2\ \io @ xs @ [x] \ L M2\ append.assoc snoc.IH) qed show ?case proof - show ?thesis by (metis (no_types) \(\q\nodes M2. card (R M2 q io (xs @ [x]))) = Suc (\q\nodes M2. card (R M2 q io xs))\ \(\q\nodes M2. card (R M2 q io xs)) = length xs\ length_append_singleton) qed qed lemma R_state_repetition_via_long_sequence : assumes "OFSM M" and "card (nodes M) \ m" and "Suc (m * m) \ length xs" and "vs@xs \ L M" shows "\ q \ nodes M . card (R M q vs xs) > m" proof (rule ccontr) assume "\ (\q\nodes M. m < card (R M q vs xs))" then have "\ q \ nodes M . card (R M q vs xs) \ m" by auto then have "sum (\ q . card (R M q vs xs)) (nodes M) \ sum (\ q . m) (nodes M)" by (meson sum_mono) moreover have "sum (\ q . m) (nodes M) \ m * m" using assms(2) by auto ultimately have "sum (\ q . card (R M q vs xs)) (nodes M) \ m * m" by presburger moreover have "Suc (m*m) \ sum (\ q . card (R M q vs xs)) (nodes M)" using R_union_card_is_suffix_length[OF assms(1), of vs xs] assms(4,3) by auto ultimately show "False" by simp qed lemma R_state_repetition_distribution : assumes "OFSM M" and "Suc (card (nodes M) * m) \ length xs" and "vs@xs \ L M" shows "\ q \ nodes M . card (R M q vs xs) > m" proof (rule ccontr) assume "\ (\q\nodes M. m < card (R M q vs xs))" then have "\ q \ nodes M . card (R M q vs xs) \ m" by auto then have "sum (\ q . card (R M q vs xs)) (nodes M) \ sum (\ q . m) (nodes M)" by (meson sum_mono) moreover have "sum (\ q . m) (nodes M) \ card (nodes M) * m" using assms(2) by auto ultimately have "sum (\ q . card (R M q vs xs)) (nodes M) \ card (nodes M) * m" by presburger moreover have "Suc (card (nodes M)*m) \ sum (\ q . card (R M q vs xs)) (nodes M)" using R_union_card_is_suffix_length[OF assms(1), of vs xs] assms(3,2) by auto ultimately show "False" by simp qed subsection \ Function RP \ text \ Function @{verbatim RP} extends function @{verbatim MR} by adding all elements from a set of IO-sequences that also reach the given state. \ fun RP :: "('in, 'out, 'state) FSM \ 'state \ ('in \ 'out) list \ ('in \ 'out) list \ ('in \ 'out) list set \ ('in \ 'out) list set" where "RP M s vs xs V'' = R M s vs xs \ {vs' \ V'' . io_targets M (initial M) vs' = {s}}" lemma RP_from_R: assumes "is_det_state_cover M2 V" and "V'' \ Perm V M1" shows "RP M2 s vs xs V'' = R M2 s vs xs \ (\ vs' \ V'' . vs' \ R M2 s vs xs \ RP M2 s vs xs V'' = insert vs' (R M2 s vs xs))" proof (rule ccontr) assume assm : "\ (RP M2 s vs xs V'' = R M2 s vs xs \ (\vs'\V''. vs' \ R M2 s vs xs \ RP M2 s vs xs V'' = insert vs' (R M2 s vs xs)))" moreover have "R M2 s vs xs \ RP M2 s vs xs V''" by simp moreover have "RP M2 s vs xs V'' \ R M2 s vs xs \ V''" by auto ultimately obtain vs1 vs2 where vs_def : "vs1 \ vs2 \ vs1 \ V'' \ vs2 \ V'' \ vs1 \ R M2 s vs xs \ vs2 \ R M2 s vs xs \ vs1 \ RP M2 s vs xs V'' \ vs2 \ RP M2 s vs xs V''" by blast then have "io_targets M2 (initial M2) vs1 = {s} \ io_targets M2 (initial M2) vs2 = {s}" by (metis (mono_tags, lifting) RP.simps Un_iff mem_Collect_eq) then have "io_targets M2 (initial M2) vs1 = io_targets M2 (initial M2) vs2" by simp obtain f where f_def : "is_det_state_cover_ass M2 f \ V = f ` d_reachable M2 (initial M2)" using assms by auto moreover have "V = image f (d_reachable M2 (initial M2))" using f_def by blast moreover have "map fst vs1 \ V \ map fst vs2 \ V" using assms(2) perm_inputs vs_def by blast ultimately obtain r1 r2 where r_def : "f r1 = map fst vs1 \ r1 \ d_reachable M2 (initial M2)" "f r2 = map fst vs2 \ r2 \ d_reachable M2 (initial M2)" by force then have "d_reaches M2 (initial M2) (map fst vs1) r1" "d_reaches M2 (initial M2) (map fst vs2) r2" by (metis f_def is_det_state_cover_ass.elims(2))+ then have "io_targets M2 (initial M2) vs1 \ {r1}" using d_reaches_io_target[of M2 "initial M2" "map fst vs1" r1 "map snd vs1"] by simp moreover have "io_targets M2 (initial M2) vs2 \ {r2}" using d_reaches_io_target[of M2 "initial M2" "map fst vs2" r2 "map snd vs2"] \d_reaches M2 (initial M2) (map fst vs2) r2\ by auto ultimately have "r1 = r2" using \io_targets M2 (initial M2) vs1 = {s} \ io_targets M2 (initial M2) vs2 = {s}\ by auto have "map fst vs1 \ map fst vs2" using assms(2) perm_inputs_diff vs_def by blast then have "r1 \ r2" using r_def(1) r_def(2) by force then show "False" using \r1 = r2\ by auto qed lemma finite_RP : assumes "is_det_state_cover M2 V" and "V'' \ Perm V M1" shows "finite (RP M2 s vs xs V'')" using assms RP_from_R finite_R by (metis finite_insert) lemma RP_count : assumes "(vs @ xs) \ L M1 \ L M2" and "observable M1" and "observable M2" and "well_formed M1" and "well_formed M2" and "s \ nodes M2" and "productF M2 M1 FAIL PM" and "io_targets PM (initial PM) vs = {(q2,q1)}" and "path PM (xs || tr) (q2,q1)" and "length xs = length tr" and "distinct (states (xs || tr) (q2,q1))" and "is_det_state_cover M2 V" and "V'' \ Perm V M1" and "\ s' \ set (states (xs || map fst tr) q2) . \ (\ v \ V . d_reaches M2 (initial M2) v s')" shows "card (\ (image (io_targets M1 (initial M1)) (RP M2 s vs xs V''))) = card (RP M2 s vs xs V'')" \ \each sequence in the set calculated by RP reaches a different state in M1\ proof - \ \Proof sketch: - RP calculates either the same set as R or the set of R and an additional element - in the first case, the result for R applies - in the second case, the additional element is not contained in the set calcualted by R due to the assumption that no state reached by a non-empty prefix of xs after vs is also reached by some sequence in V (see the last two assumptions)\ have RP_cases : "RP M2 s vs xs V'' = R M2 s vs xs \ (\ vs' \ V'' . vs' \ R M2 s vs xs \ RP M2 s vs xs V'' = insert vs' (R M2 s vs xs))" using RP_from_R assms by metis show ?thesis proof (cases "RP M2 s vs xs V'' = R M2 s vs xs") case True then show ?thesis using R_count assms by metis next case False then obtain vs' where vs'_def : "vs' \ V'' \ vs' \ R M2 s vs xs \ RP M2 s vs xs V'' = insert vs' (R M2 s vs xs)" using RP_cases by auto have obs_PM : "observable PM" using observable_productF assms(2) assms(3) assms(7) by blast have state_component_2 : "\ io \ (R M2 s vs xs) . io_targets M2 (initial M2) io = {s}" proof fix io assume "io \ R M2 s vs xs" then have "s \ io_targets M2 (initial M2) io" by auto moreover have "io \ language_state M2 (initial M2)" using calculation by auto ultimately show "io_targets M2 (initial M2) io = {s}" using assms(3) io_targets_observable_singleton_ex by (metis singletonD) qed have "vs' \ L M1" using assms(13) perm_language vs'_def by blast then obtain s' where s'_def : "io_targets M1 (initial M1) vs' = {s'}" by (meson assms(2) io_targets_observable_singleton_ob) moreover have "s' \ \ (image (io_targets M1 (initial M1)) (R M2 s vs xs))" proof (rule ccontr) assume "\ s' \ \(io_targets M1 (initial M1) ` R M2 s vs xs)" then obtain xs' where xs'_def : "vs @ xs' \ R M2 s vs xs \ s' \ io_targets M1 (initial M1) (vs @ xs')" proof - assume a1: "\xs'. vs @ xs' \ R M2 s vs xs \ s' \ io_targets M1 (initial M1) (vs @ xs') \ thesis" obtain pps :: "('a \ 'b) list set \ (('a \ 'b) list \ 'c set) \ 'c \ ('a \ 'b) list" where "\x0 x1 x2. (\v3. v3 \ x0 \ x2 \ x1 v3) = (pps x0 x1 x2 \ x0 \ x2 \ x1 (pps x0 x1 x2))" by moura then have f2: "pps (R M2 s vs xs) (io_targets M1 (initial M1)) s' \ R M2 s vs xs \ s' \ io_targets M1 (initial M1) (pps (R M2 s vs xs) (io_targets M1 (initial M1)) s')" using \\ s' \ \(io_targets M1 (initial M1) ` R M2 s vs xs)\ by blast then have "\ps. pps (R M2 s vs xs) (io_targets M1 (initial M1)) s' = vs @ ps \ ps \ [] \ prefix ps xs \ s \ io_targets M2 (initial M2) (vs @ ps)" by simp then show ?thesis using f2 a1 by (metis (no_types)) qed then obtain tr' where tr'_def : "path M2 (vs @ xs' || tr') (initial M2) \ length tr' = length (vs @ xs')" by auto then obtain trV' trX' where tr'_split : "trV' = take (length vs) tr'" "trX' = drop (length vs) tr'" "tr' = trV' @ trX'" by fastforce then have "path M2 (vs || trV') (initial M2) \ length trV' = length vs" by (metis (no_types) FSM.path_append_elim \trV' = take (length vs) tr'\ append_eq_conv_conj length_take tr'_def zip_append1) have "initial PM = (initial M2, initial M1)" using assms(7) by simp moreover have "vs \ L M2" "vs \ L M1" using assms(1) language_state_prefix by auto ultimately have "io_targets M1 (initial M1) vs = {q1}" "io_targets M2 (initial M2) vs = {q2}" using productF_path_io_targets[of M2 M1 FAIL PM "initial M2" "initial M1" vs q2 q1] by (metis FSM.nodes.initial assms(7) assms(8) assms(2) assms(3) assms(4) assms(5) io_targets_observable_singleton_ex singletonD)+ then have "target (vs || trV') (initial M2) = q2" using \path M2 (vs || trV') (initial M2) \ length trV' = length vs\ io_target_target by metis then have path_xs' : "path M2 (xs' || trX') q2 \ length trX' = length xs'" by (metis (no_types) FSM.path_append_elim \path M2 (vs || trV') (initial M2) \ length trV' = length vs\ \target (vs || trV') (initial M2) = q2\ append_eq_conv_conj length_drop tr'_def tr'_split(1) tr'_split(2) zip_append2) have "io_targets M2 (initial M2) (vs @ xs') = {s}" using state_component_2 xs'_def by blast then have "io_targets M2 q2 xs' = {s}" by (meson assms(3) observable_io_targets_split \io_targets M2 (initial M2) vs = {q2}\) then have target_xs' : "target (xs' || trX') q2 = s" using io_target_target path_xs' by metis moreover have "length xs' > 0" using xs'_def by auto ultimately have "last (states (xs' || trX') q2) = s" using path_xs' target_in_states by metis moreover have "length (states (xs' || trX') q2) > 0" using \0 < length xs'\ path_xs' by auto ultimately have states_xs' : "s \ set (states (xs' || trX') q2)" using last_in_set by blast have "vs @ xs \ L M2" using assms by simp then obtain q' where "io_targets M2 (initial M2) (vs@xs) = {q'}" using io_targets_observable_singleton_ob[of M2 "vs@xs" "initial M2"] assms(3) by auto then have "xs \ language_state M2 q2" using assms(3) \io_targets M2 (initial M2) vs = {q2}\ observable_io_targets_split[of M2 "initial M2" vs xs q' q2] by auto moreover have "path PM (xs || map fst tr || map snd tr) (q2,q1) \ length xs = length (map fst tr)" using assms(7) assms(9) assms(10) productF_path_unzip by simp moreover have "xs \ language_state PM (q2,q1)" using assms(9) assms(10) by auto moreover have "q2 \ nodes M2" using \io_targets M2 (initial M2) vs = {q2}\ io_targets_nodes by (metis FSM.nodes.initial insertI1) moreover have "q1 \ nodes M1" using \io_targets M1 (initial M1) vs = {q1}\ io_targets_nodes by (metis FSM.nodes.initial insertI1) ultimately have path_xs : "path M2 (xs || map fst tr) q2" using productF_path_reverse_ob_2(1)[of xs "map fst tr" "map snd tr" M2 M1 FAIL PM q2 q1] assms(2,3,4,5,7) by simp moreover have "prefix xs' xs" using xs'_def by auto ultimately have "trX' = take (length xs') (map fst tr)" using \path PM (xs || map fst tr || map snd tr) (q2, q1) \ length xs = length (map fst tr)\ assms(3) path_xs' by (metis observable_path_prefix) then have states_xs : "s \ set (states (xs || map fst tr) q2)" by (metis assms(10) in_set_takeD length_map map_snd_zip path_xs' states_alt_def states_xs') have "d_reaches M2 (initial M2) (map fst vs') s" proof - obtain fV where fV_def : "is_det_state_cover_ass M2 fV \ V = fV ` d_reachable M2 (initial M2)" using assms(12) by auto moreover have "V = image fV (d_reachable M2 (initial M2))" using fV_def by blast moreover have "map fst vs' \ V" using perm_inputs vs'_def assms(13) by metis ultimately obtain qv where qv_def : "fV qv = map fst vs' \ qv\ d_reachable M2 (initial M2)" by force then have "d_reaches M2 (initial M2) (map fst vs') qv" by (metis fV_def is_det_state_cover_ass.elims(2)) then have "io_targets M2 (initial M2) vs' \ {qv}" using d_reaches_io_target[of M2 "initial M2" "map fst vs'" qv "map snd vs'"] by simp moreover have "io_targets M2 (initial M2) vs' = {s}" using vs'_def by (metis (mono_tags, lifting) RP.simps Un_iff insertI1 mem_Collect_eq) ultimately have "qv = s" by simp then show ?thesis using \d_reaches M2 (initial M2) (map fst vs') qv\ by blast qed then show "False" by (meson assms(14) assms(13) perm_inputs states_xs vs'_def) qed moreover have "\ (image (io_targets M1 (initial M1)) (insert vs' (R M2 s vs xs))) = insert s' (\ (image (io_targets M1 (initial M1)) (R M2 s vs xs)))" using s'_def by simp moreover have "finite (\ (image (io_targets M1 (initial M1)) (R M2 s vs xs)))" proof show "finite (R M2 s vs xs)" using finite_R by simp show "\a. a \ R M2 s vs xs \ finite (io_targets M1 (initial M1) a)" proof - fix a assume "a \ R M2 s vs xs" then have "prefix a (vs@xs)" by auto then have "a \ L M1" using language_state_prefix by (metis IntD1 assms(1) prefix_def) then obtain p where "io_targets M1 (initial M1) a = {p}" using assms(2) io_targets_observable_singleton_ob by metis then show "finite (io_targets M1 (initial M1) a)" by simp qed qed ultimately have "card (\ (image (io_targets M1 (initial M1)) (insert vs' (R M2 s vs xs)))) = Suc (card (\ (image (io_targets M1 (initial M1)) (R M2 s vs xs))))" by (metis (no_types) card_insert_disjoint) moreover have "card (\ (image (io_targets M1 (initial M1)) (RP M2 s vs xs V''))) = card (\ (image (io_targets M1 (initial M1)) (insert vs' (R M2 s vs xs))))" using vs'_def by simp ultimately have "card (\ (image (io_targets M1 (initial M1)) (RP M2 s vs xs V''))) = Suc (card (\ (image (io_targets M1 (initial M1)) (R M2 s vs xs))))" by linarith then have "card (\ (image (io_targets M1 (initial M1)) (RP M2 s vs xs V''))) = Suc (card (R M2 s vs xs))" using R_count[of vs xs M1 M2 s FAIL PM q2 q1 tr] assms(1,10,11,2-9) by linarith moreover have "card (RP M2 s vs xs V'') = Suc (card (R M2 s vs xs))" using vs'_def by (metis card_insert_if finite_R) ultimately show ?thesis by linarith qed qed lemma RP_state_component_2 : assumes "io \ (RP M2 s vs xs V'')" and "observable M2" shows "io_targets M2 (initial M2) io = {s}" by (metis (mono_tags, lifting) RP.simps R_state_component_2 Un_iff assms mem_Collect_eq) lemma RP_io_targets_split : assumes "(vs @ xs) \ L M1 \ L M2" and "observable M1" and "observable M2" and "well_formed M1" and "well_formed M2" and "productF M2 M1 FAIL PM" and "is_det_state_cover M2 V" and "V'' \ Perm V M1" and "io \ RP M2 s vs xs V''" shows "io_targets PM (initial PM) io = io_targets M2 (initial M2) io \ io_targets M1 (initial M1) io" proof - have RP_cases : "RP M2 s vs xs V'' = R M2 s vs xs \ (\ vs' \ V'' . vs' \ R M2 s vs xs \ RP M2 s vs xs V'' = insert vs' (R M2 s vs xs))" using RP_from_R assms by metis show ?thesis proof (cases "io \ R M2 s vs xs") case True then have io_prefix : "prefix io (vs @ xs)" by auto then have io_lang_subs : "io \ L M1 \ io \ L M2" using assms(1) unfolding prefix_def by (metis IntE language_state language_state_split) then have io_lang_inter : "io \ L M1 \ L M2" by simp then have io_lang_pm : "io \ L PM" using productF_language assms by blast moreover obtain p2 p1 where "(p2,p1) \ io_targets PM (initial PM) io" by (metis assms(2) assms(3) assms(6) calculation insert_absorb insert_ident insert_not_empty io_targets_observable_singleton_ob observable_productF singleton_insert_inj_eq subrelI) ultimately have targets_pm : "io_targets PM (initial PM) io = {(p2,p1)}" using assms io_targets_observable_singleton_ex singletonD by (metis observable_productF) then obtain trP where trP_def : "target (io || trP) (initial PM) = (p2,p1) \ path PM (io || trP) (initial PM) \ length io = length trP" proof - assume a1: "\trP. target (io || trP) (initial PM) = (p2, p1) \ path PM (io || trP) (initial PM) \ length io = length trP \ thesis" have "\ps. target (io || ps) (initial PM) = (p2, p1) \ path PM (io || ps) (initial PM) \ length io = length ps" using \(p2, p1) \ io_targets PM (initial PM) io\ by auto then show ?thesis using a1 by blast qed then have trP_unique : "{tr . path PM (io || tr) (initial PM) \ length io = length tr} = {trP}" using observable_productF observable_path_unique_ex[of PM io "initial PM"] io_lang_pm assms(2) assms(3) assms(7) proof - obtain pps :: "('d \ 'c) list" where f1: "{ps. path PM (io || ps) (initial PM) \ length io = length ps} = {pps} \ \ observable PM" by (metis (no_types) \\thesis. \observable PM; io \ L PM; \tr. {t. path PM (io || t) (initial PM) \ length io = length t} = {tr} \ thesis\ \ thesis\ io_lang_pm) have f2: "observable PM" by (meson \observable M1\ \observable M2\ \productF M2 M1 FAIL PM\ observable_productF) then have "trP \ {pps}" using f1 trP_def by blast then show ?thesis using f2 f1 by force qed obtain trIO2 where trIO2_def : "{tr . path M2 (io || tr) (initial M2) \ length io = length tr} = { trIO2 }" using observable_path_unique_ex[of M2 io "initial M2"] io_lang_subs assms(3) by blast obtain trIO1 where trIO1_def : "{tr . path M1 (io || tr) (initial M1) \ length io = length tr} = { trIO1 }" using observable_path_unique_ex[of M1 io "initial M1"] io_lang_subs assms(2) by blast have "path PM (io || trIO2 || trIO1) (initial M2, initial M1) \ length io = length trIO2 \ length trIO2 = length trIO1" proof - have f1: "path M2 (io || trIO2) (initial M2) \ length io = length trIO2" using trIO2_def by auto have f2: "path M1 (io || trIO1) (initial M1) \ length io = length trIO1" using trIO1_def by auto then have "length trIO2 = length trIO1" using f1 by presburger then show ?thesis using f2 f1 assms(4) assms(5) assms(6) by blast qed then have trP_split : "path PM (io || trIO2 || trIO1) (initial PM) \ length io = length trIO2 \ length trIO2 = length trIO1" using assms(6) by auto then have trP_zip : "trIO2 || trIO1 = trP" using trP_def trP_unique length_zip by fastforce have "target (io || trIO2) (initial M2) = p2 \ path M2 (io || trIO2) (initial M2) \ length io = length trIO2" using trP_zip trP_split assms(6) trP_def trIO2_def by auto then have "p2 \ io_targets M2 (initial M2) io" by auto then have targets_2 : "io_targets M2 (initial M2) io = {p2}" by (meson assms(3) observable_io_target_is_singleton) have "target (io || trIO1) (initial M1) = p1 \ path M1 (io || trIO1) (initial M1) \ length io = length trIO1" using trP_zip trP_split assms(6) trP_def trIO1_def by auto then have "p1 \ io_targets M1 (initial M1) io" by auto then have targets_1 : "io_targets M1 (initial M1) io = {p1}" by (metis io_lang_subs assms(2) io_targets_observable_singleton_ex singletonD) have "io_targets M2 (initial M2) io \ io_targets M1 (initial M1) io = {(p2,p1)}" using targets_2 targets_1 by simp then show "io_targets PM (initial PM) io = io_targets M2 (initial M2) io \ io_targets M1 (initial M1) io" using targets_pm by simp next case False then have "io \ R M2 s vs xs \ RP M2 s vs xs V'' = insert io (R M2 s vs xs)" using RP_cases assms(9) by (metis insertE) have "io \ L M1" using assms(8) perm_language assms(9) using False by auto then obtain s' where s'_def : "io_targets M1 (initial M1) io = {s'}" by (meson assms(2) io_targets_observable_singleton_ob) then obtain tr1 where tr1_def : "target (io || tr1) (initial M1) = s' \ path M1 (io || tr1) (initial M1) \ length tr1 = length io" by (metis io_targets_elim singletonI) have "io_targets M2 (initial M2) io = {s}" using assms(9) assms(3) RP_state_component_2 by simp then obtain tr2 where tr2_def : "target (io || tr2) (initial M2) = s \ path M2 (io || tr2) (initial M2) \ length tr2 = length io" by (metis io_targets_elim singletonI) then have paths : "path M2 (io || tr2) (initial M2) \ path M1 (io || tr1) (initial M1)" using tr1_def by simp have "length io = length tr2" using tr2_def by simp moreover have "length tr2 = length tr1" using tr1_def tr2_def by simp ultimately have "path PM (io || tr2 || tr1) (initial M2, initial M1)" using assms(6) assms(5) assms(4) paths productF_path_forward[of io tr2 tr1 M2 M1 FAIL PM "initial M2" "initial M1"] by blast moreover have "target (io || tr2 || tr1) (initial M2, initial M1) = (s,s')" by (simp add: tr1_def tr2_def) moreover have "length (tr2 || tr2) = length io" using tr1_def tr2_def by simp moreover have "(initial M2, initial M1) = initial PM" using assms(6) by simp ultimately have "(s,s') \ io_targets PM (initial PM) io" by (metis io_target_from_path length_zip tr1_def tr2_def) moreover have "observable PM" using assms(2) assms(3) assms(6) observable_productF by blast then have "io_targets PM (initial PM) io = {(s,s')}" by (meson calculation observable_io_target_is_singleton) then show ?thesis using \io_targets M2 (initial M2) io = {s}\ \io_targets M1 (initial M1) io = {s'}\ by simp qed qed lemma RP_io_targets_finite_M1 : assumes "(vs @ xs) \ L M1 \ L M2" and "observable M1" and "is_det_state_cover M2 V" and "V'' \ Perm V M1" shows "finite (\ (image (io_targets M1 (initial M1)) (RP M2 s vs xs V'')))" proof show "finite (RP M2 s vs xs V'')" using finite_RP assms(3) assms(4) by simp show "\a. a \ RP M2 s vs xs V'' \ finite (io_targets M1 (initial M1) a)" proof - fix a assume "a \ RP M2 s vs xs V''" have RP_cases : "RP M2 s vs xs V'' = R M2 s vs xs \ (\ vs' \ V'' . vs' \ R M2 s vs xs \ RP M2 s vs xs V'' = insert vs' (R M2 s vs xs))" using RP_from_R assms by metis have "a \ L M1" proof (cases "a \ R M2 s vs xs") case True then have "prefix a (vs@xs)" by auto then show "a \ L M1" using language_state_prefix by (metis IntD1 assms(1) prefix_def) next case False then have "a \ V'' \ RP M2 s vs xs V'' = insert a (R M2 s vs xs)" using RP_cases \a \ RP M2 s vs xs V''\ by (metis insertE) then show "a \ L M1" by (meson assms(4) perm_language) qed then obtain p where "io_targets M1 (initial M1) a = {p}" using assms(2) io_targets_observable_singleton_ob by metis then show "finite (io_targets M1 (initial M1) a)" by simp qed qed lemma RP_io_targets_finite_PM : assumes "(vs @ xs) \ L M1 \ L M2" and "observable M1" and "observable M2" and "well_formed M1" and "well_formed M2" and "productF M2 M1 FAIL PM" and "is_det_state_cover M2 V" and "V'' \ Perm V M1" shows "finite (\ (image (io_targets PM (initial PM)) (RP M2 s vs xs V'')))" proof - have "\ io \ RP M2 s vs xs V'' . io_targets PM (initial PM) io = {s} \ io_targets M1 (initial M1) io" proof fix io assume "io \ RP M2 s vs xs V''" then have "io_targets PM (initial PM) io = io_targets M2 (initial M2) io \ io_targets M1 (initial M1) io" using assms RP_io_targets_split[of vs xs M1 M2 FAIL PM V V'' io s] by simp moreover have "io_targets M2 (initial M2) io = {s}" using \io \ RP M2 s vs xs V''\ assms(3) RP_state_component_2[of io M2 s vs xs V''] by blast ultimately show "io_targets PM (initial PM) io = {s} \ io_targets M1 (initial M1) io" by auto qed then have "\ (image (io_targets PM (initial PM)) (RP M2 s vs xs V'')) = \ (image (\ io . {s} \ io_targets M1 (initial M1) io) (RP M2 s vs xs V''))" by simp moreover have "\ (image (\ io . {s} \ io_targets M1 (initial M1) io) (RP M2 s vs xs V'')) = {s} \ \ (image (\ io . io_targets M1 (initial M1) io) (RP M2 s vs xs V''))" by blast ultimately have "\ (image (io_targets PM (initial PM)) (RP M2 s vs xs V'')) = {s} \ \ (image (io_targets M1 (initial M1)) (RP M2 s vs xs V''))" by auto moreover have "finite ({s} \ \ (image (io_targets M1 (initial M1)) (RP M2 s vs xs V'')))" using assms(1,2,7,8) RP_io_targets_finite_M1[of vs xs M1 M2 V V'' s] by simp ultimately show ?thesis by simp qed lemma RP_union_card_is_suffix_length : assumes "OFSM M2" and "io@xs \ L M2" and "is_det_state_cover M2 V" and "V'' \ Perm V M1" shows "\ q . card (R M2 q io xs) \ card (RP M2 q io xs V'')" "sum (\ q . card (RP M2 q io xs V'')) (nodes M2) \ length xs" proof - have "sum (\ q . card (R M2 q io xs)) (nodes M2) = length xs" using R_union_card_is_suffix_length[OF assms(1,2)] by assumption show "\ q . card (R M2 q io xs) \ card (RP M2 q io xs V'')" by (metis RP_from_R assms(3) assms(4) card_insert_le eq_iff finite_R) show "sum (\ q . card (RP M2 q io xs V'')) (nodes M2) \ length xs" by (metis (no_types, lifting) \(\q\nodes M2. card (R M2 q io xs)) = length xs\ \\q. card (R M2 q io xs) \ card (RP M2 q io xs V'')\ sum_mono) qed lemma RP_state_repetition_distribution_productF : assumes "OFSM M2" and "OFSM M1" and "(card (nodes M2) * m) \ length xs" and "card (nodes M1) \ m" and "vs@xs \ L M2 \ L M1" and "is_det_state_cover M2 V" and "V'' \ Perm V M1" shows "\ q \ nodes M2 . card (RP M2 q vs xs V'') > m" proof - have "finite (nodes M1)" "finite (nodes M2)" using assms(1,2) by auto then have "card(nodes M2 \ nodes M1) = card (nodes M2) * card (nodes M1)" using card_cartesian_product by blast have "nodes (product M2 M1) \ nodes M2 \ nodes M1" using product_nodes by auto have "card (nodes (product M2 M1)) \ card (nodes M2) * card (nodes M1)" by (metis (no_types) \card (nodes M2 \ nodes M1) = |M2| * |M1|\ \finite (nodes M1)\ \finite (nodes M2)\ \nodes (product M2 M1) \ nodes M2 \ nodes M1\ card_mono finite_cartesian_product) have "(\ q \ nodes M2 . card (R M2 q vs xs) = m) \ (\ q \ nodes M2 . card (R M2 q vs xs) > m)" proof (rule ccontr) assume "\ ((\q\nodes M2. card (R M2 q vs xs) = m) \ (\q\nodes M2. m < card (R M2 q vs xs)))" then have "\ q \ nodes M2 . card (R M2 q vs xs) \ m" by auto moreover obtain q' where "q'\nodes M2" "card (R M2 q' vs xs) < m" using \\ ((\q\nodes M2. card (R M2 q vs xs) = m) \ (\q\nodes M2. m < card (R M2 q vs xs)))\ nat_neq_iff by blast have "sum (\ q . card (R M2 q vs xs)) (nodes M2) = sum (\ q . card (R M2 q vs xs)) (nodes M2 - {q'}) + sum (\ q . card (R M2 q vs xs)) {q'}" using \q'\nodes M2\ by (meson \finite (nodes M2)\ empty_subsetI insert_subset sum.subset_diff) moreover have "sum (\ q . card (R M2 q vs xs)) (nodes M2 - {q'}) \ sum (\ q . m) (nodes M2 - {q'})" using \\ q \ nodes M2 . card (R M2 q vs xs) \ m\ by (meson sum_mono DiffD1) moreover have "sum (\ q . card (R M2 q vs xs)) {q'} < m" using \card (R M2 q' vs xs) < m\ by auto ultimately have "sum (\ q . card (R M2 q vs xs)) (nodes M2) < sum (\ q . m) (nodes M2)" proof - have "\C c f. infinite C \ (c::'c) \ C \ sum f C = (f c::nat) + sum f (C - {c})" by (meson sum.remove) then have "(\c\nodes M2. m) = m + (\c\nodes M2 - {q'}. m)" using \finite (nodes M2)\ \q' \ nodes M2\ by blast then show ?thesis using \(\q\nodes M2 - {q'}. card (R M2 q vs xs)) \ (\q\nodes M2 - {q'}. m)\ \(\q\nodes M2. card (R M2 q vs xs)) = (\q\nodes M2 - {q'}. card (R M2 q vs xs)) + (\q\{q'}. card (R M2 q vs xs))\ \(\q\{q'}. card (R M2 q vs xs)) < m\ by linarith qed moreover have "sum (\ q . m) (nodes M2) \ card (nodes M2) * m" using assms(2) by auto ultimately have "sum (\ q . card (R M2 q vs xs)) (nodes M2) < card (nodes M2) * m" by presburger moreover have "Suc (card (nodes M2)*m) \ sum (\ q . card (R M2 q vs xs)) (nodes M2)" using R_union_card_is_suffix_length[OF assms(1), of vs xs] assms(5,3) by (metis Int_iff \vs @ xs \ L M2 \ (\q\nodes M2. card (R M2 q vs xs)) = length xs\ \vs @ xs \ L M2 \ L M1\ \|M2| * m \ length xs\ calculation less_eq_Suc_le not_less_eq_eq) ultimately show "False" by simp qed then show ?thesis proof let ?q = "initial M2" assume "\q\nodes M2. card (R M2 q vs xs) = m" then have "card (R M2 ?q vs xs) = m" by auto have "[] \ V''" by (meson assms(6) assms(7) perm_empty) then have "[] \ RP M2 ?q vs xs V''" by auto have "[] \ R M2 ?q vs xs" by auto have "card (RP M2 ?q vs xs V'') \ card (R M2 ?q vs xs)" using finite_R[of M2 ?q vs xs] finite_RP[OF assms(6,7),of ?q vs xs] unfolding RP.simps by (simp add: card_mono) have "card (RP M2 ?q vs xs V'') > card (R M2 ?q vs xs)" proof - have f1: "\n na. (\ (n::nat) \ na \ n = na) \ n < na" by (meson le_neq_trans) have "RP M2 (initial M2) vs xs V'' \ R M2 (initial M2) vs xs" using \[] \ RP M2 (initial M2) vs xs V''\ \[] \ R M2 (initial M2) vs xs\ by blast then show ?thesis using f1 by (metis (no_types) RP_from_R \card (R M2 (initial M2) vs xs) \ card (RP M2 (initial M2) vs xs V'')\ assms(6) assms(7) card_insert_disjoint finite_R le_simps(2)) qed then show ?thesis using \card (R M2 ?q vs xs) = m\ by blast next assume "\q\nodes M2. m < card (R M2 q vs xs)" then obtain q where "q\nodes M2" "m < card (R M2 q vs xs)" by blast moreover have "card (RP M2 q vs xs V'') \ card (R M2 q vs xs)" using finite_R[of M2 q vs xs] finite_RP[OF assms(6,7),of q vs xs] unfolding RP.simps by (simp add: card_mono) ultimately have "m < card (RP M2 q vs xs V'')" by simp show ?thesis using \q \ nodes M2\ \m < card (RP M2 q vs xs V'')\ by blast qed qed subsection \ Conditions for the result of LB to be a valid lower bound \ text \ The following predicates describe the assumptions necessary to show that the value calculated by @{verbatim LB} is a lower bound on the number of states of a given FSM. \ fun Prereq :: "('in, 'out, 'state1) FSM \ ('in, 'out, 'state2) FSM \ ('in \ 'out) list \ ('in \ 'out) list \ 'in list set \ 'state1 set \ ('in, 'out) ATC set \ ('in \ 'out) list set \ bool" where "Prereq M2 M1 vs xs T S \ V'' = ( (finite T) \ (vs @ xs) \ L M2 \ L M1 \ S \ nodes M2 \ (\ s1 \ S . \ s2 \ S . s1 \ s2 \ (\ io1 \ RP M2 s1 vs xs V'' . \ io2 \ RP M2 s2 vs xs V'' . B M1 io1 \ \ B M1 io2 \ )))" fun Rep_Pre :: "('in, 'out, 'state1) FSM \ ('in, 'out, 'state2) FSM \ ('in \ 'out) list \ ('in \ 'out) list \ bool" where "Rep_Pre M2 M1 vs xs = (\ xs1 xs2 . prefix xs1 xs2 \ prefix xs2 xs \ xs1 \ xs2 \ (\ s2 . io_targets M2 (initial M2) (vs @ xs1) = {s2} \ io_targets M2 (initial M2) (vs @ xs2) = {s2}) \ (\ s1 . io_targets M1 (initial M1) (vs @ xs1) = {s1} \ io_targets M1 (initial M1) (vs @ xs2) = {s1}))" fun Rep_Cov :: "('in, 'out, 'state1) FSM \ ('in, 'out, 'state2) FSM \ ('in \ 'out) list set \ ('in \ 'out) list \ ('in \ 'out) list \ bool" where "Rep_Cov M2 M1 V'' vs xs = (\ xs' vs' . xs' \ [] \ prefix xs' xs \ vs' \ V'' \ (\ s2 . io_targets M2 (initial M2) (vs @ xs') = {s2} \ io_targets M2 (initial M2) (vs') = {s2}) \ (\ s1 . io_targets M1 (initial M1) (vs @ xs') = {s1} \ io_targets M1 (initial M1) (vs') = {s1}))" lemma distinctness_via_Rep_Pre : assumes "\ Rep_Pre M2 M1 vs xs" and "productF M2 M1 FAIL PM" and "observable M1" and "observable M2" and "io_targets PM (initial PM) vs = {(q2,q1)}" and "path PM (xs || tr) (q2,q1)" and "length xs = length tr" and "(vs @ xs) \ L M1 \ L M2" and "well_formed M1" and "well_formed M2" shows "distinct (states (xs || tr) (q2, q1))" proof (rule ccontr) assume assm : "\ distinct (states (xs || tr) (q2, q1))" then obtain i1 i2 where index_def : "i1 \ 0 \ i1 \ i2 \ i1 < length (states (xs || tr) (q2, q1)) \ i2 < length (states (xs || tr) (q2, q1)) \ (states (xs || tr) (q2, q1)) ! i1 = (states (xs || tr) (q2, q1)) ! i2" by (metis distinct_conv_nth) then have "length xs > 0" by auto let ?xs1 = "take (Suc i1) xs" let ?xs2 = "take (Suc i2) xs" let ?tr1 = "take (Suc i1) tr" let ?tr2 = "take (Suc i2) tr" let ?st = "(states (xs || tr) (q2, q1)) ! i1" have obs_PM : "observable PM" using observable_productF assms(2) assms(3) assms(4) by blast have "initial PM = (initial M2, initial M1)" using assms(2) by simp moreover have "vs \ L M2" "vs \ L M1" using assms(8) language_state_prefix by auto ultimately have "io_targets M1 (initial M1) vs = {q1}" "io_targets M2 (initial M2) vs = {q2}" using productF_path_io_targets[of M2 M1 FAIL PM "initial M2" "initial M1" vs q2 q1] by (metis FSM.nodes.initial assms(2) assms(3) assms(4) assms(5) assms(9) assms(10) io_targets_observable_singleton_ex singletonD)+ \ \paths for ?xs1\ have "(states (xs || tr) (q2, q1)) ! i1 \ io_targets PM (q2, q1) ?xs1" by (metis \0 < length xs\ assms(6) assms(7) index_def map_snd_zip states_alt_def states_index_io_target) then have "io_targets PM (q2, q1) ?xs1 = {?st}" using obs_PM by (meson observable_io_target_is_singleton) have "path PM (?xs1 || ?tr1) (q2,q1)" by (metis FSM.path_append_elim append_take_drop_id assms(6) assms(7) length_take zip_append) then have "path PM (?xs1 || map fst ?tr1 || map snd ?tr1) (q2,q1)" by auto have "vs @ ?xs1 \ L M2" by (metis (no_types) IntD2 append_assoc append_take_drop_id assms(8) language_state_prefix) then obtain q2' where "io_targets M2 (initial M2) (vs@?xs1) = {q2'}" using io_targets_observable_singleton_ob[of M2 "vs@?xs1" "initial M2"] assms(4) by auto then have "q2' \ io_targets M2 q2 ?xs1" using assms(4) \io_targets M2 (initial M2) vs = {q2}\ observable_io_targets_split[of M2 "initial M2" vs ?xs1 q2' q2] by simp then have "?xs1 \ language_state M2 q2" by auto moreover have "length ?xs1 = length (map snd ?tr1)" using assms(7) by auto moreover have "length (map fst ?tr1) = length (map snd ?tr1)" by auto moreover have "q2 \ nodes M2" using \io_targets M2 (initial M2) vs = {q2}\ io_targets_nodes by (metis FSM.nodes.initial insertI1) moreover have "q1 \ nodes M1" using \io_targets M1 (initial M1) vs = {q1}\ io_targets_nodes by (metis FSM.nodes.initial insertI1) ultimately have "path M1 (?xs1 || map snd ?tr1) q1" "path M2 (?xs1 || map fst ?tr1) q2" "target (?xs1 || map snd ?tr1) q1 = snd (target (?xs1 || map fst ?tr1 || map snd ?tr1) (q2,q1))" "target (?xs1 || map fst ?tr1) q2 = fst (target (?xs1 || map fst ?tr1 || map snd ?tr1) (q2,q1))" using assms(2) assms(9) assms(10) \path PM (?xs1 || map fst ?tr1 || map snd ?tr1) (q2,q1)\ assms(4) productF_path_reverse_ob_2[of ?xs1 "map fst ?tr1" "map snd ?tr1" M2 M1 FAIL PM q2 q1] by simp+ moreover have "target (?xs1 || map fst ?tr1 || map snd ?tr1) (q2,q1) = ?st" by (metis (no_types) index_def scan_nth take_zip zip_map_fst_snd) ultimately have "target (?xs1 || map snd ?tr1) q1 = snd ?st" "target (?xs1 || map fst ?tr1) q2 = fst ?st" by simp+ \ \paths for ?xs2\ have "(states (xs || tr) (q2, q1)) ! i2 \ io_targets PM (q2, q1) ?xs2" by (metis \0 < length xs\ assms(6) assms(7) index_def map_snd_zip states_alt_def states_index_io_target) then have "io_targets PM (q2, q1) ?xs2 = {?st}" using obs_PM by (metis index_def observable_io_target_is_singleton) have "path PM (?xs2 || ?tr2) (q2,q1)" by (metis FSM.path_append_elim append_take_drop_id assms(6) assms(7) length_take zip_append) then have "path PM (?xs2 || map fst ?tr2 || map snd ?tr2) (q2,q1)" by auto have "vs @ ?xs2 \ L M2" by (metis (no_types) IntD2 append_assoc append_take_drop_id assms(8) language_state_prefix) then obtain q2'' where "io_targets M2 (initial M2) (vs@?xs2) = {q2''}" using io_targets_observable_singleton_ob[of M2 "vs@?xs2" "initial M2"] assms(4) by auto then have "q2'' \ io_targets M2 q2 ?xs2" using assms(4) \io_targets M2 (initial M2) vs = {q2}\ observable_io_targets_split[of M2 "initial M2" vs ?xs2 q2'' q2] by simp then have "?xs2 \ language_state M2 q2" by auto moreover have "length ?xs2 = length (map snd ?tr2)" using assms(7) by auto moreover have "length (map fst ?tr2) = length (map snd ?tr2)" by auto moreover have "q2 \ nodes M2" using \io_targets M2 (initial M2) vs = {q2}\ io_targets_nodes by (metis FSM.nodes.initial insertI1) moreover have "q1 \ nodes M1" using \io_targets M1 (initial M1) vs = {q1}\ io_targets_nodes by (metis FSM.nodes.initial insertI1) ultimately have "path M1 (?xs2 || map snd ?tr2) q1" "path M2 (?xs2 || map fst ?tr2) q2" "target (?xs2 || map snd ?tr2) q1 = snd(target (?xs2 || map fst ?tr2 || map snd ?tr2) (q2,q1))" "target (?xs2 || map fst ?tr2) q2 = fst(target (?xs2 || map fst ?tr2 || map snd ?tr2) (q2,q1))" using assms(2) assms(9) assms(10) \path PM (?xs2 || map fst ?tr2 || map snd ?tr2) (q2,q1)\ assms(4) productF_path_reverse_ob_2[of ?xs2 "map fst ?tr2" "map snd ?tr2" M2 M1 FAIL PM q2 q1] by simp+ moreover have "target (?xs2 || map fst ?tr2 || map snd ?tr2) (q2,q1) = ?st" by (metis (no_types) index_def scan_nth take_zip zip_map_fst_snd) ultimately have "target (?xs2 || map snd ?tr2) q1 = snd ?st" "target (?xs2 || map fst ?tr2) q2 = fst ?st" by simp+ have "io_targets M1 q1 ?xs1 = {snd ?st}" using \path M1 (?xs1 || map snd ?tr1) q1\ \target (?xs1 || map snd ?tr1) q1 = snd ?st\ \length ?xs1 = length (map snd ?tr1)\ assms(3) obs_target_is_io_targets[of M1 ?xs1 "map snd ?tr1" q1] by simp then have tgt_1_1 : "io_targets M1 (initial M1) (vs @ ?xs1) = {snd ?st}" by (meson \io_targets M1 (initial M1) vs = {q1}\ assms(3) observable_io_targets_append) have "io_targets M2 q2 ?xs1 = {fst ?st}" using \path M2 (?xs1 || map fst ?tr1) q2\ \target (?xs1 || map fst ?tr1) q2 = fst ?st\ \length ?xs1 = length (map snd ?tr1)\ assms(4) obs_target_is_io_targets[of M2 ?xs1 "map fst ?tr1" q2] by simp then have tgt_1_2 : "io_targets M2 (initial M2) (vs @ ?xs1) = {fst ?st}" by (meson \io_targets M2 (initial M2) vs = {q2}\ assms(4) observable_io_targets_append) have "io_targets M1 q1 ?xs2 = {snd ?st}" using \path M1 (?xs2 || map snd ?tr2) q1\ \target (?xs2 || map snd ?tr2) q1 = snd ?st\ \length ?xs2 = length (map snd ?tr2)\ assms(3) obs_target_is_io_targets[of M1 ?xs2 "map snd ?tr2" q1] by simp then have tgt_2_1 : "io_targets M1 (initial M1) (vs @ ?xs2) = {snd ?st}" by (meson \io_targets M1 (initial M1) vs = {q1}\ assms(3) observable_io_targets_append) have "io_targets M2 q2 ?xs2 = {fst ?st}" using \path M2 (?xs2 || map fst ?tr2) q2\ \target (?xs2 || map fst ?tr2) q2 = fst ?st\ \length ?xs2 = length (map snd ?tr2)\ assms(4) obs_target_is_io_targets[of M2 ?xs2 "map fst ?tr2" q2] by simp then have tgt_2_2 : "io_targets M2 (initial M2) (vs @ ?xs2) = {fst ?st}" by (meson \io_targets M2 (initial M2) vs = {q2}\ assms(4) observable_io_targets_append) have "?xs1 \ []" using \0 < length xs\ by auto have "prefix ?xs1 xs" using take_is_prefix by blast have "prefix ?xs2 xs" using take_is_prefix by blast have "?xs1 \ ?xs2" proof - have f1: "\n na. \ n < na \ Suc n \ na" by presburger have f2: "Suc i1 \ length xs" using index_def by force have "Suc i2 \ length xs" using f1 by (metis index_def length_take map_snd_zip_take min_less_iff_conj states_alt_def) then show ?thesis using f2 by (metis (no_types) index_def length_take min.absorb2 nat.simps(1)) qed have "Rep_Pre M2 M1 vs xs" proof (cases "length ?xs1 < length ?xs2") case True then have "prefix ?xs1 ?xs2" by (meson \prefix (take (Suc i1) xs) xs\ \prefix (take (Suc i2) xs) xs\ leD prefix_length_le prefix_same_cases) show ?thesis by (meson Rep_Pre.elims(3) \prefix (take (Suc i1) xs) (take (Suc i2) xs)\ \prefix (take (Suc i2) xs) xs\ \take (Suc i1) xs \ take (Suc i2) xs\ tgt_1_1 tgt_1_2 tgt_2_1 tgt_2_2) next case False moreover have "length ?xs1 \ length ?xs2" by (metis (no_types) \take (Suc i1) xs \ take (Suc i2) xs\ append_eq_conv_conj append_take_drop_id) ultimately have "length ?xs2 < length ?xs1" by auto then have "prefix ?xs2 ?xs1" using \prefix (take (Suc i1) xs) xs\ \prefix (take (Suc i2) xs) xs\ less_imp_le_nat prefix_length_prefix by blast show ?thesis by (metis Rep_Pre.elims(3) \prefix (take (Suc i1) xs) xs\ \prefix (take (Suc i2) xs) (take (Suc i1) xs)\ \take (Suc i1) xs \ take (Suc i2) xs\ tgt_1_1 tgt_1_2 tgt_2_1 tgt_2_2) qed then show "False" using assms(1) by simp qed lemma RP_count_via_Rep_Cov : assumes "(vs @ xs) \ L M1 \ L M2" and "observable M1" and "observable M2" and "well_formed M1" and "well_formed M2" and "s \ nodes M2" and "productF M2 M1 FAIL PM" and "io_targets PM (initial PM) vs = {(q2,q1)}" and "path PM (xs || tr) (q2,q1)" and "length xs = length tr" and "distinct (states (xs || tr) (q2,q1))" and "is_det_state_cover M2 V" and "V'' \ Perm V M1" and "\ Rep_Cov M2 M1 V'' vs xs" shows "card (\(image (io_targets M1 (initial M1)) (RP M2 s vs xs V''))) = card (RP M2 s vs xs V'')" proof - have RP_cases : "RP M2 s vs xs V'' = R M2 s vs xs \ (\ vs' \ V'' . vs' \ R M2 s vs xs \ RP M2 s vs xs V'' = insert vs' (R M2 s vs xs))" using RP_from_R assms by metis show ?thesis proof (cases "RP M2 s vs xs V'' = R M2 s vs xs") case True then show ?thesis using R_count assms by metis next case False then obtain vs' where vs'_def : "vs' \ V'' \ vs' \ R M2 s vs xs \ RP M2 s vs xs V'' = insert vs' (R M2 s vs xs)" using RP_cases by auto have state_component_2 : "\ io \ (R M2 s vs xs) . io_targets M2 (initial M2) io = {s}" proof fix io assume "io \ R M2 s vs xs" then have "s \ io_targets M2 (initial M2) io" by auto moreover have "io \ language_state M2 (initial M2)" using calculation by auto ultimately show "io_targets M2 (initial M2) io = {s}" using assms(3) io_targets_observable_singleton_ex by (metis singletonD) qed have "vs' \ L M1" using assms(13) perm_language vs'_def by blast then obtain s' where s'_def : "io_targets M1 (initial M1) vs' = {s'}" by (meson assms(2) io_targets_observable_singleton_ob) moreover have "s' \ \ (image (io_targets M1 (initial M1)) (R M2 s vs xs))" proof (rule ccontr) assume "\ s' \ \(io_targets M1 (initial M1) ` R M2 s vs xs)" then obtain xs' where xs'_def : "vs @ xs' \ R M2 s vs xs \ s' \ io_targets M1 (initial M1) (vs @ xs')" proof - assume a1: "\xs'. vs @ xs' \ R M2 s vs xs \ s' \ io_targets M1 (initial M1) (vs @ xs') \ thesis" obtain pps :: "('a \ 'b) list set \ (('a \ 'b) list \ 'c set) \ 'c \ ('a \ 'b) list" where "\x0 x1 x2. (\v3. v3 \ x0 \ x2 \ x1 v3) = (pps x0 x1 x2 \ x0 \ x2 \ x1 (pps x0 x1 x2))" by moura then have f2: "pps (R M2 s vs xs) (io_targets M1 (initial M1)) s' \ R M2 s vs xs \ s' \ io_targets M1 (initial M1) (pps (R M2 s vs xs) (io_targets M1 (initial M1)) s')" using \\ s' \ \(io_targets M1 (initial M1) ` R M2 s vs xs)\ by blast then have "\ps. pps (R M2 s vs xs) (io_targets M1 (initial M1)) s' = vs @ ps \ ps \ [] \ prefix ps xs \ s \ io_targets M2 (initial M2) (vs @ ps)" by simp then show ?thesis using f2 a1 by (metis (no_types)) qed have "vs @ xs' \ L M1" using xs'_def by blast then have "io_targets M1 (initial M1) (vs@xs') = {s'}" by (metis assms(2) io_targets_observable_singleton_ob singletonD xs'_def) moreover have "io_targets M1 (initial M1) (vs') = {s'}" using s'_def by blast moreover have "io_targets M2 (initial M2) (vs @ xs') = {s}" using state_component_2 xs'_def by blast moreover have "io_targets M2 (initial M2) (vs') = {s}" by (metis (mono_tags, lifting) RP.simps Un_iff insertI1 mem_Collect_eq vs'_def) moreover have "xs' \ []" using xs'_def by simp moreover have "prefix xs' xs" using xs'_def by simp moreover have "vs' \ V''" using vs'_def by simp ultimately have "Rep_Cov M2 M1 V'' vs xs" by auto then show "False" using assms(14) by simp qed moreover have "\ (image (io_targets M1 (initial M1)) (insert vs' (R M2 s vs xs))) = insert s' (\ (image (io_targets M1 (initial M1)) (R M2 s vs xs)))" using s'_def by simp moreover have "finite (\ (image (io_targets M1 (initial M1)) (R M2 s vs xs)))" proof show "finite (R M2 s vs xs)" using finite_R by simp show "\a. a \ R M2 s vs xs \ finite (io_targets M1 (initial M1) a)" proof - fix a assume "a \ R M2 s vs xs" then have "prefix a (vs@xs)" by auto then have "a \ L M1" using language_state_prefix by (metis IntD1 assms(1) prefix_def) then obtain p where "io_targets M1 (initial M1) a = {p}" using assms(2) io_targets_observable_singleton_ob by metis then show "finite (io_targets M1 (initial M1) a)" by simp qed qed ultimately have "card (\ (image (io_targets M1 (initial M1)) (insert vs' (R M2 s vs xs)))) = Suc (card (\ (image (io_targets M1 (initial M1)) (R M2 s vs xs))))" by (metis (no_types) card_insert_disjoint) moreover have "card (\ (image (io_targets M1 (initial M1)) (RP M2 s vs xs V''))) = card (\ (image (io_targets M1 (initial M1)) (insert vs' (R M2 s vs xs))))" using vs'_def by simp ultimately have "card (\ (image (io_targets M1 (initial M1)) (RP M2 s vs xs V''))) = Suc (card (\ (image (io_targets M1 (initial M1)) (R M2 s vs xs))))" by linarith then have "card (\ (image (io_targets M1 (initial M1)) (RP M2 s vs xs V''))) = Suc (card (R M2 s vs xs))" using R_count[of vs xs M1 M2 s FAIL PM q2 q1 tr] using assms(1,10,11,2-9) by linarith moreover have "card (RP M2 s vs xs V'') = Suc (card (R M2 s vs xs))" using vs'_def by (metis card_insert_if finite_R) ultimately show ?thesis by linarith qed qed lemma RP_count_alt_def : assumes "(vs @ xs) \ L M1 \ L M2" and "observable M1" and "observable M2" and "well_formed M1" and "well_formed M2" and "s \ nodes M2" and "productF M2 M1 FAIL PM" and "io_targets PM (initial PM) vs = {(q2,q1)}" and "path PM (xs || tr) (q2,q1)" and "length xs = length tr" and "\ Rep_Pre M2 M1 vs xs" and "is_det_state_cover M2 V" and "V'' \ Perm V M1" and "\ Rep_Cov M2 M1 V'' vs xs" shows "card (\(image (io_targets M1 (initial M1)) (RP M2 s vs xs V''))) = card (RP M2 s vs xs V'')" proof - have "distinct (states (xs || tr) (q2,q1))" using distinctness_via_Rep_Pre[of M2 M1 vs xs FAIL PM q2 q1 tr] assms by simp then show ?thesis using RP_count_via_Rep_Cov[of vs xs M1 M2 s FAIL PM q2 q1 tr V V''] using assms(1,10,12-14,2-9) by blast qed subsection \ Function LB \ text \ @{verbatim LB} adds together the number of elements in sets calculated via RP for a given set of states and the number of ATC-reaction known to exist but not produced by a state reached by any of the above elements. \ fun LB :: "('in, 'out, 'state1) FSM \ ('in, 'out, 'state2) FSM \ ('in \ 'out) list \ ('in \ 'out) list \ 'in list set \ 'state1 set \ ('in, 'out) ATC set \ ('in \ 'out) list set \ nat" where "LB M2 M1 vs xs T S \ V'' = (sum (\ s . card (RP M2 s vs xs V'')) S) + card ((D M1 T \) - {B M1 xs' \ | xs' s' . s' \ S \ xs' \ RP M2 s' vs xs V''})" lemma LB_count_helper_RP_disjoint_and_cards : assumes "(vs @ xs) \ L M1 \ L M2" and "observable M1" and "observable M2" and "well_formed M1" and "well_formed M2" and "productF M2 M1 FAIL PM" and "is_det_state_cover M2 V" and "V'' \ Perm V M1" and "s1 \ s2" shows "\ (image (io_targets PM (initial PM)) (RP M2 s1 vs xs V'')) \ \ (image (io_targets PM (initial PM)) (RP M2 s2 vs xs V'')) = {}" "card (\ (image (io_targets PM (initial PM)) (RP M2 s1 vs xs V''))) = card (\ (image (io_targets M1 (initial M1)) (RP M2 s1 vs xs V'')))" "card (\ (image (io_targets PM (initial PM)) (RP M2 s2 vs xs V''))) = card (\ (image (io_targets M1 (initial M1)) (RP M2 s2 vs xs V'')))" proof - have "\ io \ RP M2 s1 vs xs V'' . io_targets PM (initial PM) io = {s1} \ io_targets M1 (initial M1) io" proof fix io assume "io \ RP M2 s1 vs xs V''" then have "io_targets PM (initial PM) io = io_targets M2 (initial M2) io \ io_targets M1 (initial M1) io" using assms RP_io_targets_split[of vs xs M1 M2 FAIL PM V V'' io s1] by simp moreover have "io_targets M2 (initial M2) io = {s1}" using \io \ RP M2 s1 vs xs V''\ assms(3) RP_state_component_2[of io M2 s1 vs xs V''] by blast ultimately show "io_targets PM (initial PM) io = {s1} \ io_targets M1 (initial M1) io" by auto qed then have "\ (image (io_targets PM (initial PM)) (RP M2 s1 vs xs V'')) = \ (image (\ io . {s1} \ io_targets M1 (initial M1) io) (RP M2 s1 vs xs V''))" by simp moreover have "\ (image (\ io . {s1} \ io_targets M1 (initial M1) io) (RP M2 s1 vs xs V'')) = {s1} \ \ (image (\ io . io_targets M1 (initial M1) io) (RP M2 s1 vs xs V''))" by blast ultimately have image_split_1 : "\ (image (io_targets PM (initial PM)) (RP M2 s1 vs xs V'') ) = {s1} \ \ (image (io_targets M1 (initial M1)) (RP M2 s1 vs xs V''))" by simp then show "card (\ (image (io_targets PM (initial PM)) (RP M2 s1 vs xs V''))) = card (\ (image (io_targets M1 (initial M1)) (RP M2 s1 vs xs V'')))" by (metis (no_types) card_cartesian_product_singleton) have "\ io \ RP M2 s2 vs xs V'' . io_targets PM (initial PM) io = {s2} \ io_targets M1 (initial M1) io" proof fix io assume "io \ RP M2 s2 vs xs V''" then have "io_targets PM (initial PM) io = io_targets M2 (initial M2) io \ io_targets M1 (initial M1) io" using assms RP_io_targets_split[of vs xs M1 M2 FAIL PM V V'' io s2] by simp moreover have "io_targets M2 (initial M2) io = {s2}" using \io \ RP M2 s2 vs xs V''\ assms(3) RP_state_component_2[of io M2 s2 vs xs V''] by blast ultimately show "io_targets PM (initial PM) io = {s2} \ io_targets M1 (initial M1) io" by auto qed then have "\ (image (io_targets PM (initial PM)) (RP M2 s2 vs xs V'')) = \ (image (\ io . {s2} \ io_targets M1 (initial M1) io) (RP M2 s2 vs xs V''))" by simp moreover have "\ (image (\ io . {s2} \ io_targets M1 (initial M1) io) (RP M2 s2 vs xs V'')) = {s2} \ \ (image (\ io . io_targets M1 (initial M1) io) (RP M2 s2 vs xs V''))" by blast ultimately have image_split_2 : "\ (image (io_targets PM (initial PM)) (RP M2 s2 vs xs V'')) = {s2} \ \ (image (io_targets M1 (initial M1)) (RP M2 s2 vs xs V''))" by simp then show "card (\ (image (io_targets PM (initial PM)) (RP M2 s2 vs xs V''))) = card (\ (image (io_targets M1 (initial M1)) (RP M2 s2 vs xs V'')))" by (metis (no_types) card_cartesian_product_singleton) have "\ (image (io_targets PM (initial PM)) (RP M2 s1 vs xs V'')) \ \ (image (io_targets PM (initial PM)) (RP M2 s2 vs xs V'')) = {s1} \ \ (image (io_targets M1 (initial M1)) (RP M2 s1 vs xs V'')) \ {s2} \ \ (image (io_targets M1 (initial M1)) (RP M2 s2 vs xs V''))" using image_split_1 image_split_2 by blast moreover have "{s1} \ \ (image (io_targets M1 (initial M1)) (RP M2 s1 vs xs V'')) \ {s2} \ \ (image (io_targets M1 (initial M1)) (RP M2 s2 vs xs V'')) = {}" using assms(9) by auto ultimately show "\ (image (io_targets PM (initial PM)) (RP M2 s1 vs xs V'')) \ \ (image (io_targets PM (initial PM)) (RP M2 s2 vs xs V'')) = {}" by presburger qed lemma LB_count_helper_RP_disjoint_card_M1 : assumes "(vs @ xs) \ L M1 \ L M2" and "observable M1" and "observable M2" and "well_formed M1" and "well_formed M2" and "productF M2 M1 FAIL PM" and "is_det_state_cover M2 V" and "V'' \ Perm V M1" and "s1 \ s2" shows "card (\ (image (io_targets PM (initial PM)) (RP M2 s1 vs xs V'')) \ \ (image (io_targets PM (initial PM)) (RP M2 s2 vs xs V''))) = card (\ (image (io_targets M1 (initial M1)) (RP M2 s1 vs xs V''))) + card (\ (image (io_targets M1 (initial M1)) (RP M2 s2 vs xs V'')))" proof - have "finite (\ (image (io_targets PM (initial PM)) (RP M2 s1 vs xs V'')))" using RP_io_targets_finite_PM[OF assms(1-8)] by simp moreover have "finite (\ (image (io_targets PM (initial PM)) (RP M2 s2 vs xs V'')))" using RP_io_targets_finite_PM[OF assms(1-8)] by simp ultimately show ?thesis using LB_count_helper_RP_disjoint_and_cards[OF assms] by (metis (no_types) card_Un_disjoint) qed lemma LB_count_helper_RP_disjoint_M1_pair : assumes "(vs @ xs) \ L M1 \ L M2" and "observable M1" and "observable M2" and "well_formed M1" and "well_formed M2" and "productF M2 M1 FAIL PM" and "io_targets PM (initial PM) vs = {(q2,q1)}" and "path PM (xs || tr) (q2,q1)" and "length xs = length tr" and "\ Rep_Pre M2 M1 vs xs" and "is_det_state_cover M2 V" and "V'' \ Perm V M1" and "\ Rep_Cov M2 M1 V'' vs xs" and "Prereq M2 M1 vs xs T S \ V''" and "s1 \ s2" and "s1 \ S" and "s2 \ S" and "applicable_set M1 \" and "completely_specified M1" shows "card (RP M2 s1 vs xs V'') + card (RP M2 s2 vs xs V'') = card (\ (image (io_targets M1 (initial M1)) (RP M2 s1 vs xs V''))) + card (\ (image (io_targets M1 (initial M1)) (RP M2 s2 vs xs V'')))" "\ (image (io_targets M1 (initial M1)) (RP M2 s1 vs xs V'')) \ \ (image (io_targets M1 (initial M1)) (RP M2 s2 vs xs V'')) = {}" proof - have "s1 \ nodes M2" using assms(14,16) unfolding Prereq.simps by blast have "s2 \ nodes M2" using assms(14,17) unfolding Prereq.simps by blast have "card (RP M2 s1 vs xs V'') = card (\ (image (io_targets M1 (initial M1)) (RP M2 s1 vs xs V'')))" using RP_count_alt_def[OF assms(1-5) \s1 \ nodes M2\ assms(6-13)] by linarith moreover have "card (RP M2 s2 vs xs V'') = card (\ (image (io_targets M1 (initial M1)) (RP M2 s2 vs xs V'')))" using RP_count_alt_def[OF assms(1-5) \s2 \ nodes M2\ assms(6-13)] by linarith moreover show "\ (image (io_targets M1 (initial M1)) (RP M2 s1 vs xs V'')) \ \ (image (io_targets M1 (initial M1)) (RP M2 s2 vs xs V'')) = {}" proof (rule ccontr) assume "\ (image (io_targets M1 (initial M1)) (RP M2 s1 vs xs V'')) \ \ (image (io_targets M1 (initial M1)) (RP M2 s2 vs xs V'')) \ {}" then obtain io1 io2 t where shared_elem_def : "io1 \ (RP M2 s1 vs xs V'')" "io2 \ (RP M2 s2 vs xs V'')" "t \ io_targets M1 (initial M1) io1" "t \ io_targets M1 (initial M1) io2" by blast have dist_prop: "(\ s1 \ S . \ s2 \ S . s1 \ s2 \ (\ io1 \ RP M2 s1 vs xs V'' . \ io2 \ RP M2 s2 vs xs V'' . B M1 io1 \ \ B M1 io2 \ ))" using assms(14) by simp have "io_targets M1 (initial M1) io1 \ io_targets M1 (initial M1) io2 = {}" proof (rule ccontr) assume "io_targets M1 (initial M1) io1 \ io_targets M1 (initial M1) io2 \ {}" then have "io_targets M1 (initial M1) io1 \ {}" "io_targets M1 (initial M1) io2 \ {}" by blast+ then obtain s1 s2 where "s1 \ io_targets M1 (initial M1) io1" "s2 \ io_targets M1 (initial M1) io2" by blast then have "io_targets M1 (initial M1) io1 = {s1}" "io_targets M1 (initial M1) io2 = {s2}" by (meson assms(2) observable_io_target_is_singleton)+ then have "s1 = s2" using \io_targets M1 (initial M1) io1 \ io_targets M1 (initial M1) io2 \ {}\ by auto then have "B M1 io1 \ = B M1 io2 \" using \io_targets M1 (initial M1) io1 = {s1}\ \io_targets M1 (initial M1) io2 = {s2}\ by auto then show "False" using assms(15-17) dist_prop shared_elem_def(1,2) by blast qed then show "False" using shared_elem_def(3,4) by blast qed ultimately show "card (RP M2 s1 vs xs V'') + card (RP M2 s2 vs xs V'') = card (\ (image (io_targets M1 (initial M1)) (RP M2 s1 vs xs V''))) + card (\ (image (io_targets M1 (initial M1)) (RP M2 s2 vs xs V'')))" by linarith qed lemma LB_count_helper_RP_card_union : assumes "observable M2" and "s1 \ s2" shows "RP M2 s1 vs xs V'' \ RP M2 s2 vs xs V'' = {}" proof (rule ccontr) assume "RP M2 s1 vs xs V'' \ RP M2 s2 vs xs V'' \ {}" then obtain io where "io \ RP M2 s1 vs xs V'' \ io \ RP M2 s2 vs xs V''" by blast then have "s1 \ io_targets M2 (initial M2) io" "s2 \ io_targets M2 (initial M2) io" by auto then have "s1 = s2" using assms(1) by (metis observable_io_target_is_singleton singletonD) then show "False" using assms(2) by simp qed lemma LB_count_helper_RP_inj : obtains f where "\ q \ (\ (image (\ s . \ (image (io_targets M1 (initial M1)) (RP M2 s vs xs V''))) S)) . f q \ nodes M1" "inj_on f (\ (image (\ s . \ (image (io_targets M1 (initial M1)) (RP M2 s vs xs V''))) S))" proof - let ?f = "\ q . if (q \ (\ (image (\ s . \ (image (io_targets M1 (initial M1)) (RP M2 s vs xs V''))) S))) then q else (initial M1)" have "(\ (image (\ s . \ (image (io_targets M1 (initial M1)) (RP M2 s vs xs V''))) S)) \ nodes M1" by blast then have "\ q \ (\ (image (\ s . \ (image (io_targets M1 (initial M1)) (RP M2 s vs xs V''))) S)) . ?f q \ nodes M1" by (metis Un_iff sup.order_iff) moreover have "inj_on ?f (\ (image (\ s . \ (image (io_targets M1 (initial M1)) (RP M2 s vs xs V''))) S))" proof fix x assume "x \ (\ (image (\ s . \ (image (io_targets M1 (initial M1)) (RP M2 s vs xs V''))) S))" then have "?f x = x" by presburger fix y assume "y \ (\ (image (\ s . \ (image (io_targets M1 (initial M1)) (RP M2 s vs xs V''))) S))" then have "?f y = y" by presburger assume "?f x = ?f y" then show "x = y" using \?f x = x\ \?f y = y\ by presburger qed ultimately show ?thesis using that by presburger qed abbreviation (input) UNION :: "'a set \ ('a \ 'b set) \ 'b set" where "UNION A f \ \ (f ` A)" lemma LB_count_helper_RP_card_union_sum : assumes "(vs @ xs) \ L M2 \ L M1" and "OFSM M1" and "OFSM M2" and "asc_fault_domain M2 M1 m" and "test_tools M2 M1 FAIL PM V \" and "V'' \ Perm V M1" and "Prereq M2 M1 vs xs T S \ V''" and "\ Rep_Pre M2 M1 vs xs" and "\ Rep_Cov M2 M1 V'' vs xs" shows "sum (\ s . card (RP M2 s vs xs V'')) S = sum (\ s . card (\ (image (io_targets M1 (initial M1)) (RP M2 s vs xs V'')))) S" using assms proof - have "finite (nodes M2)" using assms(3) by auto moreover have "S \ nodes M2" using assms(7) by simp ultimately have "finite S" using infinite_super by blast then have "sum (\ s . card (RP M2 s vs xs V'')) S = sum (\ s . card (\ (image (io_targets M1 (initial M1)) (RP M2 s vs xs V'')))) S" using assms proof (induction S) case empty show ?case by simp next case (insert s S) have "(insert s S) \ nodes M2" using insert.prems(7) by simp then have "s \ nodes M2" by simp have "Prereq M2 M1 vs xs T S \ V''" using \Prereq M2 M1 vs xs T (insert s S) \ V''\ by simp then have "(\s\S. card (RP M2 s vs xs V'')) = (\s\S. card (\a\RP M2 s vs xs V''. io_targets M1 (initial M1) a))" using insert.IH[OF insert.prems(1-6) _ assms(8,9)] by metis moreover have "(\s'\(insert s S). card (RP M2 s' vs xs V'')) = (\s'\S. card (RP M2 s' vs xs V'')) + card (RP M2 s vs xs V'')" by (simp add: add.commute insert.hyps(1) insert.hyps(2)) ultimately have S_prop : "(\s'\(insert s S). card (RP M2 s' vs xs V'')) = (\s\S. card (\a\RP M2 s vs xs V''. io_targets M1 (initial M1) a)) + card (RP M2 s vs xs V'')" by presburger have "vs@xs \ L M1 \ L M2" using insert.prems(1) by simp obtain q2 q1 tr where suffix_path : "io_targets PM (initial PM) vs = {(q2,q1)}" "path PM (xs || tr) (q2,q1)" "length xs = length tr" using productF_language_state_intermediate[OF insert.prems(1) test_tools_props(1)[OF insert.prems(5,4)] OFSM_props(2,1)[OF insert.prems(3)] OFSM_props(2,1)[OF insert.prems(2)]] by blast have "card (RP M2 s vs xs V'') = card (\ (image (io_targets M1 (initial M1)) (RP M2 s vs xs V'')))" using OFSM_props(2,1)[OF insert.prems(3)] OFSM_props(2,1)[OF insert.prems(2)] RP_count_alt_def[OF \vs@xs \ L M1 \ L M2\ _ _ _ _ \s\nodes M2\ test_tools_props(1)[OF insert.prems(5,4)] suffix_path insert.prems(8) test_tools_props(2)[OF insert.prems(5,4)] assms(6) insert.prems(9)] by linarith show "(\s\insert s S. card (RP M2 s vs xs V'')) = (\s\insert s S. card (UNION (RP M2 s vs xs V'') (io_targets M1 (initial M1))))" proof - have "(\c\insert s S. card (UNION (RP M2 c vs xs V'') (io_targets M1 (initial M1)))) = card (UNION (RP M2 s vs xs V'') (io_targets M1 (initial M1))) + (\c\S. card (UNION (RP M2 c vs xs V'') (io_targets M1 (initial M1))))" by (meson insert.hyps(1) insert.hyps(2) sum.insert) then show ?thesis using \(\s'\insert s S. card (RP M2 s' vs xs V'')) = (\s\S. card (\a\RP M2 s vs xs V''. io_targets M1 (initial M1) a)) + card (RP M2 s vs xs V'')\ \card (RP M2 s vs xs V'') = card (UNION (RP M2 s vs xs V'') (io_targets M1 (initial M1)))\ by presburger qed qed then show ?thesis using assms by blast qed lemma finite_insert_card : assumes "finite (\SS)" and "finite S" and "S \ (\SS) = {}" shows "card (\ (insert S SS)) = card (\SS) + card S" by (simp add: assms(1) assms(2) assms(3) card_Un_disjoint) lemma LB_count_helper_RP_disjoint_M1_union : assumes "(vs @ xs) \ L M2 \ L M1" and "OFSM M1" and "OFSM M2" and "asc_fault_domain M2 M1 m" and "test_tools M2 M1 FAIL PM V \" and "V'' \ Perm V M1" and "Prereq M2 M1 vs xs T S \ V''" and "\ Rep_Pre M2 M1 vs xs" and "\ Rep_Cov M2 M1 V'' vs xs" shows "sum (\ s . card (RP M2 s vs xs V'')) S = card (\ (image (\ s . \ (image (io_targets M1 (initial M1)) (RP M2 s vs xs V''))) S))" using assms proof - have "finite (nodes M2)" using assms(3) by auto moreover have "S \ nodes M2" using assms(7) by simp ultimately have "finite S" using infinite_super by blast then show "sum (\ s . card (RP M2 s vs xs V'')) S = card (\ (image (\ s . \ (image (io_targets M1 (initial M1)) (RP M2 s vs xs V''))) S))" using assms proof (induction S) case empty show ?case by simp next case (insert s S) have "(insert s S) \ nodes M2" using insert.prems(7) by simp then have "s \ nodes M2" by simp have "Prereq M2 M1 vs xs T S \ V''" using \Prereq M2 M1 vs xs T (insert s S) \ V''\ by simp then have applied_IH : "(\s\S. card (RP M2 s vs xs V'')) = card (\s\S. \a\RP M2 s vs xs V''. io_targets M1 (initial M1) a)" using insert.IH[OF insert.prems(1-6) _ insert.prems(8,9)] by metis obtain q2 q1 tr where suffix_path : "io_targets PM (initial PM) vs = {(q2,q1)}" "path PM (xs || tr) (q2,q1)" "length xs = length tr" using productF_language_state_intermediate [OF insert.prems(1) test_tools_props(1)[OF insert.prems(5,4)] OFSM_props(2,1)[OF insert.prems(3)] OFSM_props(2,1)[OF insert.prems(2)]] by blast have "s \ insert s S" by simp have "vs@xs \ L M1 \ L M2" using insert.prems(1) by simp have "\ s' \ S . (\a\RP M2 s vs xs V''. io_targets M1 (initial M1) a) \ (\a\RP M2 s' vs xs V''. io_targets M1 (initial M1) a) = {}" proof fix s' assume "s' \ S" have "s \ s'" using insert.hyps(2) \s' \ S\ by blast have "s' \ insert s S" using \s' \ S\ by simp show "(\a\RP M2 s vs xs V''. io_targets M1 (initial M1) a) \ (\a\RP M2 s' vs xs V''. io_targets M1 (initial M1) a) = {}" using OFSM_props(2,1)[OF assms(3)] OFSM_props(2,1,3)[OF assms(2)] LB_count_helper_RP_disjoint_M1_pair(2) [OF \vs@xs \ L M1 \ L M2\ _ _ _ _ test_tools_props(1)[OF insert.prems(5,4)] suffix_path insert.prems(8) test_tools_props(2)[OF insert.prems(5,4)] insert.prems(6,9,7) \s \ s'\ \s \ insert s S\ \s' \ insert s S\ test_tools_props(4)[OF insert.prems(5,4)]] by linarith qed then have disj_insert : "(\s\S. \a\RP M2 s vs xs V''. io_targets M1 (initial M1) a) \ (\a\RP M2 s vs xs V''. io_targets M1 (initial M1) a) = {}" by blast have finite_S : "finite (\a\RP M2 s vs xs V''. io_targets M1 (initial M1) a)" using RP_io_targets_finite_M1[OF insert.prems(1)] by (meson RP_io_targets_finite_M1 \vs @ xs \ L M1 \ L M2\ assms(2) assms(5) insert.prems(6)) have finite_s : "finite (\s\S. \a\RP M2 s vs xs V''. io_targets M1 (initial M1) a)" by (meson RP_io_targets_finite_M1 \vs @ xs \ L M1 \ L M2\ assms(2) assms(5) finite_UN_I insert.hyps(1) insert.prems(6)) have "card (\s\insert s S. \a\RP M2 s vs xs V''. io_targets M1 (initial M1) a) = card (\s\S. \a\RP M2 s vs xs V''. io_targets M1 (initial M1) a) + card (\a\RP M2 s vs xs V''. io_targets M1 (initial M1) a)" proof - have f1: "insert (UNION (RP M2 s vs xs V'') (io_targets M1 (initial M1))) ((\c. UNION (RP M2 c vs xs V'') (io_targets M1 (initial M1))) ` S) = (\c. UNION (RP M2 c vs xs V'') (io_targets M1 (initial M1))) ` insert s S" by blast have "\c. c \ S \ UNION (RP M2 s vs xs V'') (io_targets M1 (initial M1)) \ UNION (RP M2 c vs xs V'') (io_targets M1 (initial M1)) = {}" by (meson \\s'\S. (\a\RP M2 s vs xs V''. io_targets M1 (initial M1) a) \ (\a\RP M2 s' vs xs V''. io_targets M1 (initial M1) a) = {}\) then have "UNION (RP M2 s vs xs V'') (io_targets M1 (initial M1)) \ (\c\S. UNION (RP M2 c vs xs V'') (io_targets M1 (initial M1))) = {}" by blast then show ?thesis using f1 by (metis finite_S finite_insert_card finite_s) qed have "card (RP M2 s vs xs V'') = card (\a\RP M2 s vs xs V''. io_targets M1 (initial M1) a)" using assms(2) assms(3) RP_count_alt_def[OF \vs@xs \ L M1 \ L M2\ _ _ _ _ \s \ nodes M2\ test_tools_props(1)[OF insert.prems(5,4)] suffix_path insert.prems(8) test_tools_props(2)[OF insert.prems(5,4)] insert.prems(6,9)] by metis show ?case proof - have "(\c\insert s S. card (RP M2 c vs xs V'')) = card (RP M2 s vs xs V'') + (\c\S. card (RP M2 c vs xs V''))" by (meson insert.hyps(1) insert.hyps(2) sum.insert) then show ?thesis using \card (RP M2 s vs xs V'') = card (\a\RP M2 s vs xs V''. io_targets M1 (initial M1) a)\ \card (\s\insert s S. \a\RP M2 s vs xs V''. io_targets M1 (initial M1) a) = card (\s\S. \a\RP M2 s vs xs V''. io_targets M1 (initial M1) a) + card (\a\RP M2 s vs xs V''. io_targets M1 (initial M1) a)\ applied_IH by presburger qed qed qed lemma LB_count_helper_LB1 : assumes "(vs @ xs) \ L M2 \ L M1" and "OFSM M1" and "OFSM M2" and "asc_fault_domain M2 M1 m" and "test_tools M2 M1 FAIL PM V \" and "V'' \ Perm V M1" and "Prereq M2 M1 vs xs T S \ V''" and "\ Rep_Pre M2 M1 vs xs" and "\ Rep_Cov M2 M1 V'' vs xs" shows "(sum (\ s . card (RP M2 s vs xs V'')) S) \ card (nodes M1)" proof - have "(\s\S. UNION (RP M2 s vs xs V'') (io_targets M1 (initial M1))) \ nodes M1" by blast moreover have "finite (nodes M1)" using assms(2) OFSM_props(1) unfolding well_formed.simps finite_FSM.simps by simp ultimately have "card (\s\S. UNION (RP M2 s vs xs V'') (io_targets M1 (initial M1))) \ card (nodes M1)" by (meson card_mono) moreover have "(\s\S. card (RP M2 s vs xs V'')) = card (\s\S. UNION (RP M2 s vs xs V'') (io_targets M1 (initial M1)))" using LB_count_helper_RP_disjoint_M1_union[OF assms] by linarith ultimately show ?thesis by linarith qed lemma LB_count_helper_D_states : assumes "observable M" and "RS \ (D M T \)" obtains q where "q \ nodes M \ RS = IO_set M q \" proof - have "RS \ image (\ io . B M io \) (LS\<^sub>i\<^sub>n M (initial M) T)" using assms by simp then obtain io where "RS = B M io \" "io \ LS\<^sub>i\<^sub>n M (initial M) T" by blast then have "io \ language_state M (initial M)" using language_state_for_inputs_in_language_state[of M "initial M" T] by blast then obtain q where "{q} = io_targets M (initial M) io" by (metis assms(1) io_targets_observable_singleton_ob) then have "B M io \ = \ (image (\ s . IO_set M s \) {q})" by simp then have "B M io \ = IO_set M q \" by simp then have "RS = IO_set M q \" using \RS = B M io \\ by simp moreover have "q \ nodes M" using \{q} = io_targets M (initial M) io\ by (metis FSM.nodes.initial insertI1 io_targets_nodes) ultimately show ?thesis using that by simp qed lemma LB_count_helper_LB2 : assumes "observable M1" and "IO_set M1 q \ \ (D M1 T \) - {B M1 xs' \ | xs' s' . s' \ S \ xs' \ RP M2 s' vs xs V''}" shows "q \ (\ (image (\ s . \ (image (io_targets M1 (initial M1)) (RP M2 s vs xs V''))) S))" proof assume "q \ (\s\S. UNION (RP M2 s vs xs V'') (io_targets M1 (initial M1)))" then obtain s' where "s' \ S" "q \ (\ (image (io_targets M1 (initial M1)) (RP M2 s' vs xs V'')))" by blast then obtain xs' where "q \ io_targets M1 (initial M1) xs'" "xs' \ RP M2 s' vs xs V''" by blast then have "{q} = io_targets M1 (initial M1) xs'" by (metis assms(1) observable_io_target_is_singleton) then have "B M1 xs' \ = \ (image (\ s . IO_set M1 s \) {q})" by simp then have "B M1 xs' \ = IO_set M1 q \" by simp moreover have "B M1 xs' \ \ {B M1 xs' \ | xs' s' . s' \ S \ xs' \ RP M2 s' vs xs V''}" using \s' \ S\ \xs' \ RP M2 s' vs xs V''\ by blast ultimately have "IO_set M1 q \ \ {B M1 xs' \ | xs' s' . s' \ S \ xs' \ RP M2 s' vs xs V''}" by blast moreover have "IO_set M1 q \ \ {B M1 xs' \ | xs' s' . s' \ S \ xs' \ RP M2 s' vs xs V''}" using assms(2) by blast ultimately show "False" by simp qed subsection \ Validity of the result of LB constituting a lower bound \ lemma LB_count : assumes "(vs @ xs) \ L M1" and "OFSM M1" and "OFSM M2" and "asc_fault_domain M2 M1 m" and "test_tools M2 M1 FAIL PM V \" and "V'' \ Perm V M1" and "Prereq M2 M1 vs xs T S \ V''" and "\ Rep_Pre M2 M1 vs xs" and "\ Rep_Cov M2 M1 V'' vs xs" shows "LB M2 M1 vs xs T S \ V'' \ |M1|" proof - let ?D = "D M1 T \" let ?B = "{B M1 xs' \ | xs' s' . s' \ S \ xs' \ RP M2 s' vs xs V''}" let ?DB = "?D - ?B" let ?RP = "\s\S. \a\RP M2 s vs xs V''. io_targets M1 (initial M1) a" have "finite (nodes M1)" using OFSM_props[OF assms(2)] unfolding well_formed.simps finite_FSM.simps by simp then have "finite ?D" using OFSM_props[OF assms(2)] assms(7) D_bound[of M1 T \] unfolding Prereq.simps by linarith then have "finite ?DB" by simp \ \Proof sketch: Construct a function f (via induction) that maps each response set in ?DB to some state that produces that response set. This is then used to show that each response sets in ?DB indicates the existence of a distinct state in M1 not reached via the RP-sequences.\ have states_f : "\ DB' . DB' \ ?DB \ \ f . inj_on f DB' \ image f DB' \ (nodes M1) - ?RP \ (\ RS \ DB' . IO_set M1 (f RS) \ = RS)" proof - fix DB' assume "DB' \ ?DB" have "finite DB'" proof (rule ccontr) assume "infinite DB'" have "infinite ?DB" using infinite_super[OF \DB' \ ?DB\ \infinite DB'\ ] by simp then show "False" using \finite ?DB\ by simp qed then show "\ f . inj_on f DB' \ image f DB' \ (nodes M1) - ?RP \ (\ RS \ DB' . IO_set M1 (f RS) \ = RS)" using assms \DB' \ ?DB\ proof (induction DB') case empty show ?case by simp next case (insert RS DB') have "DB' \ ?DB" using insert.prems(10) by blast obtain f' where "inj_on f' DB'" "image f' DB' \ (nodes M1) - ?RP" "\ RS \ DB' . IO_set M1 (f' RS) \ = RS" using insert.IH[OF insert.prems(1-9) \DB' \ ?DB\] by blast have "RS \ D M1 T \" using insert.prems(10) by blast obtain q where "q \ nodes M1" "RS = IO_set M1 q \" using insert.prems(2) LB_count_helper_D_states[OF _ \RS \ D M1 T \\] by blast then have "IO_set M1 q \ \ ?DB" using insert.prems(10) by blast have "q \ ?RP" using insert.prems(2) LB_count_helper_LB2[OF _ \IO_set M1 q \ \ ?DB\] by blast let ?f = "f'(RS := q)" have "inj_on ?f (insert RS DB')" proof have "?f RS \ ?f ` (DB' - {RS})" proof assume "?f RS \ ?f ` (DB' - {RS})" then have "q \ ?f ` (DB' - {RS})" by auto have "RS \ DB'" proof - have "\P c f. \Pa. ((c::'c) \ f ` P \ (Pa::('a \ 'b) list set) \ P) \ (c \ f ` P \ f Pa = c)" by auto moreover { assume "q \ f' ` DB'" moreover { assume "q \ f'(RS := q) ` DB'" then have ?thesis using \q \ f'(RS := q) ` (DB' - {RS})\ by blast } ultimately have ?thesis by (metis fun_upd_image) } ultimately show ?thesis by (metis (no_types) \RS = IO_set M1 q \\ \\RS\DB'. IO_set M1 (f' RS) \ = RS\) qed then show "False" using insert.hyps(2) by simp qed then show "inj_on ?f DB' \ ?f RS \ ?f ` (DB' - {RS})" using \inj_on f' DB'\ inj_on_fun_updI by fastforce qed moreover have "image ?f (insert RS DB') \ (nodes M1) - ?RP" proof - have "image ?f {RS} = {q}" by simp then have "image ?f {RS} \ (nodes M1) - ?RP" using \q \ nodes M1\ \q \ ?RP\ by auto moreover have "image ?f (insert RS DB') = image ?f {RS} \ image ?f DB'" by auto ultimately show ?thesis by (metis (no_types, lifting) \image f' DB' \ (nodes M1) - ?RP\ fun_upd_other image_cong image_insert insert.hyps(2) insert_subset) qed moreover have "\ RS \ (insert RS DB') . IO_set M1 (?f RS) \ = RS" using \RS = IO_set M1 q \\ \\RS\DB'. IO_set M1 (f' RS) \ = RS\ by auto ultimately show ?case by blast qed qed have "?DB \ ?DB" by simp obtain f where "inj_on f ?DB" "image f ?DB \ (nodes M1) - ?RP" using states_f[OF \?DB \ ?DB\] by blast have "finite (nodes M1 - ?RP)" using \finite (nodes M1)\ by simp have "card ?DB \ card (nodes M1 - ?RP)" using card_inj_on_le[OF \inj_on f ?DB\ \image f ?DB \ (nodes M1) - ?RP\ \finite (nodes M1 - ?RP)\] by assumption have "?RP \ nodes M1" by blast then have "card (nodes M1 - ?RP) = card (nodes M1) - card ?RP" by (meson \finite (nodes M1)\ card_Diff_subset infinite_subset) then have "card ?DB \ card (nodes M1) - card ?RP" using \card ?DB \ card (nodes M1 - ?RP)\ by linarith have "vs @ xs \ L M2 \ L M1" using assms(7) by simp have "(sum (\ s . card (RP M2 s vs xs V'')) S) = card ?RP" using LB_count_helper_RP_disjoint_M1_union[OF \vs @ xs \ L M2 \ L M1\ assms(2-9)] by simp moreover have "card ?RP \ card (nodes M1)" using card_mono[OF \finite (nodes M1)\ \?RP \ nodes M1\] by assumption ultimately show ?thesis unfolding LB.simps using \card ?DB \ card (nodes M1) - card ?RP\ by linarith qed lemma contradiction_via_LB : assumes "(vs @ xs) \ L M1" and "OFSM M1" and "OFSM M2" and "asc_fault_domain M2 M1 m" and "test_tools M2 M1 FAIL PM V \" and "V'' \ Perm V M1" and "Prereq M2 M1 vs xs T S \ V''" and "\ Rep_Pre M2 M1 vs xs" and "\ Rep_Cov M2 M1 V'' vs xs" and "LB M2 M1 vs xs T S \ V'' > m" shows "False" proof - have "LB M2 M1 vs xs T S \ V'' \ card (nodes M1)" using LB_count[OF assms(1-9)] by assumption moreover have "card (nodes M1) \ m" using assms(4) by auto ultimately show "False" using assms(10) by linarith qed end \ No newline at end of file diff --git a/thys/Adaptive_State_Counting/FSM/FSM.thy b/thys/Adaptive_State_Counting/FSM/FSM.thy --- a/thys/Adaptive_State_Counting/FSM/FSM.thy +++ b/thys/Adaptive_State_Counting/FSM/FSM.thy @@ -1,2031 +1,2031 @@ theory FSM imports "Transition_Systems_and_Automata.Sequence_Zip" "Transition_Systems_and_Automata.Transition_System" "Transition_Systems_and_Automata.Transition_System_Extra" "Transition_Systems_and_Automata.Transition_System_Construction" begin section \ Finite state machines \ text \ We formalise finite state machines as a 4-tuples, omitting the explicit formulation of the state set,as it can easily be calculated from the successor function. This definition does not require the successor function to be restricted to the input or output alphabet, which is later expressed by the property @{verbatim well_formed}, together with the finiteness of the state set. \ record ('in, 'out, 'state) FSM = succ :: "('in \ 'out) \ 'state \ 'state set" inputs :: "'in set" outputs :: "'out set" initial :: "'state" subsection \ FSMs as transition systems \ text \ We interpret FSMs as transition systems with a singleton initial state set, based on \<^cite>\"Transition_Systems_and_Automata-AFP"\. \ global_interpretation FSM : transition_system_initial "\ a p. snd a" \ \execute\ "\ a p. snd a \ succ A (fst a) p" \ \enabled\ "\ p. p = initial A" \ \initial\ for A defines path = FSM.path and run = FSM.run and reachable = FSM.reachable and nodes = FSM.nodes by this abbreviation "size_FSM M \ card (nodes M)" notation size_FSM ("(|_|)") subsection \ Language \ text \ The following definitions establish basic notions for FSMs similarly to those of nondeterministic finite automata as defined in \<^cite>\"Transition_Systems_and_Automata-AFP"\. In particular, the language of an FSM state are the IO-parts of the paths in the FSM enabled from that state. \ abbreviation "target \ FSM.target" abbreviation "states \ FSM.states" abbreviation "trace \ FSM.trace" abbreviation successors :: "('in, 'out, 'state, 'more) FSM_scheme \ 'state \ 'state set" where "successors \ FSM.successors TYPE('in) TYPE('out) TYPE('more)" lemma states_alt_def: "states r p = map snd r" by (induct r arbitrary: p) (auto) lemma trace_alt_def: "trace r p = smap snd r" by (coinduction arbitrary: r p) (auto) definition language_state :: "('in, 'out, 'state) FSM \ 'state \ ('in \ 'out) list set" ("LS") where "language_state M q \ {map fst r |r . path M r q}" text \ The language of an FSM is the language of its initial state. \ abbreviation "L M \ LS M (initial M)" lemma language_state_alt_def : "LS M q = {io | io tr . path M (io || tr) q \ length io = length tr}" proof - have "LS M q \ { io | io tr . path M (io || tr) q \ length io = length tr }" proof fix xr assume xr_assm : "xr \ LS M q" then obtain r where r_def : "map fst r = xr" "path M r q" unfolding language_state_def by auto then obtain xs ys where xr_split : "xr = xs || ys" "length xs = length ys" "length xs = length xr" by (metis length_map zip_map_fst_snd) then have "(xs || ys) \ { io | io tr . path M (io || tr) q \ length io = length tr }" proof - have f1: "xs || ys = map fst r" by (simp add: r_def(1) xr_split(1)) then have f2: "path M ((xs || ys) || take (min (length (xs || ys)) (length (map snd r))) (map snd r)) q" by (simp add: r_def(2)) have "length (xs || ys) = length (take (min (length (xs || ys)) (length (map snd r))) (map snd r))" using f1 by force then show ?thesis using f2 by blast qed then show "xr \ { io | io tr . path M (io || tr) q \ length io = length tr }" using xr_split by metis qed moreover have "{ io | io tr . path M (io || tr) q \ length io = length tr } \ LS M q" proof fix xs assume xs_assm : "xs \ { io | io tr . path M (io || tr) q \ length io = length tr }" then obtain ys where ys_def : "path M (xs || ys) q" "length xs = length ys" by auto then have "xs = map fst (xs || ys)" by auto then show "xs \ LS M q" using ys_def unfolding language_state_def by blast qed ultimately show ?thesis by auto qed lemma language_state[intro]: assumes "path M (w || r) q" "length w = length r" shows "w \ LS M q" using assms unfolding language_state_def by force lemma language_state_elim[elim]: assumes "w \ LS M q" obtains r where "path M (w || r) q" "length w = length r" using assms unfolding language_state_def by (force iff: split_zip_ex) lemma language_state_split: assumes "w1 @ w2 \ LS M q" obtains tr1 tr2 where "path M (w1 || tr1) q" "length w1 = length tr1" "path M (w2 || tr2) (target (w1 || tr1) q)" "length w2 = length tr2" proof - obtain tr where tr_def : "path M ((w1 @ w2) || tr) q" "length (w1 @ w2) = length tr" using assms by blast let ?tr1 = "take (length w1) tr" let ?tr2 = "drop (length w1) tr" have tr_split : "?tr1 @ ?tr2 = tr" by auto then show ?thesis proof - have f1: "length w1 + length w2 = length tr" using tr_def(2) by auto then have f2: "length w2 = length tr - length w1" by presburger then have "length w1 = length (take (length w1) tr)" using f1 by (metis (no_types) tr_split diff_add_inverse2 length_append length_drop) then show ?thesis using f2 by (metis (no_types) FSM.path_append_elim length_drop that tr_def(1) zip_append1) qed qed lemma language_state_prefix : assumes "w1 @ w2 \ LS M q" shows "w1 \ LS M q" using assms by (meson language_state language_state_split) lemma succ_nodes : fixes A :: "('a,'b,'c) FSM" and w :: "('a \ 'b)" assumes "q2 \ succ A w q1" and "q1 \ nodes A" shows "q2 \ nodes A" proof - obtain x y where "w = (x,y)" by (meson surj_pair) then have "q2 \ successors A q1" using assms by auto then have "q2 \ reachable A q1" by blast then have "q2 \ reachable A (initial A)" using assms by blast then show "q2 \ nodes A" by blast qed lemma states_target_index : assumes "i < length p" shows "(states p q1) ! i = target (take (Suc i) p) q1" using assms by auto subsection \ Product machine for language intersection \ text \ The following describes the construction of a product machine from two FSMs @{verbatim M1} and @{verbatim M2} such that the language of the product machine is the intersection of the language of @{verbatim M1} and the language of @{verbatim M2}. \ definition product :: "('in, 'out, 'state1) FSM \ ('in, 'out, 'state2) FSM \ ('in, 'out, 'state1 \'state2) FSM" where "product A B \ \ succ = \ a (p\<^sub>1, p\<^sub>2). succ A a p\<^sub>1 \ succ B a p\<^sub>2, inputs = inputs A \ inputs B, outputs = outputs A \ outputs B, initial = (initial A, initial B) \" lemma product_simps[simp]: "succ (product A B) a (p\<^sub>1, p\<^sub>2) = succ A a p\<^sub>1 \ succ B a p\<^sub>2" "inputs (product A B) = inputs A \ inputs B" "outputs (product A B) = outputs A \ outputs B" "initial (product A B) = (initial A, initial B)" unfolding product_def by simp+ lemma product_target[simp]: assumes "length w = length r\<^sub>1" "length r\<^sub>1 = length r\<^sub>2" shows "target (w || r\<^sub>1 || r\<^sub>2) (p\<^sub>1, p\<^sub>2) = (target (w || r\<^sub>1) p\<^sub>1, target (w || r\<^sub>2) p\<^sub>2)" using assms by (induct arbitrary: p\<^sub>1 p\<^sub>2 rule: list_induct3) (auto) lemma product_path[iff]: assumes "length w = length r\<^sub>1" "length r\<^sub>1 = length r\<^sub>2" shows "path (product A B) (w || r\<^sub>1 || r\<^sub>2) (p\<^sub>1, p\<^sub>2) \ path A (w || r\<^sub>1) p\<^sub>1 \ path B (w || r\<^sub>2) p\<^sub>2" using assms by (induct arbitrary: p\<^sub>1 p\<^sub>2 rule: list_induct3) (auto) lemma product_language_state[simp]: "LS (product A B) (q1,q2) = LS A q1 \ LS B q2" by (fastforce iff: split_zip) lemma product_nodes : "nodes (product A B) \ nodes A \ nodes B" proof fix q assume "q \ nodes (product A B)" then show "q \ nodes A \ nodes B" proof (induction rule: FSM.nodes.induct) case (initial p) then show ?case by auto next case (execute p a) then have "fst p \ nodes A" "snd p \ nodes B" by auto have "snd a \ (succ A (fst a) (fst p)) \ (succ B (fst a) (snd p))" using execute by auto then have "fst (snd a) \ succ A (fst a) (fst p)" "snd (snd a) \ succ B (fst a) (snd p)" by auto have "fst (snd a) \ nodes A" using \fst p \ nodes A\ \fst (snd a) \ succ A (fst a) (fst p)\ by (metis FSM.nodes.simps fst_conv snd_conv) moreover have "snd (snd a) \ nodes B" using \snd p \ nodes B\ \snd (snd a) \ succ B (fst a) (snd p)\ by (metis FSM.nodes.simps fst_conv snd_conv) ultimately show ?case by (simp add: mem_Times_iff) qed qed subsection \ Required properties \ text \ FSMs used by the adaptive state counting algorithm are required to satisfy certain properties which are introduced in here. Most notably, the observability property (see function @{verbatim observable}) implies the uniqueness of certain paths and hence allows for several stronger variations of previous results. \ fun finite_FSM :: "('in, 'out, 'state) FSM \ bool" where "finite_FSM M = (finite (nodes M) \ finite (inputs M) \ finite (outputs M))" fun observable :: "('in, 'out, 'state) FSM \ bool" where "observable M = (\ t . \ s1 . ((succ M) t s1 = {}) \ (\ s2 . (succ M) t s1 = {s2}))" fun completely_specified :: "('in, 'out, 'state) FSM \ bool" where "completely_specified M = (\ s1 \ nodes M . \ x \ inputs M . \ y \ outputs M . \ s2 . s2 \ (succ M) (x,y) s1)" fun well_formed :: "('in, 'out, 'state) FSM \ bool" where "well_formed M = (finite_FSM M \ (\ s1 x y . (x \ inputs M \ y \ outputs M) \ succ M (x,y) s1 = {}) \ inputs M \ {} \ outputs M \ {})" abbreviation "OFSM M \ well_formed M \ observable M \ completely_specified M" lemma OFSM_props[elim!] : assumes "OFSM M" shows "well_formed M" "observable M" "completely_specified M" using assms by auto lemma set_of_succs_finite : assumes "well_formed M" and "q \ nodes M" shows "finite (succ M io q)" proof (rule ccontr) assume "infinite (succ M io q)" moreover have "succ M io q \ nodes M" using assms by (simp add: subsetI succ_nodes) ultimately have "infinite (nodes M)" using infinite_super by blast then show "False" using assms by auto qed lemma well_formed_path_io_containment : assumes "well_formed M" and "path M p q" shows "set (map fst p) \ (inputs M \ outputs M)" using assms proof (induction p arbitrary: q) case Nil then show ?case by auto next case (Cons a p) have "fst a \ (inputs M \ outputs M)" proof (rule ccontr) assume "fst a \ inputs M \ outputs M" then have "fst (fst a) \ inputs M \ snd (fst a) \ outputs M" by (metis SigmaI prod.collapse) then have "succ M (fst a) q = {}" using Cons by (metis prod.collapse well_formed.elims(2)) moreover have "(snd a) \ succ M (fst a) q" using Cons by auto ultimately show "False" by auto qed moreover have "set (map fst p) \ (inputs M \ outputs M)" using Cons by blast ultimately show ?case by auto qed lemma path_input_containment : assumes "well_formed M" and "path M p q" shows "set (map fst (map fst p)) \ inputs M" using assms proof (induction p arbitrary: q rule: rev_induct) case Nil then show ?case by auto next case (snoc a p) have "set (map fst (p @ [a])) \ (inputs M \ outputs M)" using well_formed_path_io_containment[OF snoc.prems] by assumption then have "(fst a) \ (inputs M \ outputs M)" by auto then have "fst (fst a) \ inputs M" by auto moreover have "set (map fst (map fst p)) \ inputs M" using snoc.IH[OF snoc.prems(1)] using snoc.prems(2) by blast ultimately show ?case by simp qed lemma path_state_containment : assumes "path M p q" and "q \ nodes M" shows "set (map snd p) \ nodes M" using assms by (metis FSM.nodes_states states_alt_def) lemma language_state_inputs : assumes "well_formed M" and "io \ language_state M q" shows "set (map fst io) \ inputs M" proof - obtain tr where "path M (io || tr) q" "length tr = length io" using assms(2) by auto show ?thesis by (metis (no_types) \\thesis. (\tr. \path M (io || tr) q; length tr = length io\ \ thesis) \ thesis\ assms(1) map_fst_zip path_input_containment) qed lemma set_of_paths_finite : assumes "well_formed M" and "q1 \ nodes M" shows "finite { p . path M p q1 \ target p q1 = q2 \ length p \ k }" proof - let ?trs = "{ tr . set tr \ nodes M \ length tr \ k }" let ?ios = "{ io . set io \ inputs M \ outputs M \ length io \ k }" let ?iotrs = "image (\ (io,tr) . io || tr) (?ios \ ?trs)" let ?paths = "{ p . path M p q1 \ target p q1 = q2 \ length p \ k }" have "finite (inputs M \ outputs M)" using assms by auto then have "finite ?ios" using assms by (simp add: finite_lists_length_le) moreover have "finite ?trs" using assms by (simp add: finite_lists_length_le) ultimately have "finite ?iotrs" by auto moreover have "?paths \ ?iotrs" proof fix p assume p_assm : "p \ { p . path M p q1 \ target p q1 = q2 \ length p \ k }" then obtain io tr where p_split : "p = io || tr \ length io = length tr" using that by (metis (no_types) length_map zip_map_fst_snd) then have "io \ ?ios" using well_formed_path_io_containment proof - have f1: "path M p q1 \ target p q1 = q2 \ length p \ k" using p_assm by force then have "set io \ inputs M \ outputs M" by (metis (no_types) assms(1) map_fst_zip p_split well_formed_path_io_containment) then show ?thesis using f1 by (simp add: p_split) qed moreover have "tr \ ?trs" using p_split proof - have f1: "path M (io || tr) q1 \ target (io || tr) q1 = q2 \ length (io || tr) \ k" using \p \ {p. path M p q1 \ target p q1 = q2 \ length p \ k}\ p_split by force then have f2: "length tr \ k" by (simp add: p_split) have "set tr \ nodes M" using f1 by (metis (no_types) assms(2) length_map p_split path_state_containment zip_eq zip_map_fst_snd) then show ?thesis using f2 by blast qed ultimately show "p \ ?iotrs" using p_split by auto qed ultimately show ?thesis using Finite_Set.finite_subset by blast qed lemma non_distinct_duplicate_indices : assumes "\ distinct xs" shows "\ i1 i2 . i1 \ i2 \ xs ! i1 = xs ! i2 \ i1 \ length xs \ i2 \ length xs" using assms by (meson distinct_conv_nth less_imp_le) lemma reaching_path_without_repetition : assumes "well_formed M" and "q2 \ reachable M q1" and "q1 \ nodes M" shows "\ p . path M p q1 \ target p q1 = q2 \ distinct (q1 # states p q1)" proof - have shorten_nondistinct : "\ p. (path M p q1 \ target p q1 = q2 \ \ distinct (q1 # states p q1)) \ (\ p' . path M p' q1 \ target p' q1 = q2 \ length p' < length p)" proof fix p show "(path M p q1 \ target p q1 = q2 \ \ distinct (q1 # states p q1)) \ (\ p' . path M p' q1 \ target p' q1 = q2 \ length p' < length p)" proof assume assm : "path M p q1 \ target p q1 = q2 \ \ distinct (q1 # states p q1)" then show "(\p'. path M p' q1 \ target p' q1 = q2 \ length p' < length p)" proof (cases "q1 \ set (states p q1)") case True have "\ i1 . target (take i1 p) q1 = q1 \ i1 \ length p \ i1 > 0" proof (rule ccontr) assume "\ (\ i1. target (take i1 p) q1 = q1 \ i1 \ length p \ i1 > 0)" then have "\ (\ i1 . (states p q1) ! i1 = q1 \ i1 \ length (states p q1))" by (metis True in_set_conv_nth less_eq_Suc_le scan_length scan_nth zero_less_Suc) then have "q1 \ set (states p q1)" by (meson in_set_conv_nth less_imp_le) then show "False" using True by auto qed then obtain i1 where i1_def : "target (take i1 p) q1 = q1 \ i1 \ length p \ i1 > 0" by auto then have "path M (take i1 p) q1" using assm by (metis FSM.path_append_elim append_take_drop_id) moreover have "path M (drop i1 p) q1" using i1_def by (metis FSM.path_append_elim append_take_drop_id assm) ultimately have "path M (drop i1 p) q1 \ (target (drop i1 p) q1 = q2)" using i1_def by (metis (no_types) append_take_drop_id assm fold_append o_apply) moreover have "length (drop i1 p) < length p" using i1_def by auto ultimately show ?thesis using assms by blast next case False then have assm' : "path M p q1 \ target p q1 = q2 \ \ distinct (states p q1)" using assm by auto have "\ i1 i2 . i1 \ i2 \ target (take i1 p) q1 = target (take i2 p) q1 \ i1 \ length p \ i2 \ length p" proof (rule ccontr) assume "\ (\ i1 i2 . i1 \ i2 \ target (take i1 p) q1 = target (take i2 p) q1 \ i1 \ length p \ i2 \ length p)" then have "\ (\ i1 i2 . i1 \ i2 \ (states p q1) ! i1 = (states p q1) ! i2 \ i1 \ length (states p q1) \ i2 \ length (states p q1))" by (metis (no_types, lifting) Suc_leI assm' distinct_conv_nth nat.inject scan_length scan_nth) then have "distinct (states p q1)" using non_distinct_duplicate_indices by blast then show "False" using assm' by auto qed then obtain i1 i2 where i_def : "i1 < i2 \ target (take i1 p) q1 = target (take i2 p) q1 \ i1 \ length p \ i2 \ length p" by (metis nat_neq_iff) then have "path M (take i1 p) q1" using assm by (metis FSM.path_append_elim append_take_drop_id) moreover have "path M (drop i2 p) (target (take i2 p) q1)" by (metis FSM.path_append_elim append_take_drop_id assm) ultimately have "path M ((take i1 p) @ (drop i2 p)) q1 \ (target ((take i1 p) @ (drop i2 p)) q1 = q2)" using i_def assm by (metis FSM.path_append append_take_drop_id fold_append o_apply) moreover have "length ((take i1 p) @ (drop i2 p)) < length p" using i_def by auto ultimately have "path M ((take i1 p) @ (drop i2 p)) q1 \ target ((take i1 p) @ (drop i2 p)) q1 = q2 \ length ((take i1 p) @ (drop i2 p)) < length p" by simp then show ?thesis using assms by blast qed qed qed obtain p where p_def : "path M p q1 \ target p q1 = q2" using assms by auto let ?paths = "{p' . (path M p' q1 \ target p' q1 = q2 \ length p' \ length p)}" let ?minPath = "arg_min length (\ io . io \ ?paths)" have "?paths \ empty" using p_def by auto moreover have "finite ?paths" using assms by (simp add: set_of_paths_finite) ultimately have minPath_def : "?minPath \ ?paths \ (\ p' \ ?paths . length ?minPath \ length p')" by (meson arg_min_nat_lemma equals0I) moreover have "distinct (q1 # states ?minPath q1)" proof (rule ccontr) assume "\ distinct (q1 # states ?minPath q1)" then have "\ p' . path M p' q1 \ target p' q1 = q2 \ length p' < length ?minPath" using shorten_nondistinct minPath_def by blast then show "False" using minPath_def using arg_min_nat_le dual_order.strict_trans1 by auto qed ultimately show ?thesis by auto qed lemma observable_path_unique[simp] : assumes "io \ LS M q" and "observable M" and "path M (io || tr1) q" "length io = length tr1" and "path M (io || tr2) q" "length io = length tr2" shows "tr1 = tr2" proof (rule ccontr) assume tr_assm : "tr1 \ tr2" then have state_diff : "(states (io || tr1) q ) \ (states (io || tr2) q)" by (metis assms(4) assms(6) map_snd_zip states_alt_def) show "False" using assms tr_assm proof (induction io arbitrary: q tr1 tr2) case Nil then show ?case using Nil by simp next case (Cons io_hd io_tl) then obtain tr1_hd tr1_tl tr2_hd tr2_tl where tr_split : "tr1 = tr1_hd # tr1_tl \ tr2 = tr2_hd # tr2_tl" by (metis length_0_conv neq_Nil_conv) have p1: "path M ([io_hd] || [tr1_hd]) q" using Cons.prems tr_split by auto have p2: "path M ([io_hd] || [tr2_hd]) q" using Cons.prems tr_split by auto have tr_hd_eq : "tr1_hd = tr2_hd" using Cons.prems unfolding observable.simps proof - assume "\t s1. succ M t s1 = {} \ (\s2. succ M t s1 = {s2})" then show ?thesis by (metis (no_types) p1 p2 FSM.path_cons_elim empty_iff prod.sel(1) prod.sel(2) singletonD zip_Cons_Cons) qed then show ?thesis using Cons.IH Cons.prems(3) Cons.prems(4) Cons.prems(5) Cons.prems(6) Cons.prems(7) assms(2) tr_split by auto qed qed lemma observable_path_unique_ex[elim] : assumes "observable M" and "io \ LS M q" obtains tr where "{ t . path M (io || t) q \ length io = length t } = { tr }" proof - obtain tr where tr_def : "path M (io || tr) q" "length io = length tr" using assms by auto then have "{ t . path M (io || t) q \ length io = length t } \ {}" by blast moreover have "\ t \ { t . path M (io || t) q \ length io = length t } . t = tr" using assms tr_def by auto ultimately show ?thesis - using that by moura + using that by auto qed lemma well_formed_product[simp] : assumes "well_formed M1" and "well_formed M2" shows "well_formed (product M2 M1)" (is "well_formed ?PM") unfolding well_formed.simps proof have "finite (nodes M1)" "finite (nodes M2)" using assms by auto then have "finite (nodes M2 \ nodes M1)" by simp moreover have "nodes ?PM \ nodes M2 \ nodes M1" using product_nodes assms by blast ultimately show "finite_FSM ?PM" using infinite_subset assms by auto next have "inputs ?PM = inputs M2 \ inputs M1" "outputs ?PM = outputs M2 \ outputs M1" by auto then show "(\s1 x y. x \ inputs ?PM \ y \ outputs ?PM \ succ ?PM (x, y) s1 = {}) \ inputs ?PM \ {} \ outputs ?PM \ {}" using assms by auto qed subsection \ States reached by a given IO-sequence \ text \ Function @{verbatim io_targets} collects all states of an FSM reached from a given state by a given IO-sequence. Notably, for any observable FSM, this set contains at most one state. \ fun io_targets :: "('in, 'out, 'state) FSM \ 'state \ ('in \ 'out) list \ 'state set" where "io_targets M q io = { target (io || tr) q | tr . path M (io || tr) q \ length io = length tr }" lemma io_target_implies_L : assumes "q \ io_targets M (initial M) io" shows "io \ L M" proof - obtain tr where "path M (io || tr) (initial M)" "length tr = length io" "target (io || tr) (initial M) = q" using assms by auto then show ?thesis by auto qed lemma io_target_from_path : assumes "path M (w || tr) q" and "length w = length tr" shows "target (w || tr) q \ io_targets M q w" using assms by auto lemma io_targets_observable_singleton_ex : assumes "observable M" and "io \ LS M q1" shows "\ q2 . io_targets M q1 io = { q2 }" proof - obtain tr where tr_def : "{ t . path M (io || t) q1 \ length io = length t } = { tr }" using assms observable_path_unique_ex by (metis (mono_tags, lifting)) then have "io_targets M q1 io = { target (io || tr) q1 }" by fastforce then show ?thesis by blast qed lemma io_targets_observable_singleton_ob : assumes "observable M" and "io \ LS M q1" obtains q2 where "io_targets M q1 io = { q2 }" proof - obtain tr where tr_def : "{ t . path M (io || t) q1 \ length io = length t } = { tr }" using assms observable_path_unique_ex by (metis (mono_tags, lifting)) then have "io_targets M q1 io = { target (io || tr) q1 }" by fastforce then show ?thesis using that by blast qed lemma io_targets_elim[elim] : assumes "p \ io_targets M q io" obtains tr where "target (io || tr) q = p \ path M (io || tr) q \ length io = length tr" using assms unfolding io_targets.simps by force lemma io_targets_reachable : assumes "q2 \ io_targets M q1 io" shows "q2 \ reachable M q1" using assms unfolding io_targets.simps by blast lemma io_targets_nodes : assumes "q2 \ io_targets M q1 io" and "q1 \ nodes M" shows "q2 \ nodes M" using assms by auto lemma observable_io_targets_split : assumes "observable M" and "io_targets M q1 (vs @ xs) = {q3}" and "io_targets M q1 vs = {q2}" shows "io_targets M q2 xs = {q3}" proof - have "vs @ xs \ LS M q1" using assms(2) by force then obtain trV trX where tr_def : "path M (vs || trV) q1" "length vs = length trV" "path M (xs || trX) (target (vs || trV) q1)" "length xs = length trX" using language_state_split[of vs xs M q1] by auto then have tgt_V : "target (vs || trV) q1 = q2" using assms(3) by auto then have path_X : "path M (xs || trX) q2 \ length xs = length trX" using tr_def by auto have tgt_all : "target (vs @ xs || trV @ trX) q1 = q3" proof - have f1: "\cs. q3 = target (vs @ xs || cs) q1 \ path M (vs @ xs || cs) q1 \ length (vs @ xs) = length cs" using assms(2) by auto have "length (vs @ xs) = length trV + length trX" by (simp add: tr_def(2) tr_def(4)) then have "length (vs @ xs) = length (trV @ trX)" by simp then show ?thesis using f1 by (metis FSM.path_append \vs @ xs \ LS M q1\ assms(1) observable_path_unique tr_def(1) tr_def(2) tr_def(3) zip_append) qed then have "target ((vs || trV) @ (xs || trX)) q1 = q3" using tr_def by simp then have "target (xs || trX) q2 = q3" using tgt_V by auto then have "q3 \ io_targets M q2 xs" using path_X by auto then show ?thesis by (metis (no_types) \observable M\ path_X insert_absorb io_targets_observable_singleton_ex language_state singleton_insert_inj_eq') qed lemma observable_io_target_unique_target : assumes "observable M" and "io_targets M q1 io = {q2}" and "path M (io || tr) q1" and "length io = length tr" shows "target (io || tr) q1 = q2" using assms by auto lemma target_in_states : assumes "length io = length tr" and "length io > 0" shows "last (states (io || tr) q) = target (io || tr) q" proof - have "0 < length tr" using assms(1) assms(2) by presburger then show ?thesis by (simp add: FSM.target_alt_def assms(1) states_alt_def) qed lemma target_alt_def : assumes "length io = length tr" shows "length io = 0 \ target (io || tr) q = q" "length io > 0 \ target (io || tr) q = last tr" proof - show "length io = 0 \ target (io || tr) q = q" by simp show "length io > 0 \ target (io || tr) q = last tr" by (metis assms last_ConsR length_greater_0_conv map_snd_zip scan_last states_alt_def) qed lemma obs_target_is_io_targets : assumes "observable M" and "path M (io || tr) q" and "length io = length tr" shows "io_targets M q io = {target (io || tr) q}" by (metis assms(1) assms(2) assms(3) io_targets_observable_singleton_ex language_state observable_io_target_unique_target) lemma io_target_target : assumes "io_targets M q1 io = {q2}" and "path M (io || tr) q1" and "length io = length tr" shows "target (io || tr) q1 = q2" proof - have "target (io || tr) q1 \ io_targets M q1 io" using assms(2) assms(3) by auto then show ?thesis using assms(1) by blast qed lemma index_last_take : assumes "i < length xs" shows "xs ! i = last (take (Suc i) xs)" by (simp add: assms take_Suc_conv_app_nth) lemma path_last_io_target : assumes "path M (xs || tr) q" and "length xs = length tr" and "length xs > 0" shows "last tr \ io_targets M q xs" proof - have "last tr = target (xs || tr) q" by (metis assms(2) assms(3) map_snd_zip states_alt_def target_in_states) then show ?thesis using assms(1) assms(2) by auto qed lemma path_prefix_io_targets : assumes "path M (xs || tr) q" and "length xs = length tr" and "length xs > 0" shows "last (take (Suc i) tr) \ io_targets M q (take (Suc i) xs)" proof - have "path M (take (Suc i) xs || take (Suc i) tr) q" by (metis (no_types) FSM.path_append_elim append_take_drop_id assms(1) take_zip) then show ?thesis using assms(2) assms(3) path_last_io_target by fastforce qed lemma states_index_io_target : assumes "i < length xs" and "path M (xs || tr) q" and "length xs = length tr" and "length xs > 0" shows "(states (xs || tr) q) ! i \ io_targets M q (take (Suc i) xs)" proof - have "(states (xs || tr) q) ! i = last (take (Suc i) (states (xs || tr) q))" by (metis assms(1) assms(3) map_snd_zip states_alt_def index_last_take) then have "(states (xs || tr) q) ! i = last (states (take (Suc i) xs || take (Suc i) tr) q)" by (simp add: take_zip) then have "(states (xs || tr) q) ! i = last (take (Suc i) tr)" by (simp add: assms(3) states_alt_def) moreover have "last (take (Suc i) tr) \ io_targets M q (take (Suc i) xs)" by (meson assms(2) assms(3) assms(4) path_prefix_io_targets) ultimately show ?thesis by simp qed lemma observable_io_targets_append : assumes "observable M" and "io_targets M q1 vs = {q2}" and "io_targets M q2 xs = {q3}" shows "io_targets M q1 (vs@xs) = {q3}" proof - obtain trV where "path M (vs || trV) q1 \ length trV = length vs \ target (vs || trV) q1 = q2" by (metis assms(2) io_targets_elim singletonI) moreover obtain trX where "path M (xs || trX) q2 \ length trX = length xs \ target (xs || trX) q2 = q3" by (metis assms(3) io_targets_elim singletonI) ultimately have "path M (vs @ xs || trV @ trX) q1 \ length (trV @ trX) = length (vs @ xs) \ target (vs @ xs || trV @ trX) q1 = q3" by auto then show ?thesis by (metis assms(1) obs_target_is_io_targets) qed lemma io_path_states_prefix : assumes "observable M" and "path M (io1 || tr1) q" and "length tr1 = length io1" and "path M (io2 || tr2) q" and "length tr2 = length io2" and "prefix io1 io2" shows "tr1 = take (length tr1) tr2" proof - let ?tr1' = "take (length tr1) tr2" let ?io1' = "take (length tr1) io2" have "path M (?io1' || ?tr1') q" by (metis FSM.path_append_elim append_take_drop_id assms(4) take_zip) have "length ?tr1' = length ?io1'" using assms (5) by auto have "?io1' = io1" proof - have "\ps psa. \ prefix (ps::('a \ 'b) list) psa \ length ps \ length psa" using prefix_length_le by blast then have "length (take (length tr1) io2) = length io1" using assms(3) assms(6) min.absorb2 by auto then show ?thesis by (metis assms(6) min.cobounded2 min_def_raw prefix_length_prefix prefix_order.dual_order.antisym take_is_prefix) qed show "tr1 = ?tr1'" by (metis \length (take (length tr1) tr2) = length (take (length tr1) io2)\ \path M (take (length tr1) io2 || take (length tr1) tr2) q\ \take (length tr1) io2 = io1\ assms(1) assms(2) assms(3) language_state observable_path_unique) qed lemma observable_io_targets_suffix : assumes "observable M" and "io_targets M q1 vs = {q2}" and "io_targets M q1 (vs@xs) = {q3}" shows "io_targets M q2 xs = {q3}" proof - have "prefix vs (vs@xs)" by auto obtain trV where "path M (vs || trV) q1 \ length trV = length vs \ target (vs || trV) q1 = q2" by (metis assms(2) io_targets_elim singletonI) moreover obtain trVX where "path M (vs@xs || trVX) q1 \ length trVX = length (vs@xs) \ target (vs@xs || trVX) q1 = q3" by (metis assms(3) io_targets_elim singletonI) ultimately have "trV = take (length trV) trVX" using io_path_states_prefix[OF assms(1) _ _ _ _ \prefix vs (vs@xs)\, of trV q1 trVX] by auto show ?thesis by (meson assms(1) assms(2) assms(3) observable_io_targets_split) qed lemma observable_io_target_is_singleton[simp] : assumes "observable M" and "p \ io_targets M q io" shows "io_targets M q io = {p}" proof - have "io \ LS M q" using assms(2) by auto then obtain p' where "io_targets M q io = {p'}" using assms(1) by (meson io_targets_observable_singleton_ex) then show ?thesis using assms(2) by simp qed lemma observable_path_prefix : assumes "observable M" and "path M (io || tr) q" and "length io = length tr" and "path M (ioP || trP) q" and "length ioP = length trP" and "prefix ioP io" shows "trP = take (length ioP) tr" proof - have ioP_def : "ioP = take (length ioP) io" using assms(6) by (metis append_eq_conv_conj prefixE) then have "take (length ioP) (io || tr) = take (length ioP) io || take (length ioP) tr" using take_zip by blast moreover have "path M (take (length ioP) (io || tr)) q" using assms by (metis FSM.path_append_elim append_take_drop_id) ultimately have "path M (take (length ioP) io || take (length ioP) tr) q \ length (take (length ioP) io) = length (take (length ioP) tr)" using assms(3) by auto then have "path M (ioP || take (length ioP) tr) q \ length ioP = length (take (length ioP) tr)" using assms(3) using ioP_def by auto then show ?thesis by (meson assms(1) assms(4) assms(5) language_state observable_path_unique) qed lemma io_targets_succ : assumes "q2 \ io_targets M q1 [xy]" shows "q2 \ succ M xy q1" proof - obtain tr where tr_def : "target ([xy] || tr) q1 = q2" "path M ([xy] || tr) q1" "length [xy] = length tr" using assms by auto have "length tr = Suc 0" using \length [xy] = length tr\ by auto then obtain q2' where "tr = [q2']" by (metis Suc_length_conv length_0_conv) then have "target ([xy] || tr) q1 = q2'" by auto then have "q2' = q2" using \target ([xy] || tr) q1 = q2\ by simp then have "path M ([xy] || [q2]) q1" using tr_def(2) \tr = [q2']\ by auto then have "path M [(xy,q2)] q1" by auto show ?thesis proof (cases rule: FSM.path.cases[of M "[(xy,q2)]" q1]) case nil show ?case using \path M [(xy,q2)] q1\ by simp next case cons show "snd (xy, q2) \ succ M (fst (xy, q2)) q1 \ path M [] (snd (xy, q2)) \ q2 \ succ M xy q1" by auto qed qed subsection \ D-reachability \ text \ A state of some FSM is d-reached (deterministically reached) by some input sequence if any sequence in the language of the FSM with this input sequence reaches that state. That state is then called d-reachable. \ abbreviation "d_reached_by M p xs q tr ys \ ((length xs = length ys \ length xs = length tr \ (path M ((xs || ys) || tr) p) \ target ((xs || ys) || tr) p = q) \ (\ ys2 tr2 . (length xs = length ys2 \ length xs = length tr2 \ path M ((xs || ys2) || tr2) p) \ target ((xs || ys2) || tr2) p = q))" fun d_reaches :: "('in, 'out, 'state) FSM \ 'state \ 'in list \ 'state \ bool" where "d_reaches M p xs q = (\ tr ys . d_reached_by M p xs q tr ys)" fun d_reachable :: "('in, 'out, 'state) FSM \ 'state \ 'state set" where "d_reachable M p = { q . (\ xs . d_reaches M p xs q) }" lemma d_reaches_unique[elim] : assumes "d_reaches M p xs q1" and "d_reaches M p xs q2" shows "q1 = q2" using assms unfolding d_reaches.simps by blast lemma d_reaches_unique_cases[simp] : "{ q . d_reaches M (initial M) xs q } = {} \ (\ q2 . { q . d_reaches M (initial M) xs q } = { q2 })" unfolding d_reaches.simps by blast lemma d_reaches_unique_obtain[simp] : assumes "d_reaches M (initial M) xs q" shows "{ p . d_reaches M (initial M) xs p } = { q }" using assms unfolding d_reaches.simps by blast lemma d_reaches_io_target : assumes "d_reaches M p xs q" and "length ys = length xs" shows "io_targets M p (xs || ys) \ {q}" proof fix q' assume "q' \ io_targets M p (xs || ys)" then obtain trQ where "path M ((xs || ys) || trQ) p \ length (xs || ys) = length trQ" by auto moreover obtain trD ysD where "d_reached_by M p xs q trD ysD" using assms(1) by auto ultimately have "target ((xs || ys) || trQ) p = q" by (simp add: assms(2)) then show "q' \ {q}" using \d_reached_by M p xs q trD ysD\ \q' \ io_targets M p (xs || ys)\ assms(2) by auto qed lemma d_reachable_reachable : "d_reachable M p \ reachable M p" unfolding d_reaches.simps d_reachable.simps by blast subsection \ Deterministic state cover \ text \ The deterministic state cover of some FSM is a minimal set of input sequences such that every d-reachable state of the FSM is d-reached by a sequence in the set and the set contains the empty sequence (which d-reaches the initial state). \ fun is_det_state_cover_ass :: "('in, 'out, 'state) FSM \ ('state \ 'in list) \ bool" where "is_det_state_cover_ass M f = (f (initial M) = [] \ (\ s \ d_reachable M (initial M) . d_reaches M (initial M) (f s) s))" lemma det_state_cover_ass_dist : assumes "is_det_state_cover_ass M f" and "s1 \ d_reachable M (initial M)" and "s2 \ d_reachable M (initial M)" and "s1 \ s2" shows "\(d_reaches M (initial M) (f s2) s1)" by (meson assms(1) assms(3) assms(4) d_reaches_unique is_det_state_cover_ass.simps) lemma det_state_cover_ass_diff : assumes "is_det_state_cover_ass M f" and "s1 \ d_reachable M (initial M)" and "s2 \ d_reachable M (initial M)" and "s1 \ s2" shows "f s1 \ f s2" by (metis assms det_state_cover_ass_dist is_det_state_cover_ass.simps) fun is_det_state_cover :: "('in, 'out, 'state) FSM \ 'in list set \ bool" where "is_det_state_cover M V = (\ f . is_det_state_cover_ass M f \ V = image f (d_reachable M (initial M)))" lemma det_state_cover_d_reachable[elim] : assumes "is_det_state_cover M V" and "v \ V" obtains q where "d_reaches M (initial M) v q" by (metis (no_types, opaque_lifting) assms(1) assms(2) image_iff is_det_state_cover.simps is_det_state_cover_ass.elims(2)) lemma det_state_cover_card[simp] : assumes "is_det_state_cover M V" and "finite (nodes M)" shows "card (d_reachable M (initial M)) = card V" proof - obtain f where f_def : "is_det_state_cover_ass M f \ V = image f (d_reachable M (initial M))" using assms unfolding is_det_state_cover.simps by blast then have card_f : "card V = card (image f (d_reachable M (initial M)))" by simp have "d_reachable M (initial M) \ nodes M" unfolding d_reachable.simps d_reaches.simps using d_reachable_reachable by blast then have dr_finite : "finite (d_reachable M (initial M))" using assms infinite_super by blast then have card_le : "card (image f (d_reachable M (initial M))) \ card (d_reachable M (initial M))" using card_image_le by blast have "card (image f (d_reachable M (initial M))) = card (d_reachable M (initial M))" by (meson card_image det_state_cover_ass_diff f_def inj_onI) then show ?thesis using card_f by auto qed lemma det_state_cover_finite : assumes "is_det_state_cover M V" and "finite (nodes M)" shows "finite V" proof - have "d_reachable M (initial M) \ nodes M" by auto show "finite V" using det_state_cover_card[OF assms] by (metis \d_reachable M (initial M) \ nodes M\ assms(1) assms(2) finite_imageI infinite_super is_det_state_cover.simps) qed lemma det_state_cover_initial : assumes "is_det_state_cover M V" shows "[] \ V" proof - have "d_reached_by M (initial M) [] (initial M) [] []" by (simp add: FSM.nil) then have "d_reaches M (initial M) [] (initial M)" by auto have "initial M \ d_reachable M (initial M)" by (metis (no_types) \d_reaches M (initial M) [] (initial M)\ d_reachable.simps mem_Collect_eq) then show ?thesis by (metis (no_types, lifting) assms image_iff is_det_state_cover.elims(2) is_det_state_cover_ass.simps) qed lemma det_state_cover_empty : assumes "is_det_state_cover M V" shows "[] \ V" proof - obtain f where f_def : "is_det_state_cover_ass M f \ V = f ` d_reachable M (initial M)" using assms by auto then have "f (initial M) = []" by auto moreover have "initial M \ d_reachable M (initial M)" proof - have "d_reaches M (initial M) [] (initial M)" by auto then show ?thesis by (metis d_reachable.simps mem_Collect_eq) qed moreover have "f (initial M) \ V" using f_def calculation by blast ultimately show ?thesis by auto qed subsection \ IO reduction \ text \ An FSM is a reduction of another, if its language is a subset of the language of the latter FSM. \ fun io_reduction :: "('in, 'out, 'state) FSM \ ('in, 'out, 'state) FSM \ bool" (infix "\" 200) where "M1 \ M2 = (LS M1 (initial M1) \ LS M2 (initial M2))" lemma language_state_inclusion_of_state_reached_by_same_sequence : assumes "LS M1 q1 \ LS M2 q2" and "observable M1" and "observable M2" and "io_targets M1 q1 io = { q1t }" and "io_targets M2 q2 io = { q2t }" shows "LS M1 q1t \ LS M2 q2t" proof fix x assume "x \ LS M1 q1t" obtain q1x where "io_targets M1 q1t x = {q1x}" by (meson \x \ LS M1 q1t\ assms(2) io_targets_observable_singleton_ex) have "io \ LS M1 q1" using assms(4) by auto have "io@x \ LS M1 q1" using observable_io_targets_append[OF assms(2) \io_targets M1 q1 io = { q1t }\ \io_targets M1 q1t x = {q1x}\] by (metis io_targets_elim language_state singletonI) then have "io@x \ LS M2 q2" using assms(1) by blast then obtain q2x where "io_targets M2 q2 (io@x) = {q2x}" by (meson assms(3) io_targets_observable_singleton_ex) show "x \ LS M2 q2t" using observable_io_targets_split[OF assms(3) \io_targets M2 q2 (io @ x) = {q2x}\ assms(5)] by auto qed subsection \ Language subsets for input sequences \ text \ The following definitions describe restrictions of languages to only those IO-sequences that exhibit a certain input sequence or whose input sequence is contained in a given set of input sequences. This allows to define the notion that some FSM is a reduction of another over a given set of input sequences, but not necessarily over the entire language of the latter FSM. \ fun language_state_for_input :: "('in, 'out, 'state) FSM \ 'state \ 'in list \ ('in \ 'out) list set" where "language_state_for_input M q xs = {(xs || ys) | ys . (length xs = length ys \ (xs || ys) \ LS M q)}" fun language_state_for_inputs :: "('in, 'out, 'state) FSM \ 'state \ 'in list set \ ('in \ 'out) list set" ("(LS\<^sub>i\<^sub>n _ _ _)" [1000,1000,1000]) where "language_state_for_inputs M q ISeqs = {(xs || ys) | xs ys . (xs \ ISeqs \ length xs = length ys \ (xs || ys) \ LS M q)}" abbreviation "L\<^sub>i\<^sub>n M TS \ LS\<^sub>i\<^sub>n M (initial M) TS" abbreviation "io_reduction_on M1 TS M2 \ (L\<^sub>i\<^sub>n M1 TS \ L\<^sub>i\<^sub>n M2 TS)" notation io_reduction_on ("(_ \\_\ _)" [1000,0,0] 61) notation (latex output) io_reduction_on ("(_ \\<^bsub>_\<^esub> _)" [1000,0,0] 61) lemma language_state_for_input_alt_def : "language_state_for_input M q xs = LS\<^sub>i\<^sub>n M q {xs}" unfolding language_state_for_input.simps language_state_for_inputs.simps by blast lemma language_state_for_inputs_alt_def : "LS\<^sub>i\<^sub>n M q ISeqs = \ (image (language_state_for_input M q) ISeqs)" by auto lemma language_state_for_inputs_in_language_state : "LS\<^sub>i\<^sub>n M q T \ language_state M q" unfolding language_state_for_inputs.simps language_state_def by blast lemma language_state_for_inputs_map_fst : assumes "io \ language_state M q" and "map fst io \ T" shows "io \ LS\<^sub>i\<^sub>n M q T" proof - let ?xs = "map fst io" let ?ys = "map snd io" have "?xs \ T \ length ?xs = length ?ys \ ?xs || ?ys \ language_state M q" using assms(2,1) by auto then have "?xs || ?ys \ LS\<^sub>i\<^sub>n M q T" unfolding language_state_for_inputs.simps by blast then show ?thesis by simp qed lemma language_state_for_inputs_nonempty : assumes "set xs \ inputs M" and "completely_specified M" and "q \ nodes M" shows "LS\<^sub>i\<^sub>n M q {xs} \ {}" using assms proof (induction xs arbitrary: q) case Nil then show ?case by auto next case (Cons x xs) then have "x \ inputs M" by simp then obtain y q' where x_step : "q' \ succ M (x,y) q" using Cons(3,4) unfolding completely_specified.simps by blast then have "path M ([(x,y)] || [q']) q \ length [q] = length [(x,y)]" "target ([(x,y)] || [q']) q = q'" by auto then have "q' \ nodes M" using Cons(4) by (metis FSM.nodes_target) then have "LS\<^sub>i\<^sub>n M q' {xs} \ {}" using Cons.prems Cons.IH by auto then obtain ys where "length xs = length ys \ (xs || ys) \ LS M q'" by auto then obtain tr where "path M ((xs || ys) || tr) q' \ length tr = length (xs || ys)" by auto then have "path M ([(x,y)] @ (xs || ys) || [q'] @ tr) q \ length ([q'] @ tr) = length ([(x,y)] @ (xs || ys))" by (simp add: FSM.path.intros(2) x_step) then have "path M ((x#xs || y#ys) || [q'] @ tr) q \ length ([q'] @ tr) = length (x#xs || y#ys)" by auto then have "(x#xs || y#ys) \ LS M q" by (metis language_state) moreover have "length (x#xs) = length (y#ys)" by (simp add: \length xs = length ys \ xs || ys \ LS M q'\) ultimately have "(x#xs || y#ys) \ LS\<^sub>i\<^sub>n M q {x # xs}" unfolding language_state_for_inputs.simps by blast then show ?case by blast qed lemma language_state_for_inputs_map_fst_contained : assumes "vs \ LS\<^sub>i\<^sub>n M q V" shows "map fst vs \ V" proof - have "(map fst vs) || (map snd vs) = vs" by auto then have "(map fst vs) || (map snd vs) \ LS\<^sub>i\<^sub>n M q V" using assms by auto then show ?thesis by auto qed lemma language_state_for_inputs_empty : assumes "[] \ V" shows "[] \ LS\<^sub>i\<^sub>n M q V" proof - have "[] \ language_state_for_input M q []" by auto then show ?thesis using language_state_for_inputs_alt_def by (metis UN_I assms) qed lemma language_state_for_input_empty[simp] : "language_state_for_input M q [] = {[]}" by auto lemma language_state_for_input_take : assumes "io \ language_state_for_input M q xs" shows "take n io \ language_state_for_input M q (take n xs)" proof - obtain ys where "io = xs || ys" "length xs = length ys" "xs || ys \ language_state M q" using assms by auto then obtain p where "length p = length xs" "path M ((xs || ys) || p) q " by auto then have "path M (take n ((xs || ys) || p)) q" by (metis FSM.path_append_elim append_take_drop_id) then have "take n (xs || ys) \ language_state M q" by (simp add: \length p = length xs\ \length xs = length ys\ language_state take_zip) then have "(take n xs) || (take n ys) \ language_state M q" by (simp add: take_zip) have "take n io = (take n xs) || (take n ys)" using \io = xs || ys\ take_zip by blast moreover have "length (take n xs) = length (take n ys)" by (simp add: \length xs = length ys\) ultimately show ?thesis using \(take n xs) || (take n ys) \ language_state M q\ unfolding language_state_for_input.simps by blast qed lemma language_state_for_inputs_prefix : assumes "vs@xs \ L\<^sub>i\<^sub>n M1 {vs'@xs'}" and "length vs = length vs'" shows "vs \ L\<^sub>i\<^sub>n M1 {vs'}" proof - have "vs@xs \ L M1" using assms(1) by auto then have "vs \ L M1" by (meson language_state_prefix) then have "vs \ L\<^sub>i\<^sub>n M1 {map fst vs}" by (meson insertI1 language_state_for_inputs_map_fst) moreover have "vs' = map fst vs" by (metis append_eq_append_conv assms(1) assms(2) language_state_for_inputs_map_fst_contained length_map map_append singletonD) ultimately show ?thesis by blast qed lemma language_state_for_inputs_union : shows "LS\<^sub>i\<^sub>n M q T1 \ LS\<^sub>i\<^sub>n M q T2 = LS\<^sub>i\<^sub>n M q (T1 \ T2)" unfolding language_state_for_inputs.simps by blast lemma io_reduction_on_subset : assumes "io_reduction_on M1 T M2" and "T' \ T" shows "io_reduction_on M1 T' M2" proof (rule ccontr) assume "\ io_reduction_on M1 T' M2" then obtain xs' where "xs' \ T'" "\ L\<^sub>i\<^sub>n M1 {xs'} \ L\<^sub>i\<^sub>n M2 {xs'}" proof - have f1: "\ps P Pa. (ps::('a \ 'b) list) \ P \ \ P \ Pa \ ps \ Pa" by blast obtain pps :: "('a \ 'b) list set \ ('a \ 'b) list set \ ('a \ 'b) list" where "\x0 x1. (\v2. v2 \ x1 \ v2 \ x0) = (pps x0 x1 \ x1 \ pps x0 x1 \ x0)" by moura then have f2: "\P Pa. pps Pa P \ P \ pps Pa P \ Pa \ P \ Pa" by (meson subsetI) have f3: "\ps f c A. (ps::('a \ 'b) list) \ LS\<^sub>i\<^sub>n f (c::'c) A \ map fst ps \ A" by (meson language_state_for_inputs_map_fst_contained) then have "L\<^sub>i\<^sub>n M1 T' \ L\<^sub>i\<^sub>n M1 T" using f2 by (meson assms(2) language_state_for_inputs_in_language_state language_state_for_inputs_map_fst set_rev_mp) then show ?thesis using f3 f2 f1 by (meson \\ io_reduction_on M1 T' M2\ assms(1) language_state_for_inputs_in_language_state language_state_for_inputs_map_fst) qed then have "xs' \ T" using assms(2) by blast have "\ io_reduction_on M1 T M2" proof - have f1: "\as. as \ T' \ as \ T" using assms(2) by auto obtain pps :: "('a \ 'b) list set \ ('a \ 'b) list set \ ('a \ 'b) list" where "\x0 x1. (\v2. v2 \ x1 \ v2 \ x0) = (pps x0 x1 \ x1 \ pps x0 x1 \ x0)" by moura then have "\P Pa. (\ P \ Pa \ (\ps. ps \ P \ ps \ Pa)) \ (P \ Pa \ pps Pa P \ P \ pps Pa P \ Pa)" by blast then show ?thesis using f1 by (meson \\ io_reduction_on M1 T' M2\ language_state_for_inputs_in_language_state language_state_for_inputs_map_fst language_state_for_inputs_map_fst_contained) qed then show "False" using assms(1) by auto qed subsection \ Sequences to failures \ text \ A sequence to a failure for FSMs @{verbatim M1} and @{verbatim M2} is a sequence such that any proper prefix of it is contained in the languages of both @{verbatim M1} and @{verbatim M2}, while the sequence itself is contained only in the language of A. That is, if a sequence to a failure for @{verbatim M1} and @{verbatim M2} exists, then @{verbatim M1} is not a reduction of @{verbatim M2}. \ fun sequence_to_failure :: "('in,'out,'state) FSM \ ('in,'out,'state) FSM \ ('in \ 'out) list \ bool" where "sequence_to_failure M1 M2 xs = ( (butlast xs) \ (language_state M2 (initial M2) \ language_state M1 (initial M1)) \ xs \ (language_state M1 (initial M1) - language_state M2 (initial M2)))" lemma sequence_to_failure_ob : assumes "\ M1 \ M2" and "well_formed M1" and "well_formed M2" obtains io where "sequence_to_failure M1 M2 io" proof - let ?diff = "{ io . io \ language_state M1 (initial M1) \ io \ language_state M2 (initial M2)}" have "?diff \ empty" using assms by auto moreover obtain io where io_def[simp] : "io = arg_min length (\ io . io \ ?diff)" using assms by auto ultimately have io_diff : "io \ ?diff" using assms by (meson all_not_in_conv arg_min_natI) then have "io \ []" using assms io_def language_state by auto then obtain io_init io_last where io_split[simp] : "io = io_init @ [io_last]" by (metis append_butlast_last_id) have io_init_inclusion : "io_init \ language_state M1 (initial M1) \ io_init \ language_state M2 (initial M2)" proof (rule ccontr) assume assm : "\ (io_init \ language_state M1 (initial M1) \ io_init \ language_state M2 (initial M2))" have "io_init @ [io_last] \ language_state M1 (initial M1)" using io_diff io_split by auto then have "io_init \ language_state M1 (initial M1)" by (meson language_state language_state_split) moreover have "io_init \ language_state M2 (initial M2)" using assm calculation by auto ultimately have "io_init \ ?diff" by auto moreover have "length io_init < length io" using io_split by auto ultimately have "io \ arg_min length (\ io . io \ ?diff)" proof - have "\ps. ps \ {ps \ language_state M1 (initial M1). ps \ language_state M2 (initial M2)} \ \ length io \ length ps" using \io_init \ {io\ language_state M1 (initial M1). io \ language_state M2 (initial M2)}\ \length io_init < length io\ linorder_not_less by blast then show ?thesis by (meson arg_min_nat_le) qed then show "False" using io_def by simp qed then have "sequence_to_failure M1 M2 io" using io_split io_diff by auto then show ?thesis using that by auto qed lemma sequence_to_failure_succ : assumes "sequence_to_failure M1 M2 io" shows "\ q \ io_targets M2 (initial M2) (butlast io) . succ M2 (last io) q = {}" proof have "io \ []" using assms by auto fix q assume "q \ io_targets M2 (initial M2) (butlast io)" then obtain tr where "q = target (butlast io || tr) (initial M2)" and "path M2 (butlast io || tr) (initial M2)" and "length (butlast io) = length tr" unfolding io_targets.simps by auto show "succ M2 (last io) q = {}" proof (rule ccontr) assume "succ M2 (last io) q \ {}" then obtain q' where "q' \ succ M2 (last io) q" by blast then have "path M2 [(last io, q')] (target (butlast io || tr) (initial M2))" using \q = target (butlast io || tr) (initial M2)\ by auto have "path M2 ((butlast io || tr) @ [(last io, q')]) (initial M2)" using \path M2 (butlast io || tr) (initial M2)\ \path M2 [(last io, q')] (target (butlast io || tr) (initial M2))\ by auto have "butlast io @ [last io] = io" by (meson \io \ []\ append_butlast_last_id) have "path M2 (io || (tr@[q'])) (initial M2)" proof - have "path M2 ((butlast io || tr) @ ([last io] || [q'])) (initial M2)" by (simp add: FSM.path_append \path M2 (butlast io || tr) (initial M2)\ \path M2 [(last io, q')] (target (butlast io || tr) (initial M2))\) then show ?thesis by (metis (no_types) \butlast io @ [last io] = io\ \length (butlast io) = length tr\ zip_append) qed have "io \ L M2" proof - have "length tr + (0 + Suc 0) = length io" by (metis \butlast io @ [last io] = io\ \length (butlast io) = length tr\ length_append list.size(3) list.size(4)) then show ?thesis using \path M2 (io || tr @ [q']) (initial M2)\ by fastforce qed then show "False" using assms by auto qed qed lemma sequence_to_failure_non_nil : assumes "sequence_to_failure M1 M2 xs" shows "xs \ []" proof assume "xs = []" then have "xs \ L M1 \ L M2" by auto then show "False" using assms by auto qed lemma sequence_to_failure_from_arbitrary_failure : assumes "vs@xs \ L M1 - L M2" and "vs \ L M2 \ L M1" shows "\ xs' . prefix xs' xs \ sequence_to_failure M1 M2 (vs@xs')" using assms proof (induction xs rule: rev_induct) case Nil then show ?case by auto next case (snoc x xs) have "vs @ xs \ L M1" using snoc.prems(1) by (metis Diff_iff append.assoc language_state_prefix) show ?case proof (cases "vs@xs \ L M2") case True have "butlast (vs@xs@[x]) \ L M2 \ L M1" using True \vs @ xs \ L M1\ by (simp add: butlast_append) then show ?thesis using sequence_to_failure.simps snoc.prems by blast next case False then have "vs@xs \ L M1 - L M2" using \vs @ xs \ L M1\ by blast then obtain xs' where "prefix xs' xs" "sequence_to_failure M1 M2 (vs@xs')" using snoc.prems(2) snoc.IH by blast then show ?thesis using prefix_snoc by auto qed qed text \ The following lemma shows that if @{verbatim M1} is not a reduction of @{verbatim M2}, then a minimal sequence to a failure exists that is of length at most the number of states in @{verbatim M1} times the number of states in @{verbatim M2}. \ lemma sequence_to_failure_length : assumes "well_formed M1" and "well_formed M2" and "observable M1" and "observable M2" and "\ M1 \ M2" shows "\ xs . sequence_to_failure M1 M2 xs \ length xs \ |M2| * |M1|" proof - obtain seq where "sequence_to_failure M1 M2 seq" using assms sequence_to_failure_ob by blast then have "seq \ []" by auto let ?bls = "butlast seq" have "?bls \ L M1" "?bls \ L M2" using \sequence_to_failure M1 M2 seq\ by auto then obtain tr1b tr2b where "path M1 (?bls || tr1b) (initial M1)" "length tr1b = length ?bls" "path M2 (?bls || tr2b) (initial M2)" "length ?bls = length tr2b" by fastforce then have "length tr2b = length tr1b" by auto let ?PM = "product M2 M1" have "well_formed ?PM" using well_formed_product[OF assms(1,2)] by assumption have "path ?PM (?bls || tr2b || tr1b) (initial M2, initial M1)" using product_path[OF \length ?bls = length tr2b\ \length tr2b = length tr1b\, of M2 M1 "initial M2" "initial M1"] using \path M1 (butlast seq || tr1b) (initial M1)\ \path M2 (butlast seq || tr2b) (initial M2)\ by blast let ?q1b = "target (?bls || tr1b) (initial M1)" let ?q2b = "target (?bls || tr2b) (initial M2)" have "io_targets M2 (initial M2) ?bls = {?q2b}" by (metis \length (butlast seq) = length tr2b\ \path M2 (butlast seq || tr2b) (initial M2)\ assms(4) obs_target_is_io_targets) have "io_targets M1 (initial M1) ?bls = {?q1b}" by (metis \length tr1b = length (butlast seq)\ \path M1 (butlast seq || tr1b) (initial M1)\ assms(3) obs_target_is_io_targets) have "(?q2b, ?q1b) \ reachable (product M2 M1) (initial M2, initial M1)" proof - have "target (butlast seq || tr2b || tr1b) (initial M2, initial M1) \ reachable (product M2 M1) (initial M2, initial M1)" using \path (product M2 M1) (butlast seq || tr2b || tr1b) (initial M2, initial M1)\ by blast then show ?thesis using \length (butlast seq) = length tr2b\ \length tr2b = length tr1b\ by auto qed have "(initial M2, initial M1) \ nodes (product M2 M1)" by (simp add: FSM.nodes.initial) obtain p where repFreePath : "path (product M2 M1) p (initial M2, initial M1) \ target p (initial M2, initial M1) = (?q2b,?q1b)" "distinct ((initial M2, initial M1) # states p (initial M2, initial M1))" using reaching_path_without_repetition[OF \well_formed ?PM\ \(?q2b, ?q1b) \ reachable (product M2 M1) (initial M2, initial M1)\ \(initial M2, initial M1) \ nodes (product M2 M1)\] by blast then have "set (states p (initial M2, initial M1)) \ nodes ?PM" by (simp add: FSM.nodes_states \(initial M2, initial M1) \ nodes (product M2 M1)\) moreover have "(initial M2, initial M1) \ set (states p (initial M2, initial M1))" using \distinct ((initial M2, initial M1) # states p (initial M2, initial M1))\ by auto ultimately have "set (states p (initial M2, initial M1)) \ nodes ?PM - {(initial M2,initial M1)}" by blast moreover have "finite (nodes ?PM)" using \well_formed ?PM\ by auto ultimately have "card (set (states p (initial M2, initial M1))) < card (nodes ?PM)" by (metis \(initial M2, initial M1) \ nodes (product M2 M1)\ \(initial M2, initial M1) \ set (states p (initial M2, initial M1))\ \set (states p (initial M2, initial M1)) \ nodes (product M2 M1)\ psubsetI psubset_card_mono) moreover have "card (set (states p (initial M2, initial M1))) = length (states p (initial M2, initial M1))" using distinct_card repFreePath(2) by fastforce ultimately have "length (states p (initial M2, initial M1)) < |?PM|" by linarith then have "length p < |?PM|" by auto let ?p1 = "map (snd \ snd) p" let ?p2 = "map (fst \ snd) p" let ?pIO = "map fst p" have "p = ?pIO || ?p2 || ?p1" by (metis map_map zip_map_fst_snd) have "path M2 (?pIO || ?p2) (initial M2)" "path M1 (?pIO || ?p1) (initial M1)" using product_path[of ?pIO ?p2 ?p1 M2 M1] using \p = ?pIO || ?p2 || ?p1\ repFreePath(1) by auto have "(?q2b, ?q1b) = (target (?pIO || ?p2 || ?p1) (initial M2, initial M1))" using \p = ?pIO || ?p2 || ?p1\ repFreePath(1) by auto then have "?q2b = target (?pIO || ?p2) (initial M2)" "?q1b = target (?pIO || ?p1) (initial M1)" by auto have "io_targets M2 (initial M2) ?pIO = {?q2b}" by (metis \path M2 (map fst p || map (fst \ snd) p) (initial M2)\ \target (?bls || tr2b) (initial M2) = target (map fst p || map (fst \ snd) p) (initial M2)\ assms(4) length_map obs_target_is_io_targets) have "io_targets M1 (initial M1) ?pIO = {?q1b}" by (metis \path M1 (map fst p || map (snd \ snd) p) (initial M1)\ \target (?bls || tr1b) (initial M1) = target (map fst p || map (snd \ snd) p) (initial M1)\ assms(3) length_map obs_target_is_io_targets) have "seq \ L M1" "seq \ L M2" using \sequence_to_failure M1 M2 seq\ by auto have "io_targets M1 (initial M1) ?bls = {?q1b}" by (metis \length tr1b = length (butlast seq)\ \path M1 (butlast seq || tr1b) (initial M1)\ assms(3) obs_target_is_io_targets) obtain q1s where "io_targets M1 (initial M1) seq = {q1s}" by (meson \seq \ L M1\ assms(3) io_targets_observable_singleton_ob) moreover have "seq = (butlast seq)@[last seq]" using \seq \ []\ by auto ultimately have "io_targets M1 (initial M1) ((butlast seq)@[last seq]) = {q1s}" by auto have "io_targets M1 ?q1b [last seq] = {q1s}" using observable_io_targets_suffix[OF assms(3) \io_targets M1 (initial M1) ?bls = {?q1b}\ \io_targets M1 (initial M1) ((butlast seq)@[last seq]) = {q1s}\] by assumption then obtain tr1s where "q1s = target ([last seq] || tr1s) ?q1b" "path M1 ([last seq] || tr1s) ?q1b" "length [last seq] = length tr1s" by auto have "path M1 ([last seq] || [q1s]) ?q1b" by (metis (no_types) \length [last seq] = length tr1s\ \path M1 ([last seq] || tr1s) (target (butlast seq || tr1b) (initial M1))\ \q1s = target ([last seq] || tr1s) (target (butlast seq || tr1b) (initial M1))\ append_Nil append_butlast_last_id butlast.simps(2) length_butlast length_greater_0_conv not_Cons_self2 target_alt_def(2)) then have "q1s \ succ M1 (last seq) ?q1b" by auto have "succ M2 (last seq) ?q2b = {}" proof (rule ccontr) assume "succ M2 (last seq) (target (butlast seq || tr2b) (initial M2)) \ {}" then obtain q2f where "q2f \ succ M2 (last seq) ?q2b" by blast then have "target ([last seq] || [q2f]) ?q2b = q2f" "path M2 ([last seq] || [q2f]) ?q2b" "length [q2f] = length [last seq]" by auto then have "q2f \ io_targets M2 ?q2b [last seq]" by (metis io_target_from_path) then have "io_targets M2 ?q2b [last seq] = {q2f}" using assms(4) by (meson observable_io_target_is_singleton) have "io_targets M2 (initial M2) (butlast seq @ [last seq]) = {q2f}" using observable_io_targets_append[OF assms(4) \io_targets M2 (initial M2) ?bls = {?q2b}\ \io_targets M2 ?q2b [last seq] = {q2f}\] by assumption then have "seq \ L M2" using \seq = butlast seq @ [last seq]\ by auto then show "False" using \seq \ L M2\ by blast qed have "?pIO \ L M1" "?pIO \ L M2" using \path M1 (?pIO || ?p1) (initial M1)\ \path M2 (?pIO || ?p2) (initial M2)\ by auto then have "butlast (?pIO@[last seq]) \ L M1 \ L M2" by auto have "?pIO@[last seq] \ L M1" using observable_io_targets_append[OF assms(3) \io_targets M1 (initial M1) ?pIO = {?q1b}\ \io_targets M1 ?q1b [last seq] = {q1s}\] by (metis all_not_in_conv insert_not_empty io_targets_elim language_state) moreover have "?pIO@[last seq] \ L M2" proof assume "?pIO@[last seq] \ L M2" then obtain q2f where "io_targets M2 (initial M2) (?pIO@[last seq]) = {q2f}" by (meson assms(4) io_targets_observable_singleton_ob) have "io_targets M2 ?q2b [last seq] = {q2f}" using observable_io_targets_split[OF assms(4) \io_targets M2 (initial M2) (?pIO@[last seq]) = {q2f}\ \io_targets M2 (initial M2) (map fst p) = {?q2b}\] by assumption then have "q2f \ succ M2 (last seq) ?q2b" by (simp add: io_targets_succ) then show "False" using \succ M2 (last seq) ?q2b = {}\ by auto qed ultimately have "?pIO@[last seq] \ L M1 - L M2" by auto have "sequence_to_failure M1 M2 (?pIO@[last seq])" using \butlast (?pIO@[last seq]) \ L M1 \ L M2\ \?pIO@[last seq] \ L M1 - L M2\ by auto have "length (?pIO@[last seq]) = Suc (length ?pIO)" by auto then have "length (?pIO@[last seq]) \ |?PM|" using \length p < |?PM|\ by auto have "card (nodes M2 \ nodes M1) \ |M2| * |M1|" by (simp add: card_cartesian_product) have "finite (nodes M2 \ nodes M1)" proof show "finite (nodes M2)" using assms by auto show "finite (nodes M1)" using assms by auto qed have "|?PM| \ |M2| * |M1|" by (meson \card (nodes M2 \ nodes M1) \ |M2| * |M1|\ \finite (nodes M2 \ nodes M1)\ card_mono dual_order.trans product_nodes) then have "length (?pIO@[last seq]) \ |M2| * |M1|" using \length (?pIO@[last seq]) \ |?PM|\ by auto then have "sequence_to_failure M1 M2 (?pIO@[last seq]) \ length (?pIO@[last seq]) \ |M2| * |M1|" using \sequence_to_failure M1 M2 (?pIO@[last seq])\ by auto then show ?thesis by blast qed subsection \ Minimal sequence to failure extending \ text \ A minimal sequence to a failure extending some some set of IO-sequences is a sequence to a failure of minimal length such that a prefix of that sequence is contained in the set. \ fun minimal_sequence_to_failure_extending :: "'in list set \ ('in,'out,'state) FSM \ ('in,'out,'state) FSM \ ('in \ 'out) list \ ('in \ 'out) list \ bool" where "minimal_sequence_to_failure_extending V M1 M2 v' io = ( v' \ L\<^sub>i\<^sub>n M1 V \ sequence_to_failure M1 M2 (v' @ io) \ \ (\ io' . \ w' \ L\<^sub>i\<^sub>n M1 V . sequence_to_failure M1 M2 (w' @ io') \ length io' < length io))" lemma minimal_sequence_to_failure_extending_det_state_cover_ob : assumes "well_formed M1" and "well_formed M2" and "observable M2" and "is_det_state_cover M2 V" and "\ M1 \ M2" obtains vs xs where "minimal_sequence_to_failure_extending V M1 M2 vs xs" proof - \ \set of all IO-sequences that extend some reaction of M1 to V to a failure\ let ?exts = "{xs. \vs' \ L\<^sub>i\<^sub>n M1 V. sequence_to_failure M1 M2 (vs'@xs)}" \ \arbitrary sequence to failure\ \ \must be contained in ?exts as V contains the empty sequence\ obtain stf where "sequence_to_failure M1 M2 stf" using assms sequence_to_failure_ob by blast then have "sequence_to_failure M1 M2 ([] @ stf)" by simp moreover have "[] \ L\<^sub>i\<^sub>n M1 V" by (meson assms(4) det_state_cover_initial language_state_for_inputs_empty) ultimately have "stf \ ?exts" by blast \ \the minimal length sequence of ?exts\ \ \is a minimal sequence to a failure extending V by construction\ let ?xsMin = "arg_min length (\xs. xs \ ?exts)" have xsMin_def : "?xsMin \ ?exts \ (\xs \ ?exts. length ?xsMin \ length xs)" by (metis (no_types, lifting) \stf \ ?exts\ arg_min_nat_lemma) then obtain vs where "vs \ L\<^sub>i\<^sub>n M1 V \ sequence_to_failure M1 M2 (vs @ ?xsMin)" by blast moreover have "\(\xs . \ws \ L\<^sub>i\<^sub>n M1 V. sequence_to_failure M1 M2 (ws@xs) \ length xs < length ?xsMin)" using leD xsMin_def by blast ultimately have "minimal_sequence_to_failure_extending V M1 M2 vs ?xsMin" by auto then show ?thesis using that by auto qed lemma mstfe_prefix_input_in_V : assumes "minimal_sequence_to_failure_extending V M1 M2 vs xs" shows "(map fst vs) \ V" proof - have "vs \ L\<^sub>i\<^sub>n M1 V" using assms by auto then show ?thesis using language_state_for_inputs_map_fst_contained by auto qed subsection \ Complete test suite derived from the product machine \ text \ The classical result of testing FSMs for language inclusion : Any failure can be observed by a sequence of length at most n*m where n is the number of states of the reference model (here FSM @{verbatim M2}) and m is an upper bound on the number of states of the SUT (here FSM @{verbatim M1}). \ lemma product_suite_soundness : assumes "well_formed M1" and "well_formed M2" and "observable M1" and "observable M2" and "inputs M2 = inputs M1" and "|M1| \ m " shows "\ M1 \ M2 \ \ M1 \\{xs . set xs \ inputs M2 \ length xs \ |M2| * m}\ M2" (is "\ M1 \ M2 \ \ M1 \\?TS\ M2") proof assume "\ M1 \ M2" obtain stf where "sequence_to_failure M1 M2 stf \ length stf \ |M2| * |M1|" using sequence_to_failure_length[OF assms(1-4) \\ M1 \ M2\] by blast then have "sequence_to_failure M1 M2 stf" "length stf \ |M2| * |M1|" by auto then have "stf \ L M1" by auto let ?xs = "map fst stf" have "set ?xs \ inputs M1" by (meson \stf \ L M1\ assms(1) language_state_inputs) then have "set ?xs \ inputs M2" using assms(5) by auto have "length ?xs \ |M2| * |M1|" using \length stf \ |M2| * |M1|\ by auto have "length ?xs \ |M2| * m" proof - show ?thesis by (metis (no_types) \length (map fst stf) \ |M2| * |M1|\ \|M1| \ m\ dual_order.trans mult.commute mult_le_mono1) qed have "stf \ L\<^sub>i\<^sub>n M1 {?xs}" by (meson \stf \ L M1\ insertI1 language_state_for_inputs_map_fst) have "?xs \ ?TS" using \set ?xs \ inputs M2\ \length ?xs \ |M2| * m\ by blast have "stf \ L\<^sub>i\<^sub>n M1 ?TS" by (metis (no_types, lifting) \map fst stf \ {xs. set xs \ inputs M2 \ length xs \ |M2| * m}\ \stf \ L M1\ language_state_for_inputs_map_fst) have "stf \ L M2" using \sequence_to_failure M1 M2 stf\ by auto then have "stf \ L\<^sub>i\<^sub>n M2 ?TS" by auto show "\ M1 \\?TS\ M2" using \stf \ L\<^sub>i\<^sub>n M1 ?TS\ \stf \ L\<^sub>i\<^sub>n M2 ?TS\ by blast qed lemma product_suite_completeness : assumes "well_formed M1" and "well_formed M2" and "observable M1" and "observable M2" and "inputs M2 = inputs M1" and "|M1| \ m " shows "M1 \ M2 \ M1 \\{xs . set xs \ inputs M2 \ length xs \ |M2| * m}\ M2" (is "M1 \ M2 \ M1 \\?TS\ M2") proof show "M1 \ M2 \ M1 \\?TS\ M2" \ \soundness holds trivially\ unfolding language_state_for_inputs.simps io_reduction.simps by blast show "M1 \\?TS\ M2 \ M1 \ M2" using product_suite_soundness[OF assms] by auto qed end \ No newline at end of file diff --git a/thys/Automated_Stateful_Protocol_Verification/Stateful_Protocol_Model.thy b/thys/Automated_Stateful_Protocol_Verification/Stateful_Protocol_Model.thy --- a/thys/Automated_Stateful_Protocol_Verification/Stateful_Protocol_Model.thy +++ b/thys/Automated_Stateful_Protocol_Verification/Stateful_Protocol_Model.thy @@ -1,9074 +1,9074 @@ (* Title: Stateful_Protocol_Model.thy Author: Andreas Viktor Hess, DTU Author: Sebastian A. Mödersheim, DTU Author: Achim D. Brucker, University of Exeter Author: Anders Schlichtkrull, DTU SPDX-License-Identifier: BSD-3-Clause *) section\Stateful Protocol Model\ theory Stateful_Protocol_Model imports Stateful_Protocol_Composition_and_Typing.Stateful_Compositionality Transactions Term_Abstraction begin subsection \Locale Setup\ locale stateful_protocol_model = fixes arity\<^sub>f::"'fun \ nat" and arity\<^sub>s::"'sets \ nat" and public\<^sub>f::"'fun \ bool" and Ana\<^sub>f::"'fun \ ((('fun,'atom::finite,'sets,'lbl) prot_fun, nat) term list \ nat list)" and \\<^sub>f::"'fun \ 'atom option" and label_witness1::"'lbl" and label_witness2::"'lbl" assumes Ana\<^sub>f_assm1: "\f. let (K, M) = Ana\<^sub>f f in (\k \ subterms\<^sub>s\<^sub>e\<^sub>t (set K). is_Fun k \ (is_Fu (the_Fun k)) \ length (args k) = arity\<^sub>f (the_Fu (the_Fun k)))" and Ana\<^sub>f_assm2: "\f. let (K, M) = Ana\<^sub>f f in \i \ fv\<^sub>s\<^sub>e\<^sub>t (set K) \ set M. i < arity\<^sub>f f" and public\<^sub>f_assm: "\f. arity\<^sub>f f > (0::nat) \ public\<^sub>f f" and \\<^sub>f_assm: "\f. arity\<^sub>f f = (0::nat) \ \\<^sub>f f \ None" and label_witness_assm: "label_witness1 \ label_witness2" begin lemma Ana\<^sub>f_assm1_alt: assumes "Ana\<^sub>f f = (K,M)" "k \ subterms\<^sub>s\<^sub>e\<^sub>t (set K)" shows "(\x. k = Var x) \ (\h T. k = Fun (Fu h) T \ length T = arity\<^sub>f h)" proof (cases k) case (Fun g T) let ?P = "\k. is_Fun k \ is_Fu (the_Fun k) \ length (args k) = arity\<^sub>f (the_Fu (the_Fun k))" let ?Q = "\K M. \k \ subterms\<^sub>s\<^sub>e\<^sub>t (set K). ?P k" have "?Q (fst (Ana\<^sub>f f)) (snd (Ana\<^sub>f f))" using Ana\<^sub>f_assm1 split_beta[of ?Q "Ana\<^sub>f f"] by meson hence "?Q K M" using assms(1) by simp hence "?P k" using assms(2) by blast thus ?thesis using Fun by (cases g) auto qed simp lemma Ana\<^sub>f_assm2_alt: assumes "Ana\<^sub>f f = (K,M)" "i \ fv\<^sub>s\<^sub>e\<^sub>t (set K) \ set M" shows "i < arity\<^sub>f f" using Ana\<^sub>f_assm2 assms by fastforce subsection \Definitions\ fun arity where "arity (Fu f) = arity\<^sub>f f" | "arity (Set s) = arity\<^sub>s s" | "arity (Val _) = 0" | "arity (Abs _) = 0" | "arity Pair = 2" | "arity (Attack _) = 0" | "arity OccursFact = 2" | "arity OccursSec = 0" | "arity (PubConst _ _) = 0" fun public where "public (Fu f) = public\<^sub>f f" | "public (Set s) = (arity\<^sub>s s > 0)" | "public (Val n) = False" | "public (Abs _) = False" | "public Pair = True" | "public (Attack _) = False" | "public OccursFact = True" | "public OccursSec = False" | "public (PubConst _ _) = True" fun Ana where "Ana (Fun (Fu f) T) = ( if arity\<^sub>f f = length T \ arity\<^sub>f f > 0 then let (K,M) = Ana\<^sub>f f in (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T, map ((!) T) M) else ([], []))" | "Ana _ = ([], [])" definition \\<^sub>v where "\\<^sub>v v \ ( if (\t \ subterms (fst v). case t of (TComp f T) \ arity f > 0 \ arity f = length T | _ \ True) then fst v else TAtom Bottom)" fun \ where "\ (Var v) = \\<^sub>v v" | "\ (Fun f T) = ( if arity f = 0 then case f of (Fu g) \ TAtom (case \\<^sub>f g of Some a \ Atom a | None \ Bottom) | (Val _) \ TAtom Value | (Abs _) \ TAtom AbsValue | (Set _) \ TAtom SetType | (Attack _) \ TAtom AttackType | OccursSec \ TAtom OccursSecType | (PubConst a _) \ TAtom a | _ \ TAtom Bottom else TComp f (map \ T))" lemma \_consts_simps[simp]: "arity\<^sub>f g = 0 \ \ (Fun (Fu g) []::('fun,'atom,'sets,'lbl) prot_term) = TAtom (case \\<^sub>f g of Some a \ Atom a | None \ Bottom)" "\ (Fun (Val n) []::('fun,'atom,'sets,'lbl) prot_term) = TAtom Value" "\ (Fun (Abs b) []::('fun,'atom,'sets,'lbl) prot_term) = TAtom AbsValue" "arity\<^sub>s s = 0 \ \ (Fun (Set s) []::('fun,'atom,'sets,'lbl) prot_term) = TAtom SetType" "\ (Fun (Attack x) []::('fun,'atom,'sets,'lbl) prot_term) = TAtom AttackType" "\ (Fun OccursSec []::('fun,'atom,'sets,'lbl) prot_term) = TAtom OccursSecType" "\ (Fun (PubConst a t) []::('fun,'atom,'sets,'lbl) prot_term) = TAtom a" by simp+ lemma \_Fu_simps[simp]: "arity\<^sub>f f \ 0 \ \ (Fun (Fu f) T) = TComp (Fu f) (map \ T)" (is "?A1 \ ?A2") "arity\<^sub>f f = 0 \ \\<^sub>f f = Some a \ \ (Fun (Fu f) T) = TAtom (Atom a)" (is "?B1 \ ?B2 \ ?B3") "arity\<^sub>f f = 0 \ \\<^sub>f f = None \ \ (Fun (Fu f) T) = TAtom Bottom" (is "?C1 \ ?C2 \ ?C3") "\ (Fun (Fu f) T) \ TAtom Value" (is ?D) "\ (Fun (Fu f) T) \ TAtom AttackType" (is ?E) "\ (Fun (Fu f) T) \ TAtom OccursSecType" (is ?F) proof - show "?A1 \ ?A2" by simp show "?B1 \ ?B2 \ ?B3" by simp show "?C1 \ ?C2 \ ?C3" by simp show ?D by (cases "\\<^sub>f f") simp_all show ?E by (cases "\\<^sub>f f") simp_all show ?F by (cases "\\<^sub>f f") simp_all qed lemma \_Set_simps[simp]: "arity\<^sub>s s \ 0 \ \ (Fun (Set s) T) = TComp (Set s) (map \ T)" "\ (Fun (Set s) T) = TAtom SetType \ \ (Fun (Set s) T) = TComp (Set s) (map \ T)" "\ (Fun (Set s) T) \ TAtom Value" "\ (Fun (Set s) T) \ TAtom (Atom a)" "\ (Fun (Set s) T) \ TAtom AttackType" "\ (Fun (Set s) T) \ TAtom OccursSecType" "\ (Fun (Set s) T) \ TAtom Bottom" by auto subsection \Locale Interpretations\ lemma Ana_Fu_cases: assumes "Ana (Fun f T) = (K,M)" and "f = Fu g" and "Ana\<^sub>f g = (K',M')" shows "(K,M) = (if arity\<^sub>f g = length T \ arity\<^sub>f g > 0 then (K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T, map ((!) T) M') else ([],[]))" (is ?A) and "(K,M) = (K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T, map ((!) T) M') \ (K,M) = ([],[])" (is ?B) proof - show ?A using assms by (cases "arity\<^sub>f g = length T \ arity\<^sub>f g > 0") auto thus ?B by metis qed lemma Ana_Fu_intro: assumes "arity\<^sub>f f = length T" "arity\<^sub>f f > 0" and "Ana\<^sub>f f = (K',M')" shows "Ana (Fun (Fu f) T) = (K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T, map ((!) T) M')" using assms by simp lemma Ana_Fu_elim: assumes "Ana (Fun f T) = (K,M)" and "f = Fu g" and "Ana\<^sub>f g = (K',M')" and "(K,M) \ ([],[])" shows "arity\<^sub>f g = length T" (is ?A) and "(K,M) = (K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T, map ((!) T) M')" (is ?B) proof - show ?A using assms by force moreover have "arity\<^sub>f g > 0" using assms by force ultimately show ?B using assms by auto qed lemma Ana_nonempty_inv: assumes "Ana t \ ([],[])" shows "\f T. t = Fun (Fu f) T \ arity\<^sub>f f = length T \ arity\<^sub>f f > 0 \ (\K M. Ana\<^sub>f f = (K, M) \ Ana t = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T, map ((!) T) M))" using assms proof (induction t rule: Ana.induct) case (1 f T) hence *: "arity\<^sub>f f = length T" "0 < arity\<^sub>f f" "Ana (Fun (Fu f) T) = (case Ana\<^sub>f f of (K, M) \ (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T, map ((!) T) M))" using Ana.simps(1)[of f T] unfolding Let_def by metis+ obtain K M where **: "Ana\<^sub>f f = (K, M)" by (metis surj_pair) hence "Ana (Fun (Fu f) T) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T, map ((!) T) M)" using *(3) by simp thus ?case using ** *(1,2) by blast qed simp_all lemma assm1: assumes "Ana t = (K,M)" shows "fv\<^sub>s\<^sub>e\<^sub>t (set K) \ fv t" using assms proof (induction t rule: term.induct) case (Fun f T) have aux: "fv\<^sub>s\<^sub>e\<^sub>t (set K \\<^sub>s\<^sub>e\<^sub>t (!) T) \ fv\<^sub>s\<^sub>e\<^sub>t (set T)" when K: "\i \ fv\<^sub>s\<^sub>e\<^sub>t (set K). i < length T" for K::"(('fun,'atom,'sets,'lbl) prot_fun, nat) term list" proof fix x assume "x \ fv\<^sub>s\<^sub>e\<^sub>t (set K \\<^sub>s\<^sub>e\<^sub>t (!) T)" - then obtain k where k: "k \ set K" "x \ fv (k \ (!) T)" by moura + then obtain k where k: "k \ set K" "x \ fv (k \ (!) T)" by auto have "\i \ fv k. i < length T" using K k(1) by simp thus "x \ fv\<^sub>s\<^sub>e\<^sub>t (set T)" by (metis (no_types, lifting) k(2) contra_subsetD fv_set_mono image_subsetI nth_mem subst_apply_fv_unfold) qed { fix g assume f: "f = Fu g" and K: "K \ []" - obtain K' M' where *: "Ana\<^sub>f g = (K',M')" by moura + obtain K' M' where *: "Ana\<^sub>f g = (K',M')" by force have "(K, M) \ ([], [])" using K by simp hence "(K, M) = (K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T, map ((!) T) M')" "arity\<^sub>f g = length T" using Ana_Fu_cases(1)[OF Fun.prems f *] by presburger+ hence ?case using aux[of K'] Ana\<^sub>f_assm2_alt[OF *] by auto } thus ?case using Fun by (cases f) fastforce+ qed simp lemma assm2: assumes "Ana t = (K,M)" and "\g S'. Fun g S' \ t \ length S' = arity g" and "k \ set K" and "Fun f T' \ k" shows "length T' = arity f" using assms proof (induction t rule: term.induct) case (Fun g T) obtain h where 2: "g = Fu h" using Fun.prems(1,3) by (cases g) auto - obtain K' M' where 1: "Ana\<^sub>f h = (K',M')" by moura + obtain K' M' where 1: "Ana\<^sub>f h = (K',M')" by force have "(K,M) \ ([],[])" using Fun.prems(3) by auto hence "(K,M) = (K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T, map ((!) T) M')" "\i. i \ fv\<^sub>s\<^sub>e\<^sub>t (set K') \ set M' \ i < length T" using Ana_Fu_cases(1)[OF Fun.prems(1) 2 1] Ana\<^sub>f_assm2_alt[OF 1] by presburger+ hence "K = K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T" and 3: "\i\fv\<^sub>s\<^sub>e\<^sub>t (set K'). i < length T" by simp_all - then obtain k' where k': "k' \ set K'" "k = k' \ (!) T" using Fun.prems(3) by moura + then obtain k' where k': "k' \ set K'" "k = k' \ (!) T" using Fun.prems(3) by force hence 4: "Fun f T' \ subterms (k' \ (!) T)" "fv k' \ fv\<^sub>s\<^sub>e\<^sub>t (set K')" using Fun.prems(4) by auto show ?case proof (cases "\i \ fv k'. Fun f T' \ subterms (T ! i)") case True hence "Fun f T' \ subterms\<^sub>s\<^sub>e\<^sub>t (set T)" using k' Fun.prems(4) 3 by auto thus ?thesis using Fun.prems(2) by auto next case False then obtain S where "Fun f S \ subterms k'" "Fun f T' = Fun f S \ (!) T" using k'(2) Fun.prems(4) subterm_subst_not_img_subterm by force thus ?thesis using Ana\<^sub>f_assm1_alt[OF 1, of "Fun f S"] k'(1) by (cases f) auto qed qed simp lemma assm4: assumes "Ana (Fun f T) = (K, M)" shows "set M \ set T" using assms proof (cases f) case (Fu g) - obtain K' M' where *: "Ana\<^sub>f g = (K',M')" by moura + obtain K' M' where *: "Ana\<^sub>f g = (K',M')" by force have "M = [] \ (arity\<^sub>f g = length T \ M = map ((!) T) M')" using Ana_Fu_cases(1)[OF assms Fu *] by (meson prod.inject) thus ?thesis using Ana\<^sub>f_assm2_alt[OF *] by auto qed auto lemma assm5: "Ana t = (K,M) \ K \ [] \ M \ [] \ Ana (t \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \, M \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" proof (induction t rule: term.induct) case (Fun f T) thus ?case proof (cases f) case (Fu g) - obtain K' M' where *: "Ana\<^sub>f g = (K',M')" by moura + obtain K' M' where *: "Ana\<^sub>f g = (K',M')" by force have **: "K = K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T" "M = map ((!) T) M'" "arity\<^sub>f g = length T" "\i \ fv\<^sub>s\<^sub>e\<^sub>t (set K') \ set M'. i < arity\<^sub>f g" "0 < arity\<^sub>f g" using Fun.prems(2) Ana_Fu_cases(1)[OF Fun.prems(1) Fu *] Ana\<^sub>f_assm2_alt[OF *] by (meson prod.inject)+ have ***: "\i \ fv\<^sub>s\<^sub>e\<^sub>t (set K'). i < length T" "\i \ set M'. i < length T" using **(3,4) by auto have "K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) (map (\t. t \ \) T)" "M \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = map ((!) (map (\t. t \ \) T)) M'" using subst_idx_map[OF ***(2), of \] subst_idx_map'[OF ***(1), of \] **(1,2) by fast+ thus ?thesis using Fu * **(3,5) by auto qed auto qed simp sublocale intruder_model arity public Ana apply unfold_locales by (metis assm1, metis assm2, rule Ana.simps, metis assm4, metis assm5) adhoc_overloading INTRUDER_SYNTH intruder_synth adhoc_overloading INTRUDER_DEDUCT intruder_deduct lemma assm6: "arity c = 0 \ \a. \X. \ (Fun c X) = TAtom a" by (cases c) auto lemma assm7: "0 < arity f \ \ (Fun f T) = TComp f (map \ T)" by auto lemma assm8: "infinite {c. \ (Fun c []::('fun,'atom,'sets,'lbl) prot_term) = TAtom a \ public c}" (is "?P a") proof - let ?T = "\f. (range f)::('fun,'atom,'sets,'lbl) prot_fun set" let ?A = "\f. \x::nat \ UNIV. \y::nat \ UNIV. (f x = f y) = (x = y)" let ?B = "\f. \x::nat \ UNIV. f x \ ?T f" let ?C = "\f. \y::('fun,'atom,'sets,'lbl) prot_fun \ ?T f. \x \ UNIV. y = f x" let ?D = "\f b. ?T f \ {c. \ (Fun c []::('fun,'atom,'sets,'lbl) prot_term) = TAtom b \ public c}" have sub_lmm: "?P b" when "?A f" "?C f" "?C f" "?D f b" for b f proof - have "\g::nat \ ('fun,'atom,'sets,'lbl) prot_fun. bij_betw g UNIV (?T f)" using bij_betwI'[of UNIV f "?T f"] that(1,2,3) by blast hence "infinite (?T f)" by (metis nat_not_finite bij_betw_finite) thus ?thesis using infinite_super[OF that(4)] by blast qed show ?thesis proof (cases a) case (Atom b) thus ?thesis using sub_lmm[of "PubConst (Atom b)" a] by force next case Value thus ?thesis using sub_lmm[of "PubConst Value" a] by force next case SetType thus ?thesis using sub_lmm[of "PubConst SetType" a] by fastforce next case AttackType thus ?thesis using sub_lmm[of "PubConst AttackType" a] by fastforce next case Bottom thus ?thesis using sub_lmm[of "PubConst Bottom" a] by fastforce next case OccursSecType thus ?thesis using sub_lmm[of "PubConst OccursSecType" a] by fastforce next case AbsValue thus ?thesis using sub_lmm[of "PubConst AbsValue" a] by force qed qed lemma assm9: "TComp f T \ \ t \ arity f > 0" proof (induction t rule: term.induct) case (Var x) hence "\ (Var x) \ TAtom Bottom" by force hence "\t \ subterms (fst x). case t of TComp f T \ arity f > 0 \ arity f = length T | _ \ True" using Var \.simps(1)[of x] unfolding \\<^sub>v_def by meson thus ?case using Var by (fastforce simp add: \\<^sub>v_def) next case (Fun g S) have "arity g \ 0" using Fun.prems Var_subtermeq assm6 by force thus ?case using Fun by (cases "TComp f T = TComp g (map \ S)") auto qed lemma assm10: "wf\<^sub>t\<^sub>r\<^sub>m (\ (Var x))" unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by (auto simp add: \\<^sub>v_def) lemma assm11: "arity f > 0 \ public f" using public\<^sub>f_assm by (cases f) auto lemma assm12: "\ (Var (\, n)) = \ (Var (\, m))" by (simp add: \\<^sub>v_def) lemma assm13: "arity c = 0 \ Ana (Fun c T) = ([],[])" by (cases c) simp_all lemma assm14: assumes "Ana (Fun f T) = (K,M)" shows "Ana (Fun f T \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \, M \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" proof - show ?thesis proof (cases "(K, M) = ([],[])") case True { fix g assume f: "f = Fu g" - obtain K' M' where "Ana\<^sub>f g = (K',M')" by moura + obtain K' M' where "Ana\<^sub>f g = (K',M')" by force hence ?thesis using assms f True by auto } thus ?thesis using True assms by (cases f) auto next case False then obtain g where **: "f = Fu g" using assms by (cases f) auto - obtain K' M' where *: "Ana\<^sub>f g = (K',M')" by moura + obtain K' M' where *: "Ana\<^sub>f g = (K',M')" by force have ***: "K = K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T" "M = map ((!) T) M'" "arity\<^sub>f g = length T" "\i \ fv\<^sub>s\<^sub>e\<^sub>t (set K') \ set M'. i < arity\<^sub>f g" using Ana_Fu_cases(1)[OF assms ** *] False Ana\<^sub>f_assm2_alt[OF *] by (meson prod.inject)+ have ****: "\i\fv\<^sub>s\<^sub>e\<^sub>t (set K'). i < length T" "\i\set M'. i < length T" using ***(3,4) by auto have "K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) (map (\t. t \ \) T)" "M \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = map ((!) (map (\t. t \ \) T)) M'" using subst_idx_map[OF ****(2), of \] subst_idx_map'[OF ****(1), of \] ***(1,2) by auto thus ?thesis using assms * ** ***(3) by auto qed qed sublocale labeled_stateful_typing' arity public Ana \ Pair label_witness1 label_witness2 apply unfold_locales subgoal by (metis assm6) subgoal by (metis assm7) subgoal by (metis assm9) subgoal by (rule assm10) subgoal by (metis assm12) subgoal by (metis assm13) subgoal by (metis assm14) subgoal by (rule label_witness_assm) subgoal by (rule arity.simps(5)) subgoal by (metis assm14) subgoal by (metis assm8) subgoal by (metis assm11) done subsection \The Protocol Transition System, Defined in Terms of the Reachable Constraints\ definition transaction_decl_subst where "transaction_decl_subst (\::('fun,'atom,'sets,'lbl) prot_subst) T \ subst_domain \ = fst ` set (transaction_decl T ()) \ (\(x,cs) \ set (transaction_decl T ()). \c \ cs. \ x = Fun (Fu c) [] \ arity (Fu c::('fun,'atom,'sets,'lbl) prot_fun) = 0) \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" definition transaction_fresh_subst where "transaction_fresh_subst \ T M \ subst_domain \ = set (transaction_fresh T) \ (\t \ subst_range \. \c. t = Fun c [] \ \public c \ arity c = 0) \ (\t \ subst_range \. t \ subterms\<^sub>s\<^sub>e\<^sub>t M) \ (\t \ subst_range \. t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms_transaction T)) \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ inj_on \ (subst_domain \)" (* NB: We need the protocol P as a parameter for this definition---even though we will only apply \ to a single transaction T of P---because we have to ensure that \(fv(T)) is disjoint from the bound variables of P and \. *) definition transaction_renaming_subst where "transaction_renaming_subst \ P X \ \n \ max_var_set (\(vars_transaction ` set P) \ X). \ = var_rename n" definition (in intruder_model) constraint_model where "constraint_model \ \ \ constr_sem_stateful \ (unlabel \) \ interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" definition (in typed_model) welltyped_constraint_model where "welltyped_constraint_model \ \ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ constraint_model \ \" text \ The set of symbolic constraints reachable in any symbolic run of the protocol \P\. \\\ instantiates the "declared variables" of transaction \T\ with ground terms. \\\ instantiates the fresh variables of transaction \T\ with fresh terms. \\\ is a variable-renaming whose range consists of fresh variables. \ inductive_set reachable_constraints:: "('fun,'atom,'sets,'lbl) prot \ ('fun,'atom,'sets,'lbl) prot_constr set" for P::"('fun,'atom,'sets,'lbl) prot" where init[simp]: "[] \ reachable_constraints P" | step: "\\ \ reachable_constraints P; T \ set P; transaction_decl_subst \ T; transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \); transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ \ \@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \) \ reachable_constraints P" subsection \Minor Lemmata\ lemma \\<^sub>v_TAtom[simp]: "\\<^sub>v (TAtom a, n) = TAtom a" unfolding \\<^sub>v_def by simp lemma \\<^sub>v_TAtom': assumes "a \ Bottom" shows "\\<^sub>v (\, n) = TAtom a \ \ = TAtom a" proof assume "\\<^sub>v (\, n) = TAtom a" thus "\ = TAtom a" by (metis (no_types, lifting) assms \\<^sub>v_def fst_conv term.inject(1)) qed simp lemma \\<^sub>v_TAtom_inv: "\\<^sub>v x = TAtom (Atom a) \ \m. x = (TAtom (Atom a), m)" "\\<^sub>v x = TAtom Value \ \m. x = (TAtom Value, m)" "\\<^sub>v x = TAtom SetType \ \m. x = (TAtom SetType, m)" "\\<^sub>v x = TAtom AttackType \ \m. x = (TAtom AttackType, m)" "\\<^sub>v x = TAtom OccursSecType \ \m. x = (TAtom OccursSecType, m)" by (metis \\<^sub>v_TAtom' surj_pair prot_atom.distinct(7), metis \\<^sub>v_TAtom' surj_pair prot_atom.distinct(18), metis \\<^sub>v_TAtom' surj_pair prot_atom.distinct(26), metis \\<^sub>v_TAtom' surj_pair prot_atom.distinct(32), metis \\<^sub>v_TAtom' surj_pair prot_atom.distinct(38)) lemma \\<^sub>v_TAtom'': "(fst x = TAtom (Atom a)) = (\\<^sub>v x = TAtom (Atom a))" (is "?A = ?A'") "(fst x = TAtom Value) = (\\<^sub>v x = TAtom Value)" (is "?B = ?B'") "(fst x = TAtom SetType) = (\\<^sub>v x = TAtom SetType)" (is "?C = ?C'") "(fst x = TAtom AttackType) = (\\<^sub>v x = TAtom AttackType)" (is "?D = ?D'") "(fst x = TAtom OccursSecType) = (\\<^sub>v x = TAtom OccursSecType)" (is "?E = ?E'") proof - have 1: "?A \ ?A'" "?B \ ?B'" "?C \ ?C'" "?D \ ?D'" "?E \ ?E'" by (metis \\<^sub>v_TAtom prod.collapse)+ have 2: "?A' \ ?A" "?B' \ ?B" "?C' \ ?C" "?D' \ ?D" "?E' \ ?E" using \\<^sub>v_TAtom \\<^sub>v_TAtom_inv(1) apply fastforce using \\<^sub>v_TAtom \\<^sub>v_TAtom_inv(2) apply fastforce using \\<^sub>v_TAtom \\<^sub>v_TAtom_inv(3) apply fastforce using \\<^sub>v_TAtom \\<^sub>v_TAtom_inv(4) apply fastforce using \\<^sub>v_TAtom \\<^sub>v_TAtom_inv(5) by fastforce show "?A = ?A'" "?B = ?B'" "?C = ?C'" "?D = ?D'" "?E = ?E'" using 1 2 by metis+ qed lemma \\<^sub>v_Var_image: "\\<^sub>v ` X = \ ` Var ` X" by force lemma \_Fu_const: assumes "arity\<^sub>f g = 0" shows "\a. \ (Fun (Fu g) T) = TAtom (Atom a)" proof - have "\\<^sub>f g \ None" using assms \\<^sub>f_assm by blast thus ?thesis using assms by force qed lemma Fun_Value_type_inv: fixes T::"('fun,'atom,'sets,'lbl) prot_term list" assumes "\ (Fun f T) = TAtom Value" shows "(\n. f = Val n) \ (\bs. f = Abs bs) \ (\n. f = PubConst Value n)" proof - have *: "arity f = 0" by (metis const_type_inv assms) show ?thesis using assms proof (cases f) case (Fu g) hence "arity\<^sub>f g = 0" using * by simp hence False using Fu \_Fu_const[of g T] assms by auto thus ?thesis by metis next case (Set s) hence "arity\<^sub>s s = 0" using * by simp hence False using Set assms by auto thus ?thesis by metis qed simp_all qed lemma Ana\<^sub>f_keys_not_val_terms: assumes "Ana\<^sub>f f = (K, T)" and "k \ set K" and "g \ funs_term k" shows "\is_Val g" and "\is_PubConstValue g" and "\is_Abs g" proof - { assume "is_Val g" then obtain n S where *: "Fun (Val n) S \ subterms\<^sub>s\<^sub>e\<^sub>t (set K)" using assms(2) funs_term_Fun_subterm[OF assms(3)] by (cases g) auto hence False using Ana\<^sub>f_assm1_alt[OF assms(1) *] by simp } moreover { assume "is_PubConstValue g" then obtain n S where *: "Fun (PubConst Value n) S \ subterms\<^sub>s\<^sub>e\<^sub>t (set K)" using assms(2) funs_term_Fun_subterm[OF assms(3)] unfolding is_PubConstValue_def by (cases g) auto hence False using Ana\<^sub>f_assm1_alt[OF assms(1) *] by simp } moreover { assume "is_Abs g" then obtain a S where *: "Fun (Abs a) S \ subterms\<^sub>s\<^sub>e\<^sub>t (set K)" using assms(2) funs_term_Fun_subterm[OF assms(3)] by (cases g) auto hence False using Ana\<^sub>f_assm1_alt[OF assms(1) *] by simp } ultimately show "\is_Val g" "\is_PubConstValue g" "\is_Abs g" by metis+ qed lemma Ana\<^sub>f_keys_not_pairs: assumes "Ana\<^sub>f f = (K, T)" and "k \ set K" and "g \ funs_term k" shows "g \ Pair" proof assume "g = Pair" then obtain S where *: "Fun Pair S \ subterms\<^sub>s\<^sub>e\<^sub>t (set K)" using assms(2) funs_term_Fun_subterm[OF assms(3)] by (cases g) auto show False using Ana\<^sub>f_assm1_alt[OF assms(1) *] by simp qed lemma Ana_Fu_keys_funs_term_subset: fixes K::"('fun,'atom,'sets,'lbl) prot_term list" assumes "Ana (Fun (Fu f) S) = (K, T)" and "Ana\<^sub>f f = (K', T')" shows "\(funs_term ` set K) \ \(funs_term ` set K') \ funs_term (Fun (Fu f) S)" proof - { fix k assume k: "k \ set K" then obtain k' where k': "k' \ set K'" "k = k' \ (!) S" "arity\<^sub>f f = length S" "subterms k' \ subterms\<^sub>s\<^sub>e\<^sub>t (set K')" using assms Ana_Fu_elim[OF assms(1) _ assms(2)] by fastforce have 1: "funs_term k' \ \(funs_term ` set K')" using k'(1) by auto have "i < length S" when "i \ fv k'" for i using that Ana\<^sub>f_assm2_alt[OF assms(2), of i] k'(1,3) by auto hence 2: "funs_term (S ! i) \ funs_term (Fun (Fu f) S)" when "i \ fv k'" for i using that by force have "funs_term k \ \(funs_term ` set K') \ funs_term (Fun (Fu f) S)" using funs_term_subst[of k' "(!) S"] k'(2) 1 2 by fast } thus ?thesis by blast qed lemma Ana_Fu_keys_not_pubval_terms: fixes k::"('fun,'atom,'sets,'lbl) prot_term" assumes "Ana (Fun (Fu f) S) = (K, T)" and "Ana\<^sub>f f = (K', T')" and "k \ set K" and "\g \ funs_term (Fun (Fu f) S). \is_PubConstValue g" shows "\g \ funs_term k. \is_PubConstValue g" using assms(3,4) Ana\<^sub>f_keys_not_val_terms(1,2)[OF assms(2)] Ana_Fu_keys_funs_term_subset[OF assms(1,2)] by blast lemma Ana_Fu_keys_not_abs_terms: fixes k::"('fun,'atom,'sets,'lbl) prot_term" assumes "Ana (Fun (Fu f) S) = (K, T)" and "Ana\<^sub>f f = (K', T')" and "k \ set K" and "\g \ funs_term (Fun (Fu f) S). \is_Abs g" shows "\g \ funs_term k. \is_Abs g" using assms(3,4) Ana\<^sub>f_keys_not_val_terms(3)[OF assms(2)] Ana_Fu_keys_funs_term_subset[OF assms(1,2)] by blast lemma Ana_Fu_keys_not_pairs: fixes k::"('fun,'atom,'sets,'lbl) prot_term" assumes "Ana (Fun (Fu f) S) = (K, T)" and "Ana\<^sub>f f = (K', T')" and "k \ set K" and "\g \ funs_term (Fun (Fu f) S). g \ Pair" shows "\g \ funs_term k. g \ Pair" using assms(3,4) Ana\<^sub>f_keys_not_pairs[OF assms(2)] Ana_Fu_keys_funs_term_subset[OF assms(1,2)] by blast lemma Ana_Fu_keys_length_eq: assumes "length T = length S" shows "length (fst (Ana (Fun (Fu f) T))) = length (fst (Ana (Fun (Fu f) S)))" proof (cases "arity\<^sub>f f = length T \ arity\<^sub>f f > 0") case True thus ?thesis using assms by (cases "Ana\<^sub>f f") auto next case False thus ?thesis using assms by force qed lemma Ana_key_PubConstValue_subterm_in_term: fixes k::"('fun,'atom,'sets,'lbl) prot_term" assumes KR: "Ana t = (K, R)" and k: "k \ set K" and n: "Fun (PubConst Value n) [] \ k" shows "Fun (PubConst Value n) [] \ t" proof (cases t) case (Var x) thus ?thesis using KR k n by force next case (Fun f ts) note t = this then obtain g where f: "f = Fu g" using KR k by (cases f) auto obtain K' R' where KR': "Ana\<^sub>f g = (K', R')" by fastforce have K: "K = K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) ts" using k Ana_Fu_elim(2)[OF KR[unfolded t] f KR'] by force obtain k' where k': "k' \ set K'" "k = k' \ (!) ts" using k K by auto have 0: "\(Fun (PubConst Value n) [] \ k')" proof assume *: "Fun (PubConst Value n) [] \ k'" have **: "PubConst Value n \ funs_term k'" using funs_term_Fun_subterm'[OF *] by (cases k') auto show False using Ana\<^sub>f_keys_not_val_terms(2)[OF KR' k'(1) **] unfolding is_PubConstValue_def by force qed hence "\i \ fv k'. Fun (PubConst Value n) [] \ ts ! i" by (metis n const_subterm_subst_var_obtain k'(2)) then obtain i where i: "i \ fv k'" "Fun (PubConst Value n) [] \ ts ! i" by blast have "i < length ts" using i(1) KR' k'(1) Ana\<^sub>f_assm2_alt[OF KR', of i] Ana_Fu_elim(1)[OF KR[unfolded t] f KR'] k by fastforce thus ?thesis using i(2) unfolding t by force qed lemma deduct_occurs_in_ik: fixes t::"('fun,'atom,'sets,'lbl) prot_term" assumes t: "M \ occurs t" and M: "\s \ subterms\<^sub>s\<^sub>e\<^sub>t M. OccursFact \ \(funs_term ` set (snd (Ana s)))" "\s \ subterms\<^sub>s\<^sub>e\<^sub>t M. OccursSec \ \(funs_term ` set (snd (Ana s)))" "Fun OccursSec [] \ M" shows "occurs t \ M" using private_fun_deduct_in_ik''[of M OccursFact "[Fun OccursSec [], t]" OccursSec] t M by fastforce lemma deduct_val_const_swap: fixes \ \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "M \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" and "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv t. (\n. \ x = Fun (Val n) []) \ (\n. \ x = Fun (PubConst Value n) [])" and "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv t. (\n. \ x = Fun (Val n) [])" and "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv t. (\n. \ x = Fun (PubConst Value n) []) \ \ x \ M \ N" and "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv t. (\n. \ x = Fun (Val n) []) \ \ x = \ x" and "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv t. \y \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv t. \ x = \ y \ \ x = \ y" and "\n. \(Fun (PubConst Value n) [] \\<^sub>s\<^sub>e\<^sub>t insert t M)" shows "(M \\<^sub>s\<^sub>e\<^sub>t \) \ N \ t \ \" proof - obtain n where n: "intruder_deduct_num (M \\<^sub>s\<^sub>e\<^sub>t \) n (t \ \)" using assms(1) deduct_num_if_deduct by blast hence "\m \ n. intruder_deduct_num ((M \\<^sub>s\<^sub>e\<^sub>t \) \ N) m (t \ \)" using assms(2-) proof (induction n arbitrary: t rule: nat_less_induct) case (1 n) note prems = "1.prems" note IH = "1.IH" show ?case proof (cases "t \ \ \ M \\<^sub>s\<^sub>e\<^sub>t \") case True note 2 = this have 3: "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv t. \c. \ x = Fun c []" "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv t. \c. \ x = Fun c []" using prems(2,3) by (blast, blast) have "t \ \ \ M \\<^sub>s\<^sub>e\<^sub>t \" using subst_const_swap_eq_mem[OF 2 _ 3 prems(6)] prems(2,5,7) by metis thus ?thesis using intruder_deduct_num.AxiomN by auto next case False then obtain n' where n: "n = Suc n'" using prems(1) deduct_zero_in_ik by (cases n) fast+ have M_subterms_eq: "subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \) = subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" "subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \) = subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" subgoal using prems(2) subterms_subst''[of M \] by blast subgoal using prems(3) subterms_subst''[of M \] by blast done from deduct_inv[OF prems(1)] show ?thesis proof (elim disjE) assume "t \ \ \ M \\<^sub>s\<^sub>e\<^sub>t \" thus ?thesis using False by argo next assume "\f ts. t \ \ = Fun f ts \ public f \ length ts = arity f \ (\t \ set ts. \l < n. intruder_deduct_num (M \\<^sub>s\<^sub>e\<^sub>t \) l t)" then obtain f ts where t: "t \ \ = Fun f ts" "public f" "length ts = arity f" "\t \ set ts. \l < n. intruder_deduct_num (M \\<^sub>s\<^sub>e\<^sub>t \) l t" by blast show ?thesis proof (cases t) case (Var x) hence ts: "ts = []" and f: "\c. f = PubConst Value c" using t(1,2) prems(2) by (force, auto) have "\ x \ M \ N" using prems(4) Var f ts t(1) by auto moreover have "fv (\ x) = {}" using prems(3) Var by auto hence "\ x \ M \\<^sub>s\<^sub>e\<^sub>t \" when "\ x \ M" using that subst_ground_ident[of "\ x" \] by force ultimately have "\ x \ (M \\<^sub>s\<^sub>e\<^sub>t \) \ N" by fast thus ?thesis using intruder_deduct_num.AxiomN Var by force next case (Fun g ss) hence f: "f = g" and ts: "ts = ss \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" using t(1) by auto have ss: "\l < n. intruder_deduct_num (M \\<^sub>s\<^sub>e\<^sub>t \) l (s \ \)" when s: "s \ set ss" for s using t(4) ts s by auto have IH': "\l < n. intruder_deduct_num ((M \\<^sub>s\<^sub>e\<^sub>t \) \ N) l (s \ \)" when s: "s \ set ss" for s proof - obtain l where l: "l < n" "intruder_deduct_num (M \\<^sub>s\<^sub>e\<^sub>t \) l (s \ \)" using ss s by blast have *: "fv s \ fv t" "subterms\<^sub>s\<^sub>e\<^sub>t (insert s M) \ subterms\<^sub>s\<^sub>e\<^sub>t (insert t M)" using s unfolding Fun f ts by auto have "\l' \ l. intruder_deduct_num ((M \\<^sub>s\<^sub>e\<^sub>t \) \ N) l' (s \ \)" proof - have "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv s. (\n. \ x = Fun (Val n) []) \ (\n. \ x = Fun (PubConst Value n) [])" "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv s. \n. \ x = Fun (Val n) []" "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv s. (\n. \ x = Fun (PubConst Value n) []) \ \ x \ M \ N" "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv s. (\n. \ x = Fun (Val n) []) \ \ x = \ x" "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv s. \y \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv s. \ x = \ y \ \ x = \ y" "\n. Fun (PubConst Value n) [] \ subterms\<^sub>s\<^sub>e\<^sub>t (insert s M)" subgoal using prems(2) *(1) by blast subgoal using prems(3) *(1) by blast subgoal using prems(4) *(1) by blast subgoal using prems(5) *(1) by blast subgoal using prems(6) *(1) by blast subgoal using prems(7) *(2) by blast done thus ?thesis using IH l by presburger qed then obtain l' where l': "l' \ l" "intruder_deduct_num ((M \\<^sub>s\<^sub>e\<^sub>t \) \ N) l' (s \ \)" by blast have "l' < n" using l'(1) l(1) by linarith thus ?thesis using l'(2) by blast qed have g: "length (ss \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \) = arity g" "public g" using t(2,3) unfolding f ts by auto let ?P = "\s l. l < n \ intruder_deduct_num ((M \\<^sub>s\<^sub>e\<^sub>t \) \ N) l s" define steps where "steps \ \s. SOME l. ?P s l" have 2: "steps (s \ \) < n" "intruder_deduct_num ((M \\<^sub>s\<^sub>e\<^sub>t \) \ N) (steps (s \ \)) (s \ \)" when s: "s \ set ss" for s using someI_ex[OF IH'[OF s]] unfolding steps_def by (blast, blast) have 3: "Suc (Max (insert 0 (steps ` set (ss \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)))) \ n" proof (cases "ss = []") case True show ?thesis unfolding True n by simp next case False thus ?thesis using 2 Max_nat_finite_lt[of "set (ss \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" steps n] by (simp add: Suc_leI) qed show ?thesis using intruder_deduct_num.ComposeN[OF g, of "(M \\<^sub>s\<^sub>e\<^sub>t \) \ N" steps] 2(2) 3 unfolding Fun by auto qed next assume "\s \ subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \). (\l < n. intruder_deduct_num (M \\<^sub>s\<^sub>e\<^sub>t \) l s) \ (\k \ set (fst (Ana s)). \l < n. intruder_deduct_num (M \\<^sub>s\<^sub>e\<^sub>t \) l k) \ t \ \ \ set (snd (Ana s))" then obtain s l where s: "s \ subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" "\k \ set (fst (Ana s)). \l < n. intruder_deduct_num (M \\<^sub>s\<^sub>e\<^sub>t \) l k" "t \ \ \ set (snd (Ana s))" and l: "l < n" "intruder_deduct_num (M \\<^sub>s\<^sub>e\<^sub>t \) l s" by (metis (no_types, lifting) M_subterms_eq(1)) obtain u where u: "u \\<^sub>s\<^sub>e\<^sub>t M" "s = u \ \" using s(1) by blast have u_fv: "fv u \ fv\<^sub>s\<^sub>e\<^sub>t M" by (metis fv_subset_subterms u(1)) have "\x. u = Var x" proof assume "\x. u = Var x" then obtain x where x: "u = Var x" by blast then obtain c where c: "s = Fun c []" using u prems(2) u_fv by auto thus False using s(3) Ana_subterm by (cases "Ana s") force qed then obtain f ts where u': "u = Fun f ts" by (cases u) auto obtain K R where KR: "Ana u = (K,R)" by (metis surj_pair) have KR': "Ana s = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \, R \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using KR Ana_subst'[OF KR[unfolded u'], of \] unfolding u(2) u' by blast hence s': "\k \ set K. \l < n. intruder_deduct_num (M \\<^sub>s\<^sub>e\<^sub>t \) l (k \ \)" "t \ \ \ set (R \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using s(2,3) by auto have IH1: "\l < n. intruder_deduct_num ((M \\<^sub>s\<^sub>e\<^sub>t \) \ N) l (u \ \)" proof - have "subterms u \ subterms\<^sub>s\<^sub>e\<^sub>t M" using u(1) subterms_subset by auto hence "subterms\<^sub>s\<^sub>e\<^sub>t (insert u M) = subterms\<^sub>s\<^sub>e\<^sub>t M" by blast hence *: "subterms\<^sub>s\<^sub>e\<^sub>t (insert u M) \ subterms\<^sub>s\<^sub>e\<^sub>t (insert t M)" by auto have "\l' \ l. intruder_deduct_num ((M \\<^sub>s\<^sub>e\<^sub>t \) \ N) l' (u \ \)" proof - have "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv u. (\n. \ x = Fun (Val n) []) \ (\n. \ x = Fun (PubConst Value n) [])" "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv u. \n. \ x = Fun (Val n) []" "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv u. (\n. \ x = Fun (PubConst Value n) []) \ \ x \ M \ N" "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv u. (\n. \ x = Fun (Val n) []) \ \ x = \ x" "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv u. \y \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv u. \ x = \ y \ \ x = \ y" "\n. Fun (PubConst Value n) [] \ subterms\<^sub>s\<^sub>e\<^sub>t (insert u M)" subgoal using prems(2) u_fv by blast subgoal using prems(3) u_fv by blast subgoal using prems(4) u_fv by blast subgoal using prems(5) u_fv by blast subgoal using prems(6) u_fv by blast subgoal using prems(7) * by blast done thus ?thesis using IH l unfolding u(2) by presburger qed then obtain l' where l': "l' \ l" "intruder_deduct_num ((M \\<^sub>s\<^sub>e\<^sub>t \) \ N) l' (u \ \)" by blast have "l' < n" using l'(1) l(1) by linarith thus ?thesis using l'(2) by blast qed have IH2: "\l < n. intruder_deduct_num ((M \\<^sub>s\<^sub>e\<^sub>t \) \ N) l (k \ \)" when k: "k \ set K" for k using k IH prems(2-) Ana\<^sub>f_keys_not_val_terms s'(1) KR u(1) proof - have *: "fv k \ fv\<^sub>s\<^sub>e\<^sub>t M" using k KR Ana_keys_fv u(1) fv_subset_subterms by blast have **: "Fun (PubConst Value n) [] \\<^sub>s\<^sub>e\<^sub>t M" when "Fun (PubConst Value n) [] \ k" for n using in_subterms_subset_Union[OF u(1)] Ana_key_PubConstValue_subterm_in_term[OF KR k that] by fast obtain lk where lk: "lk < n" "intruder_deduct_num (M \\<^sub>s\<^sub>e\<^sub>t \) lk (k \ \)" using s'(1) k by fast have "\l' \ lk. intruder_deduct_num ((M \\<^sub>s\<^sub>e\<^sub>t \) \ N) l' (k \ \)" proof - have "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv k. (\n. \ x = Fun (Val n) []) \ (\n. \ x = Fun (PubConst Value n) [])" "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv k. \n. \ x = Fun (Val n) []" "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv k. (\n. \ x = Fun (PubConst Value n) []) \ \ x \ M \ N" "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv k. (\n. \ x = Fun (Val n) []) \ \ x = \ x" "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv k. \y \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv k. \ x = \ y \ \ x = \ y" "\n. Fun (PubConst Value n) [] \ subterms\<^sub>s\<^sub>e\<^sub>t (insert k M)" subgoal using prems(2) * by blast subgoal using prems(3) * by blast subgoal using prems(4) * by blast subgoal using prems(5) * by blast subgoal using prems(6) * by blast subgoal using prems(7) ** by blast done thus ?thesis using IH lk by presburger qed then obtain lk' where lk': "lk' \ lk" "intruder_deduct_num ((M \\<^sub>s\<^sub>e\<^sub>t \) \ N) lk' (k \ \)" by blast have "lk' < n" using lk'(1) lk(1) by linarith thus ?thesis using lk'(2) by blast qed have KR'': "Ana (u \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \, R \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using Ana_subst' KR unfolding u' by blast obtain r where r: "r \ set R" "t \ \ = r \ \" using s'(2) by fastforce have r': "t \ \ \ set (R \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" proof - have r_subterm_u: "r \ u" using r(1) KR Ana_subterm by blast have r_fv: "fv r \ fv\<^sub>s\<^sub>e\<^sub>t M" by (meson r_subterm_u u(1) fv_subset_subterms in_mono in_subterms_subset_Union) have t_subterms_M: "subterms t \ subterms\<^sub>s\<^sub>e\<^sub>t (insert t M)" by blast have r_subterm_M: "subterms r \ subterms\<^sub>s\<^sub>e\<^sub>t (insert t M)" using subterms_subset[OF r_subterm_u] in_subterms_subset_Union[OF u(1)] by (auto intro: subterms\<^sub>s\<^sub>e\<^sub>t_mono) have *: "\x \ fv t \ fv r. \ x = \ x \ \(\ x \ t) \ \(\ x \ r)" proof fix x assume "x \ fv t \ fv r" hence "x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv t" using r_fv by blast thus "\ x = \ x \ \(\ x \ t) \ \(\ x \ r)" using prems(2,5,7) r_subterm_M t_subterms_M by (metis (no_types, opaque_lifting) in_mono) qed have **: "\x \ fv t \ fv r. \c. \ x = Fun c []" "\x \ fv t \ fv r. \c. \ x = Fun c []" "\x \ fv t \ fv r. \y \ fv t \ fv r. \ x = \ y \ \ x = \ y" subgoal using prems(2) r_fv by blast subgoal using prems(3) r_fv by blast subgoal using prems(6) r_fv by blast done have "t \ \ = r \ \" by (rule subst_const_swap_eq'[OF r(2) * **]) thus ?thesis using r(1) by simp qed obtain l1 where l1: "l1 < n" "intruder_deduct_num ((M \\<^sub>s\<^sub>e\<^sub>t \) \ N) l1 (u \ \)" using IH1 by blast let ?P = "\s l. l < n \ intruder_deduct_num ((M \\<^sub>s\<^sub>e\<^sub>t \) \ N) l s" define steps where "steps \ \s. SOME l. ?P s l" have 2: "steps (k \ \) < n" "intruder_deduct_num ((M \\<^sub>s\<^sub>e\<^sub>t \) \ N) (steps (k \ \)) (k \ \)" when k: "k \ set K" for k using someI_ex[OF IH2[OF k]] unfolding steps_def by (blast, blast) have 3: "Suc (Max (insert l1 (steps ` set (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)))) \ n" proof (cases "K = []") case True show ?thesis using l1(1) unfolding True n by simp next case False thus ?thesis using l1(1) 2 Max_nat_finite_lt[of "set (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" steps n] by (simp add: Suc_leI) qed have IH2': "intruder_deduct_num ((M \\<^sub>s\<^sub>e\<^sub>t \) \ N) (steps k) k" when k: "k \ set (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" for k using IH2 k 2 by auto show ?thesis using l1(1) intruder_deduct_num.DecomposeN[OF l1(2) KR'' IH2' r'] 3 by fast qed qed qed thus ?thesis using deduct_if_deduct_num by blast qed lemma constraint_model_Nil: assumes I: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)" shows "constraint_model I []" using I unfolding constraint_model_def by simp lemma welltyped_constraint_model_Nil: assumes I: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)" shows "welltyped_constraint_model I []" using I(1) constraint_model_Nil[OF I(2,3)] unfolding welltyped_constraint_model_def by simp lemma constraint_model_prefix: assumes "constraint_model I (A@B)" shows "constraint_model I A" by (metis assms strand_sem_append_stateful unlabel_append constraint_model_def) lemma welltyped_constraint_model_prefix: assumes "welltyped_constraint_model I (A@B)" shows "welltyped_constraint_model I A" by (metis assms constraint_model_prefix welltyped_constraint_model_def) lemma welltyped_constraint_model_deduct_append: assumes "welltyped_constraint_model I A" and "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I \ s \ I" shows "welltyped_constraint_model I (A@[(l,send\[s]\)])" using assms strand_sem_append_stateful[of "{}" "{}" "unlabel A" _ I] unfolding welltyped_constraint_model_def constraint_model_def by simp lemma welltyped_constraint_model_deduct_split: assumes "welltyped_constraint_model I (A@[(l,send\[s]\)])" shows "welltyped_constraint_model I A" and "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I \ s \ I" using assms strand_sem_append_stateful[of "{}" "{}" "unlabel A" _ I] unfolding welltyped_constraint_model_def constraint_model_def by simp_all lemma welltyped_constraint_model_deduct_iff: "welltyped_constraint_model I (A@[(l,send\[s]\)]) \ welltyped_constraint_model I A \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I \ s \ I" by (metis welltyped_constraint_model_deduct_append welltyped_constraint_model_deduct_split) lemma welltyped_constraint_model_attack_if_receive_attack: assumes I: "welltyped_constraint_model \ \" and rcv_attack: "receive\ts\ \ set (unlabel \)" "attack\n\ \ set ts" shows "welltyped_constraint_model \ (\@[(l, send\[attack\n\]\)])" proof - have "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ attack\n\" using rcv_attack in_ik\<^sub>s\<^sub>s\<^sub>t_iff[of "attack\n\" "unlabel \"] ideduct_subst[OF intruder_deduct.Axiom[of "attack\n\" "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"], of \] by auto thus ?thesis using I strand_sem_append_stateful[of "{}" "{}" "unlabel \" "[send\[attack\n\]\]" \] unfolding welltyped_constraint_model_def constraint_model_def by auto qed lemma constraint_model_Val_is_Value_term: assumes "welltyped_constraint_model I A" and "t \ I = Fun (Val n) []" shows "t = Fun (Val n) [] \ (\m. t = Var (TAtom Value, m))" proof - have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" using assms(1) unfolding welltyped_constraint_model_def by simp moreover have "\ (Fun (Val n) []) = TAtom Value" by auto ultimately have *: "\ t = TAtom Value" by (metis (no_types) assms(2) wt_subst_trm'') show ?thesis proof (cases t) case (Var x) obtain \ m where x: "x = (\, m)" by (metis surj_pair) have "\\<^sub>v x = TAtom Value" using * Var by auto hence "\ = TAtom Value" using x \\<^sub>v_TAtom'[of Value \ m] by simp thus ?thesis using x Var by metis next case (Fun f T) thus ?thesis using assms(2) by auto qed qed lemma wellformed_transaction_sem_receives: fixes T::"('fun,'atom,'sets,'lbl) prot_transaction" assumes T_valid: "wellformed_transaction T" and \: "strand_sem_stateful IK DB (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) \" and s: "receive\ts\ \ set (unlabel (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" shows "\t \ set ts. IK \ t \ \" proof - let ?R = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" let ?S = "\A. unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" let ?S' = "?S (transaction_receive T)" obtain l B s where B: "(l,send\ts\) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "prefix ((B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@[(l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \]) (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using s dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(2)[of ts "transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_in_set_prefix_obtain_subst[of "send\ts\" "transaction_receive T" \] by blast have 1: "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@[(l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \])) = unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))@[send\ts\]" using B(1) unlabel_append dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst singleton_lst_proj(4) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst_snoc subst_lsst_append subst_lsst_singleton by (metis (no_types, lifting) subst_apply_labeled_stateful_strand_step.simps) have "strand_sem_stateful IK DB ?S' \" using \ strand_sem_append_stateful[of IK DB _ _ \] transaction_dual_subst_unfold[of T \] by fastforce hence "strand_sem_stateful IK DB (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))@[send\ts\]) \" using B 1 unfolding prefix_def unlabel_def by (metis dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def map_append strand_sem_append_stateful) hence t_deduct: "\t \ set ts. IK \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \\<^sub>s\<^sub>e\<^sub>t \) \ t \ \" using strand_sem_append_stateful[of IK DB "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" "[send\ts\]" \] by simp have "\s \ set (unlabel (transaction_receive T)). \t. s = receive\t\" using T_valid wellformed_transaction_unlabel_cases(1)[OF T_valid] by auto moreover { fix A::"('fun,'atom,'sets,'lbl) prot_strand" and \ assume "\s \ set (unlabel A). \t. s = receive\t\" hence "\s \ set (unlabel (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)). \t. s = receive\t\" proof (induction A) case (Cons a A) thus ?case using subst_lsst_cons[of a A \] by (cases a) auto qed simp hence "\s \ set (unlabel (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)). \t. s = receive\t\" by (simp add: list.pred_set is_Receive_def) hence "\s \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))). \t. s = send\t\" by (metis dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_memberD dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_inv(2) unlabel_in unlabel_mem_has_label) } ultimately have "\s \ set ?R. \ts. s = send\ts\" by simp hence "ik\<^sub>s\<^sub>s\<^sub>t ?R = {}" unfolding unlabel_def ik\<^sub>s\<^sub>s\<^sub>t_def by fast hence "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) = {}" using B(2) 1 ik\<^sub>s\<^sub>s\<^sub>t_append dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_append by (metis (no_types, lifting) Un_empty map_append prefix_def unlabel_def) thus ?thesis using t_deduct by simp qed lemma wellformed_transaction_sem_pos_checks: assumes T_valid: "wellformed_transaction T" and \: "strand_sem_stateful IK DB (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) \" shows "\ac: t \ u\ \ set (unlabel (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ (t \ \, u \ \) \ DB" and "\ac: t \ u\ \ set (unlabel (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ t \ \ = u \ \" proof - let ?s = "\ac: t \ u\" let ?s' = "\ac: t \ u\" let ?C = "set (unlabel (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" let ?R = "transaction_receive T@transaction_checks T" let ?R' = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" let ?S = "\A. unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" let ?S' = "?S (transaction_receive T)@?S (transaction_checks T)" let ?P = "\a. is_Receive a \ is_Check_or_Assignment a" let ?Q = "\a. is_Send a \ is_Check_or_Assignment a" let ?dbupd = "\B. dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) \ DB" have s_in: "?s \ ?C \ ?s \ set (unlabel (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" "?s' \ ?C \ ?s' \ set (unlabel (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using subst_lsst_append[of "transaction_receive T"] unlabel_append[of "transaction_receive T"] by auto have 1: "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@[(l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \])) = unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))@[s']" when B: "(l,s') = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "s' = \ac: t \ u\ \ s' = \ac: t \ u\" for l s s' and B::"('fun,'atom,'sets,'lbl) prot_strand" using B unlabel_append dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst singleton_lst_proj(4) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst_snoc subst_lsst_append subst_lsst_singleton by (metis (no_types, lifting) subst_apply_labeled_stateful_strand_step.simps) have 2: "strand_sem_stateful IK DB (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))@[s']) \" when B: "(l,s') = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "prefix ((B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@[(l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \]) (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "s' = \ac: t \ u\ \ s' = \ac: t \ u\" for l s s' and B::"('fun,'atom,'sets,'lbl) prot_strand" proof - have "strand_sem_stateful IK DB ?S' \" using \ strand_sem_append_stateful[of IK DB _ _ \] transaction_dual_subst_unfold[of T \] by fastforce thus ?thesis using B(2) 1[OF B(1,3)] strand_sem_append_stateful subst_lsst_append unfolding prefix_def unlabel_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by (metis (no_types) map_append) qed have s_sem: "?s \ ?C \ (l,?s) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \ (t \ \, u \ \) \ ?dbupd B" "?s' \ ?C \ (l,?s') = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \ t \ \ = u \ \" when B: "prefix ((B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@[(l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \]) (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" for l s and B::"('fun,'atom,'sets,'lbl) prot_strand" using 2[OF _ B] strand_sem_append_stateful[of IK DB "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" _ \] by (fastforce, fastforce) have 3: "\a \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))). \is_Insert a \ \is_Delete a" when B: "prefix ((B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@[(l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \]) (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" for l s and B::"('fun,'atom,'sets,'lbl) prot_strand" proof - have "\a \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))). ?Q a" proof fix a assume a: "a \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" have "?P a" when a: "a \ set (unlabel ?R)" for a using a wellformed_transaction_unlabel_cases(1,2)[OF T_valid] unfolding unlabel_def by fastforce hence "?P a" when a: "a \ set (unlabel (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" for a using a stateful_strand_step_cases_subst(2,11)[of _ \] subst_lsst_unlabel[of ?R \] unfolding subst_apply_stateful_strand_def by auto hence B_P: "\a \ set (unlabel (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)). ?P a" using unlabel_mono[OF set_mono_prefix[OF append_prefixD[OF B]]] by blast obtain l where "(l,a) \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using a by (meson unlabel_mem_has_label) then obtain b where b: "(l,b) \ set (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (l,b) = (l,a)" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_memberD by blast hence "?P b" using B_P unfolding unlabel_def by fastforce thus "?Q a" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_inv[OF b(2)] by (cases b) auto qed thus ?thesis by fastforce qed show "(t \ \, u \ \) \ DB" when s: "?s \ ?C" proof - obtain l B s where B: "(l,?s) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "prefix ((B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@[(l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \]) (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using s_in(1)[OF s] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(6)[of _ t u] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_in_set_prefix_obtain_subst[of ?s ?R \] by blast show ?thesis using 3[OF B(2)] s_sem(1)[OF B(2) s B(1)] dbupd\<^sub>s\<^sub>s\<^sub>t_no_upd[of "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" \ DB] by simp qed show "t \ \ = u \ \" when s: "?s' \ ?C" proof - obtain l B s where B: "(l,?s') = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "prefix ((B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@[(l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \]) (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using s_in(2)[OF s] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(3)[of _ t u] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_in_set_prefix_obtain_subst[of ?s' ?R \] by blast show ?thesis using 3[OF B(2)] s_sem(2)[OF B(2) s B(1)] dbupd\<^sub>s\<^sub>s\<^sub>t_no_upd[of "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" \ DB] by simp qed qed lemma wellformed_transaction_sem_neg_checks: assumes T_valid: "wellformed_transaction T" and \: "strand_sem_stateful IK DB (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) \" and "NegChecks X F G \ set (unlabel (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" shows "negchecks_model \ DB X F G" proof - let ?s = "NegChecks X F G" let ?R = "transaction_receive T@transaction_checks T" let ?R' = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" let ?S = "\A. unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" let ?S' = "?S (transaction_receive T)@?S (transaction_checks T)" let ?P = "\a. is_Receive a \ is_Check_or_Assignment a" let ?Q = "\a. is_Send a \ is_Check_or_Assignment a" let ?U = "\\. subst_domain \ = set X \ ground (subst_range \)" have s: "?s \ set (unlabel (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using assms(3) subst_lsst_append[of "transaction_receive T"] unlabel_append[of "transaction_receive T"] by auto obtain l B s where B: "(l,?s) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "prefix ((B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@[(l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \]) (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using s dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(7)[of X F G] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_in_set_prefix_obtain_subst[of ?s ?R \] by blast have 1: "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@[(l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \])) = unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))@[?s]" using B(1) unlabel_append dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst singleton_lst_proj(4) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst_snoc subst_lsst_append subst_lsst_singleton by (metis (no_types, lifting) subst_apply_labeled_stateful_strand_step.simps) have "strand_sem_stateful IK DB ?S' \" using \ strand_sem_append_stateful[of IK DB _ _ \] transaction_dual_subst_unfold[of T \] by fastforce hence "strand_sem_stateful IK DB (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))@[?s]) \" using B 1 strand_sem_append_stateful subst_lsst_append unfolding prefix_def unlabel_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by (metis (no_types) map_append) hence s_sem: "negchecks_model \ (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) \ DB) X F G" using strand_sem_append_stateful[of IK DB "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" "[?s]" \] by fastforce have "\a \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))). ?Q a" proof fix a assume a: "a \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" have "?P a" when a: "a \ set (unlabel ?R)" for a using a wellformed_transaction_unlabel_cases(1,2,3)[OF T_valid] unfolding unlabel_def by fastforce hence "?P a" when a: "a \ set (unlabel (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" for a using a stateful_strand_step_cases_subst(2,11)[of _ \] subst_lsst_unlabel[of ?R \] unfolding subst_apply_stateful_strand_def by auto hence B_P: "\a \ set (unlabel (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)). ?P a" using unlabel_mono[OF set_mono_prefix[OF append_prefixD[OF B(2)]]] by blast obtain l where "(l,a) \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using a by (meson unlabel_mem_has_label) then obtain b where b: "(l,b) \ set (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (l,b) = (l,a)" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_memberD by blast hence "?P b" using B_P unfolding unlabel_def by fastforce thus "?Q a" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_inv[OF b(2)] by (cases b) auto qed hence "\a \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))). \is_Insert a \ \is_Delete a" by fastforce thus ?thesis using dbupd\<^sub>s\<^sub>s\<^sub>t_no_upd[of "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" \ DB] s_sem by simp qed lemma wellformed_transaction_sem_neg_checks': assumes T_valid: "wellformed_transaction T" and \: "strand_sem_stateful IK DB (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) \" and c: "NegChecks X [] [(t,u)] \ set (unlabel (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" shows "\\. subst_domain \ = set X \ ground (subst_range \) \ (t \ \ \ \, u \ \ \ \) \ DB" (is ?A) and "X = [] \ (t \ \, u \ \) \ DB" (is "?B \ ?B'") proof - show ?A using wellformed_transaction_sem_neg_checks[OF T_valid \ c] unfolding negchecks_model_def by auto moreover have "\ = Var" "t \ \ = t" when "subst_domain \ = set []" for t and \::"('fun, 'atom, 'sets, 'lbl) prot_subst" using that by auto moreover have "subst_domain Var = set []" "range_vars Var = {}" by simp_all ultimately show "?B \ ?B'" unfolding range_vars_alt_def by metis qed lemma wellformed_transaction_sem_iff: fixes T \ defines "A \ unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" and "rm \ \X. rm_vars (set X)" assumes T: "wellformed_transaction T" and I: " interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)" shows "strand_sem_stateful M D A I \ ( (\l ts. (l,receive\ts\) \ set (transaction_receive T) \ (\t \ set ts. M \ t \ \ \ I)) \ (\l ac t s. (l,\ac: t \ s\) \ set (transaction_checks T) \ t \ \ \ I = s \ \ \ I) \ (\l ac t s. (l,\ac: t \ s\) \ set (transaction_checks T) \ (t \ \ \ I, s \ \ \ I) \ D) \ (\l X F G. (l,\X\\\: F \\: G\) \ set (transaction_checks T) \ (\\. subst_domain \ = set X \ ground (subst_range \) \ (\(t,s) \ set F. t \ rm X \ \ \ \ I \ s \ rm X \ \ \ \ I) \ (\(t,s) \ set G. (t \ rm X \ \ \ \ I, s \ rm X \ \ \ \ I) \ D))))" (is "?A \ ?B") proof note 0 = A_def transaction_dual_subst_unlabel_unfold note 1 = wellformed_transaction_sem_receives[OF T, of M D \ I, unfolded A_def[symmetric]] wellformed_transaction_sem_pos_checks[OF T, of M D \ I, unfolded A_def[symmetric]] wellformed_transaction_sem_neg_checks[OF T, of M D \ I, unfolded A_def[symmetric]] note 2 = stateful_strand_step_subst_inI[OF unlabel_in] note 3 = unlabel_subst note 4 = strand_sem_append_stateful[of M D _ _ I] let ?C = "\T. unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" let ?P = "\X \. subst_domain \ = set X \ ground (subst_range \)" let ?sem = "\M D T. strand_sem_stateful M D (?C T) I" let ?negchecks = "\X F G. \\. ?P X \ \ (\(t,s) \ set F. t \ rm X \ \ \ \ I \ s \ rm X \ \ \ \ I) \ (\(t,s) \ set G. (t \ rm X \ \ \ \ I, s \ rm X \ \ \ \ I) \ D)" have "list_all is_Receive (unlabel (transaction_receive T))" "list_all is_Check_or_Assignment (unlabel (transaction_checks T))" "list_all is_Update (unlabel (transaction_updates T))" "list_all is_Send (unlabel (transaction_send T))" using T unfolding wellformed_transaction_def by (blast, blast, blast, blast) hence 5: "list_all is_Send (?C (transaction_receive T))" "list_all is_Check_or_Assignment (?C (transaction_checks T))" "list_all is_Update (?C (transaction_updates T))" "list_all is_Receive (?C (transaction_send T))" by (metis (no_types) subst_sst_list_all(2) unlabel_subst dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_list_all(1), metis (no_types) subst_sst_list_all(11) unlabel_subst dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_list_all(11), metis (no_types) subst_sst_list_all(10) unlabel_subst dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_list_all(10), metis (no_types) subst_sst_list_all(1) unlabel_subst dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_list_all(2)) have "\a \ set (?C (transaction_receive T)). \is_Receive a \ \is_Insert a \ \is_Delete a" "\a \ set (?C (transaction_checks T)). \is_Receive a \ \is_Insert a \ \is_Delete a" using 5(1,2) unfolding list_all_iff by (blast,blast) hence 6: "M \ (ik\<^sub>s\<^sub>s\<^sub>t (?C (transaction_receive T)) \\<^sub>s\<^sub>e\<^sub>t I) = M" "dbupd\<^sub>s\<^sub>s\<^sub>t (?C (transaction_receive T)) I D = D" "M \ (ik\<^sub>s\<^sub>s\<^sub>t (?C (transaction_checks T)) \\<^sub>s\<^sub>e\<^sub>t I) = M" "dbupd\<^sub>s\<^sub>s\<^sub>t (?C (transaction_checks T)) I D = D" by (metis ik\<^sub>s\<^sub>s\<^sub>t_snoc_no_receive_empty sup_bot.right_neutral, metis dbupd\<^sub>s\<^sub>s\<^sub>t_no_upd, metis ik\<^sub>s\<^sub>s\<^sub>t_snoc_no_receive_empty sup_bot.right_neutral, metis dbupd\<^sub>s\<^sub>s\<^sub>t_no_upd) have ?B when A: ?A proof - have "M \ t \ \ \ I" when "(l, receive\ts\) \ set (transaction_receive T)" "t \ set ts" for l ts t using that(2) 1(1)[OF A, of "ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \"] 2(2)[OF that(1)] unfolding 3 by auto moreover have "t \ \ \ I = s \ \ \ I" when "(l, \ac: t \ s\) \ set (transaction_checks T)" for l ac t s using 1(3)[OF A] 2(3)[OF that] unfolding 3 by blast moreover have "(t \ \ \ I, s \ \ \ I) \ D" when "(l, \ac: t \ s\) \ set (transaction_checks T)" for l ac t s using 1(2)[OF A] 2(6)[OF that] unfolding 3 by blast moreover have "?negchecks X F G" when "(l, \X\\\: F \\: G\) \ set (transaction_checks T)" for l X F G using 1(4)[OF A 2(7)[OF that, of \, unfolded 3]] unfolding negchecks_model_def rm_def subst_apply_pairs_def by fastforce ultimately show ?B by blast qed thus "?A \ ?B" by fast have ?A when B: ?B proof - have 7: "\t \ set ts. M \ t \ I" when ts: "send\ts\ \ set (?C (transaction_receive T))" for ts proof - obtain l ss where "(l,receive\ss\) \ set (transaction_receive T)" "ts = ss \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" by (metis ts dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(2) subst_lsst_memD(1) unlabel_mem_has_label) thus ?thesis using B by auto qed have 8: "t \ I = s \ I" when ts: "\ac: t \ s\ \ set (?C (transaction_checks T))" for ac t s proof - obtain l t' s' where "(l,\ac: t' \ s'\) \ set (transaction_checks T)" "t = t' \ \" "s = s' \ \" by (metis ts dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(3) subst_lsst_memD(3) unlabel_mem_has_label) thus ?thesis using B by auto qed have 9: "(t \ I, s \ I) \ D" when ts: "\ac: t \ s\ \ set (?C (transaction_checks T))" for ac t s proof - obtain l t' s' where "(l,\ac: t' \ s'\) \ set (transaction_checks T)" "t = t' \ \" "s = s' \ \" by (metis ts dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(6) subst_lsst_memD(6) unlabel_mem_has_label) thus ?thesis using B by auto qed have 10: "negchecks_model I D X F G" when ts: "\X\\\: F \\: G\ \ set (?C (transaction_checks T))" for X F G proof - obtain l F' G' where *: "(l,\X\\\: F' \\: G'\) \ set (transaction_checks T)" "F = F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \" "G = G' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \" using unlabel_mem_has_label[OF iffD2[OF dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(7) ts]] subst_lsst_memD(7)[of _ X F G "transaction_checks T" \] by fast have "?negchecks X F' G'" using *(1) B by blast moreover have "\(t,s) \ set F. t \ \ \\<^sub>s I \ s \ \ \\<^sub>s I" when "(t,s) \ set F'" "t \ rm X \ \ \ \ I \ s \ rm X \ \ \ \ I" for \ t s using that unfolding rm_def *(2) subst_apply_pairs_def by force moreover have "\(t,s) \ set G. (t,s) \\<^sub>p \ \\<^sub>s I \ D" when "(t,s) \ set G'" "(t \ rm X \ \ \ \ I, s \ rm X \ \ \ \ I) \ D" for \ t s using that unfolding rm_def *(3) subst_apply_pairs_def by force ultimately show ?thesis unfolding negchecks_model_def by auto qed have "?sem M D (transaction_receive T)" using 7 strand_sem_stateful_if_sends_deduct[OF 5(1)] by blast moreover have "?sem M D (transaction_checks T)" using 8 9 10 strand_sem_stateful_if_checks[OF 5(2)] by blast moreover have "?sem M D (transaction_updates T)" for M D using 5(3) strand_sem_stateful_if_no_send_or_check unfolding list_all_iff by blast moreover have "?sem M D (transaction_send T)" for M D using 5(4) strand_sem_stateful_if_no_send_or_check unfolding list_all_iff by blast ultimately show ?thesis using 4[of "?C (transaction_receive T)" "?C (transaction_checks T)@?C (transaction_updates T)@?C (transaction_send T)"] 4[of "?C (transaction_checks T)" "?C (transaction_updates T)@?C (transaction_send T)"] 4[of "?C (transaction_updates T)" "?C (transaction_send T)"] unfolding 0 6 by blast qed thus "?B \ ?A" by fast qed lemma wellformed_transaction_unlabel_sem_iff: fixes T \ defines "A \ unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" and "rm \ \X. rm_vars (set X)" assumes T: "wellformed_transaction T" and I: " interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)" shows "strand_sem_stateful M D A I \ ( (\ts. receive\ts\ \ set (unlabel (transaction_receive T)) \ (\t \ set ts. M \ t \ \ \ I)) \ (\ac t s. \ac: t \ s\ \ set (unlabel (transaction_checks T)) \ t \ \ \ I = s \ \ \ I) \ (\ac t s. \ac: t \ s\ \ set (unlabel (transaction_checks T)) \ (t \ \ \ I, s \ \ \ I) \ D) \ (\X F G. \X\\\: F \\: G\ \ set (unlabel (transaction_checks T)) \ (\\. subst_domain \ = set X \ ground (subst_range \) \ (\(t,s) \ set F. t \ rm X \ \ \ \ I \ s \ rm X \ \ \ \ I) \ (\(t,s) \ set G. (t \ rm X \ \ \ \ I, s \ rm X \ \ \ \ I) \ D))))" using wellformed_transaction_sem_iff[OF T I, of M D \] unlabel_in[of _ _ "transaction_receive T"] unlabel_mem_has_label[of _ "transaction_receive T"] unlabel_in[of _ _ "transaction_checks T"] unlabel_mem_has_label[of _ "transaction_checks T"] unfolding A_def[symmetric] rm_def by meson lemma dual_transaction_ik_is_transaction_send'': fixes \ \::"('a,'b,'c,'d) prot_subst" assumes "wellformed_transaction T" shows "(ik\<^sub>s\<^sub>s\<^sub>t (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a = (trms\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a" (is "?A = ?B") using dual_transaction_ik_is_transaction_send[OF assms] subst_lsst_unlabel[of "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)" \] ik\<^sub>s\<^sub>s\<^sub>t_subst[of "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T))" \] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst[of "transaction_strand T" \] by (auto simp add: abs_apply_terms_def) lemma while_prot_terms_fun_mono: "mono (\M'. M \ \(subterms ` M') \ \((set \ fst \ Ana) ` M'))" unfolding mono_def by fast lemma while_prot_terms_SMP_overapprox: fixes M::"('fun,'atom,'sets,'lbl) prot_terms" assumes N_supset: "M \ \(subterms ` N) \ \((set \ fst \ Ana) ` N) \ N" and Value_vars_only: "\x \ fv\<^sub>s\<^sub>e\<^sub>t N. \\<^sub>v x = TAtom Value" shows "SMP M \ {a \ \ | a \. a \ N \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)}" proof - define f where "f \ \M'. M \ \(subterms ` M') \ \((set \ fst \ Ana) ` M')" define S where "S \ {a \ \ | a \. a \ N \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)}" note 0 = Value_vars_only have "t \ S" when "t \ SMP M" for t using that proof (induction t rule: SMP.induct) case (MP t) hence "t \ N" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t Var" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range Var)" using N_supset by auto hence "t \ Var \ S" unfolding S_def by blast thus ?case by simp next case (Subterm t t') then obtain \ a where a: "a \ \ = t" "a \ N" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (auto simp add: S_def) hence "\x \ fv a. \\. \ (Var x) = TAtom \" using 0 by auto hence *: "\x \ fv a. (\f. \ x = Fun f []) \ (\y. \ x = Var y)" using a(3) TAtom_term_cases[OF wf_trm_subst_rangeD[OF a(4)]] by (metis wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) obtain b where b: "b \ \ = t'" "b \ subterms a" using subterms_subst_subterm[OF *, of t'] Subterm.hyps(2) a(1) by fast hence "b \ N" using N_supset a(2) by blast thus ?case using a b(1) unfolding S_def by blast next case (Substitution t \) then obtain \ a where a: "a \ \ = t" "a \ N" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (auto simp add: S_def) have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \))" by (fact wt_subst_compose[OF a(3) Substitution.hyps(2)], fact wf_trms_subst_compose[OF a(4) Substitution.hyps(3)]) moreover have "t \ \ = a \ \ \\<^sub>s \" using a(1) subst_subst_compose[of a \ \] by simp ultimately show ?case using a(2) unfolding S_def by blast next case (Ana t K T k) then obtain \ a where a: "a \ \ = t" "a \ N" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (auto simp add: S_def) - obtain Ka Ta where a': "Ana a = (Ka,Ta)" by moura + obtain Ka Ta where a': "Ana a = (Ka,Ta)" by force have *: "K = Ka \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" proof (cases a) case (Var x) then obtain g U where gU: "t = Fun g U" using a(1) Ana.hyps(2,3) Ana_var by (cases t) simp_all have "\ (Var x) = TAtom Value" using Var a(2) 0 by auto hence "\ (Fun g U) = TAtom Value" using a(1,3) Var gU wt_subst_trm''[OF a(3), of a] by argo thus ?thesis using gU Fun_Value_type_inv Ana.hyps(2,3) by fastforce next case (Fun g U) thus ?thesis using a(1) a' Ana.hyps(2) Ana_subst'[of g U] by simp qed then obtain ka where ka: "k = ka \ \" "ka \ set Ka" using Ana.hyps(3) by auto have "ka \ set ((fst \ Ana) a)" using ka(2) a' by simp hence "ka \ N" using a(2) N_supset by auto thus ?case using ka a(3,4) unfolding S_def by blast qed thus ?thesis unfolding S_def by blast qed subsection \Admissible Transactions\ definition admissible_transaction_checks where "admissible_transaction_checks T \ \x \ set (unlabel (transaction_checks T)). (is_InSet x \ is_Var (the_elem_term x) \ is_Fun_Set (the_set_term x) \ fst (the_Var (the_elem_term x)) = TAtom Value) \ (is_NegChecks x \ bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p x = [] \ ((the_eqs x = [] \ length (the_ins x) = 1) \ (the_ins x = [] \ length (the_eqs x) = 1))) \ (is_NegChecks x \ the_eqs x = [] \ (let h = hd (the_ins x) in is_Var (fst h) \ is_Fun_Set (snd h) \ fst (the_Var (fst h)) = TAtom Value))" definition admissible_transaction_updates where "admissible_transaction_updates T \ \x \ set (unlabel (transaction_updates T)). is_Update x \ is_Var (the_elem_term x) \ is_Fun_Set (the_set_term x) \ fst (the_Var (the_elem_term x)) = TAtom Value" definition admissible_transaction_terms where "admissible_transaction_terms T \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)) \ (\f \ \(funs_term ` trms_transaction T). \is_Val f \ \is_Abs f \ \is_PubConst f \ f \ Pair) \ (\r \ set (unlabel (transaction_strand T)). (\f \ \(funs_term ` (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p r)). is_Attack f) \ transaction_fresh T = [] \ is_Send r \ length (the_msgs r) = 1 \ is_Fun_Attack (hd (the_msgs r)))" definition admissible_transaction_send_occurs_form where "admissible_transaction_send_occurs_form T \ ( let snds = transaction_send T; frsh = transaction_fresh T in \t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t snds. OccursFact \ funs_term t \ OccursSec \ funs_term t \ (\x \ set frsh. t = occurs (Var x)) )" definition admissible_transaction_occurs_checks where "admissible_transaction_occurs_checks T \ ( let occ_in = \x S. occurs (Var x) \ set (the_msgs (hd (unlabel S))); rcvs = transaction_receive T; snds = transaction_send T; frsh = transaction_fresh T; fvs = fv_transaction T in admissible_transaction_send_occurs_form T \ ((\x \ fvs - set frsh. fst x = TAtom Value) \ ( rcvs \ [] \ is_Receive (hd (unlabel rcvs)) \ (\x \ fvs - set frsh. fst x = TAtom Value \ occ_in x rcvs))) \ (frsh \ [] \ ( snds \ [] \ is_Send (hd (unlabel snds)) \ (\x \ set frsh. occ_in x snds))) )" definition admissible_transaction_no_occurs_msgs where "admissible_transaction_no_occurs_msgs T \ ( let no_occ = \t. is_Fun t \ the_Fun t \ OccursFact; rcvs = transaction_receive T; snds = transaction_send T in list_all (\a. is_Receive (snd a) \ list_all no_occ (the_msgs (snd a))) rcvs \ list_all (\a. is_Send (snd a) \ list_all no_occ (the_msgs (snd a))) snds )" definition admissible_transaction' where "admissible_transaction' T \ ( wellformed_transaction T \ transaction_decl T () = [] \ list_all (\x. fst x = TAtom Value) (transaction_fresh T) \ (\x \ vars_transaction T. is_Var (fst x) \ (the_Var (fst x) = Value)) \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T) = {} \ set (transaction_fresh T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (filter (is_Insert \ snd) (transaction_updates T)) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \ (\x \ fv_transaction T - set (transaction_fresh T). \y \ fv_transaction T - set (transaction_fresh T). x \ y \ \Var x != Var y\ \ set (unlabel (transaction_checks T)) \ \Var y != Var x\ \ set (unlabel (transaction_checks T))) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) - set (transaction_fresh T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) \ (\r \ set (unlabel (transaction_checks T)). is_Equality r \ fv (the_rhs r) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (filter (\s. is_InSet (snd s) \ the_check (snd s) = Assign) (transaction_checks T)) \ list_all (\a. is_Receive (snd a) \ the_msgs (snd a) \ []) (transaction_receive T) \ list_all (\a. is_Send (snd a) \ the_msgs (snd a) \ []) (transaction_send T) \ admissible_transaction_checks T \ admissible_transaction_updates T \ admissible_transaction_terms T \ admissible_transaction_send_occurs_form T )" definition admissible_transaction where "admissible_transaction T \ admissible_transaction' T \ admissible_transaction_no_occurs_msgs T" definition has_initial_value_producing_transaction where "has_initial_value_producing_transaction P \ let f = \s. list_all (\T. list_all (\a. ((is_Delete a \ is_InSet a) \ the_set_term a \ \s\\<^sub>s) \ (is_NegChecks a \ list_all (\(_,t). t \ \s\\<^sub>s) (the_ins a))) (unlabel (transaction_checks T@transaction_updates T))) P in list_ex (\T. length (transaction_fresh T) = 1 \ transaction_receive T = [] \ transaction_checks T = [] \ length (transaction_send T) = 1 \ (let x = hd (transaction_fresh T); a = hd (transaction_send T); u = transaction_updates T in is_Send (snd a) \ Var x \ set (the_msgs (snd a)) \ fv\<^sub>s\<^sub>e\<^sub>t (set (the_msgs (snd a))) = {x} \ (u \ [] \ ( let b = hd u; c = snd b in tl u = [] \ is_Insert c \ the_elem_term c = Var x \ is_Fun_Set (the_set_term c) \ f (the_Set (the_Fun (the_set_term c)))))) ) P" lemma admissible_transaction_is_wellformed_transaction: assumes "admissible_transaction' T" shows "wellformed_transaction T" and "admissible_transaction_checks T" and "admissible_transaction_updates T" and "admissible_transaction_terms T" and "admissible_transaction_send_occurs_form T" using assms unfolding admissible_transaction'_def by blast+ lemma admissible_transaction_no_occurs_msgsE: assumes T: "admissible_transaction' T" "admissible_transaction_no_occurs_msgs T" shows "\ts. send\ts\ \ set (unlabel (transaction_strand T)) \ receive\ts\ \ set (unlabel (transaction_strand T)) \ (\t s. t \ set ts \ t \ occurs s)" proof - note 1 = admissible_transaction_is_wellformed_transaction(1)[OF T(1)] have 2: "send\ts\ \ set (unlabel (transaction_send T))" when "send\ts\ \ set (unlabel (transaction_strand T))" for ts using wellformed_transaction_strand_unlabel_memberD(8)[OF 1 that] by fast have 3: "receive\ts\ \ set (unlabel (transaction_receive T))" when "receive\ts\ \ set (unlabel (transaction_strand T))" for ts using wellformed_transaction_strand_unlabel_memberD(1)[OF 1 that] by fast show ?thesis using T(2) 2 3 wellformed_transaction_unlabel_cases(1,4)[OF 1] unfolding admissible_transaction_no_occurs_msgs_def Let_def list_all_iff by (metis sndI stateful_strand_step.discI(1,2) stateful_strand_step.sel(1,2) term.discI(2) term.sel(2) unlabel_mem_has_label) qed lemma admissible_transactionE: assumes T: "admissible_transaction' T" shows "transaction_decl T () = []" (is ?A) and "\x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" (is ?B) and "\x \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T). \\<^sub>v x = TAtom Value" (is ?C) and "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T) = {}" (is ?D1) and "fv_transaction T \ bvars_transaction T = {}" (is ?D2) and "set (transaction_fresh T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (filter (is_Insert \ snd) (transaction_updates T)) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" (is ?E) and "set (transaction_fresh T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" (is ?F) and "\x \ fv_transaction T - set (transaction_fresh T). \y \ fv_transaction T - set (transaction_fresh T). x \ y \ \Var x != Var y\ \ set (unlabel (transaction_checks T)) \ \Var y != Var x\ \ set (unlabel (transaction_checks T))" (is ?G) and "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T). x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ (\t s. select\t,s\ \ set (unlabel (transaction_checks T)) \ x \ fv t \ fv s)" (is ?H) and "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) - set (transaction_fresh T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" (is ?I) and "\x \ set (unlabel (transaction_checks T)). is_Equality x \ fv (the_rhs x) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" (is ?J) (* TODO: why do we need this requirement? *) and "set (transaction_fresh T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) = {}" (is ?K1) and "set (transaction_fresh T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) = {}" (is ?K2) and "list_all (\x. fst x = Var Value) (transaction_fresh T)" (is ?K3) and "\x \ vars_transaction T. \TAtom AttackType \ \\<^sub>v x" (is ?K4) and "\l ts. (l,receive\ts\) \ set (transaction_receive T) \ ts \ []" (is ?L1) and "\l ts. (l,send\ts\) \ set (transaction_send T) \ ts \ []" (is ?L2) proof - show ?A ?D1 ?D2 ?G ?I ?J ?K3 using T unfolding admissible_transaction'_def by (blast, blast, blast, blast, blast, blast, blast) have "list_all (\a. is_Receive (snd a) \ the_msgs (snd a) \ []) (transaction_receive T)" "list_all (\a. is_Send (snd a) \ the_msgs (snd a) \ []) (transaction_send T)" using T unfolding admissible_transaction'_def by auto thus ?L1 ?L2 unfolding list_all_iff by (force,force) have "list_all (\x. fst x = Var Value) (transaction_fresh T)" "\x\vars_transaction T. is_Var (fst x) \ the_Var (fst x) = Value" using T unfolding admissible_transaction'_def by (blast, blast) thus ?B ?C ?K4 using \\<^sub>v_TAtom''(2) unfolding list_all_iff by (blast, force, force) show ?E using T unfolding admissible_transaction'_def by argo thus ?F unfolding unlabel_def by auto show ?K1 ?K2 using T unfolding admissible_transaction'_def wellformed_transaction_def by (argo, argo) let ?selects = "filter (\s. is_InSet (snd s) \ the_check (snd s) = Assign) (transaction_checks T)" show ?H proof fix x assume "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" hence "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?selects" using T unfolding admissible_transaction'_def by blast thus "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ (\t s. select\t,s\ \ set (unlabel (transaction_checks T)) \ x \ fv t \ fv s)" proof assume "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?selects" then obtain r where r: "x \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p r" "r \ set (unlabel (transaction_checks T))" "is_InSet r" "the_check r = Assign" unfolding unlabel_def by force thus ?thesis by (cases r) auto qed simp qed qed lemma admissible_transactionE': assumes T: "admissible_transaction T" shows "admissible_transaction' T" (is ?A) and "admissible_transaction_no_occurs_msgs T" (is ?B) and "\ts. send\ts\ \ set (unlabel (transaction_strand T)) \ receive\ts\ \ set (unlabel (transaction_strand T)) \ (\t s. t \ set ts \ t \ occurs s)" (is ?C) proof - show 0: ?A ?B using T unfolding admissible_transaction_def by (blast, blast) show ?C using admissible_transaction_no_occurs_msgsE[OF 0] by blast qed lemma transaction_inserts_are_Value_vars: assumes T_valid: "wellformed_transaction T" and "admissible_transaction_updates T" and "insert\t,s\ \ set (unlabel (transaction_strand T))" shows "\n. t = Var (TAtom Value, n)" and "\u. s = Fun (Set u) []" proof - let ?x = "insert\t,s\" have "?x \ set (unlabel (transaction_updates T))" using assms(3) wellformed_transaction_unlabel_cases[OF T_valid, of ?x] by (auto simp add: transaction_strand_def unlabel_def) hence *: "is_Var (the_elem_term ?x)" "fst (the_Var (the_elem_term ?x)) = TAtom Value" "is_Fun (the_set_term ?x)" "args (the_set_term ?x) = []" "is_Set (the_Fun (the_set_term ?x))" using assms(2) unfolding admissible_transaction_updates_def is_Fun_Set_def by fastforce+ show "\n. t = Var (TAtom Value, n)" using *(1,2) by (cases t) auto show "\u. s = Fun (Set u) []" using *(3,4,5) unfolding is_Set_def by (cases s) auto qed lemma transaction_deletes_are_Value_vars: assumes T_valid: "wellformed_transaction T" and "admissible_transaction_updates T" and "delete\t,s\ \ set (unlabel (transaction_strand T))" shows "\n. t = Var (TAtom Value, n)" and "\u. s = Fun (Set u) []" proof - let ?x = "delete\t,s\" have "?x \ set (unlabel (transaction_updates T))" using assms(3) wellformed_transaction_unlabel_cases[OF T_valid, of ?x] by (auto simp add: transaction_strand_def unlabel_def) hence *: "is_Var (the_elem_term ?x)" "fst (the_Var (the_elem_term ?x)) = TAtom Value" "is_Fun (the_set_term ?x)" "args (the_set_term ?x) = []" "is_Set (the_Fun (the_set_term ?x))" using assms(2) unfolding admissible_transaction_updates_def is_Fun_Set_def by fastforce+ show "\n. t = Var (TAtom Value, n)" using *(1,2) by (cases t) auto show "\u. s = Fun (Set u) []" using *(3,4,5) unfolding is_Set_def by (cases s) auto qed lemma transaction_selects_are_Value_vars: assumes T_valid: "wellformed_transaction T" and "admissible_transaction_checks T" and "select\t,s\ \ set (unlabel (transaction_strand T))" shows "\n. t = Var (TAtom Value, n) \ (TAtom Value, n) \ set (transaction_fresh T)" (is ?A) and "\u. s = Fun (Set u) []" (is ?B) proof - let ?x = "select\t,s\" have *: "?x \ set (unlabel (transaction_checks T))" using assms(3) wellformed_transaction_unlabel_cases[OF T_valid, of ?x] by (auto simp add: transaction_strand_def unlabel_def) have **: "is_Var (the_elem_term ?x)" "fst (the_Var (the_elem_term ?x)) = TAtom Value" "is_Fun (the_set_term ?x)" "args (the_set_term ?x) = []" "is_Set (the_Fun (the_set_term ?x))" using * assms(2) unfolding admissible_transaction_checks_def is_Fun_Set_def by fastforce+ have "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p ?x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" using * by force hence ***: "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p ?x \ set (transaction_fresh T) = {}" using T_valid unfolding wellformed_transaction_def by fast show ?A using **(1,2) *** by (cases t) auto show ?B using **(3,4,5) unfolding is_Set_def by (cases s) auto qed lemma transaction_inset_checks_are_Value_vars: assumes T_valid: "admissible_transaction' T" and t: "\t in s\ \ set (unlabel (transaction_strand T))" shows "\n. t = Var (TAtom Value, n) \ (TAtom Value, n) \ set (transaction_fresh T)" (is ?A) and "\u. s = Fun (Set u) []" (is ?B) proof - let ?x = "\t in s\" note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_valid] note T_adm_checks = admissible_transaction_is_wellformed_transaction(2)[OF T_valid] have *: "?x \ set (unlabel (transaction_checks T))" using t wellformed_transaction_unlabel_cases[OF T_wf, of ?x] unfolding transaction_strand_def unlabel_def by fastforce have **: "is_Var (the_elem_term ?x)" "fst (the_Var (the_elem_term ?x)) = TAtom Value" "is_Fun (the_set_term ?x)" "args (the_set_term ?x) = []" "is_Set (the_Fun (the_set_term ?x))" using * T_adm_checks unfolding admissible_transaction_checks_def is_Fun_Set_def by fastforce+ have "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p ?x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" using * by force hence ***: "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p ?x \ set (transaction_fresh T) = {}" using T_wf unfolding wellformed_transaction_def by fast show ?A using **(1,2) *** by (cases t) auto show ?B using **(3,4,5) unfolding is_Set_def by (cases s) auto qed lemma transaction_notinset_checks_are_Value_vars: assumes T_adm: "admissible_transaction' T" and FG: "\X\\\: F \\: G\ \ set (unlabel (transaction_strand T))" and t: "(t,s) \ set G" shows "\n. t = Var (TAtom Value, n) \ (TAtom Value, n) \ set (transaction_fresh T)" (is ?A) and "\u. s = Fun (Set u) []" (is ?B) and "F = []" (is ?C) and "G = [(t,s)]" (is ?D) proof - let ?x = "\X\\\: F \\: G\" note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] note T_adm_checks = admissible_transaction_is_wellformed_transaction(2)[OF T_adm] have 0: "?x \ set (unlabel (transaction_checks T))" using FG wellformed_transaction_unlabel_cases[OF T_wf, of ?x] by (auto simp add: transaction_strand_def unlabel_def) hence 1: "F = [] \ length G = 1" using T_adm_checks t unfolding admissible_transaction_checks_def by fastforce hence "hd G = (t,s)" using t by (cases "the_ins ?x") auto hence **: "is_Var t" "fst (the_Var t) = TAtom Value" "is_Fun s" "args s = []" "is_Set (the_Fun s)" using 1 Set.bspec[OF T_adm_checks[unfolded admissible_transaction_checks_def] 0] unfolding is_Fun_Set_def by auto show ?C using 1 by blast show ?D using 1 t by force have "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p ?x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p ?x) \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" using 0 by force+ moreover have "set (transaction_fresh T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) = {}" "set (transaction_fresh T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) = {}" using T_wf unfolding wellformed_transaction_def by fast+ ultimately have "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p ?x \ set (transaction_fresh T) = {}" "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p ?x) \ set (transaction_fresh T) = {}" using admissible_transactionE(7)[OF T_adm] wellformed_transaction_wf\<^sub>s\<^sub>s\<^sub>t(2)[OF T_wf] fv_transaction_unfold[of T] bvars_transaction_unfold[of T] by (blast, blast) hence ***: "fv t \ set (transaction_fresh T) = {}" using t by auto show ?A using **(1,2) *** by (cases t) auto show ?B using **(3,4,5) unfolding is_Set_def by (cases s) auto qed lemma transaction_noteqs_checks_case: assumes T_adm: "admissible_transaction' T" and FG: "\X\\\: F \\: G\ \ set (unlabel (transaction_strand T))" and G: "G = []" shows "\t s. F = [(t,s)]" (is ?A) proof - let ?x = "\X\\\: F \\: G\" note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] note T_adm_checks = admissible_transaction_is_wellformed_transaction(2)[OF T_adm] have "?x \ set (unlabel (transaction_checks T))" using FG wellformed_transaction_unlabel_cases[OF T_wf, of ?x] by (auto simp add: transaction_strand_def unlabel_def) hence "length F = 1" using T_adm_checks unfolding admissible_transaction_checks_def G by fastforce thus ?thesis by fast qed lemma admissible_transaction_fresh_vars_notin: assumes T: "admissible_transaction' T" and x: "x \ set (transaction_fresh T)" shows "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" (is ?A) and "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" (is ?B) and "x \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" (is ?C) and "x \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" (is ?D) and "x \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" (is ?E) and "x \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" (is ?F) proof - note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T] have 0: "set (transaction_fresh T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "set (transaction_fresh T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) = {}" "set (transaction_fresh T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) = {}" "fv_transaction T \ bvars_transaction T = {}" using admissible_transactionE[OF T] by argo+ have 1: "set (transaction_fresh T) \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) = {}" using 0(1,4) fv_transaction_unfold[of T] bvars_transaction_unfold[of T] by blast have 2: "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) = {}" using bvars_wellformed_transaction_unfold[OF T_wf] vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_receive T)"] by blast+ show ?A ?B ?C ?E ?F using 0 1 2 x by (fast, fast, fast, fast, fast) show ?D using 0(3) 1 x vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_checks T)"] by fast qed lemma admissible_transaction_fv_in_receives_or_selects: assumes T: "admissible_transaction' T" and x: "x \ fv_transaction T" "x \ set (transaction_fresh T)" shows "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ (x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) \ (\t s. select\t,s\ \ set (unlabel (transaction_checks T)) \ x \ fv t \ fv s))" proof - have "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" using x(1) fv\<^sub>s\<^sub>s\<^sub>t_append unlabel_append by (metis transaction_strand_def append_assoc) thus ?thesis using x(2) admissible_transactionE(9,10)[OF T] by blast qed lemma admissible_transaction_fv_in_receives_or_selects': assumes T: "admissible_transaction' T" and x: "x \ fv_transaction T" "x \ set (transaction_fresh T)" shows "(\ts. receive\ts\ \ set (unlabel (transaction_receive T)) \ x \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)) \ (\s. select\Var x, s\ \ set (unlabel (transaction_checks T)))" proof (cases "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)") case True thus ?thesis using wellformed_transaction_unlabel_cases(1)[ OF admissible_transaction_is_wellformed_transaction(1)[OF T]] by force next case False then obtain t s where t: "select\t,s\ \ set (unlabel (transaction_checks T))" "x \ fv t \ fv s" using admissible_transaction_fv_in_receives_or_selects[OF T x] by blast have t': "select\t,s\ \ set (unlabel (transaction_strand T))" using t(1) unfolding transaction_strand_def by simp show ?thesis using t transaction_selects_are_Value_vars[ OF admissible_transaction_is_wellformed_transaction(1,2)[OF T] t'] by force qed lemma admissible_transaction_fv_in_receives_or_selects_subst: assumes T: "admissible_transaction' T" and x: "x \ fv_transaction T" "x \ set (transaction_fresh T)" shows "(\ts. receive\ts\ \ set (unlabel (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ \ x \\<^sub>s\<^sub>e\<^sub>t set ts) \ (\s. select\\ x, s\ \ set (unlabel (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" proof - note 0 = admissible_transaction_fv_in_receives_or_selects'[OF T x] have 1: "\ x \\<^sub>s\<^sub>e\<^sub>t set (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" when ts: "x \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" for ts using that subst_mono_fv[of x _ \] by auto have 2: "receive\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\ \ set (A \\<^sub>s\<^sub>s\<^sub>t \)" when "receive\ts\ \ set A" for ts A using that by fast have 3: "select\t \ \,s \ \\ \ set (A \\<^sub>s\<^sub>s\<^sub>t \)" when "select\t,s\ \ set A" for t s A using that by fast show ?thesis using 0 1 2[of _ "unlabel (transaction_receive T)"] 3[of _ _ "unlabel (transaction_checks T)"] unfolding unlabel_subst by (metis eval_term.simps(1)) qed lemma admissible_transaction_fv_in_receives_or_selects_dual_subst: defines "f \ \S. unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S)" assumes T: "admissible_transaction' T" and x: "x \ fv_transaction T" "x \ set (transaction_fresh T)" shows "(\ts. send\ts\ \ set (f (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ \ x \\<^sub>s\<^sub>e\<^sub>t set ts) \ (\s. select\\ x, s\ \ set (f (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" using admissible_transaction_fv_in_receives_or_selects_subst[OF T x, of \] by (metis (no_types, lifting) f_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(2) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(6)) lemma admissible_transaction_decl_subst_empty': assumes T: "transaction_decl T () = []" and \: "transaction_decl_subst \ T" shows "\ = Var" proof - have "subst_domain \ = {}" using \ T unfolding transaction_decl_subst_def by auto thus ?thesis by auto qed lemma admissible_transaction_decl_subst_empty: assumes T: "admissible_transaction' T" and \: "transaction_decl_subst \ T" shows "\ = Var" by (rule admissible_transaction_decl_subst_empty'[OF admissible_transactionE(1)[OF T] \]) lemma admissible_transaction_no_bvars: assumes "admissible_transaction' T" shows "fv_transaction T = vars_transaction T" and "bvars_transaction T = {}" using admissible_transactionE(4)[OF assms] bvars_wellformed_transaction_unfold vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t by (fast, fast) lemma admissible_transactions_fv_bvars_disj: assumes "\T \ set P. admissible_transaction' T" shows "(\T \ set P. fv_transaction T) \ (\T \ set P. bvars_transaction T) = {}" using assms admissible_transaction_no_bvars(2) by fast lemma admissible_transaction_occurs_fv_types: assumes "admissible_transaction' T" and "x \ vars_transaction T" shows "\a. \ (Var x) = TAtom a \ \ (Var x) \ TAtom OccursSecType" proof - have "is_Var (fst x)" "the_Var (fst x) = Value" using assms unfolding admissible_transaction'_def by blast+ thus ?thesis using \\<^sub>v_TAtom''(2)[of x] by force qed lemma admissible_transaction_Value_vars_are_fv: assumes "admissible_transaction' T" and "x \ vars_transaction T" and "\\<^sub>v x = TAtom Value" shows "x \ fv_transaction T" using assms(2,3) \\<^sub>v_TAtom''(2)[of x] vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_strand T)"] admissible_transactionE(4)[OF assms(1)] by fast lemma transaction_receive_deduct: assumes T_wf: "wellformed_transaction T" and \: "constraint_model \ (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" and \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T M" and \: "transaction_renaming_subst \ P X" and t: "receive\ts\ \ set (unlabel (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" shows "\t \ set ts. ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" proof - define \ where "\ \ \ \\<^sub>s \ \\<^sub>s \" have t': "send\ts\ \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" using t dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(2) unfolding \_def by blast then obtain T1 T2 where T: "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) = T1@send\ts\#T2" using t' by (meson split_list) have "constr_sem_stateful \ (unlabel A@unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" using \ unlabel_append[of A] unfolding constraint_model_def \_def by simp hence "constr_sem_stateful \ (unlabel A@T1@[send\ts\])" using strand_sem_append_stateful[of "{}" "{}" "unlabel A@T1@[send\ts\]" _ \] transaction_dual_subst_unlabel_unfold[of T \] T by (metis append.assoc append_Cons append_Nil) hence "\t \ set ts. ik\<^sub>s\<^sub>s\<^sub>t (unlabel A@T1) \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" using strand_sem_append_stateful[of "{}" "{}" "unlabel A@T1" "[send\ts\]" \] T by force moreover have "\is_Receive x" when x: "x \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" for x proof - have *: "is_Receive a" when "a \ set (unlabel (transaction_receive T))" for a using T_wf Ball_set[of "unlabel (transaction_receive T)" is_Receive] that unfolding wellformed_transaction_def by blast obtain l where l: "(l,x) \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using x unfolding unlabel_def by fastforce then obtain ly where ly: "ly \ set (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "(l,x) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ly" unfolding dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by auto obtain j y where j: "ly = (j,y)" by (metis surj_pair) hence "j = l" using ly(2) by (cases y) auto hence y: "(l,y) \ set (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "(l,x) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (l,y)" by (metis j ly(1), metis j ly(2)) obtain z where z: "z \ set (unlabel (transaction_receive T))" "(l,z) \ set (transaction_receive T)" "(l,y) = (l,z) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" using y(1) unfolding subst_apply_labeled_stateful_strand_def unlabel_def by force have "is_Receive y" using *[OF z(1)] z(3) by (cases z) auto thus "\is_Receive x" using l y by (cases y) auto qed hence "\is_Receive x" when "x \ set T1" for x using T that by simp hence "ik\<^sub>s\<^sub>s\<^sub>t T1 = {}" unfolding ik\<^sub>s\<^sub>s\<^sub>t_def is_Receive_def by fast hence "ik\<^sub>s\<^sub>s\<^sub>t (unlabel A@T1) = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" using ik\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel A" T1] by simp ultimately show ?thesis by (simp add: \_def) qed lemma transaction_checks_db: assumes T: "admissible_transaction' T" and \: "constraint_model \ (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" and \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T M" and \: "transaction_renaming_subst \ P X" shows "\Var (TAtom Value, n) in Fun (Set s) []\ \ set (unlabel (transaction_checks T)) \ (\ (TAtom Value, n) \ \, Fun (Set s) []) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \)" (is "?A \ ?B") and "\Var (TAtom Value, n) not in Fun (Set s) []\ \ set (unlabel (transaction_checks T)) \ (\ (TAtom Value, n) \ \, Fun (Set s) []) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \)" (is "?C \ ?D") proof - let ?x = "\n. (TAtom Value, n)" let ?s = "Fun (Set s) []" let ?T = "transaction_receive T@transaction_checks T" let ?T' = "?T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \" let ?S = "\S. transaction_receive T@S" let ?S' = "\S. ?S S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \" note \_empty = admissible_transaction_decl_subst_empty[OF T \] note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T] have "constr_sem_stateful \ (unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)))" using \ unfolding constraint_model_def by simp moreover have "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S (T1@[c]) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T2@transaction_updates T@transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" when "transaction_checks T = T1@c#T2" for T1 T2 c \ using that dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_append subst_lsst_append unfolding transaction_strand_def by (metis append.assoc append_Cons append_Nil) ultimately have T'_model: "constr_sem_stateful \ (unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S' (T1@[(l,c)]))))" when "transaction_checks T = T1@(l,c)#T2" for T1 T2 l c using strand_sem_append_stateful[of _ _ _ _ \] by (simp add: that transaction_strand_def) show "?A \ ?B" proof - assume a: ?A hence *: "\Var (?x n) in ?s\ \ set (unlabel ?T)" unfolding transaction_strand_def unlabel_def by simp then obtain l T1 T2 where T1: "transaction_checks T = T1@(l,\Var (?x n) in ?s\)#T2" by (metis a split_list unlabel_mem_has_label) have "?x n \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" using a by force hence "?x n \ set (transaction_fresh T)" using a admissible_transaction_fresh_vars_notin[OF T] by fast hence "unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S' (T1@[(l,\Var (?x n) in ?s\)]))) = unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S' T1))@[\\ (?x n) in ?s\]" using T a \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_append subst_lsst_append unlabel_append \_empty by (fastforce simp add: transaction_fresh_subst_def unlabel_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def subst_apply_labeled_stateful_strand_def subst_compose) moreover have "db\<^sub>s\<^sub>s\<^sub>t (unlabel A) = db\<^sub>s\<^sub>s\<^sub>t (unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S' T1)))" by (simp add: T1 db\<^sub>s\<^sub>s\<^sub>t_transaction_prefix_eq[OF T_wf] del: unlabel_append) ultimately have "\M. strand_sem_stateful M (set (db\<^sub>s\<^sub>s\<^sub>t (unlabel A) \)) [\\ (?x n) in ?s\] \" using T'_model[OF T1] db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t[of _ \] strand_sem_append_stateful[of _ _ _ _ \] by (simp add: db\<^sub>s\<^sub>s\<^sub>t_def del: unlabel_append) thus ?B by simp qed show "?C \ ?D" proof - assume a: ?C hence *: "\Var (?x n) not in ?s\ \ set (unlabel ?T)" unfolding transaction_strand_def unlabel_def by simp then obtain l T1 T2 where T1: "transaction_checks T = T1@(l,\Var (?x n) not in ?s\)#T2" by (metis a split_list unlabel_mem_has_label) have "?x n \ vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p \Var (?x n) not in ?s\" using vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_cases(9)[of "[]" "Var (?x n)" ?s] by auto hence "?x n \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" using a unfolding vars\<^sub>s\<^sub>s\<^sub>t_def by force hence "?x n \ set (transaction_fresh T)" using a admissible_transaction_fresh_vars_notin[OF T] by fast hence "unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S' (T1@[(l,\Var (?x n) not in ?s\)]))) = unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S' T1))@[\\ (?x n) not in ?s\]" using T a \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_append subst_lsst_append unlabel_append \_empty by (fastforce simp add: transaction_fresh_subst_def unlabel_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def subst_apply_labeled_stateful_strand_def subst_compose) moreover have "db\<^sub>s\<^sub>s\<^sub>t (unlabel A) = db\<^sub>s\<^sub>s\<^sub>t (unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S' T1)))" by (simp add: T1 db\<^sub>s\<^sub>s\<^sub>t_transaction_prefix_eq[OF T_wf] del: unlabel_append) ultimately have "\M. strand_sem_stateful M (set (db\<^sub>s\<^sub>s\<^sub>t (unlabel A) \)) [\\ (?x n) not in ?s\] \" using T'_model[OF T1] db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t[of _ \] strand_sem_append_stateful[of _ _ _ _ \] by (simp add: db\<^sub>s\<^sub>s\<^sub>t_def del: unlabel_append) thus ?D using stateful_strand_sem_NegChecks_no_bvars(1)[of _ _ _ ?s \] by simp qed qed lemma transaction_selects_db: assumes T: "admissible_transaction' T" and \: "constraint_model \ (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" and \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T M" and \: "transaction_renaming_subst \ P X" shows "select\Var (TAtom Value, n), Fun (Set s) []\ \ set (unlabel (transaction_checks T)) \ (\ (TAtom Value, n) \ \, Fun (Set s) []) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \)" (is "?A \ ?B") proof - let ?x = "\n. (TAtom Value, n)" let ?s = "Fun (Set s) []" let ?T = "transaction_receive T@transaction_checks T" let ?T' = "?T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \" let ?S = "\S. transaction_receive T@S" let ?S' = "\S. ?S S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \" note \_empty = admissible_transaction_decl_subst_empty[OF T \] note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T] have "constr_sem_stateful \ (unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)))" using \ unfolding constraint_model_def by simp moreover have "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S (T1@[c]) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T2@transaction_updates T@transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" when "transaction_checks T = T1@c#T2" for T1 T2 c \ using that dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_append subst_lsst_append unfolding transaction_strand_def by (metis append.assoc append_Cons append_Nil) ultimately have T'_model: "constr_sem_stateful \ (unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S' (T1@[(l,c)]))))" when "transaction_checks T = T1@(l,c)#T2" for T1 T2 l c using strand_sem_append_stateful[of _ _ _ _ \] by (simp add: that transaction_strand_def) show "?A \ ?B" proof - assume a: ?A hence *: "select\Var (?x n), ?s\ \ set (unlabel ?T)" unfolding transaction_strand_def unlabel_def by simp then obtain l T1 T2 where T1: "transaction_checks T = T1@(l,select\Var (?x n), ?s\)#T2" by (metis a split_list unlabel_mem_has_label) have "?x n \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" using a by force hence "?x n \ set (transaction_fresh T)" using a admissible_transaction_fresh_vars_notin[OF T] by fast hence "unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S' (T1@[(l,select\Var (?x n), ?s\)]))) = unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S' T1))@[select\\ (?x n), ?s\]" using T a \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_append subst_lsst_append unlabel_append \_empty by (fastforce simp add: transaction_fresh_subst_def unlabel_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def subst_apply_labeled_stateful_strand_def subst_compose) moreover have "db\<^sub>s\<^sub>s\<^sub>t (unlabel A) = db\<^sub>s\<^sub>s\<^sub>t (unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S' T1)))" by (simp add: T1 db\<^sub>s\<^sub>s\<^sub>t_transaction_prefix_eq[OF T_wf] del: unlabel_append) ultimately have "\M. strand_sem_stateful M (set (db\<^sub>s\<^sub>s\<^sub>t (unlabel A) \)) [\\ (?x n) in ?s\] \" using T'_model[OF T1] db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t[of _ \] strand_sem_append_stateful[of _ _ _ _ \] by (simp add: db\<^sub>s\<^sub>s\<^sub>t_def del: unlabel_append) thus ?B by simp qed qed lemma admissible_transaction_terms_no_Value_consts: assumes "admissible_transaction_terms T" and "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T))" shows "\a T. t = Fun (Val a) T" (is ?A) and "\a T. t = Fun (Abs a) T" (is ?B) and "\a T. t = Fun (PubConst Value a) T" (is ?C) proof - have "\is_Val f" "\is_Abs f" "\is_PubConstValue f" when "f \ \(funs_term ` (trms_transaction T))" for f using that assms(1)[unfolded admissible_transaction_terms_def] unfolding is_PubConstValue_def by (blast,blast,blast) moreover have "f \ \(funs_term ` (trms_transaction T))" when "f \ funs_term t" for f using that assms(2) funs_term_subterms_eq(2)[of "trms_transaction T"] by blast+ ultimately have *: "\is_Val f" "\is_Abs f" "\is_PubConstValue f" when "f \ funs_term t" for f using that by presburger+ show ?A using *(1) by force show ?B using *(2) by force show ?C using *(3) unfolding is_PubConstValue_def by force qed lemma admissible_transactions_no_Value_consts: assumes "admissible_transaction' T" and "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T))" shows "\a T. t = Fun (Val a) T" (is ?A) and "\a T. t = Fun (Abs a) T" (is ?B) and "\a T. t = Fun (PubConst Value a) T" (is ?C) using admissible_transaction_terms_no_Value_consts[OF admissible_transaction_is_wellformed_transaction(4)[OF assms(1)] assms(2)] by auto lemma admissible_transactions_no_Value_consts': assumes "admissible_transaction' T" and "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)" shows "\a T. Fun (Val a) T \ subterms t" and "\a T. Fun (Abs a) T \ subterms t" using admissible_transactions_no_Value_consts[OF assms(1)] assms(2) by fast+ lemma admissible_transactions_no_Value_consts'': assumes "admissible_transaction' T" shows "\n. PubConst Value n \ \(funs_term ` trms_transaction T)" and "\n. Abs n \ \(funs_term ` trms_transaction T)" using assms unfolding admissible_transaction'_def admissible_transaction_terms_def by (meson prot_fun.discI(6), meson prot_fun.discI(4)) lemma admissible_transactions_no_PubConsts: assumes "admissible_transaction' T" and "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T))" shows "\a n T. t = Fun (PubConst a n) T" proof - have "\is_PubConst f" when "f \ \(funs_term ` (trms_transaction T))" for f using that conjunct1[OF conjunct2[OF admissible_transaction_is_wellformed_transaction(4)[ OF assms(1), unfolded admissible_transaction_terms_def]]] by blast moreover have "f \ \(funs_term ` (trms_transaction T))" when "f \ funs_term t" for f using that assms(2) funs_term_subterms_eq(2)[of "trms_transaction T"] by blast+ ultimately have *: "\is_PubConst f" when "f \ funs_term t" for f using that by presburger+ show ?thesis using * by force qed lemma admissible_transactions_no_PubConsts': assumes "admissible_transaction' T" and "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)" shows "\a n T. Fun (PubConst a n) T \ subterms t" using admissible_transactions_no_PubConsts[OF assms(1)] assms(2) by fast+ lemma admissible_transaction_strand_step_cases: assumes T_adm: "admissible_transaction' T" shows "r \ set (unlabel (transaction_receive T)) \ \t. r = receive\t\" (is "?A \ ?A'") and "r \ set (unlabel (transaction_checks T)) \ (\x s. (r = \Var x in Fun (Set s) []\ \ r = select\Var x, Fun (Set s) []\ \ r = \Var x not in Fun (Set s) []\) \ fst x = TAtom Value \ x \ fv_transaction T - set (transaction_fresh T)) \ (\s t. r = \s == t\ \ r = \s := t\ \ r = \s != t\)" (is "?B \ ?B'") and "r \ set (unlabel (transaction_updates T)) \ \x s. (r = insert\Var x, Fun (Set s) []\ \ r = delete\Var x, Fun (Set s) []\) \ fst x = TAtom Value" (is "?C \ ?C'") and "r \ set (unlabel (transaction_send T)) \ \t. r = send\t\" (is "?D \ ?D'") proof - note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] show "?A \ ?A'" using T_wf Ball_set[of "unlabel (transaction_receive T)" is_Receive] unfolding wellformed_transaction_def is_Receive_def by blast show "?D \ ?D'" using T_wf Ball_set[of "unlabel (transaction_send T)" is_Send] unfolding wellformed_transaction_def is_Send_def by blast show "?B \ ?B'" proof - assume r: ?B note adm_checks = admissible_transaction_is_wellformed_transaction(1,2)[OF T_adm] have fv_r1: "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p r \ fv_transaction T" using r fv_transaction_unfold[of T] by auto have fv_r2: "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p r \ set (transaction_fresh T) = {}" using r T_wf unfolding wellformed_transaction_def by fastforce have "list_all is_Check_or_Assignment (unlabel (transaction_checks T))" using adm_checks(1) unfolding wellformed_transaction_def by blast hence "is_InSet r \ is_Equality r \ is_NegChecks r" using r unfolding list_all_iff by blast thus ?B' proof (elim disjE conjE) assume *: "is_InSet r" hence **: "is_Var (the_elem_term r)" "is_Fun (the_set_term r)" "is_Set (the_Fun (the_set_term r))" "args (the_set_term r) = []" "fst (the_Var (the_elem_term r)) = TAtom Value" using r adm_checks unfolding admissible_transaction_checks_def is_Fun_Set_def by fast+ obtain ac rt rs where r': "r = \ac: rt \ rs\" using * by (cases r) auto obtain x where x: "rt = Var x" "fst x = TAtom Value" using **(1,5) r' by auto obtain f S where fS: "rs = Fun f S" using **(2) r' by auto obtain s where s: "f = Set s" using **(3) fS r' by (cases f) auto hence S: "S = []" using **(4) fS r' by auto show ?B' using r' x fS s S fv_r1 fv_r2 by (cases ac) simp_all next assume *: "is_NegChecks r" hence **: "bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p r = []" "(the_eqs r = [] \ length (the_ins r) = 1) \ (the_ins r = [] \ length (the_eqs r) = 1)" using r adm_checks unfolding admissible_transaction_checks_def by fast+ show ?B' using **(2) proof (elim disjE conjE) assume ***: "the_eqs r = []" "length (the_ins r) = 1" then obtain t s where ts: "the_ins r = [(t,s)]" by (cases "the_ins r") auto hence "hd (the_ins r) = (t,s)" by simp hence ****: "is_Var (fst (t,s))" "is_Fun (snd (t,s))" "is_Set (the_Fun (snd (t,s)))" "args (snd (t,s)) = []" "fst (the_Var (fst (t,s))) = TAtom Value" using * ***(1) Set.bspec[OF adm_checks(2)[unfolded admissible_transaction_checks_def] r] unfolding is_Fun_Set_def by simp_all obtain x where x: "t = Var x" "fst x = TAtom Value" using ts ****(1,5) by (cases t) simp_all obtain f S where fS: "s = Fun f S" using ts ****(2) by (cases s) simp_all obtain ss where ss: "f = Set ss" using fS ****(3) by (cases f) simp_all have S: "S = []" using ts fS ss ****(4) by simp show ?B' using ts x fS ss S *** **(1) * fv_r1 fv_r2 by (cases r) auto next assume ***: "the_ins r = []" "length (the_eqs r) = 1" then obtain t s where "the_eqs r = [(t,s)]" by (cases "the_eqs r") auto thus ?B' using *** **(1) * by (cases r) auto qed qed (auto simp add: is_Equality_def the_check_def intro: poscheckvariant.exhaust) qed show "?C \ ?C'" proof - assume r: ?C note adm_upds = admissible_transaction_is_wellformed_transaction(3)[OF T_adm] have *: "is_Update r" "is_Var (the_elem_term r)" "is_Fun (the_set_term r)" "is_Set (the_Fun (the_set_term r))" "args (the_set_term r) = []" "fst (the_Var (the_elem_term r)) = TAtom Value" using r adm_upds unfolding admissible_transaction_updates_def is_Fun_Set_def by fast+ obtain t s where ts: "r = insert\t,s\ \ r = delete\t,s\" using *(1) by (cases r) auto obtain x where x: "t = Var x" "fst x = TAtom Value" using ts *(2,6) by (cases t) auto obtain f T where fT: "s = Fun f T" using ts *(3) by (cases s) auto obtain ss where ss: "f = Set ss" using ts fT *(4) by (cases f) fastforce+ have T: "T = []" using ts fT *(5) ss by (cases T) auto show ?C' using ts x fT ss T by blast qed qed lemma protocol_transaction_vars_TAtom_typed: assumes T_adm: "admissible_transaction' T" shows "\x \ vars_transaction T. \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" and "\x \ fv_transaction T. \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" and "\x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" proof - note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] show "\x \ vars_transaction T. \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" using admissible_transactionE(3)[OF T_adm] by fast thus "\x \ fv_transaction T. \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" using vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t by fast show "\x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" using admissible_transactionE(2)[OF T_adm] by argo qed lemma protocol_transactions_no_pubconsts: assumes "admissible_transaction' T" shows "Fun (Val n) S \ subterms\<^sub>s\<^sub>e\<^sub>t (trms_transaction T)" and "Fun (PubConst Value n) S \ subterms\<^sub>s\<^sub>e\<^sub>t (trms_transaction T)" using assms admissible_transactions_no_Value_consts(1,3) by (blast, blast) lemma protocol_transactions_no_abss: assumes "admissible_transaction' T" shows "Fun (Abs n) S \ subterms\<^sub>s\<^sub>e\<^sub>t (trms_transaction T)" using assms admissible_transactions_no_Value_consts(2) by fast lemma admissible_transaction_strand_sem_fv_ineq: assumes T_adm: "admissible_transaction' T" and \: "strand_sem_stateful IK DB (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) \" and x: "x \ fv_transaction T - set (transaction_fresh T)" and y: "y \ fv_transaction T - set (transaction_fresh T)" and x_not_y: "x \ y" shows "\ x \ \ \ \ y \ \" proof - have "\Var x != Var y\ \ set (unlabel (transaction_checks T)) \ \Var y != Var x\ \ set (unlabel (transaction_checks T))" using x y x_not_y admissible_transactionE(8)[OF T_adm] by auto hence "\Var x != Var y\ \ set (unlabel (transaction_strand T)) \ \Var y != Var x\ \ set (unlabel (transaction_strand T))" unfolding transaction_strand_def unlabel_def by auto hence "\\ x != \ y\ \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) \ \\ y != \ x\ \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" using stateful_strand_step_subst_inI(8)[of _ _ "unlabel (transaction_strand T)" \] subst_lsst_unlabel[of "transaction_strand T" \] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(7)[of "[]" _ "[]"] by force then obtain B where B: "prefix (B@[\\ x != \ y\]) (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) \ prefix (B@[\\ y != \ x\]) (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" unfolding prefix_def by (metis (no_types, opaque_lifting) append.assoc append_Cons append_Nil split_list) thus ?thesis using \ strand_sem_append_stateful[of IK DB _ _ \] stateful_strand_sem_NegChecks_no_bvars(2) unfolding prefix_def by metis qed lemma admissible_transaction_sem_iff: fixes \ and T::"('fun,'atom,'sets,'lbl) prot_transaction" defines "A \ unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" assumes T: "admissible_transaction' T" and I: " interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)" shows "strand_sem_stateful M D A I \ ( (\l ts. (l,receive\ts\) \ set (transaction_receive T) \ (\t \ set ts. M \ t \ \ \ I)) \ (\l ac t s. (l,\ac: t \ s\) \ set (transaction_checks T) \ t \ \ \ I = s \ \ \ I) \ (\l ac t s. (l,\ac: t \ s\) \ set (transaction_checks T) \ (t \ \ \ I, s \ \ \ I) \ D) \ (\l t s. (l,\t != s\) \ set (transaction_checks T) \ t \ \ \ I \ s \ \ \ I) \ (\l t s. (l,\t not in s\) \ set (transaction_checks T) \ (t \ \ \ I, s \ \ \ I) \ D))" (is "?A \ ?B") proof - define P where "P \ \X. \\::('fun,'atom,'sets,'lbl) prot_subst. subst_domain \ = set X \ ground (subst_range \)" define rm where "rm \ \X. \\::('fun,'atom,'sets,'lbl) prot_subst. rm_vars (set X) \" define chks where "chks \ transaction_checks T" define q1 where "q1 \ \t X \. t \ rm X \ \ \ \ I" define q2 where "q2 \ \t. t \ \ \ I" note 0 = admissible_transaction_is_wellformed_transaction[OF T] note 1 = wellformed_transaction_sem_iff[OF 0(1) I, of M D \, unfolded A_def[symmetric]] note 2 = admissible_transactionE[OF T] have 3: "rm X \ = \" when "X = []" for X using that unfolding rm_def by auto have 4: "P X \ \ \ = Var" when "X = []" for X and \ using that unfolding P_def by auto have 5: "\t s. \X\\\: F \\: G\ = \t != s\ \ \X\\\: F \\: G\ = \t not in s\" when X: "(l, \X\\\: F \\: G\) \ set chks" for l X F G proof - have *: "\X\\\: F \\: G\ \ set (unlabel (transaction_strand T))" using X transaction_strand_subsets(2)[of T] unlabel_in unfolding chks_def by fast hence **: "X = []" using 2(4) by auto note *** = transaction_notinset_checks_are_Value_vars(3,4)[OF T *] transaction_noteqs_checks_case[OF T *] show ?thesis proof (cases "G = []") case True thus ?thesis using ** ***(3) by blast next case False then obtain t s where g: "(t,s) \ set G" by (cases G) auto show ?thesis using ** ***(1,2)[OF g] by blast qed qed have 6: "q1 t X \ = q2 t" when "P X \" "X = []" for X \ t using that 3 4 unfolding q1_def q2_def by simp let ?negcheck_sem = "\X F G. \\. P X \ \ (\(t,s) \ set F. q1 t X \ \ q1 s X \) \ (\(t,s) \ set G. (q1 t X \, q1 s X \) \ D)" have "(\l X F G. (l,\X\\\: F \\: G\) \ set chks \ ?negcheck_sem X F G) \ ((\l t s. (l,\t != s\) \ set chks \ q2 t \ q2 s) \ (\l t s. (l,\t not in s\) \ set chks \ (q2 t, q2 s) \ D))" (is "?A \ ?B") proof have "q2 t \ q2 s" when t: "(l,\t != s\) \ set chks" ?A for l t s proof - have "?negcheck_sem [] [(t,s)] []" using t by blast thus ?thesis using 4[of "[]"] 6[of "[]"] by force qed moreover have "(q2 t, q2 s) \ D" when t: "(l,\t not in s\) \ set chks" ?A for l t s proof - have "?negcheck_sem [] [] [(t,s)]" using t by blast thus ?thesis using 4[of "[]"] 6[of "[]"] by force qed ultimately show "?A \ ?B" by blast have "?negcheck_sem X F G" when t: "(l,\X\\\: F \\: G\) \ set chks" ?B for l X F G proof - obtain t s where ts: "(X = [] \ F = [(t,s)] \ G = []) \ (X = [] \ F = [] \ G = [(t,s)])" using 5[OF t(1)] by blast hence "(X = [] \ F = [(t,s)] \ G = [] \ q2 t \ q2 s) \ (X = [] \ F = [] \ G = [(t,s)] \ (q2 t, q2 s) \ D)" using t by blast thus ?thesis using 4[of "[]"] 6[of "[]"] by fastforce qed thus "?B \ ?A" by simp qed thus ?thesis using 1 unfolding rm_def chks_def P_def q1_def q2_def by simp qed lemma admissible_transaction_terms_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: assumes "admissible_transaction_terms T" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms_transaction T)" by (rule conjunct1[OF assms[unfolded admissible_transaction_terms_def wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code[symmetric]]]) lemma admissible_transactions_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: assumes "admissible_transaction' T" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms_transaction T)" proof - have "admissible_transaction_terms T" using assms[unfolded admissible_transaction'_def] by fast thus ?thesis by (metis admissible_transaction_terms_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s) qed lemma admissible_transaction_no_Ana_Attack: assumes "admissible_transaction_terms T" and "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms_transaction T)" shows "attack\n\ \ set (snd (Ana t))" proof - obtain r where r: "r \ set (unlabel (transaction_strand T))" "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p r)" using assms(2) by force obtain K M where t: "Ana t = (K, M)" by (metis surj_pair) show ?thesis proof assume n: "attack\n\ \ set (snd (Ana t))" hence "attack\n\ \ set M" using t by simp hence n': "attack\n\ \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p r)" using Ana_subterm[OF t] r(2) subterms_subset by fast hence "\f \ \(funs_term ` trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p r). is_Attack f" using funs_term_Fun_subterm' unfolding is_Attack_def by fast hence "is_Send r" "length (the_msgs r) = 1" "is_Fun (hd (the_msgs r))" "is_Attack (the_Fun (hd (the_msgs r)))" "args (hd (the_msgs r)) = []" using assms(1) r(1) unfolding admissible_transaction_terms_def is_Fun_Attack_def by metis+ hence "t = attack\n\" using n' r(2) unfolding is_Send_def is_Attack_def by (cases "the_msgs r") auto thus False using n by fastforce qed qed lemma admissible_transaction_Value_vars: assumes T: "admissible_transaction' T" and x: "x \ fv_transaction T" shows "\\<^sub>v x = TAtom Value" proof - have "x \ vars_transaction T" using x vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_strand T)"] by blast thus "\\<^sub>v x = TAtom Value" using admissible_transactionE(3)[OF T] by simp qed lemma admissible_transaction_occurs_checksE1: assumes T: "admissible_transaction_occurs_checks T" and x: "x \ fv_transaction T - set (transaction_fresh T)" "\\<^sub>v x = TAtom Value" obtains l ts S where "transaction_receive T = (l,receive\ts\)#S" "occurs (Var x) \ set ts" proof - let ?rcvs = "transaction_receive T" let ?frsh = "transaction_fresh T" let ?fvs = "fv_transaction T" have *: "?rcvs \ []" "is_Receive (hd (unlabel ?rcvs))" "\x \ ?fvs - set ?frsh. \\<^sub>v x = TAtom Value \ occurs (Var x) \ set (the_msgs (hd (unlabel ?rcvs)))" using T x unfolding admissible_transaction_occurs_checks_def \\<^sub>v_TAtom''(2) by meson+ obtain r S where S: "?rcvs = r#S" using *(1) by (cases ?rcvs) auto obtain l ts where r: "r = (l,receive\ts\)" by (metis *(1,2) S list.map_sel(1) list.sel(1) prod.collapse is_Receive_def unlabel_def) have 0: "occurs (Var x) \ set ts" using *(3) x S r by fastforce show ?thesis using that[unfolded S r, of l ts S] 0 by blast qed lemma admissible_transaction_occurs_checksE2: assumes T: "admissible_transaction_occurs_checks T" and x: "x \ set (transaction_fresh T)" obtains l ts S where "transaction_send T = (l,send\ts\)#S" "occurs (Var x) \ set ts" proof - let ?snds = "transaction_send T" let ?frsh = "transaction_fresh T" let ?fvs = "fv_transaction T" define ts where "ts \ the_msgs (hd (unlabel ?snds))" let ?P = "\x \ set ?frsh. occurs (Var x) \ set ts" have "?frsh \ []" using x by auto hence *: "?snds \ []" "is_Send (hd (unlabel ?snds))" "?P" using T x unfolding admissible_transaction_occurs_checks_def ts_def by meson+ obtain r S where S: "?snds = r#S" using *(1) by (cases ?snds) auto obtain l where r: "r = (l,send\ts\)" by (metis *(1,2) S list.map_sel(1) list.sel(1) prod.collapse unlabel_def ts_def stateful_strand_step.collapse(1)) have ts: "occurs (Var x) \ set ts" using x *(3) unfolding S by auto show ?thesis using that[unfolded S r, of l ts S] ts by blast qed lemma admissible_transaction_occurs_checksE3: assumes T: "admissible_transaction_occurs_checks T" and t: "OccursFact \ funs_term t \ OccursSec \ funs_term t" "t \ set ts" and ts: "send\ts\ \ set (unlabel (transaction_send T))" obtains x where "t = occurs (Var x)" "x \ set (transaction_fresh T)" proof - let ?P = "\t. \x \ set (transaction_fresh T). t = occurs (Var x)" have "?P t" when "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "OccursFact \ funs_term t \ OccursSec \ funs_term t" for t using assms that unfolding admissible_transaction_occurs_checks_def admissible_transaction_send_occurs_form_def by metis moreover have "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" using t(2) ts unfolding trms\<^sub>s\<^sub>s\<^sub>t_def by fastforce ultimately have "?P t" using t(1) by blast thus thesis using that by blast qed lemma admissible_transaction_occurs_checksE4: assumes T: "admissible_transaction_occurs_checks T" and ts: "send\ts\ \ set (unlabel (transaction_send T))" and t: "occurs t \ set ts" obtains x where "t = Var x" "x \ set (transaction_fresh T)" using admissible_transaction_occurs_checksE3[OF T _ t ts] by auto lemma admissible_transaction_occurs_checksE5: assumes T: "admissible_transaction_occurs_checks T" shows "Fun OccursSec [] \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" proof - have "\x \ set (transaction_fresh T). t = occurs (Var x)" when "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "OccursFact \ funs_term t \ OccursSec \ funs_term t" for t using assms that unfolding admissible_transaction_occurs_checks_def admissible_transaction_send_occurs_form_def by metis thus ?thesis by fastforce qed lemma admissible_transaction_occurs_checksE6: assumes T: "admissible_transaction_occurs_checks T" and t: "t \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" shows "Fun OccursSec [] \ set (snd (Ana t))" (is ?A) and "occurs k \ set (snd (Ana t))" (is ?B) proof - obtain t' where t': "t' \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "t \ t'" using t by blast have "?A \ ?B" proof (rule ccontr) assume *: "\(?A \ ?B)" hence "OccursSec \ funs_term t' \ OccursFact \ funs_term t'" by (meson t'(2) Ana_subterm' funs_term_Fun_subterm' term.order.trans) then obtain x where x: "x \ set (transaction_fresh T)" "t' = occurs (Var x)" using t'(1) T unfolding admissible_transaction_occurs_checks_def admissible_transaction_send_occurs_form_def by metis have "t = occurs (Var x) \ t = Var x \ t = Fun OccursSec []" using x(2) t'(2) by auto thus False using * by fastforce qed thus ?A ?B by simp_all qed lemma has_initial_value_producing_transactionE: fixes P::"('fun,'atom,'sets,'lbl) prot" assumes P: "has_initial_value_producing_transaction P" and P_adm: "\T \ set P. admissible_transaction' T" obtains T x s ts upds l l' where "\\<^sub>v x = TAtom Value" "Var x \ set ts" "fv\<^sub>s\<^sub>e\<^sub>t (set ts) = {x}" "\n. \(Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t set ts)" "T \ set P" "T = Transaction (\(). []) [x] [] [] upds [(l, send\ts\)]" "upds = [] \ (upds = [(l', insert\Var x, \s\\<^sub>s\)] \ (\T \ set P. \(l,a) \ set (transaction_strand T). \t. a \ select\t, \s\\<^sub>s\ \ a \ \t in \s\\<^sub>s\ \ a \ \t not in \s\\<^sub>s\ \ a \ delete\t, \s\\<^sub>s\)) \ T = Transaction (\(). []) [x] [] [] [] [(l, send\ts\)]" proof - define f where "f \ \s. list_all (\T. list_all (\a. ((is_Delete a \ is_InSet a) \ the_set_term a \ \s\\<^sub>s) \ (is_NegChecks a \ list_all (\(_,t). t \ \s\\<^sub>s) (the_ins a))) (unlabel (transaction_checks T@transaction_updates T))) P" obtain T where T0: "T \ set P" "length (transaction_fresh T) = 1" "transaction_receive T = []" "transaction_checks T = []" "length (transaction_send T) = 1" "let x = hd (transaction_fresh T); a = hd (transaction_send T); u = transaction_updates T in is_Send (snd a) \ Var x \ set (the_msgs (snd a)) \ fv\<^sub>s\<^sub>e\<^sub>t (set (the_msgs (snd a))) = {x} \ (u \ [] \ ( let b = hd u; c = snd b in tl u = [] \ is_Insert c \ the_elem_term c = Var x \ is_Fun_Set (the_set_term c) \ f (the_Set (the_Fun (the_set_term c)))))" using P unfolding has_initial_value_producing_transaction_def Let_def list_ex_iff f_def by blast obtain x upds ts h s l l' where T1: "T = Transaction h [x] [] [] upds [(l, send\ts\)]" "Var x \ set ts" "fv\<^sub>s\<^sub>e\<^sub>t (set ts) = {x}" "upds = [] \ (upds = [(l', insert\Var x, \s\\<^sub>s\)] \ f s)" proof (cases T) case T: (Transaction A B C D E F) obtain x where B: "B = [x]" using T0(2) unfolding T by (cases B) auto have C: "C = []" using T0(3) unfolding T by simp have D: "D = []" using T0(4) unfolding T by simp obtain l a where F: "F = [(l,a)]" using T0(5) unfolding T by fastforce obtain ts where ts: "a = send\ts\" using T0(6) unfolding T F by (cases a) auto obtain k u where E: "E = [] \ E = [(k,u)]" using T0(6) unfolding T by (cases E) fastforce+ have x: "Var x \ set ts" "fv\<^sub>s\<^sub>e\<^sub>t (set ts) = {x}" using T0(6) unfolding T B F ts by auto from E show ?thesis proof assume E': "E = [(k,u)]" obtain t t' where u: "u = insert\t,t'\" using T0(6) unfolding T E' by (cases u) auto have t: "t = Var x" using T0(6) unfolding T B E' u Let_def by simp obtain s where t': "t' = \s\\<^sub>s" and s: "f s" using T0(6) unfolding T B E' u Let_def by auto show ?thesis using that[OF T[unfolded B C D F ts E' u t t'] x] s by blast qed (use that[OF T[unfolded B C D F ts] x] in blast) qed note T_adm = bspec[OF P_adm T0(1)] have "x \ set (transaction_fresh T)" using T1(1) by fastforce hence x: "\\<^sub>v x = TAtom Value" using admissible_transactionE(2)[OF T_adm] by fast have "set ts \ trms_transaction T" unfolding T1(1) trms_transaction_unfold by simp hence ts: "\n. \(Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t set ts)" using admissible_transactions_no_Value_consts[OF T_adm] by blast have "a \ select\t, \s\\<^sub>s\ \ a \ \t in \s\\<^sub>s\ \ a \ \t not in \s\\<^sub>s\ \ a \ delete\t, \s\\<^sub>s\" when upds: "upds = [(k, insert\Var x,\s\\<^sub>s\)]" and T': "T' \ set P" and la: "(l,a) \ set (transaction_strand T')" for T' l k a t proof - note T'_wf = admissible_transaction_is_wellformed_transaction(1)[OF bspec[OF P_adm T']] have "a \ set (unlabel (transaction_checks T'@transaction_updates T'))" when a': "is_Check_or_Assignment a \ is_Update a" using that wellformed_transaction_strand_unlabel_memberD[OF T'_wf unlabel_in[OF la]] by (cases a) auto note 0 = this T1(4) T' note 1 = upds f_def list_all_iff show ?thesis proof (cases a) case (Delete t' s') thus ?thesis using 0 unfolding 1 by fastforce next case (InSet ac t' s') thus ?thesis using 0 unfolding 1 by fastforce next case (NegChecks X F G) thus ?thesis using 0 unfolding 1 by fastforce qed auto qed hence s: "\T \ set P. \(l,a) \ set (transaction_strand T). \t. a \ select\t, \s\\<^sub>s\ \ a \ \t in \s\\<^sub>s\ \ a \ \t not in \s\\<^sub>s\ \ a \ delete\t, \s\\<^sub>s\" when upds: "upds = [(k, insert\Var x,\s\\<^sub>s\)]" for k using upds by force have h: "h = (\(). [])" proof - have "transaction_decl T = h" using T1(1) by fastforce hence "h a = []" for a using admissible_transactionE(1)[OF T_adm] by simp thus ?thesis using ext[of h "\(). []"] by (metis case_unit_Unity) qed show ?thesis using that[OF x T1(2,3) ts T0(1)] T1(1,4) s unfolding h by auto qed lemma has_initial_value_producing_transaction_update_send_ex_filter: fixes P::"('a,'b,'c,'d) prot" defines "f \ \T. transaction_fresh T = [] \ list_ex (\a. is_Update (snd a) \ is_Send (snd a)) (transaction_strand T)" assumes P: "has_initial_value_producing_transaction P" shows "has_initial_value_producing_transaction (filter f P)" proof - define g where "g \ \P::('a,'b,'c,'d) prot. \s. list_all (\T. list_all (\a. ((is_Delete a \ is_InSet a) \ the_set_term a \ \s\\<^sub>s) \ (is_NegChecks a \ list_all (\(_,t). t \ \s\\<^sub>s) (the_ins a))) (unlabel (transaction_checks T@transaction_updates T))) P" let ?Q = "\P T. let x = hd (transaction_fresh T); a = hd (transaction_send T); u = transaction_updates T in is_Send (snd a) \ Var x \ set (the_msgs (snd a)) \ fv\<^sub>s\<^sub>e\<^sub>t (set (the_msgs (snd a))) = {x} \ (u \ [] \ ( let b = hd u; c = snd b in tl u = [] \ is_Insert c \ the_elem_term c = Var x \ is_Fun_Set (the_set_term c) \ g P (the_Set (the_Fun (the_set_term c)))))" have "set (filter f P) \ set P" by simp hence "list_all h P \ list_all h (filter f P)" for h unfolding list_all_iff by blast hence g_f_subset: "g P s \ g (filter f P) s" for s unfolding g_def by blast obtain T where T: "T \ set P" "length (transaction_fresh T) = 1" "transaction_receive T = []" "transaction_checks T = []" "length (transaction_send T) = 1" "?Q P T" using P unfolding has_initial_value_producing_transaction_def Let_def list_ex_iff g_def by blast obtain x where x: "transaction_fresh T = [x]" using T(2) by blast obtain a where a: "transaction_send T = [a]" using T(5) by blast obtain l b where b: "a = (l,b)" by (cases a) auto obtain ts where ts: "b = send\ts\" using T(6) unfolding Let_def a b by (cases b) auto have "T \ set (filter f P)" using T(1) x a unfolding b ts f_def by auto moreover have "?Q (filter f P) T" using T(6) g_f_subset by meson ultimately show ?thesis using T(2-5) unfolding has_initial_value_producing_transaction_def Let_def list_ex_iff g_def by blast qed subsection \Lemmata: Renaming, Declaration, and Fresh Substitutions\ lemma transaction_decl_subst_empty_inv: assumes "transaction_decl_subst Var T" shows "transaction_decl T () = []" using assms unfolding transaction_decl_subst_def subst_domain_Var by blast lemma transaction_decl_subst_domain: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_decl_subst \ T" shows "subst_domain \ = fst ` set (transaction_decl T ())" using assms unfolding transaction_decl_subst_def by argo lemma transaction_decl_subst_grounds_domain: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_decl_subst \ T" and "x \ fst ` set (transaction_decl T ())" shows "fv (\ x) = {}" proof - obtain c where "\ x = Fun c []" - using assms unfolding transaction_decl_subst_def by moura + using assms unfolding transaction_decl_subst_def by force thus ?thesis by simp qed lemma transaction_decl_subst_range_vars_empty: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_decl_subst \ T" shows "range_vars \ = {}" using assms unfolding transaction_decl_subst_def range_vars_def by auto lemma transaction_decl_subst_wt: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_decl_subst \ T" shows "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using assms unfolding transaction_decl_subst_def by blast lemma transaction_decl_subst_is_wf_trm: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_decl_subst \ P" shows "wf\<^sub>t\<^sub>r\<^sub>m (\ v)" proof (cases "v \ subst_domain \") case True thus ?thesis using assms unfolding transaction_decl_subst_def by fastforce qed auto lemma transaction_decl_subst_range_wf_trms: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_decl_subst \ P" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (metis transaction_decl_subst_is_wf_trm[OF assms] wf_trm_subst_range_iff) lemma transaction_renaming_subst_is_renaming: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes \: "transaction_renaming_subst \ P A" shows "\m. \\ n. \ (\,n) = Var (\,n+Suc m)" (is ?A) and "\y. \ x = Var y" (is ?B) and "\ x \ Var x" (is ?C) and "subst_domain \ = UNIV" (is ?D) and "subst_range \ \ range Var" (is ?E) and "fv (t \ \) \ range_vars \" (is ?F) proof - show 0: ?A using \ unfolding transaction_renaming_subst_def var_rename_def by force show ?B using \ unfolding transaction_renaming_subst_def var_rename_def by blast show ?C using 0 by (cases x) auto show 1: ?D using 0 by fastforce show ?E using 0 by auto show ?F by (induct t) (auto simp add: 1 subst_dom_vars_in_subst subst_fv_imgI) qed lemma transaction_renaming_subst_is_injective: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_renaming_subst \ P A" shows "inj \" proof (intro injI) fix x y::"('fun,'atom,'sets,'lbl) prot_var" obtain \x nx where x: "x = (\x,nx)" by (metis surj_pair) obtain \y ny where y: "y = (\y,ny)" by (metis surj_pair) obtain m where m: "\\. \n. \ (\, n) = Var (\, n + Suc m)" using transaction_renaming_subst_is_renaming(1)[OF assms] by blast assume "\ x = \ y" hence "\x = \y" "nx = ny" using x y m by simp_all thus "x = y" using x y by argo qed lemma transaction_renaming_subst_vars_disj: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" shows "fv\<^sub>s\<^sub>e\<^sub>t (\ ` (\(vars_transaction ` set P))) \ (\(vars_transaction ` set P)) = {}" (is ?A) and "fv\<^sub>s\<^sub>e\<^sub>t (\ ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = {}" (is ?B) and "T \ set P \ vars_transaction T \ range_vars \ = {}" (is "T \ set P \ ?C1") and "T \ set P \ bvars_transaction T \ range_vars \ = {}" (is "T \ set P \ ?C2") and "T \ set P \ fv_transaction T \ range_vars \ = {}" (is "T \ set P \ ?C3") and "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ range_vars \ = {}" (is ?D1) and "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ range_vars \ = {}" (is ?D2) and "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ range_vars \ = {}" (is ?D3) proof - define X where "X \ \(vars_transaction ` set P) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" have 1: "finite X" by (simp add: X_def) obtain n where n: "n \ max_var_set X" "\ = var_rename n" - using assms unfolding transaction_renaming_subst_def X_def by moura + using assms unfolding transaction_renaming_subst_def X_def by force hence 2: "\x \ X. snd x < Suc n" using less_Suc_max_var_set[OF _ 1] unfolding var_rename_def by fastforce have 3: "x \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` X)" "fv (\ x) \ X = {}" "x \ range_vars \" when x: "x \ X" for x using 2 x n unfolding var_rename_def by force+ show ?A ?B using 3(1,2) unfolding X_def by auto show ?C1 when T: "T \ set P" using T 3(3) unfolding X_def by blast thus ?C2 ?C3 when T: "T \ set P" using T by (simp_all add: disjoint_iff_not_equal vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t) show ?D1 using 3(3) unfolding X_def by auto thus ?D2 ?D3 by (simp_all add: disjoint_iff_not_equal vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t) qed lemma transaction_renaming_subst_wt: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_renaming_subst \ P X" shows "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" proof - { fix x::"('fun,'atom,'sets,'lbl) prot_var" - obtain \ n where x: "x = (\,n)" by moura + obtain \ n where x: "x = (\,n)" by force then obtain m where m: "\ x = Var (\,m)" - using assms transaction_renaming_subst_is_renaming(1) by moura + using assms transaction_renaming_subst_is_renaming(1) by force hence "\ (\ x) = \\<^sub>v x" using x by (simp add: \\<^sub>v_def) } thus ?thesis by (simp add: wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) qed lemma transaction_renaming_subst_is_wf_trm: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_renaming_subst \ P X" shows "wf\<^sub>t\<^sub>r\<^sub>m (\ v)" proof - - obtain \ n where "v = (\, n)" by moura + obtain \ n where "v = (\, n)" by force then obtain m where "\ v = Var (\, n + Suc m)" using transaction_renaming_subst_is_renaming(1)[OF assms] - by moura + by force thus ?thesis by (metis wf_trm_Var) qed lemma transaction_renaming_subst_range_wf_trms: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_renaming_subst \ P X" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (metis transaction_renaming_subst_is_wf_trm[OF assms] wf_trm_subst_range_iff) lemma transaction_renaming_subst_range_notin_vars: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" shows "\y. \ x = Var y \ y \ \(vars_transaction ` set P) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" proof - obtain \ n where x: "x = (\,n)" by (metis surj_pair) define y where "y \ \m. (\,n+Suc m)" have "\m \ max_var_set (\(vars_transaction ` set P) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A). \ x = Var (y m)" using assms x by (auto simp add: y_def transaction_renaming_subst_def var_rename_def) moreover have "finite (\(vars_transaction ` set P) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" by auto ultimately show ?thesis using x unfolding y_def by force qed lemma transaction_renaming_subst_var_obtain: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes \: "transaction_renaming_subst \ P X" shows "x \ fv\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \) \ \y. \ y = Var x" (is "?A1 \ ?B1") and "x \ fv (t \ \) \ \y \ fv t. \ y = Var x" (is "?A2 \ ?B2") proof - assume x: ?A1 - obtain y where y: "y \ fv\<^sub>s\<^sub>s\<^sub>t S" "x \ fv (\ y)" using fv\<^sub>s\<^sub>s\<^sub>t_subst_obtain_var[OF x] by moura + obtain y where y: "y \ fv\<^sub>s\<^sub>s\<^sub>t S" "x \ fv (\ y)" using fv\<^sub>s\<^sub>s\<^sub>t_subst_obtain_var[OF x] by force thus ?B1 using transaction_renaming_subst_is_renaming(2)[OF \, of y] by fastforce next assume x: ?A2 - obtain y where y: "y \ fv t" "x \ fv (\ y)" using fv_subst_obtain_var[OF x] by moura + obtain y where y: "y \ fv t" "x \ fv (\ y)" using fv_subst_obtain_var[OF x] by force thus ?B2 using transaction_renaming_subst_is_renaming(2)[OF \, of y] by fastforce qed lemma transaction_renaming_subst_set_eq: assumes "set P1 = set P2" shows "transaction_renaming_subst \ P1 X = transaction_renaming_subst \ P2 X" (is "?A = ?B") using assms unfolding transaction_renaming_subst_def by presburger lemma transaction_renaming_subst_vars_transaction_neq: assumes T: "T \ set P" and \: "transaction_renaming_subst \ P vars" and vars:"finite vars" and x: "x \ vars_transaction T" shows "\ y \ Var x" proof - have "\n. \ = var_rename n \ n \ max_var_set (\(vars_transaction ` set P))" using T \ vars x unfolding transaction_renaming_subst_def by auto then obtain n where n_p: "\ = var_rename n" "n \ max_var_set (\(vars_transaction ` set P))" by blast moreover have "\(vars_transaction ` set P) \ vars_transaction T" using T by blast ultimately have n_gt: "n \ max_var_set (vars_transaction T)" by auto obtain a b where ab: "x = (a,b)" by (cases x) auto obtain c d where cd: "y = (c,d)" by (cases y) auto have nb: "n \ b" using n_gt x ab by auto have "\ y = \ (c, d)" using cd by auto moreover have "... = Var (c, Suc (d + n))" unfolding n_p(1) unfolding var_rename_def by simp moreover have "... \ Var x" using nb ab by auto ultimately show ?thesis by auto qed lemma transaction_renaming_subst_fv_disj: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" shows "fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = {}" proof - have "fv\<^sub>s\<^sub>e\<^sub>t (\ ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = {}" using assms transaction_renaming_subst_vars_disj(2) by blast moreover have "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" by (simp add: vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t) ultimately show ?thesis by auto qed lemma transaction_fresh_subst_is_wf_trm: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_fresh_subst \ T X" shows "wf\<^sub>t\<^sub>r\<^sub>m (\ v)" proof (cases "v \ subst_domain \") case True then obtain c where "\ v = Fun c []" "arity c = 0" using assms unfolding transaction_fresh_subst_def - by moura + by force thus ?thesis by auto qed auto lemma transaction_fresh_subst_wt: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_fresh_subst \ T X" shows "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using assms unfolding transaction_fresh_subst_def by blast lemma transaction_fresh_subst_domain: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_fresh_subst \ T X" shows "subst_domain \ = set (transaction_fresh T)" using assms unfolding transaction_fresh_subst_def by fast lemma transaction_fresh_subst_range_wf_trms: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_fresh_subst \ T X" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (metis transaction_fresh_subst_is_wf_trm[OF assms] wf_trm_subst_range_iff) lemma transaction_fresh_subst_range_fresh: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_fresh_subst \ T M" shows "\t \ subst_range \. t \ subterms\<^sub>s\<^sub>e\<^sub>t M" and "\t \ subst_range \. t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T))" using assms unfolding transaction_fresh_subst_def by meson+ lemma transaction_fresh_subst_sends_to_val: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes \: "transaction_fresh_subst \ T X" and y: "y \ set (transaction_fresh T)" "\\<^sub>v y = TAtom Value" obtains n where "\ y = Fun (Val n) []" "Fun (Val n) [] \ subst_range \" proof - have "\ y \ subst_range \" using assms unfolding transaction_fresh_subst_def by simp obtain c where c: "\ y = Fun c []" "\public c" "arity c = 0" using \ y(1) unfolding transaction_fresh_subst_def by fastforce have "\ (\ y) = TAtom Value" using \ y(2) \\<^sub>v_TAtom''(2)[of y] wt_subst_trm''[of \ "Var y"] unfolding transaction_fresh_subst_def by simp then obtain n where "c = Val n" using c by (cases c) (auto split: option.splits) thus ?thesis using c that unfolding transaction_fresh_subst_def by fastforce qed lemma transaction_fresh_subst_sends_to_val': fixes \ \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_fresh_subst \ T X" and "y \ set (transaction_fresh T)" "\\<^sub>v y = TAtom Value" obtains n where "(\ \\<^sub>s \) y \ \ = Fun (Val n) []" "Fun (Val n) [] \ subst_range \" proof - obtain n where "\ y = Fun (Val n) []" "Fun (Val n) [] \ subst_range \" - using transaction_fresh_subst_sends_to_val[OF assms] by moura + using transaction_fresh_subst_sends_to_val[OF assms] by force thus ?thesis using that by (fastforce simp add: subst_compose_def) qed lemma transaction_fresh_subst_grounds_domain: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_fresh_subst \ T X" and "y \ set (transaction_fresh T)" shows "fv (\ y) = {}" proof - obtain c where "\ y = Fun c []" - using assms unfolding transaction_fresh_subst_def by moura + using assms unfolding transaction_fresh_subst_def by force thus ?thesis by simp qed lemma transaction_fresh_subst_range_vars_empty: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_fresh_subst \ T X" shows "range_vars \ = {}" proof - have "fv t = {}" when "t \ subst_range \" for t using assms that unfolding transaction_fresh_subst_def by fastforce thus ?thesis unfolding range_vars_def by blast qed lemma transaction_decl_fresh_renaming_substs_range: fixes \ \ \::"('fun,'atom,'sets,'lbl) prot_subst" assumes \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T M" and \: "transaction_renaming_subst \ P X" shows "x \ fst ` set (transaction_decl T ()) \ \c. (\ \\<^sub>s \ \\<^sub>s \) x = Fun c [] \ arity c = 0" and "x \ fst ` set (transaction_decl T ()) \ x \ set (transaction_fresh T) \ \c. (\ \\<^sub>s \ \\<^sub>s \) x = Fun c [] \ \public c \ arity c = 0" and "x \ fst ` set (transaction_decl T ()) \ x \ set (transaction_fresh T) \ fst x = TAtom Value \ \n. (\ \\<^sub>s \ \\<^sub>s \) x = Fun (Val n) []" and "x \ fst ` set (transaction_decl T ()) \ x \ set (transaction_fresh T) \ \y. (\ \\<^sub>s \ \\<^sub>s \) x = Var y" proof - assume "x \ fst ` set (transaction_decl T ())" then obtain c where c: "\ x = Fun c []" "arity c = 0" using \ unfolding transaction_decl_subst_def by fastforce thus "\c. (\ \\<^sub>s \ \\<^sub>s \) x = Fun c [] \ arity c = 0" using subst_compose[of "\ \\<^sub>s \" \ x] subst_compose[of \ \ x] by simp next assume x: "x \ fst ` set (transaction_decl T ())" "x \ set (transaction_fresh T)" have *: "(\ \\<^sub>s \) x = \ x" using x(1) \ unfolding transaction_decl_subst_def by (metis (no_types, opaque_lifting) subst_comp_notin_dom_eq) then obtain c where c: "(\ \\<^sub>s \) x = Fun c []" "\public c" "arity c = 0" using \ x(2) unfolding transaction_fresh_subst_def by fastforce thus "\c. (\ \\<^sub>s \ \\<^sub>s \) x = Fun c [] \ \public c \ arity c = 0" using subst_compose[of "\ \\<^sub>s \" \ x] subst_compose[of \ \ x] by simp assume "fst x = TAtom Value" hence "\ ((\ \\<^sub>s \) x) = TAtom Value" using * \ \\<^sub>v_TAtom''(2)[of x] wt_subst_trm''[of \ "Var x"] unfolding transaction_fresh_subst_def by simp then obtain n where "c = Val n" using c by (cases c) (auto split: option.splits) thus "\n. (\ \\<^sub>s\ \\<^sub>s \) x = Fun (Val n) []" using c subst_compose[of "\ \\<^sub>s \" \ x] subst_compose[of \ \ x] by simp next assume "x \ fst ` set (transaction_decl T ())" "x \ set (transaction_fresh T)" hence "(\ \\<^sub>s \) x = Var x" using \ \ unfolding transaction_decl_subst_def transaction_fresh_subst_def by (metis (no_types, opaque_lifting) subst_comp_notin_dom_eq subst_domI) thus "\y. (\ \\<^sub>s \ \\<^sub>s \) x = Var y" using transaction_renaming_subst_is_renaming(1)[OF \] subst_compose[of "\ \\<^sub>s \" \ x] subst_compose[of \ \ x] by (cases x) force qed lemma transaction_decl_fresh_renaming_substs_range': fixes \ \::"('fun,'atom,'sets,'lbl) prot_subst" assumes \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T M" and \: "transaction_renaming_subst \ P X" and t: "t \ subst_range (\ \\<^sub>s \ \\<^sub>s \)" shows "(\c. t = Fun c [] \ arity c = 0) \ (\x. t = Var x)" and "\ = Var \ (\c. t = Fun c [] \ \public c \ arity c = 0) \ (\x. t = Var x)" and "\ = Var \ \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value \ (\n. t = Fun (Val n) []) \ (\x. t = Var x)" and "\ = Var \ is_Fun t \ t \ subst_range \" proof - obtain x where x: "x \ subst_domain (\ \\<^sub>s \ \\<^sub>s \)" "(\ \\<^sub>s \ \\<^sub>s \) x = t" using t by auto note 0 = x transaction_decl_fresh_renaming_substs_range[OF \ \ \, of x] show "(\c. t = Fun c [] \ arity c = 0) \ (\x. t = Var x)" using 0 unfolding \\<^sub>v_TAtom'' by auto assume 1: "\ = Var" note 2 = transaction_decl_subst_empty_inv[OF \[unfolded 1]] show 3: "(\c. t = Fun c [] \ \public c \ arity c = 0) \ (\x. t = Var x)" using 0 2 unfolding \\<^sub>v_TAtom'' by auto show "(\n. t = Fun (Val n) []) \ (\x. t = Var x)" when "\x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" using 0 2 that unfolding \\<^sub>v_TAtom'' by auto show "t \ subst_range \" when t': "is_Fun t" proof - obtain x where x: "(\ \\<^sub>s \) x = t" using t 1 by auto show ?thesis proof (cases "x \ subst_domain \") case True thus ?thesis by (metis subst_dom_vars_in_subst subst_ground_ident_compose(1) subst_imgI x transaction_fresh_subst_grounds_domain[OF \] transaction_fresh_subst_domain[OF \]) next case False thus ?thesis by (metis (no_types, lifting) subst_compose_def subst_domI term.disc(1) that transaction_renaming_subst_is_renaming(5)[OF \] var_renaming_is_Fun_iff x) qed qed qed lemma transaction_decl_fresh_renaming_substs_range'': fixes \ \ \::"('fun,'atom,'sets,'lbl) prot_subst" assumes \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and y: "y \ fv ((\ \\<^sub>s \ \\<^sub>s \) x)" shows "\ x = Var x" and "\ x = Var x" and "\ x = Var y" and "(\ \\<^sub>s \ \\<^sub>s \) x = Var y" proof - have "\z. z \ fv (\ x)" by (metis y subst_compose_fv') hence "x \ subst_domain \" using y transaction_decl_subst_domain[OF \] transaction_decl_subst_grounds_domain[OF \, of x] by blast thus 0: "\ x = Var x" by blast hence "y \ fv ((\ \\<^sub>s \) x)" using y by (simp add: subst_compose) hence "\z. z \ fv (\ x)" by (metis subst_compose_fv') hence "x \ subst_domain \" using y transaction_fresh_subst_domain[OF \] transaction_fresh_subst_grounds_domain[OF \, of x] by blast thus 1: "\ x = Var x" by blast show "\ x = Var y" "(\ \\<^sub>s \ \\<^sub>s \) x = Var y" using 0 1 y transaction_renaming_subst_is_renaming(2)[OF \, of x] unfolding subst_compose_def by (fastforce,fastforce) qed lemma transaction_decl_fresh_renaming_substs_vars_subset: fixes \ \ \::"('fun,'atom,'sets,'lbl) prot_subst" assumes \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" shows "\(fv_transaction ` set P) \ subst_domain (\ \\<^sub>s \ \\<^sub>s \)" (is ?A) and "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ subst_domain (\ \\<^sub>s \ \\<^sub>s \)" (is ?B) and "T' \ set P \ fv_transaction T' \ subst_domain (\ \\<^sub>s \ \\<^sub>s \)" (is "T' \ set P \ ?C") and "T' \ set P \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\ \\<^sub>s \ \\<^sub>s \)) \ range_vars (\ \\<^sub>s \ \\<^sub>s \)" (is "T' \ set P \ ?D") proof - have *: "x \ subst_domain (\ \\<^sub>s \ \\<^sub>s \)" for x proof (cases "x \ subst_domain \") case True thus ?thesis using transaction_decl_subst_domain[OF \] transaction_decl_subst_grounds_domain[OF \] by (simp add: subst_domI subst_dom_vars_in_subst subst_ground_ident_compose(1)) next case False hence \_x_eq: "(\ \\<^sub>s \ \\<^sub>s \) x = (\ \\<^sub>s \) x" by (auto simp add: subst_compose) show ?thesis proof (cases "x \ subst_domain \") case True hence "x \ {x. \y. \ x = Var y \ \ y = Var x}" using transaction_fresh_subst_domain[OF \] transaction_fresh_subst_grounds_domain[OF \, of x] by auto hence "x \ subst_domain (\ \\<^sub>s \)" using subst_domain_subst_compose[of \ \] by blast thus ?thesis using \_x_eq subst_dom_vars_in_subst by fastforce next case False hence "(\ \\<^sub>s \) x = \ x" unfolding subst_compose_def by fastforce moreover have "\ x \ Var x" using transaction_renaming_subst_is_renaming(1)[OF \] by (cases x) auto ultimately show ?thesis using \_x_eq by fastforce qed qed show ?A ?B using * by blast+ show ?C when T: "T' \ set P" using T * by blast hence "fv\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T') \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \) \ range_vars (\ \\<^sub>s \ \\<^sub>s \)" when T: "T' \ set P" using T fv\<^sub>s\<^sub>s\<^sub>t_subst_subset_range_vars_if_subset_domain by blast thus ?D when T: "T' \ set P" by (metis T unlabel_subst) qed lemma transaction_decl_fresh_renaming_substs_vars_disj: fixes \ \ \::"('fun,'atom,'sets,'lbl) prot_subst" assumes \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" shows "fv\<^sub>s\<^sub>e\<^sub>t ((\ \\<^sub>s \ \\<^sub>s \) ` (\(vars_transaction ` set P))) \ (\(vars_transaction ` set P)) = {}" (is ?A) and "x \ \(vars_transaction ` set P) \ fv ((\ \\<^sub>s \ \\<^sub>s \) x) \ (\(vars_transaction ` set P)) = {}" (is "?B' \ ?B") and "T' \ set P \ vars_transaction T' \ range_vars (\ \\<^sub>s \ \\<^sub>s \) = {}" (is "T' \ set P \ ?C1") and "T' \ set P \ bvars_transaction T' \ range_vars (\ \\<^sub>s \ \\<^sub>s \) = {}" (is "T' \ set P \ ?C2") and "T' \ set P \ fv_transaction T' \ range_vars (\ \\<^sub>s \ \\<^sub>s \) = {}" (is "T' \ set P \ ?C3") and "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ range_vars (\ \\<^sub>s \ \\<^sub>s \) = {}" (is ?D1) and "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ range_vars (\ \\<^sub>s \ \\<^sub>s \) = {}" (is ?D2) and "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ range_vars (\ \\<^sub>s \ \\<^sub>s \) = {}" (is ?D3) and "range_vars \ = {}" (is ?E1) and "range_vars \ = {}" (is ?E2) and "range_vars (\ \\<^sub>s \ \\<^sub>s \) \ range_vars \" (is ?E3) proof - note 0 = transaction_renaming_subst_vars_disj[OF \] show ?A proof (cases "fv\<^sub>s\<^sub>e\<^sub>t ((\ \\<^sub>s \ \\<^sub>s \) ` (\(vars_transaction ` set P))) = {}") case False hence "\x \ (\(vars_transaction ` set P)). (\ \\<^sub>s \ \\<^sub>s \) x = \ x \ fv ((\ \\<^sub>s \ \\<^sub>s \) x) = {}" using transaction_decl_fresh_renaming_substs_range''[OF \ \ \] by auto thus ?thesis using 0(1) by force qed blast thus "?B' \ ?B" by auto show ?E1 ?E2 using transaction_fresh_subst_grounds_domain[OF \] transaction_decl_subst_grounds_domain[OF \] unfolding transaction_fresh_subst_domain[OF \, symmetric] transaction_decl_subst_domain[OF \, symmetric] by (fastforce, fastforce) thus 1: ?E3 using range_vars_subst_compose_subset[of \ \] range_vars_subst_compose_subset[of "\ \\<^sub>s \" \] by blast show ?C1 ?C2 ?C3 when T: "T' \ set P" using T 1 0(3,4,5)[of T'] by blast+ show ?D1 ?D2 ?D3 using 1 0(6,7,8) by blast+ qed lemma transaction_decl_fresh_renaming_substs_trms: fixes \ \ \::"('fun,'atom,'sets,'lbl) prot_subst" assumes \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ subst_domain \ = {}" and "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ subst_domain \ = {}" and "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ subst_domain \ = {}" shows "subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\ \\<^sub>s \ \\<^sub>s \))) = subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t S) \\<^sub>s\<^sub>e\<^sub>t (\ \\<^sub>s \ \\<^sub>s \)" proof - have 1: "\x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t S). (\f. (\ \\<^sub>s \ \\<^sub>s \) x = Fun f []) \ (\y. (\ \\<^sub>s \ \\<^sub>s \) x = Var y)" using transaction_decl_fresh_renaming_substs_range'[OF \ \ \] by blast have 2: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ subst_domain (\ \\<^sub>s \ \\<^sub>s \) = {}" using assms(4-6) subst_domain_compose[of \ \] subst_domain_compose[of "\ \\<^sub>s \" \] by blast show ?thesis using subterms_subst_lsst[OF 1 2] by simp qed lemma transaction_decl_fresh_renaming_substs_wt: fixes \ \ \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_decl_subst \ T" "transaction_fresh_subst \ T M" "transaction_renaming_subst \ P X" shows "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \ \\<^sub>s \)" using transaction_renaming_subst_wt[OF assms(3)] transaction_fresh_subst_wt[OF assms(2)] transaction_decl_subst_wt[OF assms(1)] by (metis wt_subst_compose) lemma transaction_decl_fresh_renaming_substs_range_wf_trms: fixes \ \ \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_decl_subst \ T" "transaction_fresh_subst \ T M" "transaction_renaming_subst \ P X" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \ \\<^sub>s \))" using transaction_renaming_subst_range_wf_trms[OF assms(3)] transaction_fresh_subst_range_wf_trms[OF assms(2)] transaction_decl_subst_range_wf_trms[OF assms(1)] wf_trms_subst_compose[of \ \] wf_trms_subst_compose[of "\ \\<^sub>s \" \] by metis lemma transaction_decl_fresh_renaming_substs_fv: fixes \ \::"('fun,'atom,'sets,'lbl) prot_subst" assumes \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T M" and \: "transaction_renaming_subst \ P X" and x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" shows "\y \ fv_transaction T - set (transaction_fresh T). (\ \\<^sub>s \ \\<^sub>s \) y = Var x" proof - have "x \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T) \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" using x fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] by argo then obtain y where "y \ fv_transaction T" "x \ fv ((\ \\<^sub>s \ \\<^sub>s \) y)" by (metis fv\<^sub>s\<^sub>s\<^sub>t_subst_obtain_var) thus ?thesis using transaction_decl_fresh_renaming_substs_range[OF \ \ \, of y] by (cases "y \ set (transaction_fresh T)") force+ qed lemma transaction_decl_fresh_renaming_substs_range_no_attack_const: fixes \ \ \::"('fun,'atom,'sets,'lbl) prot_subst" assumes \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T M" and \: "transaction_renaming_subst \ P X" and T: "\x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" and t: "t \ subst_range (\ \\<^sub>s \ \\<^sub>s \)" shows "\n. t = attack\n\" proof - note \\\_wt = transaction_decl_fresh_renaming_substs_wt[OF \ \ \] obtain x where x: "(\ \\<^sub>s \ \\<^sub>s \) x = t" using t by auto have x_type: "\ (Var x) = \ (Var x \ \)" "\ (Var x) = \ (Var x \ \ \\<^sub>s \ \\<^sub>s \)" using \ wt_subst_trm''[of \ "Var x"] wt_subst_trm''[OF \\\_wt, of "Var x"] unfolding transaction_decl_subst_def by (blast, blast) show ?thesis proof (cases t) case (Fun f S) hence "x \ set (transaction_fresh T) \ x \ fst ` set (transaction_decl T ())" using transaction_decl_fresh_renaming_substs_range[OF \ \ \, of x] x by force thus ?thesis proof assume "x \ set (transaction_fresh T)" hence "\ t = TAtom Value \ (\a. \ t = TAtom (Atom a))" using T x_type(2) x by (metis \.simps(1) eval_term.simps(1)) thus ?thesis by auto next assume "x \ fst ` set (transaction_decl T ())" then obtain c where c: "\ x = Fun (Fu c) []" "arity\<^sub>f c = 0" using \ unfolding transaction_decl_subst_def by auto have "\ t = TAtom Bottom \ (\a. \ t = TAtom (Atom a))" using c(1) \_consts_simps(1)[OF c(2)] x x_type eval_term.simps(1)[of _ x \] eval_term.simps(1)[of _ x "\ \\<^sub>s \ \\<^sub>s \"] by (cases "\\<^sub>f c") simp_all thus ?thesis by auto qed qed simp qed lemma transaction_decl_fresh_renaming_substs_occurs_fact_send_receive: fixes t::"('fun,'atom,'sets,'lbl) prot_term" assumes \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T M" and \: "transaction_renaming_subst \ P X" and T: "admissible_transaction' T" and t: "occurs t \ set ts" shows "send\ts\ \ set (unlabel (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) \ \ts' s. send\ts'\ \ set (unlabel (transaction_send T)) \ occurs s \ set ts' \ t = s \ \ \\<^sub>s \ \\<^sub>s \" (is "?A \ ?A'") and "receive\ts\ \ set (unlabel (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) \ \ts' s. receive\ts'\ \ set (unlabel (transaction_receive T)) \ occurs s \ set ts' \ t = s \ \ \\<^sub>s \ \\<^sub>s \" (is "?B \ ?B'") proof - assume ?A then obtain s ts' where s: "s \ set ts'" "send\ts'\ \ set (unlabel (transaction_strand T))" "occurs t = s \ \ \\<^sub>s \ \\<^sub>s \" using t stateful_strand_step_mem_substD(1)[ of ts "unlabel (transaction_strand T)" "\ \\<^sub>s \ \\<^sub>s \"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] by auto note \_empty = admissible_transaction_decl_subst_empty[OF T \] have T_decl_notin: "x \ fst ` set (transaction_decl T ())" for x using transaction_decl_subst_empty_inv[OF \[unfolded \_empty]] by simp note 0 = s(3) transaction_decl_fresh_renaming_substs_range[OF \ \ \] note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T] note T_fresh = admissible_transactionE(14)[OF T] have "\u. s = occurs u" proof (cases s) case (Var x) hence "(\c. s \ \ \\<^sub>s \ \\<^sub>s \ = Fun c []) \ (\y. s \ \ \\<^sub>s \ \\<^sub>s \ = Var y)" using 0(2-5)[of x] \_empty by (auto simp del: subst_subst_compose) thus ?thesis using 0(1) by simp next case (Fun f T) hence 1: "f = OccursFact" "length T = 2" "T ! 0 \ \ \\<^sub>s \ \\<^sub>s \ = Fun OccursSec []" "T ! 1 \ \ \\<^sub>s \ \\<^sub>s \ = t" using 0(1) by auto have "T ! 0 = Fun OccursSec []" proof (cases "T ! 0") case (Var x) thus ?thesis using 0(2-5)[of x] 1(3) T_fresh T_decl_notin unfolding list_all_iff by (auto simp del: subst_subst_compose) qed (use 1(3) in simp) thus ?thesis using Fun 1 0(1) by (auto simp del: subst_subst_compose) qed - then obtain u where u: "s = occurs u" by moura + then obtain u where u: "s = occurs u" by force hence "t = u \ \ \\<^sub>s \ \\<^sub>s \" using s(3) by fastforce thus ?A' using s u wellformed_transaction_strand_unlabel_memberD(8)[OF T_wf] by metis next assume ?B then obtain s ts' where s: "s \ set ts'" "receive\ts'\ \ set (unlabel (transaction_strand T))" "occurs t = s \ \ \\<^sub>s \ \\<^sub>s \" using t stateful_strand_step_mem_substD(2)[ of ts "unlabel (transaction_strand T)" "\ \\<^sub>s \ \\<^sub>s \"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] by auto note \_empty = admissible_transaction_decl_subst_empty[OF T \] have T_decl_notin: "x \ fst ` set (transaction_decl T ())" for x using transaction_decl_subst_empty_inv[OF \[unfolded \_empty]] by simp note 0 = s(3) transaction_decl_fresh_renaming_substs_range[OF \ \ \] note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T] note T_fresh = admissible_transactionE(14)[OF T] have "\u. s = occurs u" proof (cases s) case (Var x) hence "(\c. s \ \ \\<^sub>s \ \\<^sub>s \ = Fun c []) \ (\y. s \ \ \\<^sub>s \ \\<^sub>s \ = Var y)" using 0(2-5)[of x] \_empty by (auto simp del: subst_subst_compose) thus ?thesis using 0(1) by simp next case (Fun f T) hence 1: "f = OccursFact" "length T = 2" "T ! 0 \ \ \\<^sub>s \ \\<^sub>s \ = Fun OccursSec []" "T ! 1 \ \ \\<^sub>s \ \\<^sub>s \ = t" using 0(1) by auto have "T ! 0 = Fun OccursSec []" proof (cases "T ! 0") case (Var x) thus ?thesis using 0(2-5)[of x] 1(3) T_fresh T_decl_notin unfolding list_all_iff by (auto simp del: subst_subst_compose) qed (use 1(3) in simp) thus ?thesis using Fun 1 0(1) by (auto simp del: subst_subst_compose) qed - then obtain u where u: "s = occurs u" by moura + then obtain u where u: "s = occurs u" by force hence "t = u \ \ \\<^sub>s \ \\<^sub>s \" using s(3) by fastforce thus ?B' using s u wellformed_transaction_strand_unlabel_memberD(1)[OF T_wf] by metis qed lemma transaction_decl_subst_proj: assumes "transaction_decl_subst \ T" shows "transaction_decl_subst \ (transaction_proj n T)" using assms transaction_proj_decl_eq[of n T] unfolding transaction_decl_subst_def by presburger lemma transaction_fresh_subst_proj: assumes "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" shows "transaction_fresh_subst \ (transaction_proj n T) (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n A))" using assms transaction_proj_fresh_eq[of n T] contra_subsetD[OF subterms\<^sub>s\<^sub>e\<^sub>t_mono[OF transaction_proj_trms_subset[of n T]]] contra_subsetD[OF subterms\<^sub>s\<^sub>e\<^sub>t_mono[OF trms\<^sub>s\<^sub>s\<^sub>t_proj_subset(1)[of n A]]] unfolding transaction_fresh_subst_def by metis lemma transaction_renaming_subst_proj: assumes "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" shows "transaction_renaming_subst \ (map (transaction_proj n) P) (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n A))" proof - let ?X = "\P A. \(vars_transaction ` set P) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" define Y where "Y \ ?X (map (transaction_proj n) P) (proj n A)" define Z where "Z \ ?X P A" have "Y \ Z" using sst_vars_proj_subset(3)[of n A] transaction_proj_vars_subset[of n] unfolding Y_def Z_def by fastforce hence "insert 0 (snd ` Y) \ insert 0 (snd ` Z)" by blast moreover have "finite (insert 0 (snd ` Z))" "finite (insert 0 (snd ` Y))" unfolding Y_def Z_def by auto ultimately have 0: "max_var_set Y \ max_var_set Z" using Max_mono by blast have "\n\max_var_set Z. \ = var_rename n" using assms unfolding transaction_renaming_subst_def Z_def by blast hence "\n\max_var_set Y. \ = var_rename n" using 0 le_trans by fast thus ?thesis unfolding transaction_renaming_subst_def Y_def by blast qed lemma transaction_decl_fresh_renaming_substs_wf_sst: fixes \ \ \::"('fun,'atom,'sets,'lbl) prot_subst" assumes T: "wf'\<^sub>s\<^sub>s\<^sub>t (fst ` set (transaction_decl T ()) \ set (transaction_fresh T)) (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)))" and \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" shows "wf'\<^sub>s\<^sub>s\<^sub>t {} (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)))" proof - have 0: "range_vars \ \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)) = {}" "range_vars \ \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = {}" "ground (\ ` (fst ` set (transaction_decl T ())))" "ground (\ ` set (transaction_fresh T))" "ground (\ ` {})" using transaction_decl_subst_domain[OF \] transaction_decl_subst_grounds_domain[OF \] transaction_decl_subst_range_vars_empty[OF \] transaction_fresh_subst_range_vars_empty[OF \] transaction_fresh_subst_domain[OF \] transaction_fresh_subst_grounds_domain[OF \] by (simp, simp, simp, simp, simp) have 1: "fv\<^sub>s\<^sub>e\<^sub>t (\ ` set (transaction_fresh T)) \ set (transaction_fresh T)" (is "?A \ ?B") proof fix x assume x: "x \ ?A" then obtain y where y: "y \ set (transaction_fresh T)" "x \ fv (\ y)" by auto hence "y \ subst_domain \" using transaction_decl_subst_domain[OF \] transaction_decl_subst_grounds_domain[OF \] by fast thus "x \ ?B" using x y by auto qed let ?X = "fst ` set (transaction_decl T ()) \ set (transaction_fresh T)" have "fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>s\<^sub>e\<^sub>t (\ ` ?X))) = {}" using 0(3-5) 1 by auto hence "wf'\<^sub>s\<^sub>s\<^sub>t {} (((unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)) \\<^sub>s\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>s\<^sub>t \)" by (metis wf\<^sub>s\<^sub>s\<^sub>t_subst_apply[OF wf\<^sub>s\<^sub>s\<^sub>t_subst_apply[OF wf\<^sub>s\<^sub>s\<^sub>t_subst_apply[OF T]]]) thus ?thesis using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst unlabel_subst labeled_stateful_strand_subst_comp[OF 0(1), of "\ \\<^sub>s \"] labeled_stateful_strand_subst_comp[OF 0(2), of \] subst_compose_assoc[of \ \ \] by metis qed lemma admissible_transaction_decl_fresh_renaming_subst_not_occurs: fixes \ \ \ defines "\ \ \ \\<^sub>s \ \\<^sub>s \" assumes T_adm: "admissible_transaction' T" and \\\: "transaction_decl_subst \ T" "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" shows "\t. \ x = occurs t" and "\ x \ Fun OccursSec []" proof - note \_empty = admissible_transaction_decl_subst_empty[OF T_adm \\\(1)] note T_fresh_val = admissible_transactionE(2)[OF T_adm] show "\t. \ x = occurs t" for x using transaction_decl_fresh_renaming_substs_range'(1)[OF \\\] unfolding \_def[symmetric] by (cases "x \ subst_domain \") (force,force) show "\ x \ Fun OccursSec []" for x using transaction_decl_fresh_renaming_substs_range'(3)[ OF \\\ _ \_empty T_fresh_val, of "\ x"] unfolding \_def[symmetric] by (cases "x \ subst_domain \") auto qed subsection \Lemmata: Reachable Constraints\ lemma reachable_constraints_as_transaction_lists: fixes f defines "f \ \(T,\,\,\). dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" and "g \ concat \ map f" assumes A: "A \ reachable_constraints P" obtains Ts where "A = g Ts" and "\B. prefix B Ts \ g B \ reachable_constraints P" and "\B T \ \ \. prefix (B@[(T,\,\,\)]) Ts \ T \ set P \ transaction_decl_subst \ T \ transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (g B)) \ transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (g B))" proof - let ?P1 = "\A Ts. A = g Ts" let ?P2 = "\Ts. \B. prefix B Ts \ g B \ reachable_constraints P" let ?P3 = "\Ts. \B T \ \ \. prefix (B@[(T,\,\,\)]) Ts \ T \ set P \ transaction_decl_subst \ T \ transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (g B)) \ transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (g B))" have "\Ts. ?P1 A Ts \ ?P2 Ts \ ?P3 Ts" using A proof (induction A rule: reachable_constraints.induct) case init have "?P1 [] []" "?P2 []" "?P3 []" unfolding g_def f_def by simp_all thus ?case by blast next case (step A T \ \ \) let ?A' = "A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" obtain Ts where Ts: "?P1 A Ts" "?P2 Ts" "?P3 Ts" using step.IH by blast have 1: "?P1 ?A' (Ts@[(T,\,\,\)])" using Ts(1) unfolding g_def f_def by simp have 2: "?P2 (Ts@[(T,\,\,\)])" proof (intro allI impI) fix B assume "prefix B (Ts@[(T,\,\,\)])" hence "prefix B Ts \ B = Ts@[(T,\,\,\)]" by fastforce thus "g B \ reachable_constraints P " using Ts(1,2) reachable_constraints.step[OF step.hyps] unfolding g_def f_def by auto qed have 3: "?P3 (Ts@[(T,\,\,\)])" using Ts(1,3) step.hyps(2-5) unfolding g_def f_def by auto show ?case using 1 2 3 by blast qed thus thesis using that by blast qed lemma reachable_constraints_transaction_action_obtain: assumes A: "A \ reachable_constraints P" and a: "a \ set A" obtains T b B \ \ \ where "prefix (B@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) A" and "T \ set P" "transaction_decl_subst \ T" "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" and "b \ set (transaction_strand T)" "a = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p b \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \\<^sub>s \ \\<^sub>s \" "fst a = fst b" proof - define f where "f \ \(T,\,\::('fun,'atom,'sets,'lbl) prot_subst,\). dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" define g where "g \ concat \ map f" obtain Ts where Ts: "A = g Ts" "\B. prefix B Ts \ g B \ reachable_constraints P" "\B T \ \ \. prefix (B@[(T,\,\,\)]) Ts \ T \ set P \ transaction_decl_subst \ T \ transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (g B)) \ transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (g B))" using reachable_constraints_as_transaction_lists[OF A] unfolding g_def f_def by blast obtain T \ \ \ where T: "(T,\,\,\) \ set Ts" "a \ set (f (T,\,\,\))" using Ts(1) a unfolding g_def by auto obtain B where B: "prefix (B@[(T,\,\,\)]) Ts" using T(1) by (meson prefix_snoc_in_iff) obtain b where b: "b \ set (transaction_strand T)" "a = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p b \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \\<^sub>s \ \\<^sub>s \" "fst a = fst b" using T(2) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_memberD'[of a "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \" thesis] unfolding f_def by simp have 0: "prefix (g B@f (T, \, \, \)) A" using concat_map_mono_prefix[OF B, of f] unfolding g_def Ts(1) by simp have 1: "T \ set P" "transaction_decl_subst \ T" "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (g B))" "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (g B))" using B Ts(3) by (blast,blast,blast,blast) show thesis using 0[unfolded f_def] that[OF _ 1 b] by fast qed lemma reachable_constraints_unlabel_eq: defines "transaction_unlabel_eq \ \T1 T2. transaction_decl T1 = transaction_decl T2 \ transaction_fresh T1 = transaction_fresh T2 \ unlabel (transaction_receive T1) = unlabel (transaction_receive T2) \ unlabel (transaction_checks T1) = unlabel (transaction_checks T2) \ unlabel (transaction_updates T1) = unlabel (transaction_updates T2) \ unlabel (transaction_send T1) = unlabel (transaction_send T2)" assumes Peq: "list_all2 transaction_unlabel_eq P1 P2" shows "unlabel ` reachable_constraints P1 = unlabel ` reachable_constraints P2" (is "?A = ?B") proof (intro antisym subsetI) have "transaction_unlabel_eq T2 T1 = transaction_unlabel_eq T1 T2" for T1 T2 unfolding transaction_unlabel_eq_def by argo hence Peq': "list_all2 transaction_unlabel_eq P2 P1" using Peq list_all2_sym by metis have 0: "unlabel (transaction_strand T1) = unlabel (transaction_strand T2)" when "transaction_unlabel_eq T1 T2" for T1 T2 using that unfolding transaction_unlabel_eq_def transaction_strand_def by force have "vars_transaction T1 = vars_transaction T2" when "transaction_unlabel_eq T1 T2" for T1 T2 using 0[OF that] by simp hence "vars_transaction (P1 ! i) = vars_transaction (P2 ! i)" when "i < length P1" for i using that Peq list_all2_conv_all_nth by blast moreover have "length P1 = length P2" using Peq unfolding list_all2_iff by argo ultimately have 1: "\(vars_transaction ` set P1) = \(vars_transaction ` set P2)" using in_set_conv_nth[of _ P1] in_set_conv_nth[of _ P2] by fastforce have 2: "transaction_decl_subst \ T1 \ transaction_decl_subst \ T2" (is "?A1 \ ?A2") "transaction_fresh_subst \ T1 (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ transaction_fresh_subst \ T2 (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" (is "?B1 \ ?B2") "transaction_renaming_subst \ P1 (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ transaction_renaming_subst \ P2 (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" (is "?C1 \ ?C2") "transaction_renaming_subst \ P2 (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ transaction_renaming_subst \ P1 (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" (is "?D1 \ ?D2") when "transaction_unlabel_eq T1 T2" "unlabel \ = unlabel \" for T1 T2::"('fun,'atom,'sets,'lbl) prot_transaction" and \ \::"('fun,'atom,'sets,'lbl) prot_strand" and \ \ \::"('fun,'atom,'sets,'lbl) prot_subst" proof - have *: "transaction_decl T1 = transaction_decl T2" "transaction_fresh T1 = transaction_fresh T2" "trms_transaction T1 = trms_transaction T2" using that unfolding transaction_unlabel_eq_def transaction_strand_def by force+ show "?A1 \ ?A2" using *(1) unfolding transaction_decl_subst_def by argo show "?B1 \ ?B2" using that(2) *(2,3) unfolding transaction_fresh_subst_def by force show "?C1 \ ?C2" using that(2) 1 unfolding transaction_renaming_subst_def by metis show "?D1 \ ?D2" using that(2) 1 unfolding transaction_renaming_subst_def by metis qed have 3: "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T1 \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) = unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T2 \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" when "transaction_unlabel_eq T1 T2" for T1 T2 \ using 0[OF that] unlabel_subst[of _ \] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_cong by metis have "\\ \ reachable_constraints P2. unlabel \ = unlabel \" when "\ \ reachable_constraints P1" for \ using that proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \ \) obtain \ where IH: "\ \ reachable_constraints P2" "unlabel \ = unlabel \" by (meson step.IH) obtain T' where T': "T' \ set P2" "transaction_unlabel_eq T T'" using list_all2_in_set_ex[OF Peq step.hyps(2)] by auto show ?case using 3[OF T'(2), of "\ \\<^sub>s \ \\<^sub>s \"] IH(2) reachable_constraints.step[OF IH(1) T'(1)] 2[OF T'(2) IH(2)] step.hyps(3-5) by (metis unlabel_append[of \] unlabel_append[of \]) qed (simp add: unlabel_def) thus "\ \ ?A \ \ \ ?B" for \ by fast have "\\ \ reachable_constraints P1. unlabel \ = unlabel \" when "\ \ reachable_constraints P2" for \ using that proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \ \) obtain \ where IH: "\ \ reachable_constraints P1" "unlabel \ = unlabel \" by (meson step.IH) obtain T' where T': "T' \ set P1" "transaction_unlabel_eq T T'" using list_all2_in_set_ex[OF Peq' step.hyps(2)] by auto show ?case using 3[OF T'(2), of "\ \\<^sub>s \ \\<^sub>s \"] IH(2) reachable_constraints.step[OF IH(1) T'(1)] 2[OF T'(2) IH(2)] step.hyps(3-5) by (metis unlabel_append[of \] unlabel_append[of \]) qed (simp add: unlabel_def) thus "\ \ ?B \ \ \ ?A" for \ by fast qed lemma reachable_constraints_set_eq: assumes "set P1 = set P2" shows "reachable_constraints P1 = reachable_constraints P2" (is "?A = ?B") proof (intro antisym subsetI) note 0 = assms transaction_renaming_subst_set_eq[OF assms] note 1 = reachable_constraints.intros show "\ \ ?A \ \ \ ?B" for \ by (induct \ rule: reachable_constraints.induct) (auto simp add: 0 intro: 1) show "\ \ ?B \ \ \ ?A" for \ by (induct \ rule: reachable_constraints.induct) (auto simp add: 0 intro: 1) qed lemma reachable_constraints_set_subst: assumes "set P1 = set P2" and "Q (reachable_constraints P1)" shows "Q (reachable_constraints P2)" by (rule subst[of _ _ Q, OF reachable_constraints_set_eq[OF assms(1)] assms(2)]) lemma reachable_constraints_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: assumes "\T \ set P. wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms_transaction T)" and "\ \ reachable_constraints P" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using assms(2) proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \ \) have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms_transaction T)" using assms(1) step.hyps(2) by blast hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms_transaction T \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" using transaction_decl_fresh_renaming_substs_range_wf_trms[OF step.hyps(3-5)] by (metis wf_trms_subst) hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" using wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_trms\<^sub>s\<^sub>s\<^sub>t_subst unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] by metis hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)))" using trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq by blast thus ?case using step.IH unlabel_append[of \] trms\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \"] by auto qed simp lemma reachable_constraints_var_types_in_transactions: fixes \::"('fun,'atom,'sets,'lbl) prot_constr" assumes \: "\ \ reachable_constraints P" and P: "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" shows "\\<^sub>v ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ (\T \ set P. \\<^sub>v ` fv_transaction T)" (is "?A \") and "\\<^sub>v ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ (\T \ set P. \\<^sub>v ` bvars_transaction T)" (is "?B \") and "\\<^sub>v ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ (\T \ set P. \\<^sub>v ` vars_transaction T)" (is "?C \") using \ proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \ \) define T' where "T' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" note 2 = transaction_decl_fresh_renaming_substs_wt[OF step.hyps(3-5)] have 3: "\t \ subst_range (\ \\<^sub>s \ \\<^sub>s \). fv t = {} \ (\x. t = Var x)" using transaction_decl_fresh_renaming_substs_range'(1)[OF step.hyps(3-5)] by fastforce have "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' = bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" unfolding T'_def by (metis fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq, metis bvars\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq, metis vars\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq) hence "\ ` Var ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ \ ` Var ` fv_transaction T" "\ ` Var ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' = \ ` Var ` bvars_transaction T" "\ ` Var ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ \ ` Var ` vars_transaction T" using wt_subst_lsst_vars_type_subset[OF 2 3, of "transaction_strand T"] by argo+ hence "\\<^sub>v ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ \\<^sub>v ` fv_transaction T" "\\<^sub>v ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' = \\<^sub>v ` bvars_transaction T" "\\<^sub>v ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ \\<^sub>v ` vars_transaction T" by (metis \\<^sub>v_Var_image)+ hence 4: "\\<^sub>v ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ (\T \ set P. \\<^sub>v ` fv_transaction T)" "\\<^sub>v ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ (\T \ set P. \\<^sub>v ` bvars_transaction T)" "\\<^sub>v ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ (\T \ set P. \\<^sub>v ` vars_transaction T)" using step.hyps(2) by fast+ have 5: "\\<^sub>v ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\ @ T') = (\\<^sub>v ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ (\\<^sub>v ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" "\\<^sub>v ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\ @ T') = (\\<^sub>v ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ (\\<^sub>v ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" "\\<^sub>v ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\ @ T') = (\\<^sub>v ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ (\\<^sub>v ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" using unlabel_append[of \ T'] fv\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel T'"] bvars\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel T'"] vars\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel T'"] by auto { case 1 thus ?case using step.IH(1) 4(1) 5(1) unfolding T'_def by (simp del: subst_subst_compose fv\<^sub>s\<^sub>s\<^sub>t_def) } { case 2 thus ?case using step.IH(2) 4(2) 5(2) unfolding T'_def by (simp del: subst_subst_compose bvars\<^sub>s\<^sub>s\<^sub>t_def) } { case 3 thus ?case using step.IH(3) 4(3) 5(3) unfolding T'_def by (simp del: subst_subst_compose) } qed simp_all lemma reachable_constraints_no_bvars: assumes \: "\ \ reachable_constraints P" and P: "\T \ set P. bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T) = {}" shows "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = {}" using assms proof (induction) case init then show ?case unfolding unlabel_def by auto next case (step \ T \ \ \) then have "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = {}" by metis moreover have "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) = {}" using step by (metis bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst bvars\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq) ultimately show ?case using bvars\<^sub>s\<^sub>s\<^sub>t_append unlabel_append by (metis sup_bot.left_neutral) qed lemma reachable_constraints_fv_bvars_disj: fixes \::"('fun,'atom,'sets,'lbl) prot_constr" assumes \_reach: "\ \ reachable_constraints P" and P: "\S \ set P. admissible_transaction' S" shows "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = {}" proof - let ?X = "\T \ set P. bvars_transaction T" note 0 = admissible_transactions_fv_bvars_disj[OF P] have 1: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ ?X" using \_reach proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \ \) have "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) = bvars_transaction T" using bvars\<^sub>s\<^sub>s\<^sub>t_subst[of "unlabel (transaction_strand T)" "\ \\<^sub>s \ \\<^sub>s \"] bvars\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \"] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] by argo hence "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) \ ?X" using step.hyps(2) by blast thus ?case using step.IH bvars\<^sub>s\<^sub>s\<^sub>t_append by auto qed (simp add: unlabel_def) have 2: "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ ?X = {}" using \_reach proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \ \) have "x \ y" when x: "x \ ?X" and y: "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" for x y proof - obtain y' where y': "y' \ fv_transaction T" "y \ fv ((\ \\<^sub>s \ \\<^sub>s \) y')" using y unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] by (metis fv\<^sub>s\<^sub>s\<^sub>t_subst_obtain_var) have "y \ \(vars_transaction ` set P)" using transaction_decl_fresh_renaming_substs_range''[OF step.hyps(3-5) y'(2)] transaction_renaming_subst_range_notin_vars[OF step.hyps(5), of y'] by auto thus ?thesis using x vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t by fast qed hence "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \) \ ?X = {}" by blast thus ?case using step.IH fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \"] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] fv\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)"] unlabel_append[of \ "transaction_strand T"] by force qed (simp add: unlabel_def) show ?thesis using 0 1 2 by blast qed lemma reachable_constraints_vars_TAtom_typed: fixes \::"('fun,'atom,'sets,'lbl) prot_constr" assumes \_reach: "\ \ reachable_constraints P" and P: "\T \ set P. admissible_transaction' T" and x: "x \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" shows "\\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" proof - have \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by (metis reachable_constraints_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s admissible_transactions_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s P \_reach) have T_adm: "admissible_transaction' T" when "T \ set P" for T by (meson that Ball_set P) have "\T\set P. \x\set (transaction_fresh T). \\<^sub>v x = TAtom Value" using protocol_transaction_vars_TAtom_typed(3) P by blast hence *: "\\<^sub>v ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ (\T\set P. \\<^sub>v ` vars_transaction T)" using reachable_constraints_var_types_in_transactions[of \ P, OF \_reach] by auto have "\\<^sub>v ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ TAtom ` insert Value (range Atom)" proof - have "\\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" when "T \ set P" "x \ vars_transaction T" for T x using that protocol_transaction_vars_TAtom_typed(1)[of T] P admissible_transactionE(5) by blast hence "(\T\set P. \\<^sub>v ` vars_transaction T) \ TAtom ` insert Value (range Atom)" using P by blast thus "\\<^sub>v ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ TAtom ` insert Value (range Atom)" using * by auto qed thus ?thesis using x by auto qed lemma reachable_constraints_vars_not_attack_typed: fixes \::"('fun,'atom,'sets,'lbl) prot_constr" assumes \_reach: "\ \ reachable_constraints P" and P: "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" "\T \ set P. \x \ vars_transaction T. \TAtom AttackType \ \\<^sub>v x" and x: "x \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" shows "\TAtom AttackType \ \\<^sub>v x" using reachable_constraints_var_types_in_transactions(3)[OF \_reach P(1)] P(2) x by fastforce lemma reachable_constraints_Value_vars_are_fv: assumes \_reach: "\ \ reachable_constraints P" and P: "\T \ set P. admissible_transaction' T" and x: "x \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" and "\\<^sub>v x = TAtom Value" shows "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" proof - have "\T\set P. bvars_transaction T = {}" using P admissible_transactionE(4) by metis hence \_no_bvars: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = {}" using reachable_constraints_no_bvars[OF \_reach] by metis thus ?thesis using x vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel \"] by blast qed lemma reachable_constraints_subterms_subst: assumes \_reach: "\ \ reachable_constraints P" and \: "welltyped_constraint_model \ \" and P: "\T \ set P. admissible_transaction' T" shows "subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\ \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) = (subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \\<^sub>s\<^sub>e\<^sub>t \" proof - have \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by (metis reachable_constraints_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s admissible_transactions_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s P \_reach) from \ have \': "welltyped_constraint_model \ \" using welltyped_constraint_model_prefix by auto have 1: "\x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \). (\f. \ x = Fun f []) \ (\y. \ x = Var y)" proof fix x assume xa: "x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" have "\f T. \ x = Fun f T" using \ interpretation_grounds[of \ "Var x"] unfolding welltyped_constraint_model_def constraint_model_def by (cases "\ x") auto then obtain f T where fT_p: "\ x = Fun f T" by auto hence "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" using \ unfolding welltyped_constraint_model_def constraint_model_def using wf_trm_subst_rangeD by metis moreover have "x \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using xa var_subterm_trms\<^sub>s\<^sub>s\<^sub>t_is_vars\<^sub>s\<^sub>s\<^sub>t[of x "unlabel \"] vars_iff_subtermeq[of x] by auto hence "\a. \\<^sub>v x = TAtom a" using reachable_constraints_vars_TAtom_typed[OF \_reach P] by blast hence "\a. \ (Var x) = TAtom a" by simp hence "\a. \ (Fun f T) = TAtom a" by (metis (no_types, opaque_lifting) \' welltyped_constraint_model_def fT_p wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) ultimately show "(\f. \ x = Fun f []) \ (\y. \ x = Var y)" using TAtom_term_cases fT_p by metis qed have "\T\set P. bvars_transaction T = {}" using assms admissible_transactionE(4) by metis then have "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = {}" using reachable_constraints_no_bvars assms by metis then have 2: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ subst_domain \ = {}" by auto show ?thesis using subterms_subst_lsst[OF _ 2] 1 by simp qed lemma reachable_constraints_val_funs_private': fixes \::"('fun,'atom,'sets,'lbl) prot_constr" assumes \_reach: "\ \ reachable_constraints P" and P: "\T \ set P. admissible_transaction_terms T" "\T \ set P. transaction_decl T () = []" "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" and f: "f \ \(funs_term ` trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" shows "\is_PubConstValue f" and "\is_Abs f" proof - have "\is_PubConstValue f \ \is_Abs f" using \_reach f proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \ \) let ?T' = "unlabel (transaction_strand T) \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \" let ?T'' = "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \" note \_empty = admissible_transaction_decl_subst_empty'[OF bspec[OF P(2) step.hyps(2)] step.hyps(3)] have T: "admissible_transaction_terms T" using P(1) step.hyps(2) by metis have T_fresh: "\x \ set (transaction_fresh T). fst x = TAtom Value" when "T \ set P" for T using P that admissible_transactionE(14) unfolding list_all_iff \\<^sub>v_TAtom'' by fast show ?thesis using step proof (cases "f \ \(funs_term ` trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)") case False then obtain t where t: "t \ trms\<^sub>s\<^sub>s\<^sub>t ?T'" "f \ funs_term t" using step.prems trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of ?T''] trms\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?T'')"] unlabel_append[of \ "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?T''"] unlabel_subst[of "transaction_strand T"] by fastforce show ?thesis using trms\<^sub>s\<^sub>s\<^sub>t_funs_term_cases[OF t] proof assume "\u \ trms_transaction T. f \ funs_term u" thus ?thesis using conjunct1[OF conjunct2[OF T[unfolded admissible_transaction_terms_def]]] unfolding is_PubConstValue_def by blast next assume "\x \ fv_transaction T. f \ funs_term ((\ \\<^sub>s \ \\<^sub>s \) x)" - then obtain x where "x \ fv_transaction T" "f \ funs_term ((\ \\<^sub>s \ \\<^sub>s \) x)" by moura + then obtain x where "x \ fv_transaction T" "f \ funs_term ((\ \\<^sub>s \ \\<^sub>s \) x)" by force thus ?thesis using transaction_decl_fresh_renaming_substs_range'(3)[ OF step.hyps(3-5) _ \_empty T_fresh[OF step.hyps(2), unfolded \\<^sub>v_TAtom''(2)]] unfolding is_PubConstValue_def by (metis (no_types, lifting) funs_term_Fun_subterm prot_fun.disc(30,48) subst_imgI subtermeq_Var_const(2) term.distinct(1) term.inject(2) term.set_cases(1)) qed qed simp qed simp thus "\is_PubConstValue f" "\is_Abs f" by simp_all qed lemma reachable_constraints_val_funs_private: fixes \::"('fun,'atom,'sets,'lbl) prot_constr" assumes \_reach: "\ \ reachable_constraints P" and P: "\T \ set P. admissible_transaction' T" and f: "f \ \(funs_term ` trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" shows "\is_PubConstValue f" and "\is_Abs f" using P reachable_constraints_val_funs_private'[OF \_reach _ _ _ f] admissible_transaction_is_wellformed_transaction(4) admissible_transactionE(1,14) unfolding list_all_iff \\<^sub>v_TAtom'' by (blast,fast) lemma reachable_constraints_occurs_fact_ik_case: fixes A::"('fun,'atom,'sets,'lbl) prot_constr" assumes \_reach: "A \ reachable_constraints P" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and occ: "occurs t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" shows "\n. t = Fun (Val n) []" using \_reach occ proof (induction A rule: reachable_constraints.induct) case (step A T \ \ \) define \ where "\ \ \ \\<^sub>s \ \\<^sub>s \" have T_adm: "admissible_transaction' T" using P step.hyps(2) by blast hence T: "wellformed_transaction T" "admissible_transaction_occurs_checks T" using admissible_transaction_is_wellformed_transaction(1) P_occ step.hyps(2) by (blast,blast) have T_fresh: "\x \ set (transaction_fresh T). fst x = TAtom Value" using admissible_transactionE(14)[OF T_adm] unfolding list_all_iff by fast note \_empty = admissible_transaction_decl_subst_empty[OF T_adm step.hyps(3)] have \_dom_empty: "z \ fst ` set (transaction_decl T ())" for z using transaction_decl_subst_empty_inv[OF step.hyps(3)[unfolded \_empty]] by simp show ?case proof (cases "occurs t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A") case False hence "occurs t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using step.prems unfolding \_def by simp hence "\ts. occurs t \ set ts \ receive\ts\ \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" unfolding ik\<^sub>s\<^sub>s\<^sub>t_def by force hence "\ts. occurs t \ set ts \ send\ts\ \ set (unlabel (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(1) by blast then obtain ts s where s: "s \ set ts" "send\ts\ \ set (unlabel (transaction_strand T))" "s \ \ = occurs t" using stateful_strand_step_mem_substD(1)[of _ "unlabel (transaction_strand T)" \] unlabel_subst[of "transaction_strand T" \] by force note 0 = transaction_decl_fresh_renaming_substs_range[OF step.hyps(3-5)] have 1: "send\ts\ \ set (unlabel (transaction_send T))" using s(2) wellformed_transaction_strand_unlabel_memberD(8)[OF T(1)] by blast have 2: "is_Send (send\ts\)" unfolding is_Send_def by simp have 3: "\u. s = occurs u" proof - { fix z have "(\n. \ z = Fun (Val n) []) \ (\y. \ z = Var y)" using 0(3,4) T_fresh \_dom_empty unfolding \_def by blast hence "\u. \ z = occurs u" "\ z \ Fun OccursSec []" by auto } note * = this obtain u u' where T: "s = Fun OccursFact [u,u']" using *(1) s(3) by (cases s) auto thus ?thesis using *(2) s(3) by (cases u) auto qed obtain x where x: "x \ set (transaction_fresh T)" "s = occurs (Var x)" using 3 s(1) admissible_transaction_occurs_checksE4[OF T(2) 1] by metis have "t = \ x" using s(3) x(2) by auto thus ?thesis using 0(3)[OF \_dom_empty x(1)] x(1) T_fresh unfolding \_def by fast qed (simp add: step.IH) qed simp lemma reachable_constraints_occurs_fact_send_ex: fixes A::"('fun,'atom,'sets,'lbl) prot_constr" assumes \_reach: "A \ reachable_constraints P" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and x: "\\<^sub>v x = TAtom Value" "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" shows "\ts. occurs (Var x) \ set ts \ send\ts\ \ set (unlabel A)" using \_reach x(2) proof (induction A rule: reachable_constraints.induct) case (step A T \ \ \) note \_empty = admissible_transaction_decl_subst_empty[OF bspec[OF P step.hyps(2)] step.hyps(3)] note T = bspec[OF P_occ step.hyps(2)] show ?case proof (cases "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A") case True show ?thesis using step.IH[OF True] unlabel_append[of A] by auto next case False then obtain y where y: "y \ fv_transaction T - set (transaction_fresh T)" "(\ \\<^sub>s \ \\<^sub>s \) y = Var x" using transaction_decl_fresh_renaming_substs_fv[OF step.hyps(3-5), of x] step.prems(1) fv\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel A"] unlabel_append[of A] by auto have "\ y = Var y" using y(1) step.hyps(4) unfolding transaction_fresh_subst_def by auto hence "\ y = Var x" using y(2) unfolding subst_compose_def \_empty by simp hence y_val: "fst y = TAtom Value" "\\<^sub>v y = TAtom Value" using x(1) \\<^sub>v_TAtom''[of x] \\<^sub>v_TAtom''[of y] wt_subst_trm''[OF transaction_renaming_subst_wt[OF step.hyps(5)], of "Var y"] by force+ obtain ts where ts: "occurs (Var y) \ set ts" "receive\ts\ \ set (unlabel (transaction_receive T))" using admissible_transaction_occurs_checksE1[OF T y(1) y_val(2)] by (metis list.set_intros(1) unlabel_Cons(1)) hence "receive\ts\ \ set (unlabel (transaction_strand T))" using transaction_strand_subsets(5) by blast hence *: "receive\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \\ \ set (unlabel (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" using unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] stateful_strand_step_subst_inI(2)[of _ _ "\ \\<^sub>s \ \\<^sub>s \"] by force have "occurs (Var y) \ \ \\<^sub>s \ \\<^sub>s \ = occurs (Var x)" using y(2) by (auto simp del: subst_subst_compose) hence **: "occurs (Var x) \ set ts \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \" using ts(1) by force have "send\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \\ \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)))" using * dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(2) by blast thus ?thesis using ** unlabel_append[of A] by force qed qed simp lemma reachable_constraints_db\<^sub>l\<^sub>s\<^sub>s\<^sub>t_set_args_empty: assumes \: "\ \ reachable_constraints P" and PP: "list_all wellformed_transaction P" and admissible_transaction_updates: "let f = (\T. \x \ set (unlabel (transaction_updates T)). is_Update x \ is_Var (the_elem_term x) \ is_Fun_Set (the_set_term x) \ fst (the_Var (the_elem_term x)) = TAtom Value) in list_all f P" and d: "(t, s) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" shows "\ss. s = Fun (Set ss) []" using \ d proof (induction) case (step \ TT \ \ \) let ?TT = "transaction_strand TT \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \" let ?TTu = "unlabel ?TT" let ?TTd = "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?TT" let ?TTdu = "unlabel ?TTd" from step(6) have "(t, s) \ set (db'\<^sub>s\<^sub>s\<^sub>t ?TTdu \ (db'\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ []))" by (metis db\<^sub>s\<^sub>s\<^sub>t_append db\<^sub>s\<^sub>s\<^sub>t_def step.prems unlabel_append) hence "(t, s) \ set (db'\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ []) \ (\t' s'. insert\t',s'\ \ set ?TTdu \ t = t' \ \ \ s = s' \ \)" using db\<^sub>s\<^sub>s\<^sub>t_in_cases[of t "s" ?TTdu \] by metis thus ?case proof assume "\t' s'. insert\t',s'\ \ set ?TTdu \ t = t' \ \ \ s = s' \ \" then obtain t' s' where t's'_p: "insert\t',s'\ \ set ?TTdu" "t = t' \ \" "s = s' \ \" by metis then obtain lll where "(lll, insert\t',s'\) \ set ?TTd" by (meson unlabel_mem_has_label) hence "(lll, insert\t',s'\) \ set (transaction_strand TT \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_steps_iff(4) by blast hence "insert\t',s'\ \ set ?TTu" by (meson unlabel_in) hence "insert\t',s'\ \ set ((unlabel (transaction_strand TT)) \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" by (simp add: subst_lsst_unlabel) hence "insert\t',s'\ \ (\x. x \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \\<^sub>s \ \\<^sub>s \) ` set (unlabel (transaction_strand TT))" unfolding subst_apply_stateful_strand_def by auto then obtain u where "u \ set (unlabel (transaction_strand TT)) \ u \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \\<^sub>s \ \\<^sub>s \ = insert\t',s'\" by auto hence "\t'' s''. insert\t'',s''\ \ set (unlabel (transaction_strand TT)) \ t' = t'' \ \ \\<^sub>s \ \\<^sub>s \ \ s' = s'' \ \ \\<^sub>s \ \\<^sub>s \" by (cases u) auto then obtain t'' s'' where t''s''_p: "insert\t'',s''\ \ set (unlabel (transaction_strand TT)) \ t' = t'' \ \ \\<^sub>s \ \\<^sub>s \ \ s' = s'' \ \ \\<^sub>s \ \\<^sub>s \" by auto hence "insert\t'',s''\ \ set (unlabel (transaction_updates TT))" using is_Update_in_transaction_updates[of "insert\t'',s''\" TT] using PP step(2) unfolding list_all_iff by auto moreover have "\x\set (unlabel (transaction_updates TT)). is_Fun_Set (the_set_term x)" using step(2) admissible_transaction_updates unfolding is_Fun_Set_def list_all_iff by auto ultimately have "is_Fun_Set (the_set_term (insert\t'',s''\))" by auto moreover have "s' = s'' \ \ \\<^sub>s \ \\<^sub>s \" using t''s''_p by blast ultimately have "is_Fun_Set (the_set_term (insert\t',s'\))" by (auto simp add: is_Fun_Set_subst) hence "is_Fun_Set s" by (simp add: t's'_p(3) is_Fun_Set_subst) thus ?case using is_Fun_Set_exi by auto qed (auto simp add: step db\<^sub>s\<^sub>s\<^sub>t_def) qed auto lemma reachable_constraints_occurs_fact_ik_ground: fixes A::"('fun,'atom,'sets,'lbl) prot_constr" assumes \_reach: "A \ reachable_constraints P" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and t: "occurs t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" shows "fv (occurs t) = {}" proof - have 0: "admissible_transaction' T" when "T \ set P" for T using P that unfolding list_all_iff by simp note 1 = admissible_transaction_is_wellformed_transaction(1)[OF 0] bspec[OF P_occ] have 2: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) = (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \ (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \)" when "T \ set P" for T \ and A::"('fun,'atom,'sets,'lbl) prot_constr" using dual_transaction_ik_is_transaction_send'[OF 1(1)[OF that]] by fastforce show ?thesis using \_reach t proof (induction A rule: reachable_constraints.induct) case (step A T \ \ \) note \_empty = admissible_transaction_decl_subst_empty[OF 0[OF step.hyps(2)] step.hyps(3)] from step show ?case proof (cases "occurs t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A") case False hence "occurs t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \" using 2[OF step.hyps(2)] step.prems \_empty by blast then obtain ts where ts: "occurs t \ set ts" "send\ts\ \ set (unlabel (transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" using wellformed_transaction_send_receive_subst_trm_cases(2)[OF 1(1)[OF step.hyps(2)]] by blast then obtain ts' s where s: "occurs s \ set ts'" "send\ts'\ \ set (unlabel (transaction_send T))" "t = s \ \ \\<^sub>s \ \\<^sub>s \" using transaction_decl_fresh_renaming_substs_occurs_fact_send_receive(1)[ OF step.hyps(3-5) 0[OF step.hyps(2)] ts(1)] transaction_strand_subst_subsets(8)[of T "\ \\<^sub>s \ \\<^sub>s \"] by blast obtain x where x: "x \ set (transaction_fresh T)" "s = Var x" using admissible_transaction_occurs_checksE4[OF 1(2)[OF step.hyps(2)] s(2,1)] by metis have "fv t = {}" using transaction_decl_fresh_renaming_substs_range(2)[OF step.hyps(3-5) _ x(1)] s(3) x(2) transaction_decl_subst_empty_inv[OF step.hyps(3)[unfolded \_empty]] by (auto simp del: subst_subst_compose) thus ?thesis by simp qed simp qed simp qed lemma reachable_constraints_occurs_fact_ik_funs_terms: fixes A::"('fun,'atom,'sets,'lbl) prot_constr" assumes \_reach: "A \ reachable_constraints P" and \: "welltyped_constraint_model I A" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" shows "\s \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I). OccursFact \ \(funs_term ` set (snd (Ana s)))" (is "?A A") and "\s \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I). OccursSec \ \(funs_term ` set (snd (Ana s)))" (is "?B A") and "Fun OccursSec [] \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I" (is "?C A") and "\x \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. I x \ Fun OccursSec []" (is "?D A") proof - have T_adm: "admissible_transaction' T" when "T \ set P" for T using P that unfolding list_all_iff by simp note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] note T_occ = bspec[OF P_occ] note \_empty = admissible_transaction_decl_subst_empty[OF T_adm] have \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" by (metis \ welltyped_constraint_model_def) have \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)" by (metis \ welltyped_constraint_model_def constraint_model_def) have \_grounds: "fv (I x) = {}" "\f T. I x = Fun f T" for x using \ interpretation_grounds[of I, of "Var x"] empty_fv_exists_fun[of "I x"] unfolding welltyped_constraint_model_def constraint_model_def by auto have 00: "fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \ vars_transaction T" "fv\<^sub>s\<^sub>e\<^sub>t (subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))) = fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))" for T::"('fun,'atom,'sets,'lbl) prot_transaction" using fv_trms\<^sub>s\<^sub>s\<^sub>t_subset(1)[of "unlabel (transaction_send T)"] vars_transaction_unfold fv_subterms_set[of "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)"] by blast+ have 0: "\x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)). \a. \ (Var x) = TAtom a" "\x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)). \ (Var x) \ TAtom OccursSecType" "\x \ fv\<^sub>s\<^sub>e\<^sub>t (subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))). \a. \ (Var x) = TAtom a" "\x \ fv\<^sub>s\<^sub>e\<^sub>t (subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))). \ (Var x) \ TAtom OccursSecType" "\x \ vars_transaction T. \a. \ (Var x) = TAtom a" "\x \ vars_transaction T. \ (Var x) \ TAtom OccursSecType" when "T \ set P" for T using admissible_transaction_occurs_fv_types[OF T_adm[OF that]] 00 by blast+ note T_fresh_type = admissible_transactionE(2)[OF T_adm] have 1: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \\<^sub>s\<^sub>e\<^sub>t I = (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I) \ (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t I)" when "T \ set P" for T \ and A::"('fun,'atom,'sets,'lbl) prot_constr" using dual_transaction_ik_is_transaction_send'[OF T_wf[OF that]] by fastforce have 2: "subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t I) = subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t I" when "T \ set P" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" for T \ using wt_subst_TAtom_subterms_set_subst[OF wt_subst_compose[OF \(1) \_wt] 0(1)[OF that(1)]] wf_trm_subst_rangeD[OF wf_trms_subst_compose[OF \(2) \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s]] by auto have 3: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \ \\<^sub>s \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \ \\<^sub>s \))" when "T \ set P" "transaction_decl_subst \ T" "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" for \ \ \ and T::"('fun,'atom,'sets,'lbl) prot_transaction" and A::"('fun,'atom,'sets,'lbl) prot_constr" using protocol_transaction_vars_TAtom_typed(3)[of T] P that(1) transaction_decl_fresh_renaming_substs_wt[OF that(2-4)] transaction_decl_fresh_renaming_substs_range_wf_trms[OF that(2-4)] wf_trms_subst_compose by simp_all have 4: "\s \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)). OccursFact \ \(funs_term ` set (snd (Ana s))) \ OccursSec \ \(funs_term ` set (snd (Ana s)))" when T: "T \ set P" for T proof fix t assume t: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))" then obtain ts s where s: "send\ts\ \ set (unlabel (transaction_send T))" "s \ set ts" "t \ subterms s" using wellformed_transaction_unlabel_cases(4)[OF T_wf[OF T]] by fastforce have s_occ: "\x. s = occurs (Var x)" when "OccursFact \ funs_term t \ OccursSec \ funs_term t" using that s(1) subtermeq_imp_funs_term_subset[OF s(3)] admissible_transaction_occurs_checksE3[OF T_occ[OF T] _ s(2)] by blast - obtain K T' where K: "Ana t = (K,T')" by moura + obtain K T' where K: "Ana t = (K,T')" by force show "OccursFact \ \(funs_term ` set (snd (Ana t))) \ OccursSec \ \(funs_term ` set (snd (Ana t)))" proof (rule ccontr) assume "\(OccursFact \ \(funs_term ` set (snd (Ana t))) \ OccursSec \ \(funs_term ` set (snd (Ana t))))" hence a: "OccursFact \ \(funs_term ` set (snd (Ana t))) \ OccursSec \ \(funs_term ` set (snd (Ana t)))" by simp hence "OccursFact \ \(funs_term ` set T') \ OccursSec \ \(funs_term ` set T')" using K by simp hence "OccursFact \ funs_term t \ OccursSec \ funs_term t" using Ana_subterm[OF K] funs_term_subterms_eq(1)[of t] by blast then obtain x where x: "t \ subterms (occurs (Var x))" using s(3) s_occ by blast thus False using a by fastforce qed qed have 5: "OccursFact \ \(funs_term ` subst_range (\ \\<^sub>s \ \\<^sub>s \))" "OccursSec \ \(funs_term ` subst_range (\ \\<^sub>s \ \\<^sub>s \))" when \\\: "transaction_decl_subst \ T" "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" and T: "T \ set P" for \ \ \ and T::"('fun,'atom,'sets,'lbl) prot_transaction" and A::"('fun,'atom,'sets,'lbl) prot_constr" proof - have "OccursFact \ funs_term t" "OccursSec \ funs_term t" when "t \ subst_range (\ \\<^sub>s \ \\<^sub>s \)" for t using transaction_decl_fresh_renaming_substs_range'(3)[ OF \\\ that \_empty[OF T \\\(1)] T_fresh_type[OF T]] by auto thus "OccursFact \ \(funs_term ` subst_range (\ \\<^sub>s \ \\<^sub>s \))" "OccursSec \ \(funs_term ` subst_range (\ \\<^sub>s \ \\<^sub>s \))" by blast+ qed have 6: "I x \ Fun OccursSec []" "\t. I x = occurs t" "\a. \ (I x) = TAtom a \ a \ OccursSecType" when T: "T \ set P" and \\\: "transaction_decl_subst \ T" "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" and x: "Var x \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \" for x \ \ \ and T::"('fun,'atom,'sets,'lbl) prot_transaction" and A::"('fun,'atom,'sets,'lbl) prot_constr" proof - obtain t where t: "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "t \ (\ \\<^sub>s \ \\<^sub>s \) = Var x" - using x by moura + using x by force then obtain y where y: "t = Var y" by (cases t) auto have "\a. \ t = TAtom a \ a \ OccursSecType" using 0(1,2)[OF T] t(1) y by force thus "\a. \ (I x) = TAtom a \ a \ OccursSecType" using wt_subst_trm''[OF 3(1)[OF T \\\]] wt_subst_trm''[OF \_wt] t(2) by (metis eval_term.simps(1)) thus "I x \ Fun OccursSec []" "\t. I x = occurs t" by auto qed have 7: "I x \ Fun OccursSec []" "\t. I x = occurs t" "\a. \ (I x) = TAtom a \ a \ OccursSecType" when T: "T \ set P" and \\\: "transaction_decl_subst \ T" "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" and x: "x \ fv\<^sub>s\<^sub>e\<^sub>t ((\ \\<^sub>s \ \\<^sub>s \) ` vars_transaction T)" for x \ \ \ and T::"('fun,'atom,'sets,'lbl) prot_transaction" and A::"('fun,'atom,'sets,'lbl) prot_constr" proof - obtain y where y: "y \ vars_transaction T" "x \ fv ((\ \\<^sub>s \ \\<^sub>s \) y)" using x by auto hence y': "(\ \\<^sub>s \ \\<^sub>s \) y = Var x" using transaction_decl_fresh_renaming_substs_range'(3)[ OF \\\ _ \_empty[OF T \\\(1)] T_fresh_type[OF T]] by (cases "(\ \\<^sub>s \ \\<^sub>s \) y \ subst_range (\ \\<^sub>s \ \\<^sub>s \)") force+ have "\a. \ (Var y) = TAtom a \ a \ OccursSecType" using 0(5,6)[OF T] y by force thus "\a. \ (I x) = TAtom a \ a \ OccursSecType" using wt_subst_trm''[OF 3(1)[OF T \\\]] wt_subst_trm''[OF \_wt] y' by (metis eval_term.simps(1)) thus "I x \ Fun OccursSec []" "\t. I x = occurs t" by auto qed have 8: "I x \ Fun OccursSec []" "\t. I x = occurs t" "\a. \ (I x) = TAtom a \ a \ OccursSecType" when T: "T \ set P" and \\\: "transaction_decl_subst \ T" "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" and x: "Var x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \" for x \ \ \ and T::"('fun,'atom,'sets,'lbl) prot_transaction" and A::"('fun,'atom,'sets,'lbl) prot_constr" proof - obtain t where t: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))" "t \ (\ \\<^sub>s \ \\<^sub>s \) = Var x" - using x by moura + using x by force then obtain y where y: "t = Var y" by (cases t) auto have "\a. \ t = TAtom a \ a \ OccursSecType" using 0(3,4)[OF T] t(1) y by force thus "\a. \ (I x) = TAtom a \ a \ OccursSecType" using wt_subst_trm''[OF 3(1)[OF T \\\]] wt_subst_trm''[OF \_wt] t(2) by (metis eval_term.simps(1)) thus "I x \ Fun OccursSec []" "\t. I x = occurs t" by auto qed have s_fv: "fv s \ fv\<^sub>s\<^sub>e\<^sub>t ((\ \\<^sub>s \ \\<^sub>s \) ` vars_transaction T)" when s: "s \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \" and T: "T \ set P" for s and \ \ \::"('fun,'atom,'sets,'lbl) prot_subst" and T::"('fun,'atom,'sets,'lbl) prot_transaction" proof fix x assume "x \ fv s" hence "x \ fv\<^sub>s\<^sub>e\<^sub>t (subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" using s by auto hence *: "x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" using fv_subterms_set_subst' by fast have **: "list_all is_Send (unlabel (transaction_send T))" using T_wf[OF T] unfolding wellformed_transaction_def by blast have "x \ fv\<^sub>s\<^sub>e\<^sub>t ((\ \\<^sub>s \ \\<^sub>s \) ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))" proof - obtain t where t: "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "x \ fv (t \ \ \\<^sub>s \ \\<^sub>s \)" using * by fastforce hence "fv t \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" using fv_trms\<^sub>s\<^sub>s\<^sub>t_subset(1)[of "unlabel (transaction_send T)"] by auto thus ?thesis using t(2) subst_apply_fv_subset by fast qed thus "x \ fv\<^sub>s\<^sub>e\<^sub>t ((\ \\<^sub>s \ \\<^sub>s \) ` vars_transaction T)" using vars_transaction_unfold[of T] by fastforce qed show "?A A" using \_reach proof (induction A rule: reachable_constraints.induct) case (step A T \ \ \) have *: "\s \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)). OccursFact \ \(funs_term ` set (snd (Ana s)))" using 4[OF step.hyps(2)] by blast have "\s \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \ \\<^sub>s\<^sub>e\<^sub>t I. OccursFact \ \(funs_term ` set (snd (Ana s)))" proof fix t assume t: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \ \\<^sub>s\<^sub>e\<^sub>t I" then obtain s u where su: "s \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \" "s \ I = t" "u \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))" "u \ \ \\<^sub>s \ \\<^sub>s \ = s" by force - obtain Ku Tu where KTu: "Ana u = (Ku,Tu)" by moura + obtain Ku Tu where KTu: "Ana u = (Ku,Tu)" by force have *: "OccursFact \ \(funs_term ` set Tu)" "OccursFact \ \(funs_term ` subst_range (\ \\<^sub>s \ \\<^sub>s \))" "OccursFact \ \(funs_term ` \(((set \ snd \ Ana) ` subst_range (\ \\<^sub>s \ \\<^sub>s \))))" using transaction_decl_fresh_renaming_substs_range'(3)[ OF step.hyps(3-5) _ \_empty[OF step.hyps(2,3)] T_fresh_type[OF step.hyps(2)]] 4[OF step.hyps(2)] su(3) KTu by (fastforce,fastforce,fastforce) have "OccursFact \ \(funs_term ` set (Tu \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" proof - { fix f assume f: "f \ \(funs_term ` set (Tu \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" - then obtain tf where tf: "tf \ set Tu" "f \ funs_term (tf \ \ \\<^sub>s \ \\<^sub>s \)" by moura + then obtain tf where tf: "tf \ set Tu" "f \ funs_term (tf \ \ \\<^sub>s \ \\<^sub>s \)" by force hence "f \ funs_term tf \ f \ \(funs_term ` subst_range (\ \\<^sub>s \ \\<^sub>s \))" using funs_term_subst[of tf "\ \\<^sub>s \ \\<^sub>s \"] by force hence "f \ OccursFact" using *(1,2) tf(1) by blast } thus ?thesis by metis qed hence **: "OccursFact \ \(funs_term ` set (snd (Ana s)))" proof (cases u) case (Var xu) hence "s = (\ \\<^sub>s \ \\<^sub>s \) xu" using su(4) by (metis eval_term.simps(1)) thus ?thesis using *(3) by fastforce qed (use su(4) KTu Ana_subst'[of _ _ Ku Tu "\ \\<^sub>s \ \\<^sub>s \"] in simp) show "OccursFact \ \(funs_term ` set (snd (Ana t)))" proof (cases s) case (Var sx) then obtain a where a: "\ (I sx) = Var a" using su(1) 8(3)[OF step.hyps(2-5), of sx] by fast hence "Ana (I sx) = ([],[])" by (metis \_grounds(2) const_type_inv[THEN Ana_const]) thus ?thesis using Var su(2) by simp next case (Fun f S) hence snd_Ana_t: "snd (Ana t) = snd (Ana s) \\<^sub>l\<^sub>i\<^sub>s\<^sub>t I" using su(2) Ana_subst'[of f S _ "snd (Ana s)" I] by (cases "Ana s") simp_all { fix g assume "g \ \(funs_term ` set (snd (Ana t)))" hence "g \ \(funs_term ` set (snd (Ana s))) \ (\x \ fv\<^sub>s\<^sub>e\<^sub>t (set (snd (Ana s))). g \ funs_term (I x))" using snd_Ana_t funs_term_subst[of _ I] by auto hence "g \ OccursFact" proof assume "\x \ fv\<^sub>s\<^sub>e\<^sub>t (set (snd (Ana s))). g \ funs_term (I x)" - then obtain x where x: "x \ fv\<^sub>s\<^sub>e\<^sub>t (set (snd (Ana s)))" "g \ funs_term (I x)" by moura + then obtain x where x: "x \ fv\<^sub>s\<^sub>e\<^sub>t (set (snd (Ana s)))" "g \ funs_term (I x)" by force have "x \ fv s" using x(1) Ana_vars(2)[of s] by (cases "Ana s") auto hence "x \ fv\<^sub>s\<^sub>e\<^sub>t ((\ \\<^sub>s \ \\<^sub>s \) ` vars_transaction T)" using s_fv[OF su(1) step.hyps(2)] by blast then obtain a h U where h: "I x = Fun h U" "\ (I x) = Var a" "a \ OccursSecType" "arity h = 0" using \_grounds(2) 7(3)[OF step.hyps(2-5)] const_type_inv by metis hence "h \ OccursFact" by auto moreover have "U = []" using h(1,2,4) const_type_inv_wf[of h U a] \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s by fastforce ultimately show ?thesis using h(1) x(2) by auto qed (use ** in blast) } thus ?thesis by blast qed qed thus ?case using step.IH step.prems 1[OF step.hyps(2), of A "\ \\<^sub>s \ \\<^sub>s \"] 2[OF step.hyps(2) 3[OF step.hyps(2-5)]] by auto qed simp show "?B A" using \_reach proof (induction A rule: reachable_constraints.induct) case (step A T \ \ \) have "\s \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \ \\<^sub>s\<^sub>e\<^sub>t I. OccursSec \ \(funs_term ` set (snd (Ana s)))" proof fix t assume t: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \ \\<^sub>s\<^sub>e\<^sub>t I" then obtain s u where su: "s \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \" "s \ I = t" "u \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))" "u \ \ \\<^sub>s \ \\<^sub>s \ = s" by force - obtain Ku Tu where KTu: "Ana u = (Ku,Tu)" by moura + obtain Ku Tu where KTu: "Ana u = (Ku,Tu)" by force have *: "OccursSec \ \(funs_term ` set Tu)" "OccursSec \ \(funs_term ` subst_range (\ \\<^sub>s \ \\<^sub>s \))" "OccursSec \ \(funs_term ` \(((set \ snd \ Ana) ` subst_range (\ \\<^sub>s \ \\<^sub>s \))))" using transaction_decl_fresh_renaming_substs_range'(3)[ OF step.hyps(3-5) _ \_empty[OF step.hyps(2,3)] T_fresh_type[OF step.hyps(2)]] 4[OF step.hyps(2)] su(3) KTu by (fastforce,fastforce,fastforce) have "OccursSec \ \(funs_term ` set (Tu \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" proof - { fix f assume f: "f \ \(funs_term ` set (Tu \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" - then obtain tf where tf: "tf \ set Tu" "f \ funs_term (tf \ \ \\<^sub>s \ \\<^sub>s \)" by moura + then obtain tf where tf: "tf \ set Tu" "f \ funs_term (tf \ \ \\<^sub>s \ \\<^sub>s \)" by force hence "f \ funs_term tf \ f \ \(funs_term ` subst_range (\ \\<^sub>s \ \\<^sub>s \))" using funs_term_subst[of tf "\ \\<^sub>s \ \\<^sub>s \"] by force hence "f \ OccursSec" using *(1,2) tf(1) by blast } thus ?thesis by metis qed hence **: "OccursSec \ \(funs_term ` set (snd (Ana s)))" proof (cases u) case (Var xu) hence "s = (\ \\<^sub>s \ \\<^sub>s \) xu" using su(4) by (metis eval_term.simps(1)) thus ?thesis using *(3) by fastforce qed (use su(4) KTu Ana_subst'[of _ _ Ku Tu "\ \\<^sub>s \ \\<^sub>s \"] in simp) show "OccursSec \ \(funs_term ` set (snd (Ana t)))" proof (cases s) case (Var sx) then obtain a where a: "\ (I sx) = Var a" using su(1) 8(3)[OF step.hyps(2-5), of sx] by fast hence "Ana (I sx) = ([],[])" by (metis \_grounds(2) const_type_inv[THEN Ana_const]) thus ?thesis using Var su(2) by simp next case (Fun f S) hence snd_Ana_t: "snd (Ana t) = snd (Ana s) \\<^sub>l\<^sub>i\<^sub>s\<^sub>t I" using su(2) Ana_subst'[of f S _ "snd (Ana s)" I] by (cases "Ana s") simp_all { fix g assume "g \ \(funs_term ` set (snd (Ana t)))" hence "g \ \(funs_term ` set (snd (Ana s))) \ (\x \ fv\<^sub>s\<^sub>e\<^sub>t (set (snd (Ana s))). g \ funs_term (I x))" using snd_Ana_t funs_term_subst[of _ I] by auto hence "g \ OccursSec" proof assume "\x \ fv\<^sub>s\<^sub>e\<^sub>t (set (snd (Ana s))). g \ funs_term (I x)" - then obtain x where x: "x \ fv\<^sub>s\<^sub>e\<^sub>t (set (snd (Ana s)))" "g \ funs_term (I x)" by moura + then obtain x where x: "x \ fv\<^sub>s\<^sub>e\<^sub>t (set (snd (Ana s)))" "g \ funs_term (I x)" by force have "x \ fv s" using x(1) Ana_vars(2)[of s] by (cases "Ana s") auto hence "x \ fv\<^sub>s\<^sub>e\<^sub>t ((\ \\<^sub>s \ \\<^sub>s \) ` vars_transaction T)" using s_fv[OF su(1) step.hyps(2)] by blast then obtain a h U where h: "I x = Fun h U" "\ (I x) = Var a" "a \ OccursSecType" "arity h = 0" using \_grounds(2) 7(3)[OF step.hyps(2-5)] const_type_inv by metis hence "h \ OccursSec" by auto moreover have "U = []" using h(1,2,4) const_type_inv_wf[of h U a] \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s by fastforce ultimately show ?thesis using h(1) x(2) by auto qed (use ** in blast) } thus ?thesis by blast qed qed thus ?case using step.IH step.prems 1[OF step.hyps(2), of A "\ \\<^sub>s \ \\<^sub>s \"] 2[OF step.hyps(2) 3[OF step.hyps(2-5)]] by auto qed simp show "?C A" using \_reach proof (induction A rule: reachable_constraints.induct) case (step A T \ \ \) have *: "Fun OccursSec [] \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" using admissible_transaction_occurs_checksE5[OF T_occ[OF step.hyps(2)]] by blast have **: "Fun OccursSec [] \ subst_range (\ \\<^sub>s \ \\<^sub>s \)" using transaction_decl_fresh_renaming_substs_range'(3)[ OF step.hyps(3-5) _ \_empty[OF step.hyps(2,3)] T_fresh_type[OF step.hyps(2)]] by auto have "Fun OccursSec [] \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \ \\<^sub>s\<^sub>e\<^sub>t I" proof assume "Fun OccursSec [] \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \ \\<^sub>s\<^sub>e\<^sub>t I" then obtain s where "s \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \" "s \ I = Fun OccursSec []" - by moura + by force moreover have "Fun OccursSec [] \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \" proof assume "Fun OccursSec [] \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \" then obtain u where "u \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "u \ \ \\<^sub>s \ \\<^sub>s \ = Fun OccursSec []" - by moura + by force thus False using * ** by (cases u) (force simp del: subst_subst_compose)+ qed ultimately show False using 6[OF step.hyps(2-5)] by (cases s) auto qed thus ?case using step.IH step.prems 1[OF step.hyps(2), of A "\ \\<^sub>s \ \\<^sub>s \"] by fast qed simp show "?D A" using \_reach proof (induction A rule: reachable_constraints.induct) case (step A T \ \ \) { fix x assume x: "x \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" hence x': "x \ vars\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T) \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" by (metis vars\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq unlabel_subst) hence "x \ vars_transaction T \ x \ fv\<^sub>s\<^sub>e\<^sub>t ((\ \\<^sub>s \ \\<^sub>s \) ` vars_transaction T)" using vars\<^sub>s\<^sub>s\<^sub>t_subst_cases[OF x'] by metis moreover have "I x \ Fun OccursSec []" when "x \ vars_transaction T" using that 0(5,6)[OF step.hyps(2)] wt_subst_trm''[OF \_wt, of "Var x"] by fastforce ultimately have "I x \ Fun OccursSec []" using 7(1)[OF step.hyps(2-5), of x] by blast } thus ?case using step.IH by auto qed simp qed lemma reachable_constraints_occurs_fact_ik_subst_aux: assumes \_reach: "A \ reachable_constraints P" and \: "welltyped_constraint_model I A" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and t: "t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "t \ I = occurs s" shows "\u. t = occurs u" proof - have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" using \ unfolding welltyped_constraint_model_def constraint_model_def by metis hence 0: "\ t = \ (occurs s)" using t(2) wt_subst_trm'' by metis have 1: "\\<^sub>v ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ (\T \ set P. \\<^sub>v ` fv_transaction T)" "\T \ set P. \x \ fv_transaction T. \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" using reachable_constraints_var_types_in_transactions(1)[OF \_reach] protocol_transaction_vars_TAtom_typed(2,3) P by fast+ show ?thesis proof (cases t) case (Var x) thus ?thesis using 0 1 t(1) var_subterm_ik\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t[of x "unlabel A"] by fastforce next case (Fun f T) hence 2: "f = OccursFact" "length T = Suc (Suc 0)" "T ! 0 \ I = Fun OccursSec []" using t(2) by auto have "T ! 0 = Fun OccursSec []" proof (cases "T ! 0") case (Var y) hence "I y = Fun OccursSec []" using Fun 2(3) by simp moreover have "Var y \ set T" using Var 2(2) length_Suc_conv[of T 1] by auto hence "y \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" using Fun t(1) by force hence "y \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" using fv_ik_subset_fv_sst'[of "unlabel A"] vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel A"] by blast ultimately have False using reachable_constraints_occurs_fact_ik_funs_terms(4)[OF \_reach \ P P_occ] by blast thus ?thesis by simp qed (use 2(3) in simp) moreover have "\u u'. T = [u,u']" using iffD1[OF length_Suc_conv 2(2)] iffD1[OF length_Suc_conv[of _ 0]] length_0_conv by fast ultimately show ?thesis using Fun 2(1,2) by force qed qed lemma reachable_constraints_occurs_fact_ik_subst: assumes \_reach: "A \ reachable_constraints P" and \: "welltyped_constraint_model I A" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and t: "occurs t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I" shows "occurs t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" proof - have \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" using \ unfolding welltyped_constraint_model_def constraint_model_def by metis obtain s where s: "s \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "s \ I = occurs t" using t by auto hence u: "\u. s = occurs u" using \_wt reachable_constraints_occurs_fact_ik_subst_aux[OF \_reach \ P P_occ] by blast hence "fv s = {}" using reachable_constraints_occurs_fact_ik_ground[OF \_reach P P_occ] s by fast thus ?thesis using s u subst_ground_ident[of s I] by argo qed lemma reachable_constraints_occurs_fact_send_in_ik: assumes \_reach: "A \ reachable_constraints P" and \: "welltyped_constraint_model I A" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and x: "occurs (Var x) \ set ts" "send\ts\ \ set (unlabel A)" shows "occurs (I x) \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" using \_reach \ x proof (induction A rule: reachable_constraints.induct) case (step A T \ \ \) define \ where "\ \ \ \\<^sub>s \ \\<^sub>s \" define T' where "T' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" have T_adm: "admissible_transaction' T" using P step.hyps(2) unfolding list_all_iff by blast note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] note T_adm_occ = bspec[OF P_occ] have \_is_T_model: "strand_sem_stateful (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I) (set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I)) (unlabel T') I" using step.prems unlabel_append[of A T'] db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t[of "unlabel A" I "[]"] strand_sem_append_stateful[of "{}" "{}" "unlabel A" "unlabel T'" I] by (simp add: T'_def \_def welltyped_constraint_model_def constraint_model_def db\<^sub>s\<^sub>s\<^sub>t_def) show ?case proof (cases "send\ts\ \ set (unlabel A)") case False hence "send\ts\ \ set (unlabel T')" using step.prems(3) unfolding T'_def \_def by simp hence "receive\ts\ \ set (unlabel (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(2) unfolding T'_def by blast then obtain y ts' where y: "receive\ts'\ \ set (unlabel (transaction_receive T))" "\ y = Var x" "occurs (Var y) \ set ts'" using transaction_decl_fresh_renaming_substs_occurs_fact_send_receive(2)[ OF step.hyps(3-5) T_adm] subst_to_var_is_var[of _ \ x] step.prems(2) unfolding \_def by (metis eval_term.simps(1)) hence "occurs (Var y) \ \ \ set ts' \\<^sub>s\<^sub>e\<^sub>t \" "receive\ts' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\ \ set (unlabel (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using subst_lsst_unlabel_member[of "receive\ts'\" "transaction_receive T" \] by fastforce+ hence "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I \ occurs (Var y) \ \ \ I" using wellformed_transaction_sem_receives[ OF T_wf, of "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I" "set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I)" \ I "ts' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \"] \_is_T_model unfolding T'_def list_all_iff by fastforce hence *: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I \ occurs (\ y \ I)" by auto have "occurs (\ y \ I) \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" using deduct_occurs_in_ik[OF *] reachable_constraints_occurs_fact_ik_subst[ OF step.hyps(1) welltyped_constraint_model_prefix[OF step.prems(1)] P P_occ, of "\ y \ I"] reachable_constraints_occurs_fact_ik_funs_terms[ OF step.hyps(1) welltyped_constraint_model_prefix[OF step.prems(1)] P P_occ] by blast thus ?thesis using y(2) by simp qed (simp add: step.IH[OF welltyped_constraint_model_prefix[OF step.prems(1)]] step.prems(2)) qed simp lemma reachable_constraints_occurs_fact_deduct_in_ik: assumes \_reach: "A \ reachable_constraints P" and \: "welltyped_constraint_model I A" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and k: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I \ occurs k" shows "occurs k \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I" and "occurs k \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" using reachable_constraints_occurs_fact_ik_funs_terms(1-3)[OF \_reach \ P P_occ] reachable_constraints_occurs_fact_ik_subst[OF \_reach \ P P_occ] deduct_occurs_in_ik[OF k] by (presburger, presburger) lemma reachable_constraints_fv_bvars_subset: assumes A: "A \ reachable_constraints P" shows "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ (\T \ set P. bvars_transaction T)" using assms proof (induction A rule: reachable_constraints.induct) case (step \ T \ \ \) let ?T' = "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \" show ?case using step.IH step.hyps(2) bvars\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of ?T'] bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] bvars\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?T')"] unlabel_append[of \ "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?T'"] by (metis (no_types, lifting) SUP_upper Un_subset_iff) qed simp lemma reachable_constraints_fv_disj: fixes A::"('fun,'atom,'sets,'lbl) prot_constr" assumes A: "A \ reachable_constraints P" shows "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ (\T \ set P. bvars_transaction T) = {}" using A proof (induction A rule: reachable_constraints.induct) case (step \ T \ \ \) define T' where "T' \ transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \" define X where "X \ \T \ set P. bvars_transaction T" have "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ X = {}" using transaction_decl_fresh_renaming_substs_vars_disj(4)[OF step.hyps(3-5)] transaction_decl_fresh_renaming_substs_vars_subset(4)[OF step.hyps(3-5,2)] unfolding T'_def X_def by blast hence "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T') \ X = {}" using step.IH[unfolded X_def[symmetric]] fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of T'] by auto thus ?case unfolding T'_def X_def by blast qed simp (* TODO: this lemma subsumes reachable_constraints_fv_bvars_disj *) lemma reachable_constraints_fv_bvars_disj': fixes A::"('fun,'atom,'sets,'lbl) prot_constr" assumes P: "\T \ set P. wellformed_transaction T" and A: "A \ reachable_constraints P" shows "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = {}" using A proof (induction A rule: reachable_constraints.induct) case (step \ T \ \ \) define T' where "T' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" note 0 = transaction_decl_fresh_renaming_substs_vars_disj[OF step.hyps(3-5)] note 1 = transaction_decl_fresh_renaming_substs_vars_subset[OF step.hyps(3-5)] have 2: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' = {}" using 0(7) 1(4)[OF step.hyps(2)] fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq unfolding T'_def by (metis (no_types) disjoint_iff_not_equal subset_iff) have "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ \(bvars_transaction ` set P)" "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ \(bvars_transaction ` set P) = {}" using reachable_constraints_fv_bvars_subset[OF reachable_constraints.step[OF step.hyps]] reachable_constraints_fv_disj[OF reachable_constraints.step[OF step.hyps]] unfolding T'_def by auto hence 3: "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' = {}" by blast have "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \) \ bvars_transaction T = {}" using 0(4)[OF step.hyps(2)] 1(4)[OF step.hyps(2)] by blast hence 4: "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' = {}" by (metis (no_types) T'_def fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq bvars\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq unlabel_subst bvars\<^sub>s\<^sub>s\<^sub>t_subst) have "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') = {}" using 2 3 4 step.IH unfolding unlabel_append[of \ T'] fv\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel T'"] bvars\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel T'"] by fast thus ?case unfolding T'_def by blast qed simp lemma reachable_constraints_wf: assumes P: "\T \ set P. wellformed_transaction T" "\T \ set P. wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (trms_transaction T)" and A: "A \ reachable_constraints P" shows "wf\<^sub>s\<^sub>s\<^sub>t (unlabel A)" and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" proof - let ?X = "\T. fst ` set (transaction_decl T ()) \ set (transaction_fresh T)" have "wellformed_transaction T" when "T \ set P" for T using P(1) that by fast+ hence 0: "wf'\<^sub>s\<^sub>s\<^sub>t (?X T) (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)))" "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)) \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)) = {}" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms_transaction T)" when T: "T \ set P" for T unfolding admissible_transaction_terms_def by (metis T wellformed_transaction_wf\<^sub>s\<^sub>s\<^sub>t(1), metis T wellformed_transaction_wf\<^sub>s\<^sub>s\<^sub>t(2) fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq bvars\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq, metis T wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code P(2)) from A have "wf\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" proof (induction A rule: reachable_constraints.induct) case (step A T \ \ \) let ?T' = "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" have IH: "wf'\<^sub>s\<^sub>s\<^sub>t {} (unlabel A)" "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = {}" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" using step.IH by metis+ have 1: "wf'\<^sub>s\<^sub>s\<^sub>t {} (unlabel (A@?T'))" using transaction_decl_fresh_renaming_substs_wf_sst[OF 0(1)[OF step.hyps(2)] step.hyps(3-5)] wf\<^sub>s\<^sub>s\<^sub>t_vars_mono[of "{}"] wf\<^sub>s\<^sub>s\<^sub>t_append[OF IH(1)] by simp have 2: "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@?T') \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@?T') = {}" using reachable_constraints_fv_bvars_disj'[OF P(1)] reachable_constraints.step[OF step.hyps] by blast have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?T')" using trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq unlabel_subst wf_trms_subst[ OF transaction_decl_fresh_renaming_substs_range_wf_trms[OF step.hyps(3-5)], THEN wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_trms\<^sub>s\<^sub>s\<^sub>t_subst, OF 0(3)[OF step.hyps(2)]] by metis hence 3: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@?T'))" using IH(3) by auto show ?case using 1 2 3 by force qed simp thus "wf\<^sub>s\<^sub>s\<^sub>t (unlabel A)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" by metis+ qed lemma reachable_constraints_no_Ana_attack: assumes \: "\ \ reachable_constraints P" and P: "\T \ set P. wellformed_transaction T" "\T \ set P. admissible_transaction_terms T" "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" and t: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" shows "attack\n\ \ set (snd (Ana t))" proof - have T_adm_term: "admissible_transaction_terms T" when "T \ set P" for T using P that by blast have T_wf: "wellformed_transaction T" when "T \ set P" for T using P that by blast have T_fresh: "\x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" when "T \ set P" for T using P(3) that by fast show ?thesis using \ t proof (induction \ rule: reachable_constraints.induct) case (step A T \ \ \) thus ?case proof (cases "t \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)") case False hence "t \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)))" using step.prems by simp hence "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" using dual_transaction_ik_is_transaction_send'[OF T_wf[OF step.hyps(2)]] by metis hence "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \" using transaction_decl_fresh_renaming_substs_trms[ OF step.hyps(3-5), of "transaction_send T"] wellformed_transaction_unlabel_cases(4)[OF T_wf[OF step.hyps(2)]] by fastforce then obtain s where s: "s \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))" "t = s \ \ \\<^sub>s \ \\<^sub>s \" - by moura + by force hence s': "attack\n\ \ set (snd (Ana s))" using admissible_transaction_no_Ana_Attack[OF T_adm_term[OF step.hyps(2)]] trms_transaction_unfold[of T] by blast note * = transaction_decl_fresh_renaming_substs_range'(1-3)[OF step.hyps(3-5)] transaction_decl_fresh_renaming_substs_range_no_attack_const[ OF step.hyps(3-5) T_fresh[OF step.hyps(2)]] show ?thesis proof assume n: "attack\n\ \ set (snd (Ana t))" thus False proof (cases s) case (Var x) hence "(\c. t = Fun c []) \ (\y. t = Var y)" using *(1)[of t] n s(2) by (force simp del: subst_subst_compose) thus ?thesis using n Ana_subterm' by fastforce next case (Fun f S) hence "attack\n\ \ set (snd (Ana s)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \" using Ana_subst'[of f S _ "snd (Ana s)" "\ \\<^sub>s \ \\<^sub>s \"] s(2) s' n by (cases "Ana s") auto hence "attack\n\ \ set (snd (Ana s)) \ attack\n\ \ subst_range (\ \\<^sub>s \ \\<^sub>s \)" using const_mem_subst_cases' by fast thus ?thesis using *(4) s' by fast qed qed qed simp qed simp qed lemma reachable_constraints_receive_attack_if_attack: assumes \: "\ \ reachable_constraints P" and P: "\T \ set P. wellformed_transaction T" "\T \ set P. admissible_transaction_terms T" "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" "\T \ set P. \x \ vars_transaction T. \TAtom AttackType \ \\<^sub>v x" and \: "welltyped_constraint_model \ \" and l: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ attack\l\" shows "attack\l\ \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" and "receive\[attack\l\]\ \ set (unlabel \)" and "\T \ set P. \s \ set (transaction_strand T). is_Send (snd s) \ length (the_msgs (snd s)) = 1 \ is_Fun_Attack (hd (the_msgs (snd s))) \ the_Attack_label (the_Fun (hd (the_msgs (snd s)))) = fst s \ (l,receive\[attack\l\]\) \ set \" (is "?Q \ (l,receive\[attack\l\]\) \ set \") proof - have \': "constr_sem_stateful \ (unlabel \)" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using \ unfolding welltyped_constraint_model_def constraint_model_def by metis+ have 0: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \)" when \: "\ \ reachable_constraints P" for \ using reachable_constraints_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s[OF _ \] admissible_transaction_terms_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s P(2) ik\<^sub>s\<^sub>s\<^sub>t_trms\<^sub>s\<^sub>s\<^sub>t_subset[of "unlabel \"] wf_trms_subst[OF \'(3)] by fast have 1: "\x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \). \TAtom AttackType \ \\<^sub>v x" when \: "\ \ reachable_constraints P" for \ using reachable_constraints_vars_not_attack_typed[OF \ P(3,4)] fv_ik_subset_vars_sst'[of "unlabel \"] by fast have 2: "attack\l\ \ set (snd (Ana t)) \\<^sub>s\<^sub>e\<^sub>t \" when t: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" for t proof assume "attack\l\ \ set (snd (Ana t)) \\<^sub>s\<^sub>e\<^sub>t \" - then obtain s where s: "s \ set (snd (Ana t))" "s \ \ = attack\l\" by moura + then obtain s where s: "s \ set (snd (Ana t))" "s \ \ = attack\l\" by force obtain x where x: "s = Var x" by (cases s) (use s reachable_constraints_no_Ana_attack[OF \ P(1-3) t] in auto) have "x \ fv t" using x Ana_subterm'[OF s(1)] vars_iff_subtermeq by force hence "x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using t fv_subterms by fastforce hence "\\<^sub>v x \ TAtom AttackType" using 1[OF \] by fastforce thus False using s(2) x wt_subst_trm''[OF \'(4), of "Var x"] by fastforce qed have 3: "attack\l\ \ set (snd (Ana t))" when t: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \)" for t proof assume "attack\l\ \ set (snd (Ana t))" then obtain s where s: "s \ subterms\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" "attack\l\ \ set (snd (Ana s))" using Ana_subst_subterms_cases[OF t] 2 by fast - then obtain x where x: "x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "s \ \ x" by moura + then obtain x where x: "x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "s \ \ x" by force hence "\ x \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \)" using var_is_subterm[of x] subterms_subst_subset'[of \ "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] by force hence *: "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" "wf\<^sub>t\<^sub>r\<^sub>m s" using wf_trms_subterms[OF 0[OF \]] wf_trm_subtermeq[OF _ x(2)] by auto show False using term.order_trans[ OF subtermeq_imp_subtermtypeeq[OF *(2) Ana_subterm'[OF s(2)]] subtermeq_imp_subtermtypeeq[OF *(1) x(2)]] 1[OF \] x(1) wt_subst_trm''[OF \'(4), of "Var x"] by force qed have 4: "t = attack\n\" when t: "t \ \ \\<^sub>s \ \\<^sub>s \ = attack\n\" and hyps: "transaction_decl_subst \ T" "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and T: "\x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" for t n and T::"('fun, 'atom, 'sets, 'lbl) prot_transaction" and \ \ \::"('fun, 'atom, 'sets, 'lbl) prot_subst" and \::"('fun, 'atom, 'sets, 'lbl) prot_strand" proof (cases t) case (Var x) hence "attack\n\ \ subst_range (\ \\<^sub>s \ \\<^sub>s \)" by (metis (no_types, lifting) t eval_term.simps(1) subst_imgI term.distinct(1)) thus ?thesis using transaction_decl_fresh_renaming_substs_range_no_attack_const[OF hyps T] by blast qed (use t in simp) have 5: "\ts'. ts = ts' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \ (l,send\ts'\) \ set (transaction_strand T)" when ts: "(l,receive\ts\) \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" for l ts \ and T::"('fun, 'atom, 'sets, 'lbl) prot_transaction" using subst_lsst_memD(2)[OF ts[unfolded dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_steps_iff(1)[symmetric]]] by auto have 6: "l' = l" when "(l',receive\[attack\l\]\) \ set \" and Q: "?Q" for l' using \ that(1) proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \ \) show ?case proof (cases "(l',receive\[attack\l\]\) \ set \") case False hence *: "(l',receive\[attack\l\]\) \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" using step.prems by simp have "(l',send\[attack\l\]\) \ set (transaction_strand T)" using 4[OF _ step.hyps(3-5)] P(3) step.hyps(2) 5[OF *] by force thus ?thesis using Q step.hyps(2) unfolding is_Fun_Attack_def by fastforce qed (use step.IH in simp) qed simp have 7: "\t. ts = [t] \ t = attack\l\" when ts: "receive\ts\ \ set (unlabel \)" "attack\l\ \ set ts \\<^sub>s\<^sub>e\<^sub>t \" for ts using \ ts(1) proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \ \) obtain t where t: "t \ set ts" "attack\l\ = t \ \" using ts(2) by blast hence t_in_ik: "t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\ @ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" using step.prems(1) in_ik\<^sub>s\<^sub>s\<^sub>t_iff[of t] by blast have t_attack_eq: "t = attack\l\" proof (cases t) case (Var x) hence "TAtom AttackType \ subterms (\ t)" using t_in_ik 1[OF reachable_constraints.step[OF step.hyps]] by fastforce thus ?thesis using t(2) wt_subst_trm''[OF \'(4), of t] by force qed (use t(2) in simp) show ?case proof (cases "receive\ts\ \ set (unlabel \)") case False then obtain l' where l': "(l', receive\ts\) \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" using step.prems(1) unfolding unlabel_def by force then obtain ts' where ts': "ts = ts' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \" "(l', send\ts'\) \ set (transaction_strand T)" using 5 by meson then obtain t' where t': "t' \ set ts'" "t' \ \ \\<^sub>s \ \\<^sub>s \ = attack\l\" using t(1) t_attack_eq by force note * = t'(1) 4[OF t'(2) step.hyps(3-5)] have "send\ts'\ \ set (unlabel (transaction_strand T))" using ts'(2) step.hyps(2) P(2) unfolding unlabel_def by force hence "length ts' = 1" using step.hyps(2) P(2,3) * unfolding admissible_transaction_terms_def by fastforce hence "ts' = [attack\l\]" using * P(3) step.hyps(2) by (cases ts') auto thus ?thesis by (simp add: ts'(1)) qed (use step.IH in simp) qed simp show "attack\l\ \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" using private_const_deduct[OF _ l] 3 by simp then obtain ts where ts: "receive\ts\ \ set (unlabel \)" "attack\l\ \ set ts \\<^sub>s\<^sub>e\<^sub>t \" using in_ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t_iff[of _ \] unfolding unlabel_def by force then obtain t where "ts = [t]" "t = attack\l\" using 7 by blast thus "receive\[attack\l\]\ \ set (unlabel \)" using ts(1) by blast hence "\l'. (l', receive\[attack\l\]\) \ set \" unfolding unlabel_def by fastforce thus "(l,receive\[attack\l\]\) \ set \" when ?Q using that 6 by fast qed lemma reachable_constraints_receive_attack_if_attack': assumes \: "\ \ reachable_constraints P" and P: "\T \ set P. admissible_transaction' T" and \: "welltyped_constraint_model \ \" and n: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ attack\n\" shows "attack\n\ \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" and "receive\[attack\n\]\ \ set (unlabel \)" proof - have P': "\T \ set P. wellformed_transaction T" "\T \ set P. admissible_transaction_terms T" "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" "\T \ set P. \x \ vars_transaction T. \TAtom AttackType \ \\<^sub>v x" using admissible_transaction_is_wellformed_transaction(1,4)[OF bspec[OF P]] admissible_transactionE(2,15)[OF bspec[OF P]] by (blast, blast, blast, blast) show "attack\n\ \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" "receive\[attack\n\]\ \ set (unlabel \)" using reachable_constraints_receive_attack_if_attack(1,2)[OF \ P'(1,2) _ P'(4) \ n] P'(3) by (metis, metis) qed lemma constraint_model_Value_term_is_Val: assumes \_reach: "A \ reachable_constraints P" and \: "welltyped_constraint_model I A" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and x: "\\<^sub>v x = TAtom Value" "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" shows "\n. I x = Fun (Val n) []" using reachable_constraints_occurs_fact_send_ex[OF \_reach P P_occ x] reachable_constraints_occurs_fact_send_in_ik[OF \_reach \ P P_occ] reachable_constraints_occurs_fact_ik_case[OF \_reach P P_occ] by fast lemma constraint_model_Value_term_is_Val': assumes \_reach: "A \ reachable_constraints P" and \: "welltyped_constraint_model I A" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and x: "(TAtom Value, m) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" shows "\n. I (TAtom Value, m) = Fun (Val n) []" using constraint_model_Value_term_is_Val[OF \_reach \ P P_occ _ x] by simp (* We use this lemma to show that fresh constants first occur in \(\) at the point where they were generated *) lemma constraint_model_Value_var_in_constr_prefix: assumes \_reach: "\ \ reachable_constraints P" and \: "welltyped_constraint_model \ \" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" shows "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. \\<^sub>v x = TAtom Value \ (\B. prefix B \ \ x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t B \ \ x \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" (is "\x \ ?X \. ?R x \ ?Q x \") using \_reach \ proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \ \) let ?P = "\\. \x \ ?X \. ?R x \ ?Q x \" define T' where "T' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" have IH: "?P \" using step welltyped_constraint_model_prefix by fast note \_empty = admissible_transaction_decl_subst_empty[OF bspec[OF P step.hyps(2)] step.hyps(3)] have T_adm: "admissible_transaction' T" by (metis P step.hyps(2)) note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] have \_is_T_model: "strand_sem_stateful (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \) (set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)) (unlabel T') \" using step.prems unlabel_append[of \ T'] db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t[of "unlabel \" \ "[]"] strand_sem_append_stateful[of "{}" "{}" "unlabel \" "unlabel T'" \] by (simp add: T'_def welltyped_constraint_model_def constraint_model_def db\<^sub>s\<^sub>s\<^sub>t_def) have \_interp: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (metis \ welltyped_constraint_model_def constraint_model_def, metis \ welltyped_constraint_model_def, metis \ welltyped_constraint_model_def constraint_model_def) have 1: "?Q x \" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" "\\<^sub>v x = TAtom Value" for x proof - obtain n where n: "\ x = Fun n []" "is_Val n" "\public n" using constraint_model_Value_term_is_Val[ OF reachable_constraints.step[OF step.hyps] step.prems P P_occ x(2)] x(1) fv\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel T'"] unlabel_append[of \ T'] - unfolding T'_def by moura + unfolding T'_def by force have "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" using x(1) fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq unfolding T'_def by fastforce then obtain y where y: "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)" "x \ fv ((\ \\<^sub>s \ \\<^sub>s \) y)" using fv\<^sub>s\<^sub>s\<^sub>t_subst_obtain_var[of x "unlabel (transaction_strand T)" "\ \\<^sub>s \ \\<^sub>s \"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] by auto have y_x: "(\ \\<^sub>s \ \\<^sub>s \) y = Var x" and y_not_fresh: "y \ set (transaction_fresh T)" using y(2) transaction_decl_fresh_renaming_substs_range[OF step.hyps(3-5), of y] by (force, fastforce) have "\ ((\ \\<^sub>s \ \\<^sub>s \) y) = TAtom Value" using x(2) y_x by simp moreover have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \ \\<^sub>s \)" by (rule transaction_decl_fresh_renaming_substs_wt[OF step.hyps(3-5)]) ultimately have y_val: "\\<^sub>v y = TAtom Value" by (metis wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def \.simps(1)) have "Fun n [] = (\ \\<^sub>s \ \\<^sub>s \) y \ \" using n y_x by simp hence y_n: "Fun n [] = (\ \\<^sub>s \ \\<^sub>s \ \\<^sub>s \) y" by (metis subst_subst_compose[of "Var y" "\ \\<^sub>s \ \\<^sub>s \" \] eval_term.simps(1)) have \_ik_\_vals: "\x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \). \f. \ x = Fun f []" proof - have "\a. \ (\ x) = Var a" when "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" for x using that reachable_constraints_vars_TAtom_typed[OF step.hyps(1) P, of x] vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel \"] wt_subst_trm''[OF \_wt, of "Var x"] by force hence "\f. \ x = Fun f []" when "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" for x using that wf_trm_subst[OF \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s, of "Var x"] wf_trm_Var[of x] const_type_inv_wf empty_fv_exists_fun[OF interpretation_grounds[OF \_interp], of "Var x"] by (metis eval_term.simps(1)[of _ x \]) thus ?thesis using fv_ik_subset_fv_sst'[of "unlabel \"] vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel \"] by blast qed hence \_subterms_subst_cong: "subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \) = subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>e\<^sub>t \" by (metis ik\<^sub>s\<^sub>s\<^sub>t_subst[of "unlabel \" \] unlabel_subst[of \ \] subterms_subst_lsst_ik[of \ \]) have x_nin_\: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" proof - have "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" using x(1) fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq unfolding T'_def by fast hence "x \ fv\<^sub>s\<^sub>s\<^sub>t ((unlabel (transaction_strand T) \\<^sub>s\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>s\<^sub>t \)" using transaction_fresh_subst_grounds_domain[OF step.hyps(4)] step.hyps(4) labeled_stateful_strand_subst_comp[of \ "transaction_strand T" \] unlabel_subst[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" \] unlabel_subst[of "transaction_strand T" \] by (simp add: \_empty transaction_fresh_subst_def range_vars_alt_def) then obtain y where "\ y = Var x" using transaction_renaming_subst_var_obtain(1)[OF step.hyps(5)] by blast thus ?thesis using transaction_renaming_subst_range_notin_vars[OF step.hyps(5), of y] vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel \"] by auto qed from admissible_transaction_fv_in_receives_or_selects[OF T_adm y(1) y_not_fresh] have n_cases: "Fun n [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ (\z \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. \\<^sub>v z = TAtom Value \ \ z = Fun n [])" proof assume y_in: "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" then obtain ts where ts: "receive\ts\ \ set (unlabel (transaction_receive T))" "y \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" using admissible_transaction_strand_step_cases(1)[OF T_adm] by force hence ts_deduct: "list_all (\t. ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \ \\<^sub>s \ \\<^sub>s \ \ \) ts" using wellformed_transaction_sem_receives[ OF T_wf, of "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" "set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" "\ \\<^sub>s \ \\<^sub>s \" \ "ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \"] \_is_T_model subst_lsst_unlabel_member[of "receive\ts\" "transaction_receive T" "\ \\<^sub>s \ \\<^sub>s \"] unfolding T'_def list_all_iff by force obtain ty where ty: "ty \ set ts" "y \ fv ty" using ts(2) by fastforce have "Fun n [] \\<^sub>s\<^sub>e\<^sub>t ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ (\z \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \). \\<^sub>v z = TAtom Value \ \ z = Fun n [])" proof - have "Fun n [] \ ty \ \ \\<^sub>s \ \\<^sub>s \ \ \" using imageI[of "Var y" "subterms ty" "\x. x \ \ \\<^sub>s \ \\<^sub>s \ \\<^sub>s \"] var_is_subterm[OF ty(2)] subterms_subst_subset[of "\ \\<^sub>s \ \\<^sub>s \ \\<^sub>s \" ty] subst_subst_compose[of ty "\ \\<^sub>s \ \\<^sub>s \" \] y_n by (auto simp del: subst_subst_compose) hence "Fun n [] \\<^sub>s\<^sub>e\<^sub>t ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" using ty(1) private_fun_deduct_in_ik[of _ _ n "[]"] n(2,3) ts_deduct unfolding is_Val_def is_Abs_def list_all_iff by fast hence "Fun n [] \\<^sub>s\<^sub>e\<^sub>t ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ (\z \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \). \ z = Fun n [])" using const_subterm_subst_cases[of n _ \] \_ik_\_vals by fastforce thus ?thesis using \_wt n(2) unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def is_Val_def is_Abs_def by force qed thus ?thesis using fv_ik_subset_fv_sst' ik\<^sub>s\<^sub>s\<^sub>t_trms\<^sub>s\<^sub>s\<^sub>t_subset[of "unlabel \"] \_subterms_subst_cong by fast next assume y_in: "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) \ (\t s. select\t,s\ \ set (unlabel (transaction_checks T)) \ y \ fv t \ fv s)" then obtain s where s: "select\Var y,Fun (Set s) []\ \ set (unlabel (transaction_checks T))" using admissible_transaction_strand_step_cases(2)[OF T_adm] by force hence "select\(\ \\<^sub>s \ \\<^sub>s \) y, Fun (Set s) []\ \ set (unlabel (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" using subst_lsst_unlabel_member by fastforce hence n_in_db: "(Fun n [], Fun (Set s) []) \ set (db'\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ [])" using wellformed_transaction_sem_pos_checks(1)[ OF T_wf, of "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" "set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" "\ \\<^sub>s \ \\<^sub>s \" \ assign "(\ \\<^sub>s \ \\<^sub>s \) y" "Fun (Set s) []"] \_is_T_model n y_x unfolding T'_def db\<^sub>s\<^sub>s\<^sub>t_def by fastforce obtain tn sn where tsn: "insert\tn,sn\ \ set (unlabel \)" "Fun n [] = tn \ \" using db\<^sub>s\<^sub>s\<^sub>t_in_cases[OF n_in_db] by force have "Fun n [] = tn \ (\z. \\<^sub>v z = TAtom Value \ tn = Var z)" using \_wt tsn(2) n(2) unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def is_Val_def is_Abs_def by (cases tn) auto moreover have "tn \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "fv tn \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using tsn(1) in_subterms_Union by force+ ultimately show ?thesis using tsn(2) by auto qed from n_cases show ?thesis proof assume "\z \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. \\<^sub>v z = TAtom Value \ \ z = Fun n []" then obtain B where B: "prefix B \" "Fun n [] \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" by (metis IH n(1)) thus ?thesis using n x_nin_\ trms\<^sub>s\<^sub>s\<^sub>t_unlabel_prefix_subset(1)[of B] by (metis (no_types, opaque_lifting) self_append_conv subset_iff subterms\<^sub>s\<^sub>e\<^sub>t_mono prefix_def) qed (use n x_nin_\ in fastforce) qed have "?P (\@T')" proof (intro ballI impI) fix x assume x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T')" "\\<^sub>v x = TAtom Value" show "?Q x (\@T')" proof (cases "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \") case False hence "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" using x(1) unlabel_append[of \] fv\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \"] by simp then obtain B where B: "prefix B \" "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t B" "\ x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" - using x(2) 1 by moura + using x(2) 1 by blast thus ?thesis using prefix_prefix by fast qed (use x(2) IH prefix_prefix in fast) qed thus ?case unfolding T'_def by blast qed simp lemma constraint_model_Val_const_in_constr_prefix: assumes \_reach: "\ \ reachable_constraints P" and \: "welltyped_constraint_model \ \" and P: "\T \ set P. wellformed_transaction T" "\T \ set P. admissible_transaction_terms T" and n: "Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" shows "Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" proof - have *: "wf\<^sub>s\<^sub>s\<^sub>t (unlabel \)" "constr_sem_stateful \ (unlabel \)" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using reachable_constraints_wf(1)[OF P(1) _ \_reach] admissible_transaction_terms_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s \ P(2) n unfolding welltyped_constraint_model_def constraint_model_def wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code by fast+ show ?thesis using constraint_model_priv_const_in_constr_prefix[OF * _ _ n] by simp qed lemma constraint_model_Val_const_in_constr_prefix': assumes \_reach: "\ \ reachable_constraints P" and \: "welltyped_constraint_model \ \" and P: "\T \ set P. admissible_transaction' T" and n: "Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" shows "Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using constraint_model_Val_const_in_constr_prefix[OF \_reach \ _ _ n] P admissible_transaction_is_wellformed_transaction(1,4) by fast lemma constraint_model_Value_in_constr_prefix_fresh_action': fixes P::"('fun, 'atom, 'sets, 'lbl) prot_transaction list" assumes A: "A \ reachable_constraints P" and P: "\T \ set P. admissible_transaction_terms T" "\T \ set P. transaction_decl T () = []" "\T \ set P. bvars_transaction T = {}" and n: "Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" obtains B T \ \ \ where "prefix (B@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) A" and "B \ reachable_constraints P" "T \ set P" "transaction_decl_subst \ T" "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" and "Fun (Val n) [] \ subst_range \" proof - define f where "f \ \(T::('fun, 'atom, 'sets, 'lbl) prot_transaction, \::('fun, 'atom, 'sets, 'lbl) prot_subst, \::('fun, 'atom, 'sets, 'lbl) prot_subst, \::('fun, 'atom, 'sets, 'lbl) prot_subst). dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" define g where "g \ concat \ map f" obtain Ts where Ts: "A = g Ts" "\B. prefix B Ts \ g B \ reachable_constraints P" "\B T \ \ \. prefix (B@[(T,\,\,\)]) Ts \ T \ set P \ transaction_decl_subst \ T \ transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (g B)) \ transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (g B))" using reachable_constraints_as_transaction_lists[OF A] unfolding g_def f_def by blast obtain T \ \ \ where T: "(T, \, \, \) \ set Ts" "Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" using n trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq unlabel_subst unfolding Ts(1) g_def f_def unlabel_def trms\<^sub>s\<^sub>s\<^sub>t_def by fastforce obtain B where B: "prefix (B@[(T, \, \, \)]) Ts" "g B \ reachable_constraints P" "T \ set P" "transaction_decl_subst \ T" "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (g B))" "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (g B))" proof - obtain B where "\C. B@(T, \, \, \)#C = Ts" by (metis T(1) split_list) thus ?thesis using Ts(2-) that[of B] by auto qed note T_adm_terms = bspec[OF P(1) B(3)] note T_decl_empty = bspec[OF P(2) B(3)] note T_no_bvars = bspec[OF P(3) B(3)] note \_empty = admissible_transaction_decl_subst_empty'[OF T_decl_empty B(4)] have "trms\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T) \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \) = trms_transaction T \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \" using trms\<^sub>s\<^sub>s\<^sub>t_subst[of _ "\ \\<^sub>s \ \\<^sub>s \"] T_no_bvars by blast hence "Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms_transaction T \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \" by (metis T(2) unlabel_subst) hence "Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t subst_range (\ \\<^sub>s \ \\<^sub>s \)" using admissible_transaction_terms_no_Value_consts(1)[OF T_adm_terms] const_subterms_subst_cases'[of "Val n" "\ \\<^sub>s \ \\<^sub>s \" "trms_transaction T"] by blast then obtain tn where tn: "tn \ subst_range (\ \\<^sub>s \ \\<^sub>s \)" "Fun (Val n) [] \ tn" "is_Fun tn" by fastforce have "Fun (Val n) [] \ subst_range \" using tn(1-) transaction_decl_fresh_renaming_substs_range'(2,4)[OF B(4-6) tn(1) \_empty] by fastforce moreover have "prefix (g B@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) A" using Ts(1) B(1) unfolding g_def f_def prefix_def by fastforce ultimately show thesis using that B(2-) by blast qed lemma constraint_model_Value_in_constr_prefix_fresh_action: fixes P::"('fun, 'atom, 'sets, 'lbl) prot_transaction list" assumes A: "A \ reachable_constraints P" and P_adm: "\T \ set P. admissible_transaction' T" and n: "Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" obtains B T \ \ \ where "prefix (B@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) A" and "B \ reachable_constraints P" "T \ set P" "transaction_decl_subst \ T" "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" and "Fun (Val n) [] \ subst_range \" proof - have P: "\T \ set P. admissible_transaction_terms T" "\T \ set P. transaction_decl T () = []" "\T \ set P. bvars_transaction T = {}" using P_adm admissible_transactionE(1) admissible_transaction_no_bvars(2) admissible_transaction_is_wellformed_transaction(4) by (blast,blast,blast) show ?thesis using that constraint_model_Value_in_constr_prefix_fresh_action'[OF A P n] by blast qed lemma reachable_constraints_occurs_fact_ik_case': fixes A::"('fun,'atom,'sets,'lbl) prot_constr" assumes \_reach: "A \ reachable_constraints P" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and val: "Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" shows "occurs (Fun (Val n) []) \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" proof - obtain B T \ \ \ where B: "prefix (B@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) A" "B \ reachable_constraints P" "T \ set P" "transaction_decl_subst \ T" "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" "Fun (Val n) [] \ subst_range \" using constraint_model_Value_in_constr_prefix_fresh_action[OF \_reach P val] by blast define \ where "\ \ \ \\<^sub>s \ \\<^sub>s \" have T_adm: "admissible_transaction' T" using P B(3) by blast hence T_wf: "wellformed_transaction T" "admissible_transaction_occurs_checks T" using admissible_transaction_is_wellformed_transaction(1) bspec[OF P_occ B(3)] by (blast,blast) obtain x where x: "x \ set (transaction_fresh T)" "\ x = Fun (Val n) []" using transaction_fresh_subst_domain[OF B(5)] B(7) admissible_transaction_decl_subst_empty[OF T_adm B(4)] by (force simp add: subst_compose \_def) obtain ts where ts: "send\ts\ \ set (unlabel (transaction_send T))" "occurs (Var x) \ set ts" using admissible_transaction_occurs_checksE2[OF T_wf(2) x(1)] by (metis (mono_tags, lifting) list.set_intros(1) unlabel_Cons(1)) have "occurs (Var x) \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" using ts by force hence "occurs (Var x) \ \ \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using dual_transaction_ik_is_transaction_send'[OF T_wf(1), of \] by fast hence "occurs (Fun (Val n) []) \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using x(2) by simp thus ?thesis using B(1)[unfolded \_def[symmetric]] unlabel_append[of B "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)"] ik\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel B" "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))"] unfolding prefix_def by force qed lemma reachable_constraints_occurs_fact_ik_case'': fixes A::"('fun,'atom,'sets,'lbl) prot_constr" assumes \_reach: "A \ reachable_constraints P" and \: "welltyped_constraint_model \ A" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and val: "Fun (Val n) [] \ t" "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \ \ t" shows "occurs (Fun (Val n) []) \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" proof - obtain f ts where t: "t = Fun f ts" using val(1) by (cases t) simp_all show ?thesis using private_fun_deduct_in_ik[OF val(2,1)[unfolded t]] constraint_model_Val_const_in_constr_prefix'[OF \_reach \ P, of n] reachable_constraints_occurs_fact_ik_case'[OF \_reach P P_occ, of n] by fastforce qed lemma admissible_transaction_occurs_checks_prop: assumes \_reach: "\ \ reachable_constraints P" and \: "welltyped_constraint_model \ \" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and f: "f \ \(funs_term ` (\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" shows "\is_PubConstValue f" and "\is_Abs f" proof - - obtain x where x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" "f \ funs_term (\ x)" using f by moura - obtain T where T: "Fun f T \ \ x" using funs_term_Fun_subterm[OF x(2)] by moura + obtain x where x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" "f \ funs_term (\ x)" using f by force + obtain T where T: "Fun f T \ \ x" using funs_term_Fun_subterm[OF x(2)] by force have \_interp: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (metis \ welltyped_constraint_model_def constraint_model_def, metis \ welltyped_constraint_model_def, metis \ welltyped_constraint_model_def constraint_model_def) note 0 = x(1) reachable_constraints_vars_TAtom_typed[OF \_reach P, of x] vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel \"] have 1: "\ (Var x) = \ (\ x)" using wt_subst_trm''[OF \_wt, of "Var x"] by simp hence "\a. \ (\ x) = Var a" using 0 by force hence "\f. \ x = Fun f []" using x(1) wf_trm_subst[OF \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s, of "Var x"] wf_trm_Var[of x] const_type_inv_wf empty_fv_exists_fun[OF interpretation_grounds[OF \_interp], of "Var x"] by (metis eval_term.simps(1)[of _ x \]) hence 2: "\ x = Fun f []" using x(2) by force have 3: "\\<^sub>v x \ TAtom AbsValue" using 0 by force have "\is_PubConstValue f \ \is_Abs f" proof (cases "\\<^sub>v x = TAtom Value") case True then obtain B where B: "prefix B \" "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t B" "\ x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" using constraint_model_Value_var_in_constr_prefix[OF \_reach \ P P_occ] x(1) by fast have "\ x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using B(1,3) trms\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel B"] unlabel_append[of B] unfolding prefix_def by auto hence "f \ \(funs_term ` trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using x(2) funs_term_subterms_eq(2)[of "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] by blast thus ?thesis using reachable_constraints_val_funs_private[OF \_reach P] by blast+ next case False thus ?thesis using x 1 2 3 unfolding is_PubConstValue_def by (cases f) auto qed thus "\is_PubConstValue f" "\is_Abs f" by metis+ qed lemma admissible_transaction_occurs_checks_prop': assumes \_reach: "\ \ reachable_constraints P" and \: "welltyped_constraint_model \ \" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and f: "f \ \(funs_term ` (\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" shows "\n. f = PubConst Value n" and "\n. f = Abs n" using admissible_transaction_occurs_checks_prop[OF \_reach \ P P_occ f] unfolding is_PubConstValue_def by auto lemma transaction_var_becomes_Val: assumes \_reach: "\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \) \ reachable_constraints P" and \: "welltyped_constraint_model \ (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" and \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and T: "T \ set P" and x: "x \ fv_transaction T" "fst x = TAtom Value" shows "\n. Fun (Val n) [] = (\ \\<^sub>s \ \\<^sub>s \) x \ \" proof - obtain m where m: "x = (TAtom Value, m)" by (metis x(2) eq_fst_iff) note \_empty = admissible_transaction_decl_subst_empty[OF bspec[OF P T] \] have x_not_bvar: "x \ bvars_transaction T" "fv ((\ \\<^sub>s \ \\<^sub>s \) x) \ bvars_transaction T = {}" using x(1) admissible_transactions_fv_bvars_disj[OF P] T transaction_decl_fresh_renaming_substs_vars_disj(2)[OF \ \ \, of x] vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_strand T)"] by (blast, blast) have \x_type: "\ (\ x) = TAtom Value" using \ x \\<^sub>v_TAtom''(2)[of x] wt_subst_trm''[of \ "Var x"] unfolding transaction_fresh_subst_def by simp show ?thesis proof (cases "x \ subst_domain \") case True then obtain c where c: "\ x = Fun c []" "\public c" "arity c = 0" using \ unfolding transaction_fresh_subst_def by fastforce then obtain n where n: "c = Val n" using \x_type by (cases c) (auto split: option.splits) show ?thesis using c n subst_compose[of \ \ x] \_empty by simp next case False hence "\ x = Var x" by auto then obtain n where n: "(\ \\<^sub>s \) x = Var (TAtom Value, n)" using m transaction_renaming_subst_is_renaming(1)[OF \] subst_compose[of \ \ x] by force hence "(TAtom Value, n) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" using x_not_bvar fv\<^sub>s\<^sub>s\<^sub>t_subst_fv_subset[OF x(1), of "\ \\<^sub>s \ \\<^sub>s \"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] \_empty by force hence "\n'. \ (TAtom Value, n) = Fun (Val n') []" using constraint_model_Value_term_is_Val'[OF \_reach \ P P_occ, of n] x fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \"] fv\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \"] unlabel_append[of \] by fastforce thus ?thesis using n \_empty by simp qed qed lemma reachable_constraints_SMP_subset: assumes \: "\ \ reachable_constraints P" shows "SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ SMP (\T \ set P. trms_transaction T)" (is "?A \") and "SMP (pair`setops\<^sub>s\<^sub>s\<^sub>t (unlabel \)) \ SMP (\T\set P. pair`setops_transaction T)" (is "?B \") proof - have "?A \ \ ?B \" using \ proof (induction \ rule: reachable_constraints.induct) case (step A T \ \ \) define T' where "T' \ transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \" define M where "M \ \T \ set P. trms_transaction T" define N where "N \ \T \ set P. pair ` setops_transaction T" let ?P = "\t. \s \. s \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = s \ \" let ?Q = "\t. \s \. s \ N \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = s \ \" have IH: "SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \ SMP M" "SMP (pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)) \ SMP N" using step.IH by (metis M_def, metis N_def) note \\\_wt = transaction_decl_fresh_renaming_substs_wt[OF step.hyps(3-5)] note \\\_wf = transaction_decl_fresh_renaming_substs_range_wf_trms[OF step.hyps(3-5)] have 0: "SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')) = SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \ SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" "SMP (pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'))) = SMP (pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)) \ SMP (pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel T'))" using trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of T'] setops\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of T'] trms\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel A" "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')"] setops\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel A" "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')"] unlabel_append[of A "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'"] image_Un[of pair "setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" "setops\<^sub>s\<^sub>s\<^sub>t (unlabel T')"] SMP_union[of "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'"] SMP_union[of "pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" "pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel T')"] by argo+ have 1: "SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T') \ SMP M" proof (intro SMP_subset_I ballI) fix t show "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ ?P t" using trms\<^sub>s\<^sub>s\<^sub>t_wt_subst_ex[OF \\\_wt \\\_wf, of t "unlabel (transaction_strand T)"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] step.hyps(2) unfolding T'_def M_def by auto qed have 2: "SMP (pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel T')) \ SMP N" proof (intro SMP_subset_I ballI) fix t show "t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel T') \ ?Q t" using setops\<^sub>s\<^sub>s\<^sub>t_wt_subst_ex[OF \\\_wt \\\_wf, of t "unlabel (transaction_strand T)"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] step.hyps(2) unfolding T'_def N_def by auto qed have "SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')) \ SMP M" "SMP (pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'))) \ SMP N" using 0 1 2 IH by blast+ thus ?case unfolding T'_def M_def N_def by blast qed (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) thus "?A \" "?B \" by metis+ qed lemma reachable_constraints_no_Pair_fun': assumes A: "A \ reachable_constraints P" and P: "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" "\T \ set P. transaction_decl T () = []" "\T \ set P. admissible_transaction_terms T" "\T \ set P. \x \ vars_transaction T. \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = \a\\<^sub>\\<^sub>a)" shows "Pair \ \(funs_term ` SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))" using A proof (induction A rule: reachable_constraints.induct) case (step A T \ \ \) define T' where "T' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" note T_fresh_type = bspec[OF P(1) step.hyps(2)] note \_empty = admissible_transaction_decl_subst_empty'[OF bspec[OF P(2) step.hyps(2)] step.hyps(3)] note T_adm_terms = bspec[OF P(3) step.hyps(2)] note \\\_wt = transaction_decl_fresh_renaming_substs_wt[OF step.hyps(3-5)] note \\\_wf = transaction_decl_fresh_renaming_substs_range_wf_trms[OF step.hyps(3-5)] have 0: "SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@T')) = SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \ SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" using SMP_union[of "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'"] unlabel_append[of A T'] trms\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel A" "unlabel T'"] by simp have 1: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" using reachable_constraints_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s[OF _ reachable_constraints.step[OF step.hyps]] admissible_transaction_terms_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s P(3) trms\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel A"] unlabel_append[of A] unfolding T'_def by force have 2: "Pair \ \(funs_term ` (subst_range (\ \\<^sub>s \ \\<^sub>s \)))" using transaction_decl_fresh_renaming_substs_range'(3)[ OF step.hyps(3-5) _ \_empty T_fresh_type] by force have "Pair \ \(funs_term ` (trms_transaction T))" using T_adm_terms unfolding admissible_transaction_terms_def by blast hence "Pair \ funs_term t" when t: "t \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T) \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" for t using 2 trms\<^sub>s\<^sub>s\<^sub>t_funs_term_cases[OF t] by force hence 3: "Pair \ funs_term t" when t: "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" for t using t unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \"] unfolding T'_def by metis have "\a. \\<^sub>v x = TAtom a" when "x \ vars_transaction T" for x using that protocol_transaction_vars_TAtom_typed(1) bspec[OF P(4) step.hyps(2)] by fast hence "\a. \\<^sub>v x = TAtom a" when "x \ vars\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T) \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" for x using wt_subst_fv\<^sub>s\<^sub>e\<^sub>t_termtype_subterm[OF _ \\\_wt \\\_wf, of x "vars_transaction T"] vars\<^sub>s\<^sub>s\<^sub>t_subst_cases[OF that] by fastforce hence "\a. \\<^sub>v x = TAtom a" when "x \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" for x using that unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] vars\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \"] unfolding T'_def by simp hence "\a. \\<^sub>v x = TAtom a" when "x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" for x using that fv_trms\<^sub>s\<^sub>s\<^sub>t_subset(1) by fast hence "Pair \ funs_term (\ (Var x))" when "x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" for x using that by fastforce moreover have "Pair \ funs_term s" when s: "Ana s = (K, M)" "Pair \ \(funs_term ` set K)" for s::"('fun,'atom,'sets,'lbl) prot_term" and K M proof (cases s) case (Fun f S) thus ?thesis using s Ana_Fu_keys_not_pairs[of _ S K M] by (cases f) force+ qed (use s in simp) ultimately have "Pair \ funs_term t" when t: "t \ SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" for t using t 3 SMP_funs_term[OF t _ _ 1, of Pair] funs_term_type_iff by fastforce thus ?case using 0 step.IH(1) unfolding T'_def by blast qed simp lemma reachable_constraints_no_Pair_fun: assumes A: "A \ reachable_constraints P" and P: "\T \ set P. admissible_transaction' T" shows "Pair \ \(funs_term ` SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))" using reachable_constraints_no_Pair_fun'[OF A] P admissible_transactionE(1,2,3) admissible_transaction_is_wellformed_transaction(4) by blast lemma reachable_constraints_setops_form: assumes A: "A \ reachable_constraints P" and P: "\T \ set P. admissible_transaction' T" and t: "t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" shows "\c s. t = pair (c, Fun (Set s) []) \ \ c = TAtom Value" using A t proof (induction A rule: reachable_constraints.induct) case (step A T \ \ \) have T_adm: "admissible_transaction' T" when "T \ set P" for T using P that unfolding list_all_iff by simp note T_adm' = admissible_transaction_is_wellformed_transaction(2,3)[OF T_adm] note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] note \\\_wt = transaction_decl_fresh_renaming_substs_wt[OF step.hyps(3-5)] note \\\_wf = transaction_decl_fresh_renaming_substs_range_wf_trms[OF step.hyps(3-5)] show ?case using step.IH proof (cases "t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)") case False hence "t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T) \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" using step.prems setops\<^sub>s\<^sub>s\<^sub>t_append unlabel_append setops\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] by fastforce then obtain t' \ where t': "t' \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T))" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "t = t' \ \" using setops\<^sub>s\<^sub>s\<^sub>t_wt_subst_ex[OF \\\_wt \\\_wf] by blast then obtain s s' where s: "t' = pair (s,s')" using setops\<^sub>s\<^sub>s\<^sub>t_are_pairs by fastforce moreover have "InSet ac s s' = InSet Assign s s' \ InSet ac s s' = InSet Check s s'" for ac by (cases ac) simp_all ultimately have "\n. s = Var (Var Value, n)" "\u. s' = Fun (Set u) []" using t'(1) setops\<^sub>s\<^sub>s\<^sub>t_member_iff[of s s' "unlabel (transaction_strand T)"] pair_in_pair_image_iff[of s s'] transaction_inserts_are_Value_vars[ OF T_wf[OF step.hyps(2)] T_adm'(2)[OF step.hyps(2)], of s s'] transaction_deletes_are_Value_vars[ OF T_wf[OF step.hyps(2)] T_adm'(2)[OF step.hyps(2)], of s s'] transaction_selects_are_Value_vars[ OF T_wf[OF step.hyps(2)] T_adm'(1)[OF step.hyps(2)], of s s'] transaction_inset_checks_are_Value_vars[ OF T_adm[OF step.hyps(2)], of s s'] transaction_notinset_checks_are_Value_vars[ OF T_adm[OF step.hyps(2)], of _ _ _ s s'] by metis+ then obtain ss n where ss: "t = pair (\ (Var Value, n), Fun (Set ss) [])" using t'(4) s unfolding pair_def by force have "\ (\ (Var Value, n)) = TAtom Value" "wf\<^sub>t\<^sub>r\<^sub>m (\ (Var Value, n))" using t'(2) wt_subst_trm''[OF t'(2), of "Var (Var Value, n)"] apply simp using t'(3) by (cases "(Var Value, n) \ subst_domain \") auto thus ?thesis using ss by blast qed simp qed (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) lemma reachable_constraints_insert_delete_form: assumes A: "A \ reachable_constraints P" and P: "\T \ set P. admissible_transaction' T" and t: "insert\t,s\ \ set (unlabel A) \ delete\t,s\ \ set (unlabel A)" (is "?Q t s A") shows "\k. s = Fun (Set k) []" (is ?A) and "\ t = TAtom Value" (is ?B) and "(\x. t = Var x) \ (\n. t = Fun (Val n) [])" (is ?C) proof - have 0: "pair (t,s) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" using t unfolding setops\<^sub>s\<^sub>s\<^sub>t_def by force show 1: ?A ?B using reachable_constraints_setops_form[OF A P 0] by (fast,fast) show ?C using A t proof (induction A rule: reachable_constraints.induct) case (step \ T \ \ \) let ?T' = "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \" note T_adm = bspec[OF P step.hyps(2)] note T_wf = admissible_transaction_is_wellformed_transaction(1,3)[OF T_adm] have "?Q t s \ \ ?Q t s ?T'" using step.prems dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(4,5)[of t s ?T'] unfolding unlabel_append by auto thus ?case proof assume "?Q t s ?T'" then obtain u v where u: "?Q u v (transaction_strand T)" "t = u \ \ \\<^sub>s \ \\<^sub>s \" by (metis (no_types, lifting) stateful_strand_step_mem_substD(4,5) unlabel_subst) obtain x where x: "u = Var x" using u(1) transaction_inserts_are_Value_vars(1)[OF T_wf, of u v] transaction_deletes_are_Value_vars(1)[OF T_wf, of u v] by fastforce show ?case using u(2) x transaction_decl_fresh_renaming_substs_range'(3)[ OF step.hyps(3,4,5) _ admissible_transaction_decl_subst_empty[OF T_adm step.hyps(3)] admissible_transactionE(2)[OF T_adm], of t] by (cases "t \ subst_range (\ \\<^sub>s \ \\<^sub>s \)") (blast, metis eval_term.simps(1) subst_imgI) qed (use step.IH in fastforce) qed simp qed lemma reachable_constraints_setops_type: fixes t::"('fun,'atom,'sets,'lbl) prot_term" assumes A: "A \ reachable_constraints P" and P: "\T \ set P. admissible_transaction' T" and t: "t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" shows "\ t = TComp Pair [TAtom Value, TAtom SetType]" proof - obtain s c where s: "t = pair (c, Fun (Set s) [])" "\ c = TAtom Value" - using reachable_constraints_setops_form[OF A P t] by moura + using reachable_constraints_setops_form[OF A P t] by force hence "(Fun (Set s) []::('fun,'atom,'sets,'lbl) prot_term) \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" using t setops\<^sub>s\<^sub>s\<^sub>t_member_iff[of c "Fun (Set s) []" "unlabel A"] by force hence "wf\<^sub>t\<^sub>r\<^sub>m (Fun (Set s) []::('fun,'atom,'sets,'lbl) prot_term)" using reachable_constraints_wf(2) P A admissible_transaction_is_wellformed_transaction(1,4) unfolding admissible_transaction_terms_def by blast hence "arity (Set s) = 0" unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by simp thus ?thesis using s unfolding pair_def by fastforce qed lemma reachable_constraints_setops_same_type_if_unifiable: assumes A: "A \ reachable_constraints P" and P: "\T \ set P. admissible_transaction' T" shows "\s \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A). \t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A). (\\. Unifier \ s t) \ \ s = \ t" (is "?P A") using reachable_constraints_setops_type[OF A P] by simp lemma reachable_constraints_setops_unifiable_if_wt_instance_unifiable: assumes A: "A \ reachable_constraints P" and P: "\T \ set P. admissible_transaction' T" shows "\s \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A). \t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A). (\\ \ \. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ Unifier \ (s \ \) (t \ \)) \ (\\. Unifier \ s t)" proof (intro ballI impI) fix s t assume st: "s \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" "t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" and "\\ \ \. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ Unifier \ (s \ \) (t \ \)" then obtain \ \ \ where \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "Unifier \ (s \ \) (t \ \)" - by moura + by force obtain fs ft cs ct where c: "s = pair (cs, Fun (Set fs) [])" "t = pair (ct, Fun (Set ft) [])" "\ cs = TAtom Value" "\ ct = TAtom Value" using reachable_constraints_setops_form[OF A P st(1)] reachable_constraints_setops_form[OF A P st(2)] - by moura + by force have "cs \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" "ct \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" using c(1,2) setops_subterm_trms[OF st(1), of cs] setops_subterm_trms[OF st(2), of ct] Fun_param_is_subterm[of cs "args s"] Fun_param_is_subterm[of ct "args t"] unfolding pair_def by simp_all moreover have "\T \ set P. wellformed_transaction T" "\T \ set P. wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (trms_transaction T)" using P admissible_transaction_is_wellformed_transaction(1,4) unfolding admissible_transaction_terms_def by fast+ ultimately have *: "wf\<^sub>t\<^sub>r\<^sub>m cs" "wf\<^sub>t\<^sub>r\<^sub>m ct" using reachable_constraints_wf(2)[OF _ _ A] wf_trms_subterms by blast+ have "(\x. cs = Var x) \ (\c d. cs = Fun c [])" using const_type_inv_wf c(3) *(1) by (cases cs) auto moreover have "(\x. ct = Var x) \ (\c d. ct = Fun c [])" using const_type_inv_wf c(4) *(2) by (cases ct) auto ultimately show "\\. Unifier \ s t" using reachable_constraints_setops_form[OF A P] reachable_constraints_setops_type[OF A P] st \ c unfolding pair_def by auto qed lemma reachable_constraints_tfr: assumes M: "M \ \T \ set P. trms_transaction T" "has_all_wt_instances_of \ M N" "finite N" "tfr\<^sub>s\<^sub>e\<^sub>t N" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s N" and P: "\T \ set P. admissible_transaction T" "\T \ set P. list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (transaction_strand T))" and \: "\ \ reachable_constraints P" shows "tfr\<^sub>s\<^sub>s\<^sub>t (unlabel \)" using \ proof (induction \ rule: reachable_constraints.induct) case (step A T \ \ \) define T' where "T' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" have P': "\T \ set P. admissible_transaction' T" using P(1) admissible_transactionE'(1) by blast have AT'_reach: "A@T' \ reachable_constraints P" using reachable_constraints.step[OF step.hyps] unfolding T'_def by metis note T_adm = bspec[OF P(1) step.hyps(2)] note T_adm' = bspec[OF P'(1) step.hyps(2)] note \_empty = admissible_transaction_decl_subst_empty[OF T_adm' step.hyps(3)] note \\\_wt = transaction_decl_fresh_renaming_substs_wt[OF step.hyps(3-5)] note \\\_wf = transaction_decl_fresh_renaming_substs_range_wf_trms[OF step.hyps(3-5)] have \\\_bvars_disj: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T) \ range_vars (\ \\<^sub>s \ \\<^sub>s \) = {}" using transaction_decl_fresh_renaming_substs_vars_disj(4)[OF step.hyps(3,4,5,2)] \_empty by simp have wf_trms_M: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" using admissible_transactions_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s P'(1) unfolding M(1) by blast have "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@T'))" using reachable_constraints_SMP_subset(1)[OF AT'_reach] tfr_subset(3)[OF M(4), of "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@T')"] SMP_SMP_subset[of M N] SMP_I'[OF wf_trms_M M(5,2)] unfolding M(1) by blast moreover have "\p. Ana (pair p) = ([],[])" unfolding pair_def by auto ultimately have 1: "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@T') \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel (A@T')))" using tfr_setops_if_tfr_trms[of "unlabel (A@T')"] reachable_constraints_no_Pair_fun[OF AT'_reach P'] reachable_constraints_setops_same_type_if_unifiable[OF AT'_reach P'] reachable_constraints_setops_unifiable_if_wt_instance_unifiable[OF AT'_reach P'] by blast have "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (transaction_strand T))" using step.hyps(2) P(2) tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_comp_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p unfolding comp_tfr\<^sub>s\<^sub>s\<^sub>t_def tfr\<^sub>s\<^sub>s\<^sub>t_def by fastforce hence "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel T')" using tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p_all_wt_subst_apply[OF _ \\\_wt \\\_wf \\\_bvars_disj] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] unfolding T'_def by argo hence 2: "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (A@T'))" using step.IH unlabel_append unfolding tfr\<^sub>s\<^sub>s\<^sub>t_def by auto have "tfr\<^sub>s\<^sub>s\<^sub>t (unlabel (A@T'))" using 1 2 by (metis tfr\<^sub>s\<^sub>s\<^sub>t_def) thus ?case by (metis T'_def) qed simp lemma reachable_constraints_tfr': assumes M: "M \ \T \ set P. trms_transaction T \ pair' Pair ` setops_transaction T" "has_all_wt_instances_of \ M N" "finite N" "tfr\<^sub>s\<^sub>e\<^sub>t N" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s N" and P: "\T \ set P. wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (trms_transaction T)" "\T \ set P. list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (transaction_strand T))" and \: "\ \ reachable_constraints P" shows "tfr\<^sub>s\<^sub>s\<^sub>t (unlabel \)" using \ proof (induction \ rule: reachable_constraints.induct) case (step A T \ \ \) define T' where "T' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" have AT'_reach: "A@T' \ reachable_constraints P" using reachable_constraints.step[OF step.hyps] unfolding T'_def by metis note \\\_wt = transaction_decl_fresh_renaming_substs_wt[OF step.hyps(3-5)] note \\\_wf = transaction_decl_fresh_renaming_substs_range_wf_trms[OF step.hyps(3-5)] have \\\_bvars_disj: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T) \ range_vars (\ \\<^sub>s \ \\<^sub>s \) = {}" by (rule transaction_decl_fresh_renaming_substs_vars_disj(4)[OF step.hyps(3,4,5,2)]) have wf_trms_M: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" using P(1) setops\<^sub>s\<^sub>s\<^sub>t_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s(2) unfolding M(1) pair_code wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code[symmetric] by fast have "SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@T')) \ SMP M" "SMP (pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel (A@T'))) \ SMP M" using reachable_constraints_SMP_subset[OF AT'_reach] SMP_mono[of "\T \ set P. trms_transaction T" M] SMP_mono[of "\T \ set P. pair ` setops_transaction T" M] unfolding M(1) pair_code[symmetric] by blast+ hence 1: "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@T') \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel (A@T')))" using tfr_subset(3)[OF M(4), of "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@T') \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel (A@T'))"] SMP_union[of "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@T')" "pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel (A@T'))"] SMP_SMP_subset[of M N] SMP_I'[OF wf_trms_M M(5,2)] by blast have "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (transaction_strand T))" using step.hyps(2) P(2) tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_comp_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p unfolding comp_tfr\<^sub>s\<^sub>s\<^sub>t_def tfr\<^sub>s\<^sub>s\<^sub>t_def by fastforce hence "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel T')" using tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p_all_wt_subst_apply[OF _ \\\_wt \\\_wf \\\_bvars_disj] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] unfolding T'_def by argo hence 2: "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (A@T'))" using step.IH unlabel_append unfolding tfr\<^sub>s\<^sub>s\<^sub>t_def by auto have "tfr\<^sub>s\<^sub>s\<^sub>t (unlabel (A@T'))" using 1 2 by (metis tfr\<^sub>s\<^sub>s\<^sub>t_def) thus ?case by (metis T'_def) qed simp lemma reachable_constraints_typing_cond\<^sub>s\<^sub>s\<^sub>t: assumes M: "M \ \T \ set P. trms_transaction T \ pair' Pair ` setops_transaction T" "has_all_wt_instances_of \ M N" "finite N" "tfr\<^sub>s\<^sub>e\<^sub>t N" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s N" and P: "\T \ set P. wellformed_transaction T" "\T \ set P. wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (trms_transaction T)" "\T \ set P. list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (transaction_strand T))" and \: "\ \ reachable_constraints P" shows "typing_cond\<^sub>s\<^sub>s\<^sub>t (unlabel \)" using reachable_constraints_wf[OF P(1,2) \] reachable_constraints_tfr'[OF M P(2,3) \] unfolding typing_cond\<^sub>s\<^sub>s\<^sub>t_def by blast context begin private lemma reachable_constraints_typing_result_aux: assumes 0: "wf\<^sub>s\<^sub>s\<^sub>t (unlabel \)" "tfr\<^sub>s\<^sub>s\<^sub>t (unlabel \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" shows "wf\<^sub>s\<^sub>s\<^sub>t (unlabel (\@[(l,send\[attack\n\]\)]))" "tfr\<^sub>s\<^sub>s\<^sub>t (unlabel (\@[(l,send\[attack\n\]\)]))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@[(l,send\[attack\n\]\)]))" proof - let ?n = "[(l,send\[attack\n\]\)]" let ?A = "\@?n" show "wf\<^sub>s\<^sub>s\<^sub>t (unlabel ?A)" using 0(1) wf\<^sub>s\<^sub>s\<^sub>t_append_suffix'[of "{}" "unlabel \" "unlabel ?n"] unlabel_append[of \ ?n] by simp show "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?A)" using 0(3) trms\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel ?n"] unlabel_append[of \ ?n] by fastforce have "\t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?n \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel ?n). \c. t = Fun c []" "\t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?n \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel ?n). Ana t = ([],[])" by (simp_all add: setops\<^sub>s\<^sub>s\<^sub>t_def) hence "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?n \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel ?n)))" using 0(2) tfr_consts_mono unfolding tfr\<^sub>s\<^sub>s\<^sub>t_def by blast hence "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@?n) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel (\@?n)))" using unlabel_append[of \ ?n] trms\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel ?n"] setops\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel ?n"] by (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) thus "tfr\<^sub>s\<^sub>s\<^sub>t (unlabel ?A)" using 0(2) unlabel_append[of ?A ?n] unfolding tfr\<^sub>s\<^sub>s\<^sub>t_def by auto qed lemma reachable_constraints_typing_result: fixes P assumes M: "has_all_wt_instances_of \ (\T \ set P. trms_transaction T) N" "finite N" "tfr\<^sub>s\<^sub>e\<^sub>t N" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s N" and P: "\T \ set P. admissible_transaction T" "\T \ set P. list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (transaction_strand T))" and A: "\ \ reachable_constraints P" and \: "constraint_model \ (\@[(l, send\[attack\n\]\)])" shows "\\\<^sub>\. welltyped_constraint_model \\<^sub>\ (\@[(l, send\[attack\n\]\)])" proof - have I: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "constr_sem_stateful \ (unlabel (\@[(l, send\[attack\n\]\)]))" using \ unfolding constraint_model_def by metis+ note 0 = admissible_transaction_is_wellformed_transaction(1,4)[OF admissible_transactionE'(1)] have 1: "\T \ set P. wellformed_transaction T" using P(1) 0(1) by blast have 2: "\T \ set P. wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (trms_transaction T)" using P(1) 0(2) unfolding admissible_transaction_terms_def by blast have 3: "wf\<^sub>s\<^sub>s\<^sub>t (unlabel \)" "tfr\<^sub>s\<^sub>s\<^sub>t (unlabel \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using reachable_constraints_tfr[OF _ M P A] reachable_constraints_wf[OF 1(1) 2 A] by metis+ show ?thesis using stateful_typing_result[OF reachable_constraints_typing_result_aux[OF 3] I(1,3)] by (metis welltyped_constraint_model_def constraint_model_def) qed lemma reachable_constraints_typing_result': fixes P assumes M: "M \ \T \ set P. trms_transaction T \ pair' Pair ` setops_transaction T" "has_all_wt_instances_of \ M N" "finite N" "tfr\<^sub>s\<^sub>e\<^sub>t N" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s N" and P: "\T \ set P. wellformed_transaction T" "\T \ set P. wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (trms_transaction T)" "\T \ set P. list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (transaction_strand T))" and A: "\ \ reachable_constraints P" and \: "constraint_model \ (\@[(l, send\[attack\n\]\)])" shows "\\\<^sub>\. welltyped_constraint_model \\<^sub>\ (\@[(l, send\[attack\n\]\)])" proof - have I: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "constr_sem_stateful \ (unlabel (\@[(l, send\[attack\n\]\)]))" using \ unfolding constraint_model_def by metis+ have 0: "wf\<^sub>s\<^sub>s\<^sub>t (unlabel \)" "tfr\<^sub>s\<^sub>s\<^sub>t (unlabel \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using reachable_constraints_tfr'[OF M P(2-3) A] reachable_constraints_wf[OF P(1,2) A] by metis+ show ?thesis using stateful_typing_result[OF reachable_constraints_typing_result_aux[OF 0] I(1,3)] by (metis welltyped_constraint_model_def constraint_model_def) qed end lemma reachable_constraints_transaction_proj: assumes "\ \ reachable_constraints P" shows "proj n \ \ reachable_constraints (map (transaction_proj n) P)" using assms proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \ \) show ?case using step.hyps(2) reachable_constraints.step[OF step.IH _ transaction_decl_subst_proj[OF step.hyps(3)] transaction_fresh_subst_proj[OF step.hyps(4)] transaction_renaming_subst_proj[OF step.hyps(5)]] by (simp add: proj_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t proj_subst transaction_strand_proj) qed (simp add: reachable_constraints.init) context begin private lemma reachable_constraints_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_aux: fixes P defines "Ts \ concat (map transaction_strand P)" assumes A: "A \ reachable_constraints P" shows "\b \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A). \a \ set Ts. \\. b = a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ (\t \ subst_range \. (\x. t = Var x) \ (\c. t = Fun c []))" (is "\b \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A). \a \ set Ts. ?P b a") using A proof (induction A rule: reachable_constraints.induct) case (step \ T \ \ \) define Q where "Q \ ?P" define \ where "\ \ \ \\<^sub>s \ \\<^sub>s \" let ?R = "\A Ts. \b \ set A. \a \ set Ts. Q b a" have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "\t \ subst_range \. (\x. t = Var x) \ (\c. t = Fun c [])" using transaction_decl_fresh_renaming_substs_wt[OF step.hyps(3-5)] transaction_decl_fresh_renaming_substs_range_wf_trms[OF step.hyps(3-5)] transaction_decl_fresh_renaming_substs_range'(1)[OF step.hyps(3-5)] unfolding \_def by (metis,metis,fastforce) hence "?R (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) (transaction_strand T)" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_self_inverse[of "transaction_strand T"] by (auto simp add: Q_def subst_apply_labeled_stateful_strand_def) hence "?R (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) (transaction_strand T)" by (metis dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst) hence "?R (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) Ts" using step.hyps(2) unfolding Ts_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by fastforce thus ?case using step.IH unfolding Q_def \_def by auto qed simp lemma reachable_constraints_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t: fixes P defines "f \ \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}" and "Ts \ concat (map transaction_strand P)" assumes P_pc: "comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t public arity Ana \ Pair Ts M S" and A: "A \ reachable_constraints P" shows "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t A ((f S) - {m. intruder_synth {} m})" using par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_if_comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t'[OF P_pc, of "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A", THEN par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t] reachable_constraints_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_aux[OF A, unfolded Ts_def[symmetric]] unfolding f_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_self_inverse by fast end lemma reachable_constraints_par_comp_constr: fixes P f S defines "f \ \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}" and "Ts \ concat (map transaction_strand P)" and "Sec \ f S - {m. intruder_synth {} m}" and "M \ \T \ set P. trms_transaction T \ pair' Pair ` setops_transaction T" assumes M: "has_all_wt_instances_of \ M N" "finite N" "tfr\<^sub>s\<^sub>e\<^sub>t N" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s N" and P: "\T \ set P. wellformed_transaction T" "\T \ set P. wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (trms_transaction T)" "\T \ set P. list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (transaction_strand T))" "comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t public arity Ana \ Pair Ts M_fun S" and \: "\ \ reachable_constraints P" and \: "constraint_model \ \" shows "\\\<^sub>\. welltyped_constraint_model \\<^sub>\ \ \ ((\n. welltyped_constraint_model \\<^sub>\ (proj n \)) \ (\\' l t. prefix \' \ \ suffix [(l, receive\t\)] \' \ strand_leaks\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' Sec \\<^sub>\))" proof - have \': "constr_sem_stateful \ (unlabel \)" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using \ unfolding constraint_model_def by blast+ show ?thesis using reachable_constraints_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t[OF P(4)[unfolded Ts_def] \] reachable_constraints_typing_cond\<^sub>s\<^sub>s\<^sub>t[OF M_def M P(1-3) \] par_comp_constr_stateful[OF _ _ \', of Sec] unfolding f_def Sec_def welltyped_constraint_model_def constraint_model_def by blast qed lemma reachable_constraints_component_leaks_if_composed_leaks: fixes Sec Q defines "leaks \ \\. \\\<^sub>\ \'. Q \\<^sub>\ \ prefix \' \ \ (\l' t. suffix [(l', receive\t\)] \') \ strand_leaks\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' Sec \\<^sub>\" assumes Sec: "\s \ Sec. \{} \\<^sub>c s" "ground Sec" and composed_leaks: "\\ \ reachable_constraints Ps. leaks \" shows "\l. \\ \ reachable_constraints (map (transaction_proj l) Ps). leaks \" proof - from composed_leaks obtain \ \\<^sub>\ \' s n where \: "\ \ reachable_constraints Ps" and \': "prefix \' \" "constr_sem_stateful \\<^sub>\ (proj_unl n \'@[send\[s]\])" and \\<^sub>\: "Q \\<^sub>\" and s: "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\" unfolding leaks_def strand_leaks\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by fast have "\{} \\<^sub>c s" "s \ \\<^sub>\ = s" using s Sec by auto then obtain B k' u where "constr_sem_stateful \\<^sub>\ (proj_unl n B@[send\[s]\])" "prefix (proj n B) (proj n \)" "suffix [(k', receive\u\)] (proj n B)" "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n B) \\<^sub>\" using constr_sem_stateful_proj_priv_term_prefix_obtain[OF \' s] unfolding welltyped_constraint_model_def constraint_model_def by metis thus ?thesis using \\<^sub>\ reachable_constraints_transaction_proj[OF \, of n] proj_idem[of n B] unfolding leaks_def strand_leaks\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by metis qed lemma reachable_constraints_preserves_labels: assumes \: "\ \ reachable_constraints P" shows "\a \ set \. \T \ set P. \b \ set (transaction_strand T). fst b = fst a" (is "\a \ set \. \T \ set P. ?P T a") using \ proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \ \) have "\a \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)). ?P T a" proof fix a assume a: "a \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" then obtain b where b: "b \ set (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" "a = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p b" unfolding dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by auto then obtain c where c: "c \ set (transaction_strand T)" "b = c \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \\<^sub>s \ \\<^sub>s \" unfolding subst_apply_labeled_stateful_strand_def by auto have "?P T c" using c(1) by blast hence "?P T b" using c(2) by (simp add: subst_lsstp_fst_eq) thus "?P T a" using b(2) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_fst_eq[of b] by presburger qed thus ?case using step.IH step.hyps(2) by (metis Un_iff set_append) qed simp lemma reachable_constraints_preserves_labels': assumes P: "\T \ set P. \a \ set (transaction_strand T). has_LabelN l a \ has_LabelS a" and \: "\ \ reachable_constraints P" shows "\a \ set \. has_LabelN l a \ has_LabelS a" using reachable_constraints_preserves_labels[OF \] P by fastforce lemma reachable_constraints_transaction_proj_proj_eq: assumes \: "\ \ reachable_constraints (map (transaction_proj l) P)" shows "proj l \ = \" and "prefix \' \ \ proj l \' = \'" using \ proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \ \) let ?T = "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" note A = reachable_constraints.step[OF step.hyps] have P: "\T \ set (map (transaction_proj l) P). \a \ set (transaction_strand T). has_LabelN l a \ has_LabelS a" using transaction_proj_labels[of l] unfolding list_all_iff by auto note * = reachable_constraints_preserves_labels'[OF P A] have **: "\a \ set \'. has_LabelN l a \ has_LabelS a" when "\a \ set B. has_LabelN l a \ has_LabelS a" "prefix \' B" for B using that assms unfolding prefix_def by auto note *** = proj_ident[unfolded list_all_iff] { case 1 thus ?case using *[THEN ***] by blast } { case 2 thus ?case using *[THEN **, THEN ***] by blast } qed (simp_all add: reachable_constraints.init) lemma reachable_constraints_transaction_proj_star_proj: assumes \: "\ \ reachable_constraints (map (transaction_proj l) P)" and k_neq_l: "k \ l" shows "proj k \ \ reachable_constraints (map transaction_star_proj P)" using \ proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \ \) have "map (transaction_proj k) (map (transaction_proj l) P) = map transaction_star_proj P" using transaction_star_proj_negates_transaction_proj(2)[OF k_neq_l] by fastforce thus ?case using reachable_constraints_transaction_proj[OF reachable_constraints.step[OF step.hyps], of k] by argo qed (simp add: reachable_constraints.init) lemma reachable_constraints_aligned_prefix_ex: fixes P defines "f \ \T. list_all is_Receive (unlabel (transaction_receive T)) \ list_all is_Check_or_Assignment (unlabel (transaction_checks T)) \ list_all is_Update (unlabel (transaction_updates T)) \ list_all is_Send (unlabel (transaction_send T))" assumes P: "list_all f P" "list_all ((list_all (Not \ has_LabelS)) \ tl \ transaction_send) P" and s: "\{} \\<^sub>c s" "fv s = {}" and A: "A \ reachable_constraints P" "prefix B A" and B: "\l ts. suffix [(l, receive\ts\)] B" "constr_sem_stateful \ (unlabel B@[send\[s]\])" shows "\C \ reachable_constraints P. prefix C A \ (\l ts. suffix [(l, receive\ts\)] C) \ declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t B \ = declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t C \ \ constr_sem_stateful \ (unlabel C@[send\[s]\])" using A proof (induction A rule: reachable_constraints.induct) case (step A T \ \ \) define \ where "\ \ \ \\<^sub>s \ \\<^sub>s \" let ?T = "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" note AT_reach = reachable_constraints.step[OF step.hyps] obtain lb tsb B' where B': "B = B'@[(lb, receive\tsb\)]" using B(1) unfolding suffix_def by blast define decl_ik where "decl_ik \ \S::('fun,'atom,'sets,'lbl) prot_strand. \{set ts |ts. \\, receive\ts\\ \ set S} \\<^sub>s\<^sub>e\<^sub>t \" have decl_ik_append: "decl_ik (M@N) = decl_ik M \ decl_ik N" for M N unfolding decl_ik_def by fastforce have "\\, receive\ts\\ \ set N" when "\ \ fst ` set N" for ts and N::"('fun, 'atom, 'sets, 'lbl) prot_strand" using that by force hence decl_ik_star: "decl_ik M = decl_ik (M@N)" when "\ \ fst ` set N" for M N using that unfolding decl_ik_def by simp have decl_decl_ik: "declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ = {t. decl_ik S \ t}" for S unfolding declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_alt_def decl_ik_def by blast have "f T" using P(1) step.hyps(2) by (simp add: list_all_iff) hence "list_all is_Send (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" "list_all is_Check_or_Assignment (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" "list_all is_Update (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" "list_all is_Receive (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" using subst_sst_list_all(2)[of "unlabel (transaction_receive T)" \] subst_sst_list_all(11)[of "unlabel (transaction_checks T)" \] subst_sst_list_all(10)[of "unlabel (transaction_updates T)" \] subst_sst_list_all(1)[of "unlabel (transaction_send T)" \] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_list_all(1)[of "transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_list_all(11)[of "transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_list_all(10)[of "transaction_updates T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_list_all(2)[of "transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] unfolding f_def by (metis unlabel_subst[of _ \])+ hence "\list_ex is_Receive (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" "\list_ex is_Receive (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" "\list_ex is_Receive (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" "list_all is_Receive (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" unfolding list_ex_iff list_all_iff by blast+ then obtain TA TB where T: "?T = TA@TB" "\list_ex is_Receive (unlabel TA)" "list_all is_Receive (unlabel TB)" "TB = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using transaction_dual_subst_unfold[of T \] unfolding \_def by fastforce have 0: "prefix B (A@TA@TB)" using step.prems B' T by argo have 1: "prefix B A" when "prefix B (A@TA)" using that T(2) B' prefix_prefix_inv unfolding list_ex_iff unlabel_def by fastforce have 2: "\ \ fst ` set TB2" when "TB = TB1@(l,x)#TB2" for TB1 l x TB2 proof - have "k \ \" when "k \ set (map fst (tl (transaction_send T)))" for k using that P(2) step.hyps(2) unfolding list_all_iff by auto hence "k \ \" when "k \ set (map fst (tl TB))" for k using that subst_lsst_map_fst_eq[of "tl (transaction_send T)" \] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_map_fst_eq[of "tl (transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)"] unfolding T(4) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_tl subst_lsst_tl by simp moreover have "set TB2 \ set (tl TB)" using that by (metis (no_types, lifting) append_eq_append_conv2 order.eq_iff list.sel(3) self_append_conv set_mono_suffix suffix_appendI suffix_tl tl_append2) ultimately show ?thesis by auto qed have 3: "declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t TB \ = declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t (TB1@[(l,x)]) \" when "TB = TB1@(l,x)#TB2" for TB1 l x TB2 using decl_ik_star[OF 2[OF that], of "TB1@[(l,x)]"] unfolding that decl_decl_ik by simp show ?case proof (cases "prefix B A") case False hence 4: "\prefix B (A@TA)" using 1 by blast have 5: "\l ts. suffix [(l, receive\ts\)] (A@?T)" proof - have "(lb, receive\tsb\) \ set TB" using 0 4 prefix_prefix_inv[OF _ suffixI[OF B'], of "A@TA" TB] by (metis append_assoc) hence "receive\tsb\ \ set (unlabel TB)" unfolding unlabel_def by force hence "\ts. suffix [receive\ts\] (unlabel TB)" using T(3) unfolding list_all_iff is_Receive_def suffix_def by (metis in_set_conv_decomp list.distinct(1) list.set_cases rev_exhaust) then obtain TB' ts where "unlabel TB = TB'@[receive\ts\]" unfolding suffix_def by blast then obtain TB'' x where "TB = TB''@[x]" "snd x = receive\ts\" by (metis (no_types, opaque_lifting) append1_eq_conv list.distinct(1) rev_exhaust rotate1.simps(2) rotate1_is_Nil_conv unlabel_Cons(2) unlabel_append unlabel_nil) then obtain l where "suffix [(l, receive\ts\)] TB" by (metis surj_pair prod.sel(2) suffix_def) thus ?thesis using T(4) transaction_dual_subst_unfold[of T \] suffix_append[of "[(l, receive\ts\)]"] unfolding \_def by metis qed obtain TB1 where TB: "B = A@TA@TB1@[(lb, receive\tsb\)]" "prefix (TB1@[(lb, receive\tsb\)]) TB" using 0 4 B' prefix_snoc_obtain[of B' "(lb, receive\tsb\)" "A@TA" TB thesis] by (metis append_assoc) then obtain TB2 where TB2: "TB = TB1@(lb, receive\tsb\)#TB2" unfolding prefix_def by fastforce hence TB2': "list_all is_Receive (unlabel TB2)" using T(3) unfolding list_all_iff is_Receive_def proj_def unlabel_def by auto have 6: "constr_sem_stateful \ (unlabel B)" "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t B \\<^sub>s\<^sub>e\<^sub>t \ \ s \ \" using B(2) strand_sem_append_stateful[of "{}" "{}" "unlabel B" "[send\[s]\]" \] by auto have "constr_sem_stateful \ (unlabel (A@TA@TB1@[(lb, receive\tsb\)]))" using 6(1) TB(1) by blast hence "constr_sem_stateful \ (unlabel (A@?T))" using T(1) TB2 strand_sem_receive_prepend_stateful[ of "{}" "{}" "unlabel (A@TA@TB1@[(lb, receive\tsb\)])" \, OF _ TB2'] by auto moreover have "set (unlabel B) \ set (unlabel (A@?T))" using step.prems(1) unfolding prefix_def by force hence "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@?T) \\<^sub>s\<^sub>e\<^sub>t \ \ s \ \" using ideduct_mono[OF 6(2)] subst_all_mono[of _ _ \] ik\<^sub>s\<^sub>s\<^sub>t_set_subset[of "unlabel B" "unlabel (A@?T)"] by meson ultimately have 7: "constr_sem_stateful \ (unlabel (A@?T)@[send\[s]\])" using strand_sem_append_stateful[of "{}" "{}" "unlabel (A@?T)" "[send\[s]\]" \] by auto have "declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t B \ = declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@?T) \" proof - have "declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t TB \ = declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t (TB1@[(lb, receive\tsb\)]) \" using 3[of _ lb "receive\tsb\"] TB(2) unfolding prefix_def by auto hence "(decl_ik TB \ t) \ decl_ik (TB1@[(lb,receive\tsb\)]) \ t" for t unfolding TB(1) T(1) decl_decl_ik by blast hence "(decl_ik (A@TA@TB) \ t) \ decl_ik (A@TA@TB1@[(lb,receive\tsb\)]) \ t" for t using ideduct_mono_eq[of "decl_ik TB" "decl_ik (TB1@[(lb,receive\tsb\)])" "decl_ik (A@TA)"] by (metis decl_ik_append[of "A@TA"] Un_commute[of _ "decl_ik (A@TA)"] append_assoc) thus ?thesis unfolding TB(1) T(1) decl_decl_ik by blast qed thus ?thesis using step.prems AT_reach B(1) 5 7 by force qed (use step.IH prefix_append in blast) qed (use B(1) suffix_def in simp) lemma reachable_constraints_secure_if_filter_secure_case: fixes f l n and P::"('fun,'atom,'sets,'lbl) prot_transaction list" defines "has_attack \ \P. \\ \ reachable_constraints P. \\. constraint_model \ (\@[(l, send\[attack\n\]\)])" and "f \ \T. list_ex (\a. is_Update (snd a) \ is_Send (snd a)) (transaction_strand T)" and "g \ \T. transaction_fresh T = [] \ f T" assumes att: "has_attack P" shows "has_attack (filter g P)" proof - let ?attack = "\A I. constraint_model I (A@[(l, send\[attack\n\]\)])" define constr' where "constr' \ \(T::('fun,'atom,'sets,'lbl) prot_transaction,\::('fun,'atom,'sets,'lbl) prot_subst, \::('fun,'atom,'sets,'lbl) prot_subst,\::('fun,'atom,'sets,'lbl) prot_subst). dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" define constr where "constr \ \Ts. concat (map constr' Ts)" define h where "h \ \(T::('fun,'atom,'sets,'lbl) prot_transaction,_::('fun,'atom,'sets,'lbl) prot_subst, _::('fun,'atom,'sets,'lbl) prot_subst,_::('fun,'atom,'sets,'lbl) prot_subst). g T" obtain A I where A: "A \ reachable_constraints P" "?attack A I" using att unfolding has_attack_def by blast obtain Ts where Ts: "A = constr Ts" "\B. prefix B Ts \ constr B \ reachable_constraints P" "\B T \ \ \. prefix (B@[(T,\,\,\)]) Ts \ T \ set P \ transaction_decl_subst \ T \ transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (constr B)) \ transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (constr B))" using reachable_constraints_as_transaction_lists[OF A(1)] constr_def constr'_def by auto define B where "B \ constr (filter h Ts)" have Ts': "T \ set P" when "(T,\,\,\) \ set Ts" for T \ \ \ using that Ts(3) by (meson prefix_snoc_in_iff) have constr_Cons: "constr (p#Ts) = constr' p@constr Ts" for p Ts unfolding constr_def by simp have constr_snoc: "constr (Ts@[p]) = constr Ts@constr' p" for p Ts unfolding constr_def by simp have 0: "?attack B I" when A_att: "?attack A I" proof - have not_f_T_case: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (constr' p) = {}" "\D. dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (constr' p)) I D = D" when not_f_T: "\(f T)" and p: "p = (T,\,\,\)" for p T \ \ \ proof - have constr_p: "constr' p = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" unfolding p constr'_def by fast have "\is_Receive a" when a: "(l,a) \ set (constr' p)" for l a proof assume "is_Receive a" then obtain ts where ts: "a = receive\ts\" by (cases a) auto then obtain ts' where ts': "(l,send\ts'\) \ set (transaction_strand T)" "ts = ts' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \" using a dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_steps_iff(1)[of l ts] subst_lsst_memD(2)[of l _ "transaction_strand T"] unfolding constr_p by blast thus False using not_f_T unfolding f_def list_ex_iff by fastforce qed thus "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (constr' p) = {}" using in_ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t_iff by fastforce have "\is_Update a" when a: "(l,a) \ set (constr' p)" for l a proof assume "is_Update a" then obtain t s where ts: "a = insert\t,s\ \ a = delete\t,s\" by (cases a) auto then obtain t' s' where ts': "(l,insert\t',s'\) \ set (transaction_strand T) \ (l,delete\t',s'\) \ set (transaction_strand T)" "t = t' \ \ \\<^sub>s \ \\<^sub>s \" "s = s' \ \ \\<^sub>s \ \\<^sub>s \" using a dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_steps_iff(4,5)[of l] subst_lsst_memD(4,5)[of l _ _ "transaction_strand T"] unfolding constr_p by blast thus False using not_f_T unfolding f_def list_ex_iff by fastforce qed thus "\D. dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (constr' p)) I D = D" using dbupd\<^sub>s\<^sub>s\<^sub>t_no_upd[of "unlabel (constr' p)" I] by (meson unlabel_mem_has_label) qed have *: "strand_sem_stateful M D (unlabel B) I" when "strand_sem_stateful M D (unlabel A) I" for M D using that Ts' unfolding Ts(1) B_def proof (induction Ts arbitrary: M D) case (Cons p Ts) obtain T \ \ \ where p: "p = (T,\,\,\)" by (cases p) simp have T_in: "T \ set P" using Cons.prems(2) unfolding p by fastforce let ?M' = "M \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (constr' p) \\<^sub>s\<^sub>e\<^sub>t I)" let ?D' = "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (constr' p)) I D" have p_sem: "strand_sem_stateful M D (unlabel (constr' p)) I" and IH: "strand_sem_stateful ?M' ?D' (unlabel (constr (filter h Ts))) I" using Cons.IH[of ?M' ?D'] Cons.prems strand_sem_append_stateful[of M D "unlabel (constr' p)" "unlabel (constr Ts)" I] unfolding constr_Cons unlabel_append by fastforce+ show ?case proof (cases "T \ set (filter g P)") case True hence h_p: "filter h (p#Ts) = p#filter h Ts" unfolding h_def p by simp show ?thesis using p_sem IH strand_sem_append_stateful[of M D "unlabel (constr' p)" _ I] unfolding h_p constr_Cons unlabel_append by blast next case False hence not_f: "\(f T)" and not_h: "\(h p)" using T_in unfolding g_def h_def p by auto show ?thesis using not_h not_f_T_case[OF not_f p] IH unfolding constr_Cons unlabel_append by auto qed qed simp have **: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t B = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" proof show "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t B \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" unfolding Ts(1) B_def constr_def by (induct Ts) (auto simp add: ik\<^sub>s\<^sub>s\<^sub>t_def) show "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t B" using Ts' unfolding Ts(1) B_def proof (induction Ts) case (Cons p Ts) obtain T \ \ \ where p: "p = (T,\,\,\)" by (cases p) simp have T_in: "T \ set P" using Cons.prems unfolding p by fastforce have IH: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (constr Ts) \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (constr (filter h Ts))" using Cons.IH Cons.prems by auto show ?case proof (cases "T \ set (filter g P)") case True hence h_p: "filter h (p#Ts) = p#filter h Ts" unfolding h_def p by simp show ?thesis using IH unfolding h_p constr_Cons unlabel_append ik\<^sub>s\<^sub>s\<^sub>t_append by blast next case False hence not_f: "\(f T)" and not_h: "\(h p)" using T_in unfolding g_def h_def p by auto show ?thesis using not_h not_f_T_case[OF not_f p] IH unfolding constr_Cons unlabel_append by auto qed qed simp qed show ?thesis using A_att *[of "{}" "{}"] ** strand_sem_stateful_if_sends_deduct strand_sem_append_stateful[of "{}" "{}" _ "unlabel [(l, send\[attack\n\]\)]" I] unfolding constraint_model_def unlabel_append by force qed have 1: "B \ reachable_constraints (filter g P)" using A(1) Ts(2,3) unfolding Ts(1) B_def proof (induction Ts rule: List.rev_induct) case (snoc p Ts) obtain T \ \ \ where p: "p = (T,\,\,\)" by (cases p) simp have constr_p: "constr' p = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" unfolding constr'_def p by fastforce have T_in: "T \ set P" and \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (constr Ts))" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (constr Ts))" using snoc.prems(3) unfolding p by fast+ have "transaction_fresh_subst s t bb" when "transaction_fresh_subst s t aa" "bb \ aa" for s t bb aa using that unfolding transaction_fresh_subst_def by fast have "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (constr (filter h Ts)) \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (constr Ts)" unfolding constr_def unlabel_def by fastforce hence \': "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (constr (filter h Ts)))" using \ unfolding transaction_fresh_subst_def by fast have "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (constr (filter h Ts)) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (constr Ts)" unfolding constr_def unlabel_def vars\<^sub>s\<^sub>s\<^sub>t_def by auto hence \': "transaction_renaming_subst \ (filter g P) (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (constr (filter h Ts)))" using \ unfolding transaction_renaming_subst_def by auto have IH: "constr (filter h Ts) \ reachable_constraints (filter g P)" using snoc.prems snoc.IH by simp show ?case proof (cases "h p") case True hence h_p: "filter h (Ts@[p]) = filter h Ts@[p]" by fastforce have T_in': "T \ set (filter g P)" using T_in True unfolding h_def p by fastforce show ?thesis using IH reachable_constraints.step[OF IH T_in' \ \' \'] unfolding h_p constr_snoc constr_p by fast next case False thus ?thesis using IH by fastforce qed qed (simp add: constr_def) show ?thesis using 0 1 A(2) unfolding has_attack_def by blast qed lemma reachable_constraints_fv_Value_typed: assumes P: "\T \ set P. admissible_transaction' T" and A: "\ \ reachable_constraints P" and x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" shows "\\<^sub>v x = TAtom Value" proof - obtain T where T: "T \ set P" "\\<^sub>v x \ \\<^sub>v ` fv_transaction T" using x P(1) reachable_constraints_var_types_in_transactions(1)[OF A(1)] admissible_transactionE(2) by blast show ?thesis using T(2) admissible_transactionE(3)[OF bspec[OF P(1) T(1)]] vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_strand T)"] by force qed lemma reachable_constraints_fv_Value_const_cases: assumes P: "\T \ set P. admissible_transaction' T" and A: "\ \ reachable_constraints P" and I: "welltyped_constraint_model \ \" and x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" shows "(\n. \ x = Fun (Val n) []) \ (\n. \ x = Fun (PubConst Value n) [])" proof - have x': "\ (\ x) = TAtom Value" "fv (\ x) = {}" "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" using reachable_constraints_fv_Value_typed[OF P A x] I wt_subst_trm''[of \ "Var x"] unfolding welltyped_constraint_model_def constraint_model_def by auto obtain f where f: "arity f = 0" "\ x = Fun f []" using TAtom_term_cases[OF x'(3,1)] x' const_type_inv_wf[of _ _ Value] by (cases "\ x") force+ show ?thesis proof (cases f) case (Fu g) thus ?thesis by (metis f(2) x'(1) \_Fu_simps(4)[of g "[]"]) qed (use f x'(1) in auto) qed lemma reachable_constraints_receive_attack_if_attack'': assumes P: "\T \ set P. admissible_transaction' T" and A: "\ \ reachable_constraints P" and wt_attack: "welltyped_constraint_model \ (\@[(l, send\[attack\n\]\)])" shows "receive\[attack\n\]\ \ set (unlabel \)" proof - have I: "welltyped_constraint_model \ \" using welltyped_constraint_model_prefix wt_attack by blast show ?thesis using wt_attack strand_sem_append_stateful[of "{}" "{}" "unlabel \" "[send\[attack\n\]\]" \] reachable_constraints_receive_attack_if_attack'(2)[OF A(1) P I] unfolding welltyped_constraint_model_def constraint_model_def by simp qed context begin private lemma reachable_constraints_initial_value_transaction_aux: fixes P::"('fun,'atom,'sets,'lbl) prot" and N::"nat set" assumes P: "\T \ set P. admissible_transaction' T" and A: "A \ reachable_constraints P" and P': "\T \ set P. \(l,a) \ set (transaction_strand T). \t. a \ select\t,\k\\<^sub>s\ \ a \ \t in \k\\<^sub>s\ \ a \ \t not in \k\\<^sub>s\ \ a \ delete\t,\k\\<^sub>s\" shows "(l,\ac: t \ s\) \ set A \ (\u. s = \u\\<^sub>s \ u \ k)" (is "?A A \ ?Q A") and "(l,\t not in s\) \ set A \ (\u. s = \u\\<^sub>s \ u \ k)" (is "?B A \ ?Q A") and "(l,delete\t,s\) \ set A \ (\u. s = \u\\<^sub>s \ u \ k)" (is "?C A \ ?Q A") proof - have "(?A A \ ?Q A) \ (?B A \ ?Q A) \ (?C A \ ?Q A)" (is "?D A") using A proof (induction A rule: reachable_constraints.induct) case (step \ T \ \ \) define \ where "\ \ \ \\<^sub>s \ \\<^sub>s \" let ?T' = "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" note T_adm = bspec[OF P step.hyps(2)] note T_wf = admissible_transaction_is_wellformed_transaction[OF T_adm] note T_P' = bspec[OF P' step.hyps(2)] have 0: "?Q ?T'" when A: "?A ?T'" proof - obtain t' s' where t: "(l, \ac: t' \ s'\) \ set (transaction_strand T)" "t = t' \ \" "s = s' \ \" using A dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_steps_iff(6) subst_lsst_memD(6) by blast obtain u where u: "s' = \u\\<^sub>s" using transaction_selects_are_Value_vars[OF T_wf(1,2), of t' s'] transaction_inset_checks_are_Value_vars[OF T_adm, of t' s'] unlabel_in[OF t(1)] by (cases ac) auto show ?thesis using T_P' t(1,3) unfolding u by (cases ac) auto qed have 1: "?Q ?T'" when B: "?B ?T'" proof - obtain t' s' where t: "(l, \t' not in s'\) \ set (transaction_strand T)" "t = t' \ \" "s = s' \ \" using B dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_steps_iff(7) subst_lsst_memD(9) by blast obtain u where u: "s' = \u\\<^sub>s" using transaction_notinset_checks_are_Value_vars(2)[OF T_adm unlabel_in[OF t(1)]] by fastforce show ?thesis using T_P' t(1,3) unfolding u by auto qed have 2: "?Q ?T'" when C: "?C ?T'" proof - obtain t' s' where t: "(l, delete\t', s'\) \ set (transaction_strand T)" "t = t' \ \" "s = s' \ \" using C dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_steps_iff(5) subst_lsst_memD(5) by blast obtain u where u: "s' = \u\\<^sub>s" using transaction_deletes_are_Value_vars(2)[OF T_wf(1,3) unlabel_in[OF t(1)]] by blast show ?thesis using T_P' t(1,3) unfolding u by auto qed show ?case using 0 1 2 step.IH unfolding \_def by auto qed simp thus "?A A \ ?Q A" "?B A \ ?Q A" "?C A \ ?Q A" by fast+ qed lemma reachable_constraints_initial_value_transaction: fixes P::"('fun,'atom,'sets,'lbl) prot" and N::"nat set" and k A T_upds defines "checks_not_k \ \B. T_upds \ [] \ ( (\l t s. (l,\t in s\) \ set (A@B) \ (\u. s = \u\\<^sub>s \ u \ k)) \ (\l t s. (l,\t not in s\) \ set (A@B) \ (\u. s = \u\\<^sub>s \ u \ k)) \ (\l t s. (l,delete\t,s\) \ set (A@B) \ (\u. s = \u\\<^sub>s \ u \ k)))" assumes P: "\T \ set P. admissible_transaction' T" and A: "A \ reachable_constraints P" and N: "finite N" "\n \ N. \(Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" and T: "T \ set P" "Var x \ set T_ts" "\\<^sub>v x = TAtom Value" "fv\<^sub>s\<^sub>e\<^sub>t (set T_ts) = {x}" "\n. \(Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t set T_ts)" "T = Transaction (\(). []) [x] [] [] T_upds [(l1,send\T_ts\)]" "T_upds = [] \ (T_upds = [(l2,insert\Var x, \k\\<^sub>s\)] \ (\T \ set P. \(l,a) \ set (transaction_strand T). \t. a \ select\t,\k\\<^sub>s\ \ a \ \t in \k\\<^sub>s\ \ a \ \t not in \k\\<^sub>s\ \ a \ delete\t,\k\\<^sub>s\))" shows "\B. A@B \ reachable_constraints P \ B \ reachable_constraints P \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t B = {} \ (T_upds = [] \ list_all is_Receive (unlabel B)) \ (T_upds \ [] \ list_all (\a. is_Insert a \ is_Receive a) (unlabel B)) \ (\n. Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ Fun (Val n) [] \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t B) \ (\n. Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B \ Fun (Val n) [] \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t B) \ N = {n. Fun (Val n) [] \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t B} \ checks_not_k B \ (\l a. (l,a) \ set B \ is_Insert a \ (l = l2 \ (\n. a = insert\Fun (Val n) [],\k\\<^sub>s\)))" (is "\B. A@B \ ?reach P \ B \ ?reach P \ ?Q1 B \ ?Q2 B \ ?Q3 B \ ?Q4 B \ ?Q5 B \ ?Q6 N B \ checks_not_k B \ ?Q7 B") using N proof (induction N rule: finite_induct) case empty define B where "B \ []::('fun,'atom,'sets,'lbl) prot_constr" have 0: "A@B \ ?reach P" "B \ ?reach P" using A unfolding B_def by auto have 1: "?Q1 B" "?Q2 B" "?Q3 B" "?Q4 B" "?Q6 {} B" unfolding B_def by auto have 2: "checks_not_k B" using reachable_constraints_initial_value_transaction_aux[OF P 0(1)] T(7) unfolding checks_not_k_def by presburger have 3: "?Q5 B" "?Q7 B" unfolding B_def by simp_all show ?case using 0 1 2 3 by blast next case (insert n N) obtain B where B: "A@B \ reachable_constraints P" "B \ reachable_constraints P" "?Q1 B" "?Q2 B" "?Q3 B" "?Q4 B" "?Q5 B" "?Q6 N B" "checks_not_k B" "?Q7 B" using insert.IH insert.prems by blast define \ where "\ \ Var::('fun,'atom,'sets,'lbl) prot_subst" define \ where "\ \ Var(x := Fun (Val n) [])::('fun,'atom,'sets,'lbl) prot_subst" have \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@B))" proof (unfold transaction_fresh_subst_def; intro conjI) have "subst_range \ = {Fun (Val n) []}" unfolding \_def by simp moreover have "Fun (Val n) [] \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@B))" using insert.prems insert.hyps(2) B(7,8) ik\<^sub>s\<^sub>s\<^sub>t_trms\<^sub>s\<^sub>s\<^sub>t_subset[of "unlabel B"] unfolding unlabel_append trms\<^sub>s\<^sub>s\<^sub>t_append by blast ultimately show "\t \ subst_range \. t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@B))" by fastforce next show "subst_domain \ = set (transaction_fresh T)" using T(6) unfolding \_def by auto next show "\t \ subst_range \. t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms_transaction T)" using T(5,7) unfolding \_def T(6) by fastforce qed (force simp add: T(3) \_def wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def)+ hence \': "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" unfolding transaction_fresh_subst_def by fastforce have \: "transaction_decl_subst \ T" using T(6) unfolding \_def transaction_decl_subst_def by fastforce obtain \::"('fun,'atom,'sets,'lbl) prot_subst" where \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@B))" unfolding transaction_renaming_subst_def by blast hence \': "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" unfolding transaction_renaming_subst_def by auto define \ where "\ \ \ \\<^sub>s \ \\<^sub>s \" define C where "C \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" have \x: "\ x = Fun (Val n) []" unfolding \_def \_def \_def subst_compose by force have "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = []" "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = []" using T(6) unfolding \_def \_def \_def by auto moreover have "(T_upds = [] \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = []) \ (T_upds \ [] \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = [(l2,insert\\ x, \k\\<^sub>s\)])" using subst_lsst_cons[of "(l2,insert\Var x,\k\\<^sub>s\)" "[]" \] T(6,7) unfolding dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst by auto hence "(T_upds = [] \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = []) \ (T_upds \ [] \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = [(l2,insert\Fun (Val n) [],\k\\<^sub>s\)])" unfolding \_def \_def \_def by (auto simp: subst_compose) moreover have "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = [(l1, receive\T_ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\)]" using subst_lsst_cons[of "(l1, receive\T_ts\)" "[]" \] T(6) unfolding dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst by auto hence "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = [(l1, receive\T_ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\)]" by auto ultimately have C: "(T_upds = [] \ C = [(l1, receive\T_ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\)]) \ (T_upds \ [] \ C = [(l2, insert\Fun (Val n) [],\k\\<^sub>s\), (l1, receive\T_ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\)])" unfolding C_def transaction_dual_subst_unfold by force have C': "Fun (Val n) [] \ set (T_ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" "Fun (Val n) [] \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t C" using T(2) in_ik\<^sub>s\<^sub>s\<^sub>t_iff[of _ "unlabel C"] C unfolding \_def \_def \_def by (force, force) have "fv (t \ \) = {}" when t: "t \\<^sub>s\<^sub>e\<^sub>t set T_ts" for t proof - have "fv t \ {x}" using t T(4) fv_subset_subterms by blast hence "fv (t \ \ \\<^sub>s \) = {}" unfolding \_def \_def by (induct t) auto thus ?thesis unfolding \_def by (metis subst_ground_ident_compose(2)) qed hence 1: "ground (set (T_ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \))" by auto have 2: "m = n" when m: "Fun (Val m) [] \\<^sub>s\<^sub>e\<^sub>t ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t C" for m proof - have "Fun (Val m) [] \\<^sub>s\<^sub>e\<^sub>t set (T_ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using m C in_ik\<^sub>s\<^sub>s\<^sub>t_iff[of _ "unlabel C"] by fastforce hence *: "Fun (Val m) [] \\<^sub>s\<^sub>e\<^sub>t set T_ts \\<^sub>s\<^sub>e\<^sub>t \" by simp show ?thesis using const_subterms_subst_cases[OF *] T(4,5) \x by fastforce qed have C_trms: "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t C \ {Fun (Val n) [],\k\\<^sub>s} \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t C" using C in_ik\<^sub>s\<^sub>s\<^sub>t_iff[of _ "unlabel C"] by fastforce have 3: "m = n" when m: "Fun (Val m) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t C" for m using m 2[of m] C_trms by fastforce have Q1: "?Q1 (B@C)" using B(3) C 1 by auto have Q2: "?Q2 (B@C)" using B(4) C by force have Q3: "?Q3 (B@C)" using B(5) C by force have Q4: "?Q4 (B@C)" using B(6) insert.prems C 2 unfolding unlabel_append ik\<^sub>s\<^sub>s\<^sub>t_append by blast have Q5: "?Q5 (B@C)" using B(7) C' 3 unfolding unlabel_append ik\<^sub>s\<^sub>s\<^sub>t_append trms\<^sub>s\<^sub>s\<^sub>t_append by blast have Q6: "?Q6 (insert n N) (B@C)" using B(8) C' 2 unfolding unlabel_append ik\<^sub>s\<^sub>s\<^sub>t_append by blast have Q7: "?Q7 (B@C)" using B(10) C by fastforce have Q8: "checks_not_k (B@C)" using B(9) C unfolding checks_not_k_def by force have "B@C \ reachable_constraints P" "A@B@C \ reachable_constraints P" using reachable_constraints.step[OF B(1) T(1) \ \ \] reachable_constraints.step[OF B(2) T(1) \ \' \'] unfolding \_def[symmetric] C_def[symmetric] by simp_all thus ?case using Q1 Q2 Q3 Q4 Q5 Q6 Q7 Q8 by blast qed end subsection \Equivalence Between the Symbolic Protocol Model and a Ground Protocol Model\ context begin subsubsection \Intermediate Step: Equivalence to a Ground Protocol Model with Renaming\ private definition "consts_of X = {t. t \\<^sub>s\<^sub>e\<^sub>t X \ (\c. t = Fun c [])}" private fun mk_symb where "mk_symb (\, \, I, T, \) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t((transaction_strand T) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" private fun T_symb :: " _ \ ('fun,'atom,'sets,'lbl) prot_constr" where "T_symb w = concat (map mk_symb w)" private definition "narrow \ S = (\x. if x \ S then \ x else Var x)" private fun mk_inv\I where "mk_inv\I n (\, \, I, T) = narrow ((var_rename_inv n) \\<^sub>s I) (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s var_rename n))" private fun inv\I where "inv\I ns w = foldl (\\<^sub>s) Var (map2 mk_inv\I ns w)" private fun mk_I where "mk_I (\, \, I, T, \) = narrow I (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" private fun comb_I where "comb_I w = fold (\\<^sub>s) (map mk_I w) (\x. Fun OccursSec [])" private abbreviation "ground_term t \ ground {t}" private lemma ground_term_def2: "ground_term t \ (fv t = {})" by auto private definition "ground_strand s \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t s = {}" private fun ground_step :: "(_, _) stateful_strand_step \ bool" where "ground_step s \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p s = {}" private fun ground_lstep :: "_ strand_label \ (_, _) stateful_strand_step \ bool" where "ground_lstep (l,s) \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p s = {}" private inductive_set ground_protocol_states_aux:: "('fun,'atom,'sets,'lbl) prot \ (('fun,'atom,'sets,'lbl) prot_terms \ (('fun,'atom,'sets,'lbl) prot_term \ ('fun,'atom,'sets,'lbl) prot_term) set \ _ set \ _ set \ _ list) set" for P::"('fun,'atom,'sets,'lbl) prot" where init: "({},{},{},{},[]) \ ground_protocol_states_aux P" | step: "\(IK,DB,trms,vars,w) \ ground_protocol_states_aux P; T \ set P; transaction_decl_subst \ T; transaction_fresh_subst \ T trms; transaction_renaming_subst \ P vars; A = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \); strand_sem_stateful IK DB (unlabel A) I; interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I; wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I) \ \ (IK \ ((ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \\<^sub>s\<^sub>e\<^sub>t I), dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I DB, trms \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A, vars \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A, w@[(\, \, I, T, \)]) \ ground_protocol_states_aux P" private lemma T_symb_append': " T_symb (w@w') = T_symb w @ T_symb w'" proof (induction w arbitrary: w') case Nil then show ?case by auto next case (Cons a w) then show ?case by auto qed private lemma T_symb_append: "T_symb (w@[(\, \, I, T, \)]) = T_symb w @ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t((transaction_strand T) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" using T_symb_append'[of w "[(\, \, I, T,\)]"] by auto private lemma ground_step_subst: assumes "ground_step a" shows "a = a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" using assms proof (induction a) case (NegChecks Y F F') then have FY: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set Y = {}" unfolding ground_step.simps unfolding fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p.simps by auto { have "\t s. (t,s) \ set F \ t \ (rm_vars (set Y) \) = t" proof (rule, rule, rule) fix t s assume "(t, s) \ set F" then show "t \ rm_vars (set Y) \ = t" using FY by fastforce qed moreover have "\t s. (t,s) \ set F \ s \ (rm_vars (set Y) \) = s" proof (rule, rule, rule) fix t s assume "(t, s) \ set F" then show "s \ rm_vars (set Y) \ = s" using FY by fastforce qed ultimately have "\f \ set F. f \\<^sub>p (rm_vars (set Y) \) = f" by auto then have "F = F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set Y) \" by (metis (no_types, lifting) map_idI split_cong subst_apply_pairs_def) } moreover from NegChecks have F'Y: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' - set Y = {}" unfolding ground_step.simps unfolding fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p.simps by auto { have "\t s. (t,s) \ set F' \ t \ (rm_vars (set Y) \) = t" proof (rule, rule, rule) fix t s assume "(t, s) \ set F'" then show "t \ rm_vars (set Y) \ = t" using F'Y by fastforce qed moreover have "\t s. (t,s) \ set F' \ s \ (rm_vars (set Y) \) = s" proof (rule, rule, rule) fix t s assume "(t, s) \ set F'" then show "s \ rm_vars (set Y) \ = s" using F'Y by fastforce qed ultimately have "\f \ set F'. f \\<^sub>p (rm_vars (set Y) \) = f" by auto then have "F' = F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set Y) \" by (simp add: map_idI subst_apply_pairs_def) } ultimately show ?case by simp qed (auto simp add: map_idI subst_ground_ident) private lemma ground_lstep_subst: assumes "ground_lstep a" shows "a = a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" using assms by (cases a) (auto simp add: ground_step_subst) private lemma subst_apply_term_rm_vars_swap: assumes "\x\fv t - set X. I x = I' x" shows "t \ rm_vars (set X) I = t \ rm_vars (set X) I'" using assms by (induction t) auto private lemma subst_apply_pairs_rm_vars_swap: assumes "\x\\ (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r ` set ps) - set X. I x = I' x" shows "ps \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) I = ps \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) I'" proof - have "\p \ set ps. p \\<^sub>p rm_vars (set X) I = p \\<^sub>p rm_vars (set X) I'" proof fix p assume "p \ set ps" obtain t s where "p = (t,s)" by (cases p) auto have "\x\fv t - set X. I x = I' x" by (metis DiffD1 DiffD2 DiffI \p = (t, s)\ \p \ set ps\ assms fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s.elims fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_inI(4)) then have "t \ rm_vars (set X) I = t \ rm_vars (set X) I'" using subst_apply_term_rm_vars_swap by blast have "\x\fv s - set X. I x = I' x" by (metis DiffD1 DiffD2 DiffI \p = (t, s)\ \p \ set ps\ assms fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s.elims fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_inI(5)) then have "s \ rm_vars (set X) I = s \ rm_vars (set X) I'" using subst_apply_term_rm_vars_swap by blast show "p \\<^sub>p rm_vars (set X) I = p \\<^sub>p rm_vars (set X) I'" using \p = (t, s)\ \s \ rm_vars (set X) I = s \ rm_vars (set X) I'\ \t \ rm_vars (set X) I = t \ rm_vars (set X) I'\ by fastforce qed then show ?thesis unfolding subst_apply_pairs_def by auto qed private lemma subst_apply_stateful_strand_step_swap: assumes "\x\fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p T. I x = I' x" shows "T \\<^sub>s\<^sub>s\<^sub>t\<^sub>p I = T \\<^sub>s\<^sub>s\<^sub>t\<^sub>p I'" using assms proof (induction T) case (Send ts) then show ?case using term_subst_eq by fastforce next case (NegChecks X F G) then have "\x \ \(fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r ` set F) \ \(fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r ` set G) - set X. I x = I' x" by auto then show ?case using subst_apply_pairs_rm_vars_swap[of F] subst_apply_pairs_rm_vars_swap[of G] by auto qed (simp_all add: term_subst_eq_conv) private lemma subst_apply_labeled_stateful_strand_step_swap: assumes "\x \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd T). I x = I' x" shows "T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p I = T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p I'" using assms subst_apply_stateful_strand_step_swap by (metis prod.exhaust_sel subst_apply_labeled_stateful_strand_step.simps) private lemma subst_apply_labeled_stateful_strand_swap: assumes "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T. I x = I' x" shows "T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I = T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I'" using assms proof (induction T) case Nil then show ?case by auto next case (Cons a T) then show ?case using subst_apply_labeled_stateful_strand_step_swap by (metis UnCI fv\<^sub>s\<^sub>s\<^sub>t_Cons subst_lsst_cons unlabel_Cons(2)) qed private lemma transaction_renaming_subst_not_in_fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t: fixes \ \ \::"('fun,'atom,'sets,'lbl) prot_subst" and A::"('fun,'atom,'sets,'lbl) prot_constr" assumes "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" assumes "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" shows "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" proof - have 0: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)" when x: "x \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" for x and A::"('fun,'atom,'sets,'lbl) prot_constr" and \ \::"('fun,'atom,'sets,'lbl) prot_subst" proof - have "x \ range_vars \" using x \ transaction_renaming_subst_vars_disj(6) by blast moreover have "subst_domain \ = UNIV" using \ transaction_renaming_subst_is_renaming(4) by blast ultimately show ?thesis using subst_fv_dom_img_subset[of _ \] fv\<^sub>s\<^sub>s\<^sub>t_subst_obtain_var subst_compose unlabel_subst by (metis (no_types, opaque_lifting) subset_iff top_greatest) qed have 1: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" for x and A::"('fun,'atom,'sets,'lbl) prot_constr" and \ \::"('fun,'atom,'sets,'lbl) prot_subst" using \ x 0 by (metis Un_iff vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t) show ?thesis using 1 assms by metis qed private lemma wf_comb_I_Nil: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (comb_I []))" by auto private lemma comb_I_append: "comb_I (w @ [(\, \, I, T, \)]) = (mk_I (\, \, I, T, \) \\<^sub>s (comb_I w))" by auto private lemma reachable_constraints_if_ground_protocol_states_aux: assumes "(IK, DB, trms, vars, w) \ ground_protocol_states_aux P" shows "T_symb w \ reachable_constraints P \ constr_sem_stateful (comb_I w) (unlabel (T_symb w)) \ IK = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((T_symb w) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t (comb_I w)) \ DB = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel ((T_symb w))) (comb_I w) {} \ trms = trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w) \ vars = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w) \ interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (comb_I w) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (comb_I w))" using assms proof (induction rule: ground_protocol_states_aux.induct) case init show ?case using wf_comb_I_Nil by auto next case (step IK DB trms vars w T \ \ \ A I) then have step': "T_symb w \ reachable_constraints P" "strand_sem_stateful {} {} (unlabel (T_symb w)) (comb_I w)" "IK = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w \\<^sub>l\<^sub>s\<^sub>s\<^sub>t comb_I w)" "DB = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w)) (comb_I w) {}" "trms = trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w)" "vars = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w)" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (comb_I w)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (comb_I w))" by auto define w' where "w' = w @ [(\, \, I, T, \)]" have interps_w: "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w). (comb_I w) x = (comb_I w') x" proof fix x assume "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w)" then have "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" using step(5) transaction_renaming_subst_not_in_fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t unfolding step'(6) by blast then have "mk_I (\, \, I, T, \) x = Var x" unfolding mk_I.simps narrow_def by metis then have "comb_I w x = (mk_I (\, \, I, T, \) \\<^sub>s (comb_I (w))) x" by (simp add: subst_compose) then show "comb_I w x = comb_I w' x" unfolding w'_def by auto qed have interps_T: "\x \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel (mk_symb (\, \, I, T, \))). I x = (comb_I w') x" proof fix x assume "x \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel (mk_symb (\, \, I, T, \)))" then have a: "x \ (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" by (metis fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq mk_symb.simps) have "(comb_I w') x = (mk_I (\, \, I, T, \) \\<^sub>s (comb_I (w))) x" unfolding w'_def by auto also have "... = ((mk_I (\, \, I, T, \)) x) \ comb_I w" unfolding subst_compose by auto also have "... = (narrow I (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) x) \ comb_I w" using a by auto also have "... = (I x) \ comb_I w" by (metis a narrow_def) also have "... = I x" by (metis UNIV_I ground_subst_range_empty_fv step.hyps(8) subst_compose subst_ground_ident_compose(1)) finally show "I x = (comb_I w') x" by auto qed have "T_symb w' \ reachable_constraints P" proof - have "T_symb w \ reachable_constraints P" using step'(1) . moreover have "T \ set P" using step(2) by auto moreover have "transaction_decl_subst \ T" using step(3) by auto moreover have "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w))" using step(4) step'(5) by auto moreover have "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w))" using step(5) step'(6) by auto ultimately have "(T_symb w) @ mk_symb (\, \, I, T, \) \ reachable_constraints P" using reachable_constraints.step[of "T_symb w" P T \ \ \] by auto then show "T_symb w' \ reachable_constraints P" unfolding w'_def by auto qed moreover have "strand_sem_stateful {} {} (unlabel (T_symb w')) (comb_I w')" proof - have "strand_sem_stateful {} {} (unlabel (T_symb w)) (comb_I w')" proof - have "strand_sem_stateful {} {} (unlabel (T_symb w)) (comb_I w)" using step'(2) by auto then show "strand_sem_stateful {} {} (unlabel (T_symb w)) (comb_I w')" using interps_w strand_sem_model_swap by blast qed moreover have "strand_sem_stateful (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w) \\<^sub>s\<^sub>e\<^sub>t comb_I w') (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w)) (comb_I w') {}) (unlabel (mk_symb (\, \, I, T, \))) (comb_I w')" proof - have "A = (mk_symb (\, \, I, T, \))" unfolding step(6) by auto moreover have "strand_sem_stateful (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w \\<^sub>l\<^sub>s\<^sub>s\<^sub>t comb_I w)) (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w)) (comb_I w) {}) (unlabel A) I" using step'(3) step'(4) step.hyps(7) by force moreover { have "\x\fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w)). comb_I w x = comb_I w' x" using interps_w by (metis fv_ik\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t) then have "\x\fv t. comb_I w x = comb_I w' x" when t: "t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w)" for t using t by auto then have "t \ comb_I w' = t \ comb_I w" when t: "t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w)" for t using t term_subst_eq[of t "comb_I w'" "comb_I w"] by metis then have "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w) \\<^sub>s\<^sub>e\<^sub>t comb_I w' = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w) \\<^sub>s\<^sub>e\<^sub>t comb_I w" by auto also have "... = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w \\<^sub>l\<^sub>s\<^sub>s\<^sub>t comb_I w)" by (metis ik\<^sub>s\<^sub>s\<^sub>t_subst unlabel_subst) finally have "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w) \\<^sub>s\<^sub>e\<^sub>t comb_I w' = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w \\<^sub>l\<^sub>s\<^sub>s\<^sub>t comb_I w)" by auto } moreover { have "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w)) (comb_I w) {} = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w)) (comb_I w') {}" by (metis db\<^sub>s\<^sub>s\<^sub>t_subst_swap[OF interps_w] db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t empty_set) } ultimately have "strand_sem_stateful (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w) \\<^sub>s\<^sub>e\<^sub>t comb_I w') (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w)) (comb_I w') {}) (unlabel (mk_symb (\, \, I, T, \))) I" by force then show "strand_sem_stateful (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w) \\<^sub>s\<^sub>e\<^sub>t comb_I w') (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w)) (comb_I w') {}) (unlabel (mk_symb (\, \, I, T, \))) (comb_I w')" using interps_T strand_sem_model_swap[of "unlabel (mk_symb (\, \, I, T, \))" I "comb_I w'"] by force qed ultimately show "strand_sem_stateful {} {} (unlabel (T_symb w')) (comb_I w')" using strand_sem_append_stateful[ of "{}" "{}" "unlabel (T_symb w)" "unlabel (mk_symb (\, \, I, T, \))" "(comb_I w')"] unfolding w'_def by auto qed moreover have "IK \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I) = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t comb_I w')" proof - have AI: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I) = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I" by (metis ik\<^sub>s\<^sub>s\<^sub>t_subst unlabel_subst) have "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t comb_I w') = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w \\<^sub>l\<^sub>s\<^sub>s\<^sub>t comb_I w') \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb [(\, \, I, T, \)] \\<^sub>l\<^sub>s\<^sub>s\<^sub>t comb_I w')" unfolding w'_def by (simp add: subst_lsst_append) also have "... = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w \\<^sub>l\<^sub>s\<^sub>s\<^sub>t comb_I w') \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb [(\, \, I, T, \)] \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I)" by (metis T_symb_append T_symb_append' interps_T mk_symb.simps self_append_conv2 subst_apply_labeled_stateful_strand_swap) also have "... = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w \\<^sub>l\<^sub>s\<^sub>s\<^sub>t comb_I w) \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb [(\, \, I, T, \)] \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I)" by (metis interps_w subst_apply_labeled_stateful_strand_swap) also have "... = IK \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I)" using step'(3) step.hyps(6) by auto also have "... = IK \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I)" unfolding AI by auto finally show "IK \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I) = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t comb_I w')" using step'(3) step(6) T_symb.simps mk_symb.simps by auto qed moreover have "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I DB = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w')) (comb_I w') {}" proof - have "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I DB = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w)) (comb_I w) {})" using step'(4) by auto moreover have "... = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))) I (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w)) (comb_I w) {})" using step(6) by auto moreover have "... = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (mk_symb (\, \, I, T, \))) I (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w)) (comb_I w) {})" by auto moreover have "... = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (mk_symb (\, \, I, T, \))) (comb_I w') (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w)) (comb_I w) {})" by (metis (no_types, lifting) db\<^sub>s\<^sub>s\<^sub>t_subst_swap[OF interps_T] db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t empty_set) moreover have "... = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (mk_symb (\, \, I, T, \))) (comb_I w') (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w)) (comb_I w') {})" by (metis (no_types, lifting) db\<^sub>s\<^sub>s\<^sub>t_subst_swap[OF interps_w] db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t empty_set) moreover have "... = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w) @ unlabel (mk_symb (\, \, I, T, \))) (comb_I w') {}" using dbupd\<^sub>s\<^sub>s\<^sub>t_append by metis moreover have "... = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel ((T_symb w) @ mk_symb (\, \, I, T, \))) (comb_I w') {}" by auto moreover have "... = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w')) (comb_I w') {}" unfolding w'_def by auto ultimately show "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I DB = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w')) (comb_I w') {}" by auto qed moreover have "trms \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w')" proof - have "trms \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w) \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" using step'(5) by auto moreover have "... = trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w) \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" using step(6) by auto moreover have "... = trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w) \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb [(\, \, I, T, \)])" by auto moreover have "... = trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w @ T_symb [(\, \, I, T, \)])" by auto moreover have "... = trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w')" unfolding w'_def by auto ultimately show "trms \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w')" by auto qed moreover have "vars \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w')" proof - have "vars \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" using step'(6) by fastforce moreover have "... = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" using step(6) by auto moreover have "... = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb [(\, \, I, T, \)])" by auto moreover have "... = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w @ T_symb [(\, \, I, T, \)])" by auto moreover have "... = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w')" unfolding w'_def by auto ultimately show "vars \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w')" using step(6) by auto qed moreover have interp_comb_I_w': "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (comb_I w')" using interpretation_comp(1) step'(7) unfolding w'_def by auto moreover have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (comb_I w'))" proof fix t assume "t \ subst_range (comb_I w')" then have "\x. x \ subst_domain (comb_I w') \ t = comb_I w' x" by auto then obtain x where "x \ subst_domain (comb_I w')" "t = comb_I w' x" by auto then show "wf\<^sub>t\<^sub>r\<^sub>m t" by (metis (no_types, lifting) w'_def interp_comb_I_w' comb_I_append ground_subst_dom_iff_img mk_I.simps narrow_def step'(8) step.hyps(8) step.hyps(9) subst_compose_def wf_trm_Var wf_trm_subst) qed ultimately show ?case unfolding w'_def by auto qed private lemma ground_protocol_states_aux_if_reachable_constraints: assumes "A \ reachable_constraints P" assumes "constr_sem_stateful I (unlabel A)" assumes "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" assumes "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)" shows "\w. (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I, dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I {}, trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A, vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A, w) \ ground_protocol_states_aux P" using assms proof (induction rule: reachable_constraints.induct) case init then show ?case using ground_protocol_states_aux.init by auto next case (step \ T \ \ \) have "\w. (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t I, dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) I {}, trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \, vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \, w) \ ground_protocol_states_aux P" by (metis local.step(6,7,8,9) step.IH strand_sem_append_stateful unlabel_append) then obtain w where w_p: "(ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t I, dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) I {}, trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \, vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \, w) \ ground_protocol_states_aux P" by auto define w' where "w' = w@[(\, \, I, T, \)]" define \' where "\' = \@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" let ?T = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" have "T \ set P" using step.hyps(2) . moreover have "transaction_decl_subst \ T" using step.hyps(3) . moreover have "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using step.hyps(4) . moreover have "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using step.hyps(5) . moreover have "strand_sem_stateful (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t I) (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) I {}) ?T I" using step(7) strand_sem_append_stateful[of "{}" "{}" "unlabel \" ?T I] by auto moreover have "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" using assms(3) . moreover have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)" using step.prems(3) by fastforce ultimately have "((ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t I) \ (ik\<^sub>s\<^sub>s\<^sub>t ?T \\<^sub>s\<^sub>e\<^sub>t I), dbupd\<^sub>s\<^sub>s\<^sub>t ?T I (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) I {}), trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ trms\<^sub>s\<^sub>s\<^sub>t ?T, vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ vars\<^sub>s\<^sub>s\<^sub>t ?T, w@[(\, \, I, T, \)]) \ ground_protocol_states_aux P" using ground_protocol_states_aux.step[ OF w_p, of T \ \ \ "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" I] by metis moreover have "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>s\<^sub>e\<^sub>t I = (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t I) \ (ik\<^sub>s\<^sub>s\<^sub>t ?T \\<^sub>s\<^sub>e\<^sub>t I)" unfolding \'_def by auto moreover have "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \') I {} = dbupd\<^sub>s\<^sub>s\<^sub>t ?T I (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) I {})" unfolding \'_def by (simp add: dbupd\<^sub>s\<^sub>s\<^sub>t_append) moreover have "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' = trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ trms\<^sub>s\<^sub>s\<^sub>t ?T" unfolding \'_def by auto moreover have "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ vars\<^sub>s\<^sub>s\<^sub>t ?T" unfolding \'_def by auto ultimately have "(ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>s\<^sub>e\<^sub>t I, dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \') I {}, trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \', vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \', w') \ ground_protocol_states_aux P" using w'_def by auto then show ?case unfolding \'_def w'_def by auto qed private lemma protocol_model_equivalence_aux1: "{(IK, DB) | IK DB. \w trms vars. (IK, DB, trms, vars, w) \ ground_protocol_states_aux P} = {(ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I), dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I {}) | A I. A \ reachable_constraints P \ strand_sem_stateful {} {} (unlabel A) I \ interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)}" proof (rule; rule; rule) fix IK DB assume "(IK, DB) \ {(IK, DB) | IK DB. \w trms vars. (IK, DB, trms, vars, w) \ ground_protocol_states_aux P}" then have "\w trms vars. (IK, DB, trms, vars, w) \ ground_protocol_states_aux P" by auto then obtain w trms vars where "(IK, DB, trms, vars, w) \ ground_protocol_states_aux P" by auto then have reachable: "T_symb w \ reachable_constraints P" "strand_sem_stateful {} {} (unlabel (T_symb w)) (comb_I w)" "IK = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w \\<^sub>l\<^sub>s\<^sub>s\<^sub>t comb_I w)" "DB = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w)) (comb_I w) {}" "trms = trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w)" "vars = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w)" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (comb_I w)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (comb_I w))" using reachable_constraints_if_ground_protocol_states_aux[of IK DB trms vars w P] by auto then have "IK = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w \\<^sub>l\<^sub>s\<^sub>s\<^sub>t (comb_I w))" "DB = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w)) (comb_I w) {}" "T_symb w \ reachable_constraints P" "strand_sem_stateful {} {} (unlabel (T_symb w)) (comb_I w)" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (comb_I w) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (comb_I w))" by auto then show "\A I. (IK, DB) = (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I), dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I {}) \ A \ reachable_constraints P \ strand_sem_stateful {} {} (unlabel A) I \ interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)" by blast next fix IK DB assume "(IK, DB) \ {(ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I), dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I {}) | A I. A \ reachable_constraints P \ strand_sem_stateful {} {} (unlabel A) I \ interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)}" then obtain A I where A_I_p: "IK = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I)" "DB = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I {}" "A \ reachable_constraints P" "strand_sem_stateful {} {} (unlabel A) I" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)" by auto then have "\w. (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I, dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I {}, trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A, vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A, w) \ ground_protocol_states_aux P" using ground_protocol_states_aux_if_reachable_constraints[of A P I] by auto then have "\w. (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I, DB, trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A, vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A, w) \ ground_protocol_states_aux P" using A_I_p by blast then have "\w. (ik\<^sub>s\<^sub>s\<^sub>t (unlabel A \\<^sub>s\<^sub>s\<^sub>t I), DB, trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A, vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A, w) \ ground_protocol_states_aux P" by (simp add: ik\<^sub>s\<^sub>s\<^sub>t_subst) then have "(\w trms vars. (IK, DB, trms, vars, w) \ ground_protocol_states_aux P)" by (metis (no_types) A_I_p(1) unlabel_subst) then show "\IK' DB'. (IK, DB) = (IK', DB') \ (\w trms vars. (IK', DB', trms, vars, w) \ ground_protocol_states_aux P)" by auto qed subsubsection \The Protocol Model Equivalence Proof\ private lemma subst_ground_term_ident: assumes "ground_term t" shows "t \ I = t" using assms by (simp add: subst_ground_ident) private lemma subst_comp_rm_vars_eq: fixes \ :: "('fun,'atom,'sets,'lbl) prot_subst" fixes \ :: "('fun,'atom,'sets,'lbl) prot_subst" fixes I :: "('fun,'atom,'sets,'lbl) prot_subst" assumes "subst_domain \ = set X \ ground (subst_range \)" shows "(\ \\<^sub>s \) = (\ \\<^sub>s (rm_vars (set X) \))" proof (rule ext) fix x show "(\ \\<^sub>s \) x = (\ \\<^sub>s rm_vars (set X) \) x" proof (cases "x \ set X") case True have gt: "ground_term (\ x)" using True assms by auto have "(\ \\<^sub>s \) x = (\ x) \ \" using subst_compose by metis also have "... = \ x" using gt subst_ground_term_ident by blast also have "... = (\ x) \ (rm_vars (set X) \)" using gt subst_ground_term_ident by fastforce also have "... = (\ \\<^sub>s (rm_vars (set X) \)) x" using subst_compose by metis ultimately show ?thesis by auto next case False have delta_x: "\ x = Var x" using False assms by blast have "(rm_vars (set X) \) x = \ x" using False by auto have "(\ \\<^sub>s \) x = (\ x) \ \" using subst_compose by metis also have "... = (Var x) \ \" using delta_x by presburger also have "... = (Var x) \ (rm_vars (set X) \)" using False by force also have "... = (\ x) \ (rm_vars (set X) \)" using delta_x by presburger also have "... = (\ \\<^sub>s (rm_vars (set X) \)) x" using subst_compose by metis finally show ?thesis by auto qed qed private lemma subst_comp_rm_vars_commute: assumes "\x\set X. \y. \ y \ Var x" assumes "subst_range \ \ range Var" assumes "subst_domain \ = set X" assumes "ground (subst_range \)" shows "(\ \\<^sub>s (rm_vars (set X) \)) = (rm_vars (set X) \ \\<^sub>s \)" proof (rule ext) fix x show "(\ \\<^sub>s rm_vars (set X) \) x = (rm_vars (set X) \ \\<^sub>s \) x" proof (cases "x \ set X") case True then have gt: "ground_term (\ x)" using True assms(3,4) by auto have "(\ \\<^sub>s (rm_vars (set X) \)) x = \ x \ rm_vars (set X) \" by (simp add: subst_compose) also have "... = \ x" using gt by auto also have "... = ((rm_vars (set X) \) x) \ \" by (simp add: True) also have "... = (rm_vars (set X) \ \\<^sub>s \) x" by (simp add: subst_compose) finally show ?thesis . next case False have \_x: "\ x = Var x" using False assms(3) by blast obtain y where y_p: "\ x = Var y" by (meson assms(2) image_iff subsetD subst_imgI) then have "y \ set X" using assms(1) by blast then show ?thesis using assms(3,4) subst_domI False \_x y_p by (metis (mono_tags, lifting) subst_comp_notin_dom_eq subst_compose) qed qed private lemma negchecks_model_substitution_lemma_1: fixes \ :: "('fun,'atom,'sets,'lbl) prot_subst" fixes I :: "('fun,'atom,'sets,'lbl) prot_subst" assumes "negchecks_model (\ \\<^sub>s I) DB X F F'" assumes "subst_range \ \ range Var" assumes "\x\set X. \y. \ y \ Var x" shows "negchecks_model I DB X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" unfolding negchecks_model_def proof (rule, rule) fix \ :: "('fun,'atom,'sets,'lbl) prot_subst" assume a: "subst_domain \ = set X \ ground (subst_range \)" have "(\(t, s)\set F. t \ \ \\<^sub>s (\ \\<^sub>s I) \ s \ \ \\<^sub>s (\ \\<^sub>s I)) \ (\(t, s)\set F'. (t, s) \\<^sub>p \ \\<^sub>s (\ \\<^sub>s I) \ DB)" using a assms(1) unfolding negchecks_model_def by auto then show "(\(t, s)\set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \). t \ \ \\<^sub>s I \ s \ \ \\<^sub>s I) \ (\(t, s)\set (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \). (t, s) \\<^sub>p \ \\<^sub>s I \ DB)" proof assume "\(t, s)\set F. t \ \ \\<^sub>s (\ \\<^sub>s I) \ s \ \ \\<^sub>s (\ \\<^sub>s I)" then obtain t s where t_s_p: "(t, s)\set F" "t \ \ \\<^sub>s (\ \\<^sub>s I) \ s \ \ \\<^sub>s (\ \\<^sub>s I)" by auto from this(2) have "t \ \ \\<^sub>s ((rm_vars (set X) \) \\<^sub>s I) \ s \ \ \\<^sub>s ((rm_vars (set X) \) \\<^sub>s I)" using assms(3) a using subst_comp_rm_vars_eq[of \ X \] subst_compose_assoc by (metis (no_types, lifting)) then have "t \ (rm_vars (set X) \) \\<^sub>s ( \ \\<^sub>s I) \ s \ (rm_vars (set X) \) \\<^sub>s (\ \\<^sub>s I)" using subst_comp_rm_vars_commute[of X \ \, OF assms(3) assms(2)] a by (metis (no_types, lifting) subst_compose_assoc[symmetric]) then have "t \ (rm_vars (set X) \) \ ( \ \\<^sub>s I) \ s \ (rm_vars (set X) \) \ (\ \\<^sub>s I)" by auto moreover have "(t \ rm_vars (set X) \, s \ rm_vars (set X) \) \ set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" using subst_apply_pairs_pset_subst t_s_p(1) by fastforce ultimately have "\(t, s)\set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \). t \ \ \\<^sub>s I \ s \ \ \\<^sub>s I" by auto then show ?thesis by auto next assume "\(t, s)\set F'. (t, s) \\<^sub>p \ \\<^sub>s (\ \\<^sub>s I) \ DB" then obtain t s where t_s_p: "(t, s) \ set F'" "(t, s) \\<^sub>p \ \\<^sub>s (\ \\<^sub>s I) \ DB" by auto from this(2) have "(t, s) \\<^sub>p \ \\<^sub>s (rm_vars (set X) \ \\<^sub>s I) \ DB" using assms(3) a subst_comp_rm_vars_eq[OF a] by (metis (no_types, lifting) case_prod_conv subst_subst_compose) then have "(t, s) \\<^sub>p rm_vars (set X) \ \\<^sub>s (\ \\<^sub>s I) \ DB" using a subst_comp_rm_vars_commute[of X \ \, OF assms(3) assms(2)] by (metis (no_types, lifting) case_prod_conv subst_compose_assoc) then have "(t \ rm_vars (set X) \, s \ rm_vars (set X) \) \\<^sub>p \ \\<^sub>s I \ DB" by auto moreover have "(t \ rm_vars (set X) \, s \ rm_vars (set X) \) \ set (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" using t_s_p(1) subst_apply_pairs_pset_subst by fastforce ultimately have "(\(t, s) \ set (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \). (t, s) \\<^sub>p \ \\<^sub>s I \ DB)" by auto then show ?thesis by auto qed qed private lemma negchecks_model_substitution_lemma_2: fixes \ :: "('fun,'atom,'sets,'lbl) prot_subst" fixes I :: "('fun,'atom,'sets,'lbl) prot_subst" assumes "negchecks_model I DB X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" assumes "subst_range \ \ range Var" assumes "\x\set X. \y. \ y \ Var x" shows "negchecks_model (\ \\<^sub>s I) DB X F F'" unfolding negchecks_model_def proof (rule, rule) fix \ :: "('fun,'atom,'sets,'lbl) prot_subst" assume a: "subst_domain \ = set X \ ground (subst_range \)" have "(\(t, s)\set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \). t \ \ \\<^sub>s I \ s \ \ \\<^sub>s (I)) \ (\(t, s)\set (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \). (t, s) \\<^sub>p \ \\<^sub>s I \ DB)" using a assms(1)unfolding negchecks_model_def by auto then show "(\(t, s)\set F. t \ \ \\<^sub>s (\ \\<^sub>s I) \ s \ \ \\<^sub>s (\ \\<^sub>s I)) \ (\(t, s)\set F'. (t, s) \\<^sub>p \ \\<^sub>s (\ \\<^sub>s I) \ DB)" proof assume "\(t, s)\set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \). t \ \ \\<^sub>s I \ s \ \ \\<^sub>s (I)" then obtain t s where t_s_p: "(t, s) \ set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" "t \ \ \\<^sub>s I \ s \ \ \\<^sub>s I" by auto then have "\t' s'. t = t' \ rm_vars (set X) \ \ s = s' \ rm_vars (set X) \ \ (t',s') \ set F" unfolding subst_apply_pairs_def by auto then obtain t' s' where t'_s'_p: "t = t' \ rm_vars (set X) \" "s = s' \ rm_vars (set X) \" "(t',s') \ set F" by auto then have "t' \ rm_vars (set X) \ \ \ \\<^sub>s I \ s' \ rm_vars (set X) \ \ \ \\<^sub>s I" using t_s_p by auto then have "t' \ \ \ rm_vars (set X) \ \\<^sub>s I \ s' \ \ \ rm_vars (set X) \ \\<^sub>s I" using a subst_comp_rm_vars_commute[OF assms(3,2)] by (metis (no_types, lifting) subst_subst) then have "t' \ \ \ \ \\<^sub>s I \ s' \ \ \ \ \\<^sub>s I" using subst_comp_rm_vars_eq[OF a] by (metis (no_types, lifting) subst_subst) moreover from t_s_p(1) have "(t', s') \ set F" using subst_apply_pairs_pset_subst t'_s'_p by fastforce ultimately have "\(t, s)\set F. t \ \ \\<^sub>s (\ \\<^sub>s I) \ s \ \ \\<^sub>s (\ \\<^sub>s I)" by auto then show ?thesis by auto next assume "\(t, s)\set (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \). (t, s) \\<^sub>p \ \\<^sub>s I \ DB" then obtain t s where t_s_p: "(t, s) \ set (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" "(t \ \ \\<^sub>s I, s \ \ \\<^sub>s I) \ DB" by auto then have "\t' s'. t = t' \ rm_vars (set X) \ \ s = s' \ rm_vars (set X) \ \ (t',s') \ set F'" unfolding subst_apply_pairs_def by auto then obtain t' s' where t'_s'_p: "t = t' \ rm_vars (set X) \" "s = s' \ rm_vars (set X) \" "(t',s') \ set F'" by auto then have "(t' \ rm_vars (set X) \ \ \ \\<^sub>s I, s' \ rm_vars (set X) \ \ \ \\<^sub>s I) \ DB" using t_s_p by auto then have "(t' \ \ \ rm_vars (set X) \ \\<^sub>s I, s' \ \ \ rm_vars (set X) \ \\<^sub>s I) \ DB" using a subst_comp_rm_vars_commute[OF assms(3,2)] by (metis (no_types, lifting) subst_subst) then have "(t' \ \ \ \ \\<^sub>s I , s' \ \ \ \ \\<^sub>s I) \ DB" using subst_comp_rm_vars_eq[OF a] by (metis (no_types, lifting) subst_subst) moreover from t_s_p(1) have "(t', s') \ set F'" using subst_apply_pairs_pset_subst t'_s'_p by fastforce ultimately have "\(t, s)\set F'. (t, s) \\<^sub>p \ \\<^sub>s (\ \\<^sub>s I) \ DB" by auto then show ?thesis by auto qed qed private lemma negchecks_model_substitution_lemma: fixes \ :: "('fun,'atom,'sets,'lbl) prot_subst" fixes I :: "('fun,'atom,'sets,'lbl) prot_subst" assumes "subst_range \ \ range Var" assumes "\x\set X. \y. \ y \ Var x" shows "negchecks_model (\ \\<^sub>s I) DB X F F' \ negchecks_model I DB X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" using assms negchecks_model_substitution_lemma_1[of \ I DB X F F'] negchecks_model_substitution_lemma_2[of I DB X F \ F'] assms by auto private lemma strand_sem_stateful_substitution_lemma: fixes \ :: "('fun,'atom,'sets,'lbl) prot_subst" fixes I :: "('fun,'atom,'sets,'lbl) prot_subst" assumes "subst_range \ \ range Var" assumes "\x \ bvars\<^sub>s\<^sub>s\<^sub>t T. \y. \ y \ Var x" shows "strand_sem_stateful IK DB (T \\<^sub>s\<^sub>s\<^sub>t \) I = strand_sem_stateful IK DB T (\ \\<^sub>s I)" using assms proof (induction T arbitrary: IK DB) case Nil then show ?case by auto next case (Cons a T) then show ?case proof (induction a) case (Receive ts) have "((\x. x \ \ \ I) ` (set ts)) \ IK = ((\t. t \ \) ` set ts \\<^sub>s\<^sub>e\<^sub>t I) \ IK" by blast then show ?case using Receive by (force simp add: subst_sst_cons) next case (NegChecks X F F') have bounds: "\x\bvars\<^sub>s\<^sub>s\<^sub>t T. \y. \ y \ Var x" using NegChecks by auto have "\x\bvars\<^sub>s\<^sub>s\<^sub>t ([\X\\\: F \\: F'\]). \y. \ y \ Var x" using NegChecks by auto then have bounds2: "\x\set X. \y. \ y \ Var x" by simp have "negchecks_model I DB X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \ negchecks_model (\ \\<^sub>s I) DB X F F'" using NegChecks.prems(2) bounds2 negchecks_model_substitution_lemma by blast moreover have "strand_sem_stateful IK DB (T \\<^sub>s\<^sub>s\<^sub>t \) I \ strand_sem_stateful IK DB T (\ \\<^sub>s I)" using Cons NegChecks(2) bounds by blast ultimately show ?case by (simp add: subst_sst_cons) qed (force simp add: subst_sst_cons)+ qed private lemma ground_subst_rm_vars_subst_compose_dist: assumes "ground (subst_range \\)" shows "(rm_vars (set X) (\\ \\<^sub>s \)) = (rm_vars (set X) \\ \\<^sub>s rm_vars (set X) \)" proof (rule ext) fix x show "rm_vars (set X) (\\ \\<^sub>s \) x = (rm_vars (set X) \\ \\<^sub>s rm_vars (set X) \) x" proof (cases "x \ set X") case True then show ?thesis by (simp add: subst_compose) next case False note False_outer = False show ?thesis proof (cases "x \ subst_domain \\") case True then show ?thesis by (metis (mono_tags, lifting) False assms ground_subst_range_empty_fv subst_ground_ident_compose(1)) next case False have "\\ x = Var x" using False by blast then show ?thesis using False_outer by (simp add: subst_compose) qed qed qed private lemma stateful_strand_ground_subst_comp: assumes "ground (subst_range \\)" shows "T \\<^sub>s\<^sub>s\<^sub>t \\ \\<^sub>s \ = (T \\<^sub>s\<^sub>s\<^sub>t \\) \\<^sub>s\<^sub>s\<^sub>t \" using assms by (meson disjoint_iff ground_subst_no_var stateful_strand_subst_comp) private lemma labelled_stateful_strand_ground_subst_comp: assumes "ground (subst_range \\)" shows "T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \\ \\<^sub>s \ = (T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \\) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using assms by (metis Int_empty_left ground_range_vars labeled_stateful_strand_subst_comp) private lemma transaction_fresh_subst_ground_subst_range: assumes "transaction_fresh_subst \ T trms" shows "ground (subst_range \)" using assms by (metis range_vars_alt_def transaction_fresh_subst_range_vars_empty) private lemma transaction_decl_subst_ground_subst_range: assumes "transaction_decl_subst \ T" shows "ground (subst_range \)" proof - have \_ground: "\x \ subst_domain \. ground_term (\ x)" using assms transaction_decl_subst_domain transaction_decl_subst_grounds_domain by force show ?thesis proof (rule ccontr) assume "fv\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ {}" then have "\x \ subst_domain \. fv (\ x) \ {}" by auto then obtain x where x_p: "x \ subst_domain \ \ fv (\ x) \ {}" by meson moreover have "ground_term (\ x)" using \_ground x_p by auto ultimately show "False" by auto qed qed private lemma fresh_transaction_decl_subst_ground_subst_range: assumes "transaction_fresh_subst \ T trms" assumes "transaction_decl_subst \ T" shows "ground (subst_range (\ \\<^sub>s \))" proof - have "ground (subst_range \)" using assms transaction_decl_subst_ground_subst_range by blast moreover have "ground (subst_range \)" using assms using transaction_fresh_subst_ground_subst_range by blast ultimately show "ground (subst_range (\ \\<^sub>s \))" by (metis (no_types, opaque_lifting) Diff_iff all_not_in_conv empty_iff empty_subsetI range_vars_alt_def range_vars_subst_compose_subset subset_antisym sup_bot.right_neutral) qed private lemma strand_sem_stateful_substitution_lemma': assumes "transaction_renaming_subst \ P vars" assumes "transaction_fresh_subst \ T trms" assumes "transaction_decl_subst \ T" assumes "finite vars" assumes "T \ set P" shows "strand_sem_stateful IK DB (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))) I \ strand_sem_stateful IK DB (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \))) (\ \\<^sub>s I)" proof - have \_Var: "subst_range \ \ range Var" using assms(1) transaction_renaming_subst_is_renaming(5) by blast have "(\x\vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T). \y. \ y \ Var x)" using assms(4,2) transaction_renaming_subst_vars_transaction_neq assms(1) assms(5) by blast then have "(\x\bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T). \y. \ y \ Var x)" by (metis UnCI vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t) then have T_Vars: "(\x\bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)). \y. \ y \ Var x)" by (metis bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst bvars\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq) have ground_\_\: "ground (subst_range (\ \\<^sub>s \))" using fresh_transaction_decl_subst_ground_subst_range using assms(2) assms(3) by blast from assms(1) ground_\_\ have "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \) = unlabel ((dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using stateful_strand_ground_subst_comp[of _ "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T))"] by (simp add: dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst unlabel_subst) then show ?thesis using strand_sem_stateful_substitution_lemma \_Var T_Vars by (metis dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst subst_lsst_unlabel) qed inductive_set ground_protocol_states:: "('fun,'atom,'sets,'lbl) prot \ (('fun,'atom,'sets,'lbl) prot_terms \ (('fun,'atom,'sets,'lbl) prot_term \ ('fun,'atom,'sets,'lbl) prot_term) set \ _ set ) set" (* TODO: write up the type nicer *) for P::"('fun,'atom,'sets,'lbl) prot" where init: "({},{},{}) \ ground_protocol_states P" | step: "\(IK,DB,consts) \ ground_protocol_states P; T \ set P; transaction_decl_subst \ T; transaction_fresh_subst \ T consts; A = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \); strand_sem_stateful IK DB (unlabel A) I; interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I; wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I) \ \ (IK \ ((ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \\<^sub>s\<^sub>e\<^sub>t I), dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I DB, consts \ {t. t \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ (\c. t = Fun c [])}) \ ground_protocol_states P" private lemma transaction_fresh_subst_consts_of_iff: "transaction_fresh_subst \ T (consts_of trms) \ transaction_fresh_subst \ T trms" proof (cases "transaction_fresh_subst \ T (consts_of trms) \ transaction_fresh_subst \ T trms") case True then have "\t \ subst_range \. \c. t = Fun c []" unfolding transaction_fresh_subst_def by auto have "(\t \ subst_range \. t \ subterms\<^sub>s\<^sub>e\<^sub>t (consts_of trms) \ t \ subterms\<^sub>s\<^sub>e\<^sub>t trms)" proof fix t assume "t \ subst_range \" then obtain c where c_p: "t = Fun c []" using \\t\subst_range \. \c. t = Fun c []\ by blast have "Fun c [] \ subterms\<^sub>s\<^sub>e\<^sub>t (consts_of trms) \ Fun c [] \ subterms\<^sub>s\<^sub>e\<^sub>t trms" unfolding consts_of_def by auto then show "t \ subterms\<^sub>s\<^sub>e\<^sub>t (consts_of trms) \ t \ subterms\<^sub>s\<^sub>e\<^sub>t trms" using c_p by auto qed then show ?thesis using transaction_fresh_subst_def by force next case False then show ?thesis by auto qed private lemma transaction_renaming_subst_inv: assumes "transaction_renaming_subst \ P X " shows "\\inv. \ \\<^sub>s \inv = Var \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \inv)" using var_rename_inv_comp transaction_renaming_subst_def assms subst_apply_term_empty subst_term_eqI by (metis var_rename_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_range(2)) private lemma consts_of_union_distr: "consts_of (trms1 \ trms2) = consts_of trms1 \ consts_of trms2" unfolding consts_of_def by auto private lemma ground_protocol_states_aux_finite_vars: assumes "(IK,DB,trms,vars,w) \ ground_protocol_states_aux P" shows "finite vars" using assms by (induction rule: ground_protocol_states_aux.induct) auto private lemma dbupd\<^sub>s\<^sub>s\<^sub>t_substitution_lemma: "dbupd\<^sub>s\<^sub>s\<^sub>t T (\ \\<^sub>s I) DB = dbupd\<^sub>s\<^sub>s\<^sub>t (T \\<^sub>s\<^sub>s\<^sub>t \) I DB" proof (induction T arbitrary: DB) case Nil then show ?case by auto next case (Cons a T) then show ?case by (induction a) (simp_all add: subst_apply_stateful_strand_def) qed private lemma subst_Var_const_subterm_subst: assumes "subst_range \ \ range Var" shows "Fun c [] \ t \ Fun c [] \ t \ \" using assms proof (induction t) case (Var x) then show ?case by (metis is_Var_def subtermeq_Var_const(1) term.discI(2) var_renaming_is_Fun_iff) next case (Fun f ts) then show ?case by auto qed private lemma subst_Var_consts_of: assumes "subst_range \ \ range Var" shows "consts_of T = consts_of (T \\<^sub>s\<^sub>e\<^sub>t \)" proof (rule antisym; rule subsetI) fix x assume "x \ consts_of T" then obtain t c where t_c_p: "t \ T \ x \ t \ x = Fun c []" unfolding consts_of_def by auto moreover have "x \ t \ \" using t_c_p by (meson assms subst_Var_const_subterm_subst) ultimately show "x \ consts_of (T \\<^sub>s\<^sub>e\<^sub>t \)" unfolding consts_of_def by auto next fix x assume "x \ consts_of (T \\<^sub>s\<^sub>e\<^sub>t \)" then obtain t c where t_c_p: "t \ T \ x \ t \ \ \ x = Fun c []" unfolding consts_of_def by auto moreover have "x \ t" using t_c_p by (meson assms subst_Var_const_subterm_subst) ultimately show "x \ consts_of T" unfolding consts_of_def by auto qed private lemma fst_set_subst_apply_set: "fst ` set F \\<^sub>s\<^sub>e\<^sub>t \ = fst ` (set F \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \)" by (induction F) auto private lemma snd_set_subst_apply_set: "snd ` set F \\<^sub>s\<^sub>e\<^sub>t \ = snd ` (set F \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \)" by (induction F) auto private lemma trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_fst_snd: "trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F = fst ` set F \ snd ` set F" by (auto simp add: rev_image_eqI) private lemma consts_of_trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_range_Var: assumes "subst_range \ \ range Var" shows "consts_of (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p a) = consts_of (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \))" using assms proof (induction a) case (NegChecks X F F') have \_subs_rng_Var: "subst_range (rm_vars (set X) \) \ range Var" using assms by auto have "consts_of (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) = consts_of (fst ` set F \ snd ` set F)" using trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_fst_snd by metis also have "... = consts_of (fst ` set F) \ consts_of (snd ` set F)" using consts_of_union_distr by blast also have "... = consts_of ((fst ` set F) \\<^sub>s\<^sub>e\<^sub>t rm_vars (set X) \) \ consts_of ((snd ` set F) \\<^sub>s\<^sub>e\<^sub>t rm_vars (set X) \)" using \_subs_rng_Var subst_Var_consts_of[of "rm_vars (set X) \" "fst ` set F"] subst_Var_consts_of[of "rm_vars (set X) \" "snd ` set F"] by auto also have "... = consts_of (((fst ` set F) \\<^sub>s\<^sub>e\<^sub>t rm_vars (set X) \) \ ((snd ` set F) \\<^sub>s\<^sub>e\<^sub>t rm_vars (set X) \))" using consts_of_union_distr by auto also have "... = consts_of (fst ` set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \ snd ` set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \))" unfolding subst_apply_pairs_def fst_set_subst_apply_set snd_set_subst_apply_set by simp also have "... = consts_of (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \))" using trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_fst_snd[of "F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \"] by metis finally have 1: "consts_of (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) = consts_of (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \))" by auto have "consts_of (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F') = consts_of (fst ` set F' \ snd ` set F')" using trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_fst_snd by metis also have "... = consts_of (fst ` set F') \ consts_of (snd ` set F')" using consts_of_union_distr by blast also have "... = consts_of ((fst ` set F') \\<^sub>s\<^sub>e\<^sub>t rm_vars (set X) \) \ consts_of ((snd ` set F') \\<^sub>s\<^sub>e\<^sub>t rm_vars (set X) \)" using subst_Var_consts_of[of "rm_vars (set X) \" "fst ` set F'"] \_subs_rng_Var subst_Var_consts_of[of "rm_vars (set X) \" "snd ` set F'"] by auto also have "... = consts_of ((fst ` set F' \\<^sub>s\<^sub>e\<^sub>t rm_vars (set X) \) \ (snd ` set F' \\<^sub>s\<^sub>e\<^sub>t rm_vars (set X) \))" using consts_of_union_distr by auto also have "... = consts_of (fst ` set (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \ snd ` set (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \))" unfolding subst_apply_pairs_def fst_set_subst_apply_set snd_set_subst_apply_set by simp also have "... = consts_of (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \))" using trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_fst_snd[of "F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \"] by metis finally have 2: "consts_of (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F') = consts_of (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \))" by auto show ?case using 1 2 by (simp add: consts_of_union_distr) qed (use subst_Var_consts_of[of _ "{_, _}", OF assms] subst_Var_consts_of[OF assms] in auto) private lemma consts_of_trms\<^sub>s\<^sub>s\<^sub>t_range_Var: assumes "subst_range \ \ range Var" shows "consts_of (trms\<^sub>s\<^sub>s\<^sub>t T) = consts_of (trms\<^sub>s\<^sub>s\<^sub>t (T \\<^sub>s\<^sub>s\<^sub>t \))" proof (induction T) case Nil then show ?case by auto next case (Cons a T) have "consts_of (trms\<^sub>s\<^sub>s\<^sub>t (a # T)) = consts_of (trms\<^sub>s\<^sub>s\<^sub>t [a] \ trms\<^sub>s\<^sub>s\<^sub>t T)" by simp also have "... = consts_of (trms\<^sub>s\<^sub>s\<^sub>t [a]) \ consts_of (trms\<^sub>s\<^sub>s\<^sub>t T)" using consts_of_union_distr by simp also have "... = consts_of (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p a) \ consts_of (trms\<^sub>s\<^sub>s\<^sub>t T)" by simp also have "... = consts_of (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)) \ consts_of (trms\<^sub>s\<^sub>s\<^sub>t T)" using consts_of_trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_range_Var[OF assms] by simp also have "... = consts_of (trms\<^sub>s\<^sub>s\<^sub>t ([a] \\<^sub>s\<^sub>s\<^sub>t \)) \ consts_of (trms\<^sub>s\<^sub>s\<^sub>t T)" by (simp add: subst_apply_stateful_strand_def) also have "... = consts_of (trms\<^sub>s\<^sub>s\<^sub>t ([a] \\<^sub>s\<^sub>s\<^sub>t \)) \ consts_of (trms\<^sub>s\<^sub>s\<^sub>t (T \\<^sub>s\<^sub>s\<^sub>t \))" using local.Cons by simp also have "... = consts_of (trms\<^sub>s\<^sub>s\<^sub>t (a # T \\<^sub>s\<^sub>s\<^sub>t \))" by (simp add: consts_of_union_distr subst_sst_cons) finally show ?case by simp qed private lemma consts_of_trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t_range_Var: assumes "subst_range \ \ range Var" shows "consts_of (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T) = consts_of (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using consts_of_trms\<^sub>s\<^sub>s\<^sub>t_range_Var[of \ "unlabel T"] by (metis assms unlabel_subst) private lemma transaction_renaming_subst_range: assumes "transaction_renaming_subst \ P vars" shows "subst_range \ \ range Var" using assms unfolding transaction_renaming_subst_def var_rename_def by auto private lemma protocol_models_equiv3': assumes "(IK,DB,trms,vars,w) \ ground_protocol_states_aux P" shows "(IK,DB, consts_of trms) \ ground_protocol_states P" using assms proof (induction rule: ground_protocol_states_aux.induct) case init then show ?case using ground_protocol_states.init unfolding consts_of_def by force next case (step IK DB trms vars w T \ \ \ A I) have fin_vars: "finite vars" using ground_protocol_states_aux_finite_vars step by auto have ground_\\: "ground (subst_range (\ \\<^sub>s \))" using fresh_transaction_decl_subst_ground_subst_range using step.hyps(3) step.hyps(4) by blast have \_Var: "subst_range \ \ range Var" using step(5) transaction_renaming_subst_range by metis define I' where "I' = \ \\<^sub>s I" define A' where "A' = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)" have "(IK, DB, consts_of trms) \ ground_protocol_states P" using step by force moreover have T_in_P: "T \ set P" using step by force moreover have "transaction_decl_subst \ T" using step by force moreover have "transaction_fresh_subst \ T (consts_of trms)" using step transaction_fresh_subst_consts_of_iff by force moreover have "A' = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)" using A'_def . moreover have "strand_sem_stateful IK DB (unlabel A') I'" using step(7) step(4) step(5) step(3) T_in_P fin_vars unfolding A'_def I'_def step(6) using strand_sem_stateful_substitution_lemma' by auto moreover have "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I'" using step(8) unfolding I'_def by (meson interpretation_comp(1)) moreover have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I')" using step(9) unfolding I'_def using step.hyps(5) transaction_renaming_subst_range_wf_trms wf_trms_subst_compose by blast ultimately have "(IK \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A' \\<^sub>s\<^sub>e\<^sub>t I'), dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A') I' DB, consts_of trms \ consts_of (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A')) \ ground_protocol_states P" using ground_protocol_states.step[of IK DB "consts_of trms" P T \ \ A' I'] unfolding consts_of_def by blast moreover have "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A' \\<^sub>s\<^sub>e\<^sub>t I' = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I" proof - have "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A' \\<^sub>s\<^sub>e\<^sub>t I' = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A' \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s I" unfolding A'_def I'_def step(6) by auto also have "... = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((A' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I)" unfolding unlabel_subst[symmetric] ik\<^sub>s\<^sub>s\<^sub>t_subst by auto also have "... = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I)" unfolding A'_def step(6) using labelled_stateful_strand_ground_subst_comp[of _ "transaction_strand T", OF ground_\\] by (simp add: dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst) also have "... = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I" by (metis ik\<^sub>s\<^sub>s\<^sub>t_subst unlabel_subst) finally show "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A' \\<^sub>s\<^sub>e\<^sub>t I' = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I" by auto qed moreover have "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A') I' DB = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I DB" proof - have "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A') I' DB = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A') (\ \\<^sub>s I) DB" unfolding A'_def I'_def step(6) using step by auto also have "... = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A' \\<^sub>s\<^sub>s\<^sub>t \) I DB" using dbupd\<^sub>s\<^sub>s\<^sub>t_substitution_lemma by metis also have "... = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I DB" unfolding A'_def step(6) using stateful_strand_ground_subst_comp[of _ "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T))"] ground_\\ by (simp add: dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst_unlabel) finally show "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A') I' DB = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I DB" by auto qed moreover have "consts_of (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A') = consts_of (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" by (metis (no_types, lifting) A'_def \_Var consts_of_trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t_range_Var ground_\\ labelled_stateful_strand_ground_subst_comp step.hyps(6) trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq) ultimately show ?case using consts_of_union_distr by metis qed private lemma protocol_models_equiv4': assumes "(IK, DB, csts) \ ground_protocol_states P" shows "\trms w vars. (IK,DB,trms,vars,w) \ ground_protocol_states_aux P \ csts = consts_of trms \ vars = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w)" using assms proof (induction rule: ground_protocol_states.induct) case init have "({}, {}, {}, {}, []) \ ground_protocol_states_aux P" using ground_protocol_states_aux.init by blast moreover have "{} = consts_of {}" unfolding consts_of_def by auto moreover have "{} = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb [])" by auto ultimately show ?case by metis next case (step IK DB "consts" T \ \ A I) then obtain trms w vars where trms_w_vars_p: "(IK, DB, trms, vars, w) \ ground_protocol_states_aux P" "consts = consts_of trms" "vars = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w)" by auto have "\\. transaction_renaming_subst \ P vars" unfolding transaction_renaming_subst_def by blast then obtain \ :: "('fun,'atom,'sets,'lbl) prot_subst" where \_p: "transaction_renaming_subst \ P vars" by blast then obtain \inv where \inv_p: "\ \\<^sub>s \inv = Var \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \inv)" using transaction_renaming_subst_inv[of \ P vars] by auto define A' where "A' = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" define I' where "I' = \inv \\<^sub>s I" define trms' where "trms' = trms \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A'" define vars' where "vars' = vars \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A'" define w' where "w' = w @ [(\, \, I', T, \)]" define IK' where "IK' = IK \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I)" define DB' where "DB' = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I DB" have P_state: "(IK', DB' , trms', vars', w') \ ground_protocol_states_aux P" proof - have 1: "(IK, DB, trms, vars, w) \ ground_protocol_states_aux P" using \(IK, DB, trms, vars, w) \ ground_protocol_states_aux P\ by blast moreover have "T \ set P" using step(2) . moreover have "transaction_decl_subst \ T" using step(3) . moreover have fresh_\: "transaction_fresh_subst \ T trms" using step(4) trms_w_vars_p(2) using transaction_fresh_subst_consts_of_iff by auto moreover have "transaction_renaming_subst \ P vars" using \_p . moreover have "A' = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" unfolding A'_def by auto moreover have "strand_sem_stateful IK DB (unlabel A') I'" proof - have fin_vars: "finite vars" using 1 ground_protocol_states_aux_finite_vars by blast show "strand_sem_stateful IK DB (unlabel A') I'" using step(6) strand_sem_stateful_substitution_lemma'[OF \_p fresh_\ step(3) fin_vars] step(2) unfolding A'_def step(5) by (metis (no_types, lifting) I'_def \inv_p subst_compose_assoc var_comp(2)) qed moreover have "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I'" using step(7) by (simp add: I'_def interpretation_substI subst_compose) moreover have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I')" using I'_def \inv_p step.hyps(8) wf_trms_subst_compose by blast moreover have "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I DB = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A') I' DB" proof - have "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I DB = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) (\ \\<^sub>s \inv \\<^sub>s I) DB" by (simp add: \inv_p) also have "... = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A') (\inv \\<^sub>s I) DB" unfolding A'_def step(5) by (metis (no_types, lifting) dbupd\<^sub>s\<^sub>s\<^sub>t_substitution_lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst_unlabel subst_compose_assoc) also have "... = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A') I' DB" unfolding A'_def I'_def by auto finally show ?thesis by auto qed moreover have "IK \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I) = IK \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A' \\<^sub>s\<^sub>e\<^sub>t I')" proof - have "IK \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I) = IK \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t (\ \\<^sub>s \inv) \\<^sub>s I)" using \inv_p by auto also have "... = IK \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A' \\<^sub>s\<^sub>e\<^sub>t I')" unfolding A'_def step(5) unlabel_subst[symmetric] ik\<^sub>s\<^sub>s\<^sub>t_subst dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst I'_def by auto finally show ?thesis by auto qed ultimately show "(IK', DB' , trms', vars', w') \ ground_protocol_states_aux P" using ground_protocol_states_aux.step[of IK DB trms vars w P T \ \ \ A' I'] unfolding trms'_def vars'_def w'_def IK'_def DB'_def by auto qed moreover have "consts \ consts_of (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) = consts_of trms'" proof - have \_Var: "subst_range \ \ range Var" using \_p transaction_renaming_subst_range by blast have ground_\\: "ground (subst_range (\ \\<^sub>s \))" using fresh_transaction_decl_subst_ground_subst_range using step.hyps(3) step.hyps(4) by blast have "consts \ consts_of (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) = (consts_of trms) \ consts_of (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" using trms_w_vars_p(2) by blast also have "... = (consts_of trms) \ consts_of (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using consts_of_trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t_range_Var[of \, OF \_Var, of A] by auto also have "... = (consts_of trms) \ consts_of (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A')" using step(5) A'_def ground_\\ using labelled_stateful_strand_ground_subst_comp[of _ "(dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T))"] by (simp add: dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst) also have "... = consts_of (trms \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A')" using consts_of_union_distr by blast also have "... = consts_of trms'" unfolding trms'_def by auto finally show "?thesis" by auto qed moreover have "vars' = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w')" using P_state reachable_constraints_if_ground_protocol_states_aux by auto ultimately show ?case unfolding DB'_def IK'_def consts_of_def[symmetric] by metis qed private lemma protocol_model_equivalence_aux2: "{(IK, DB) | IK DB. \csts. (IK, DB, csts) \ ground_protocol_states P} = {(IK, DB) | IK DB. \w trms vars. (IK, DB, trms, vars, w) \ ground_protocol_states_aux P}" using protocol_models_equiv4' protocol_models_equiv3' by meson theorem protocol_model_equivalence: "{(IK, DB) | IK DB. \csts. (IK, DB, csts) \ ground_protocol_states P} = {(ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I), dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I {}) | A I. A \ reachable_constraints P \ strand_sem_stateful {} {} (unlabel A) I \ interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)}" using protocol_model_equivalence_aux2 protocol_model_equivalence_aux1 by auto end end end diff --git a/thys/Automated_Stateful_Protocol_Verification/Stateful_Protocol_Verification.thy b/thys/Automated_Stateful_Protocol_Verification/Stateful_Protocol_Verification.thy --- a/thys/Automated_Stateful_Protocol_Verification/Stateful_Protocol_Verification.thy +++ b/thys/Automated_Stateful_Protocol_Verification/Stateful_Protocol_Verification.thy @@ -1,11391 +1,11391 @@ (* Title: Stateful_Protocol_Verification.thy Author: Andreas Viktor Hess, DTU Author: Sebastian A. Mödersheim, DTU Author: Achim D. Brucker, University of Exeter Author: Anders Schlichtkrull, DTU SPDX-License-Identifier: BSD-3-Clause *) section\Stateful Protocol Verification\ theory Stateful_Protocol_Verification imports Stateful_Protocol_Model Term_Implication begin subsection \Fixed-Point Intruder Deduction Lemma\ context stateful_protocol_model begin abbreviation pubval_terms::"('fun,'atom,'sets,'lbl) prot_terms" where "pubval_terms \ {t. \f \ funs_term t. is_PubConstValue f}" abbreviation abs_terms::"('fun,'atom,'sets,'lbl) prot_terms" where "abs_terms \ {t. \f \ funs_term t. is_Abs f}" definition intruder_deduct_GSMP:: "[('fun,'atom,'sets,'lbl) prot_terms, ('fun,'atom,'sets,'lbl) prot_terms, ('fun,'atom,'sets,'lbl) prot_term] \ bool" ("\_;_\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P _" 50) where "\M; T\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P t \ intruder_deduct_restricted M (\t. t \ GSMP T - (pubval_terms \ abs_terms)) t" lemma intruder_deduct_GSMP_induct[consumes 1, case_names AxiomH ComposeH DecomposeH]: assumes "\M; T\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P t" "\t. t \ M \ P M t" "\S f. \length S = arity f; public f; \s. s \ set S \ \M; T\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P s; \s. s \ set S \ P M s; Fun f S \ GSMP T - (pubval_terms \ abs_terms) \ \ P M (Fun f S)" "\t K T' t\<^sub>i. \\M; T\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P t; P M t; Ana t = (K, T'); \k. k \ set K \ \M; T\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P k; \k. k \ set K \ P M k; t\<^sub>i \ set T'\ \ P M t\<^sub>i" shows "P M t" proof - let ?Q = "\t. t \ GSMP T - (pubval_terms \ abs_terms)" show ?thesis using intruder_deduct_restricted_induct[of M ?Q t "\M Q t. P M t"] assms unfolding intruder_deduct_GSMP_def by blast qed lemma pubval_terms_subst: assumes "t \ \ \ pubval_terms" "\ ` fv t \ pubval_terms = {}" shows "t \ pubval_terms" using assms(1,2) proof (induction t) case (Fun f T) let ?P = "\f. is_PubConstValue f" from Fun show ?case proof (cases "?P f") case False then obtain t where t: "t \ set T" "t \ \ \ pubval_terms" using Fun.prems by auto hence "\ ` fv t \ pubval_terms = {}" using Fun.prems(2) by auto thus ?thesis using Fun.IH[OF t] t(1) by auto qed force qed simp lemma abs_terms_subst: assumes "t \ \ \ abs_terms" "\ ` fv t \ abs_terms = {}" shows "t \ abs_terms" using assms(1,2) proof (induction t) case (Fun f T) let ?P = "\f. is_Abs f" from Fun show ?case proof (cases "?P f") case False then obtain t where t: "t \ set T" "t \ \ \ abs_terms" using Fun.prems by auto hence "\ ` fv t \ abs_terms = {}" using Fun.prems(2) by auto thus ?thesis using Fun.IH[OF t] t(1) by auto qed force qed simp lemma pubval_terms_subst': assumes "t \ \ \ pubval_terms" "\n. PubConst Value n \ \(funs_term ` (\ ` fv t))" shows "t \ pubval_terms" proof - have False when fs: "f \ funs_term s" "s \ subterms\<^sub>s\<^sub>e\<^sub>t (\ ` fv t)" "is_PubConstValue f" for f s proof - - obtain T where T: "Fun f T \ subterms s" using funs_term_Fun_subterm[OF fs(1)] by moura + obtain T where T: "Fun f T \ subterms s" using funs_term_Fun_subterm[OF fs(1)] by force hence "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (\ ` fv t)" using fs(2) in_subterms_subset_Union by blast thus ?thesis using assms(2) funs_term_Fun_subterm'[of f T] fs(3) unfolding is_PubConstValue_def by (cases f) force+ qed thus ?thesis using pubval_terms_subst[OF assms(1)] by auto qed lemma abs_terms_subst': assumes "t \ \ \ abs_terms" "\n. Abs n \ \(funs_term ` (\ ` fv t))" shows "t \ abs_terms" proof - have "\is_Abs f" when fs: "f \ funs_term s" "s \ subterms\<^sub>s\<^sub>e\<^sub>t (\ ` fv t)" for f s proof - - obtain T where T: "Fun f T \ subterms s" using funs_term_Fun_subterm[OF fs(1)] by moura + obtain T where T: "Fun f T \ subterms s" using funs_term_Fun_subterm[OF fs(1)] by force hence "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (\ ` fv t)" using fs(2) in_subterms_subset_Union by blast thus ?thesis using assms(2) funs_term_Fun_subterm'[of f T] by (cases f) auto qed thus ?thesis using abs_terms_subst[OF assms(1)] by force qed lemma pubval_terms_subst_range_disj: "subst_range \ \ pubval_terms = {} \ \ ` fv t \ pubval_terms = {}" proof (induction t) case (Var x) thus ?case by (cases "x \ subst_domain \") auto qed auto lemma abs_terms_subst_range_disj: "subst_range \ \ abs_terms = {} \ \ ` fv t \ abs_terms = {}" proof (induction t) case (Var x) thus ?case by (cases "x \ subst_domain \") auto qed auto lemma pubval_terms_subst_range_comp: assumes "subst_range \ \ pubval_terms = {}" "subst_range \ \ pubval_terms = {}" shows "subst_range (\ \\<^sub>s \) \ pubval_terms = {}" proof - { fix t f assume t: "t \ subst_range (\ \\<^sub>s \)" "f \ funs_term t" "is_PubConstValue f" then obtain x where x: "(\ \\<^sub>s \) x = t" by auto have "\ x \ pubval_terms" using assms(1) by (cases "\ x \ subst_range \") force+ hence "(\ \\<^sub>s \) x \ pubval_terms" using assms(2) pubval_terms_subst[of "\ x" \] pubval_terms_subst_range_disj by (metis (mono_tags, lifting) subst_compose_def) hence False using t(2,3) x by blast } thus ?thesis by fast qed lemma pubval_terms_subst_range_comp': assumes "(\ ` X) \ pubval_terms = {}" "(\ ` fv\<^sub>s\<^sub>e\<^sub>t (\ ` X)) \ pubval_terms = {}" shows "((\ \\<^sub>s \) ` X) \ pubval_terms = {}" proof - { fix t f assume t: "t \ (\ \\<^sub>s \) ` X" "f \ funs_term t" "is_PubConstValue f" then obtain x where x: "(\ \\<^sub>s \) x = t" "x \ X" by auto have "\ x \ pubval_terms" using assms(1) x(2) by force moreover have "fv (\ x) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` X)" using x(2) by (auto simp add: fv_subset) hence "\ ` fv (\ x) \ pubval_terms = {}" using assms(2) by auto ultimately have "(\ \\<^sub>s \) x \ pubval_terms" using pubval_terms_subst[of "\ x" \] by (metis (mono_tags, lifting) subst_compose_def) hence False using t(2,3) x by blast } thus ?thesis by fast qed lemma abs_terms_subst_range_comp: assumes "subst_range \ \ abs_terms = {}" "subst_range \ \ abs_terms = {}" shows "subst_range (\ \\<^sub>s \) \ abs_terms = {}" proof - { fix t f assume t: "t \ subst_range (\ \\<^sub>s \)" "f \ funs_term t" "is_Abs f" then obtain x where x: "(\ \\<^sub>s \) x = t" by auto have "\ x \ abs_terms" using assms(1) by (cases "\ x \ subst_range \") force+ hence "(\ \\<^sub>s \) x \ abs_terms" using assms(2) abs_terms_subst[of "\ x" \] abs_terms_subst_range_disj by (metis (mono_tags, lifting) subst_compose_def) hence False using t(2,3) x by blast } thus ?thesis by fast qed lemma abs_terms_subst_range_comp': assumes "(\ ` X) \ abs_terms = {}" "(\ ` fv\<^sub>s\<^sub>e\<^sub>t (\ ` X)) \ abs_terms = {}" shows "((\ \\<^sub>s \) ` X) \ abs_terms = {}" proof - { fix t f assume t: "t \ (\ \\<^sub>s \) ` X" "f \ funs_term t" "is_Abs f" then obtain x where x: "(\ \\<^sub>s \) x = t" "x \ X" by auto have "\ x \ abs_terms" using assms(1) x(2) by force moreover have "fv (\ x) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` X)" using x(2) by (auto simp add: fv_subset) hence "\ ` fv (\ x) \ abs_terms = {}" using assms(2) by auto ultimately have "(\ \\<^sub>s \) x \ abs_terms" using abs_terms_subst[of "\ x" \] by (metis (mono_tags, lifting) subst_compose_def) hence False using t(2,3) x by blast } thus ?thesis by fast qed context begin private lemma Ana_abs_aux1: fixes \::"(('fun,'atom,'sets,'lbl) prot_fun, nat, ('fun,'atom,'sets,'lbl) prot_var) gsubst" and \::"nat \ 'sets set" assumes "Ana\<^sub>f f = (K,T)" shows "(K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \) \\<^sub>\\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (\n. \ n \\<^sub>\ \)" proof - { fix k assume "k \ set K" hence "k \ subterms\<^sub>s\<^sub>e\<^sub>t (set K)" by force hence "k \ \ \\<^sub>\ \ = k \ (\n. \ n \\<^sub>\ \)" proof (induction k) case (Fun g S) have "\s. s \ set S \ s \ \ \\<^sub>\ \ = s \ (\n. \ n \\<^sub>\ \)" using Fun.IH in_subterms_subset_Union[OF Fun.prems] Fun_param_in_subterms[of _ S g] by (meson contra_subsetD) thus ?case using Ana\<^sub>f_assm1_alt[OF assms Fun.prems] by (cases g) auto qed simp } thus ?thesis unfolding abs_apply_list_def by force qed private lemma Ana_abs_aux2: fixes \::"nat \ 'sets set" and K::"(('fun,'atom,'sets,'lbl) prot_fun, nat) term list" and M::"nat list" and T::"('fun,'atom,'sets,'lbl) prot_term list" assumes "\i \ fv\<^sub>s\<^sub>e\<^sub>t (set K) \ set M. i < length T" and "(K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T) \\<^sub>\\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (\n. T ! n \\<^sub>\ \)" shows "(K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T) \\<^sub>\\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) (map (\s. s \\<^sub>\ \) T)" (is "?A1 = ?A2") and "(map ((!) T) M) \\<^sub>\\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = map ((!) (map (\s. s \\<^sub>\ \) T)) M" (is "?B1 = ?B2") proof - have "T ! i \\<^sub>\ \ = (map (\s. s \\<^sub>\ \) T) ! i" when "i \ fv\<^sub>s\<^sub>e\<^sub>t (set K)" for i using that assms(1) by auto hence "k \ (\i. T ! i \\<^sub>\ \) = k \ (\i. (map (\s. s \\<^sub>\ \) T) ! i)" when "k \ set K" for k using that term_subst_eq_conv[of k "\i. T ! i \\<^sub>\ \" "\i. (map (\s. s \\<^sub>\ \) T) ! i"] by auto thus "?A1 = ?A2" using assms(2) by (force simp add: abs_apply_terms_def) have "T ! i \\<^sub>\ \ = map (\s. s \\<^sub>\ \) T ! i" when "i \ set M" for i using that assms(1) by auto thus "?B1 = ?B2" by (force simp add: abs_apply_list_def) qed private lemma Ana_abs_aux1_set: fixes \::"(('fun,'atom,'sets,'lbl) prot_fun, nat, ('fun,'atom,'sets,'lbl) prot_var) gsubst" and \::"nat \ 'sets set" assumes "Ana\<^sub>f f = (K,T)" shows "(set K \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t \ = set K \\<^sub>s\<^sub>e\<^sub>t (\n. \ n \\<^sub>\ \)" proof - { fix k assume "k \ set K" hence "k \ subterms\<^sub>s\<^sub>e\<^sub>t (set K)" by force hence "k \ \ \\<^sub>\ \ = k \ (\n. \ n \\<^sub>\ \)" proof (induction k) case (Fun g S) have "\s. s \ set S \ s \ \ \\<^sub>\ \ = s \ (\n. \ n \\<^sub>\ \)" using Fun.IH in_subterms_subset_Union[OF Fun.prems] Fun_param_in_subterms[of _ S g] by (meson contra_subsetD) thus ?case using Ana\<^sub>f_assm1_alt[OF assms Fun.prems] by (cases g) auto qed simp } thus ?thesis unfolding abs_apply_terms_def by force qed private lemma Ana_abs_aux2_set: fixes \::"nat \ 'sets set" and K::"(('fun,'atom,'sets,'lbl) prot_fun, nat) terms" and M::"nat set" and T::"('fun,'atom,'sets,'lbl) prot_term list" assumes "\i \ fv\<^sub>s\<^sub>e\<^sub>t K \ M. i < length T" and "(K \\<^sub>s\<^sub>e\<^sub>t (!) T) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t \ = K \\<^sub>s\<^sub>e\<^sub>t (\n. T ! n \\<^sub>\ \)" shows "(K \\<^sub>s\<^sub>e\<^sub>t (!) T) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t \ = K \\<^sub>s\<^sub>e\<^sub>t (!) (map (\s. s \\<^sub>\ \) T)" (is "?A1 = ?A2") and "((!) T ` M) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t \ = (!) (map (\s. s \\<^sub>\ \) T) ` M" (is "?B1 = ?B2") proof - have "T ! i \\<^sub>\ \ = (map (\s. s \\<^sub>\ \) T) ! i" when "i \ fv\<^sub>s\<^sub>e\<^sub>t K" for i using that assms(1) by auto hence "k \ (\i. T ! i \\<^sub>\ \) = k \ (\i. (map (\s. s \\<^sub>\ \) T) ! i)" when "k \ K" for k using that term_subst_eq_conv[of k "\i. T ! i \\<^sub>\ \" "\i. (map (\s. s \\<^sub>\ \) T) ! i"] by auto thus "?A1 = ?A2" using assms(2) by (force simp add: abs_apply_terms_def) have "T ! i \\<^sub>\ \ = map (\s. s \\<^sub>\ \) T ! i" when "i \ M" for i using that assms(1) by auto thus "?B1 = ?B2" by (force simp add: abs_apply_terms_def) qed lemma Ana_abs: fixes t::"('fun,'atom,'sets,'lbl) prot_term" assumes "Ana t = (K, T)" shows "Ana (t \\<^sub>\ \) = (K \\<^sub>\\<^sub>l\<^sub>i\<^sub>s\<^sub>t \, T \\<^sub>\\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using assms proof (induction t rule: Ana.induct) case (1 f S) - obtain K' T' where *: "Ana\<^sub>f f = (K',T')" by moura + obtain K' T' where *: "Ana\<^sub>f f = (K',T')" by force show ?case using 1 proof (cases "arity\<^sub>f f = length S \ arity\<^sub>f f > 0") case True hence "K = K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) S" "T = map ((!) S) T'" and **: "arity\<^sub>f f = length (map (\s. s \\<^sub>\ \) S)" "arity\<^sub>f f > 0" using 1 * by auto hence "K \\<^sub>\\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) (map (\s. s \\<^sub>\ \) S)" "T \\<^sub>\\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = map ((!) (map (\s. s \\<^sub>\ \) S)) T'" using Ana\<^sub>f_assm2_alt[OF *] Ana_abs_aux2[OF _ Ana_abs_aux1[OF *], of T' S \] unfolding abs_apply_list_def by auto moreover have "Fun (Fu f) S \\<^sub>\ \ = Fun (Fu f) (map (\s. s \\<^sub>\ \) S)" by simp ultimately show ?thesis using Ana_Fu_intro[OF ** *] by metis qed (auto simp add: abs_apply_list_def) qed (simp_all add: abs_apply_list_def) end lemma deduct_FP_if_deduct: fixes M IK FP::"('fun,'atom,'sets,'lbl) prot_terms" assumes IK: "IK \ GSMP M - (pubval_terms \ abs_terms)" "\t \ IK \\<^sub>\\<^sub>s\<^sub>e\<^sub>t \. FP \\<^sub>c t" and t: "IK \ t" "t \ GSMP M - (pubval_terms \ abs_terms)" shows "FP \ t \\<^sub>\ \" proof - let ?P = "\f. \is_PubConstValue f" let ?GSMP = "GSMP M - (pubval_terms \ abs_terms)" have 1: "\m \ IK. m \ ?GSMP" using IK(1) by blast have 2: "\t t'. t \ ?GSMP \ t' \ t \ t' \ ?GSMP" proof (intro allI impI) fix t t' assume t: "t \ ?GSMP" "t' \ t" hence "t' \ GSMP M" using ground_subterm unfolding GSMP_def by auto moreover have "\is_PubConstValue f" when "f \ funs_term t" for f using t(1) that by auto hence "\is_PubConstValue f" when "f \ funs_term t'" for f using that subtermeq_imp_funs_term_subset[OF t(2)] by auto moreover have "\is_Abs f" when "f \ funs_term t" for f using t(1) that by auto hence "\is_Abs f" when "f \ funs_term t'" for f using that subtermeq_imp_funs_term_subset[OF t(2)] by auto ultimately show "t' \ ?GSMP" by simp qed have 3: "\t K T k. t \ ?GSMP \ Ana t = (K, T) \ k \ set K \ k \ ?GSMP" proof (intro allI impI) fix t K T k assume t: "t \ ?GSMP" "Ana t = (K, T)" "k \ set K" hence "k \ GSMP M" using GSMP_Ana_key by blast moreover have "\f \ funs_term t. ?P f" using t(1) by auto with t(2,3) have "\f \ funs_term k. ?P f" proof (induction t arbitrary: k rule: Ana.induct) case 1 thus ?case by (metis Ana_Fu_keys_not_pubval_terms surj_pair) qed auto moreover have "\f \ funs_term t. \is_Abs f" using t(1) by auto with t(2,3) have "\f \ funs_term k. \is_Abs f" proof (induction t arbitrary: k rule: Ana.induct) case 1 thus ?case by (metis Ana_Fu_keys_not_abs_terms surj_pair) qed auto ultimately show "k \ ?GSMP" by simp qed have "\IK; M\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P t" unfolding intruder_deduct_GSMP_def by (rule restricted_deduct_if_deduct'[OF 1 2 3 t]) thus ?thesis proof (induction t rule: intruder_deduct_GSMP_induct) case (AxiomH t) show ?case using IK(2) abs_in[OF AxiomH.hyps] by force next case (ComposeH T f) have *: "Fun f T \\<^sub>\ \ = Fun f (map (\t. t \\<^sub>\ \) T)" using ComposeH.hyps(2,4) by (cases f) auto have **: "length (map (\t. t \\<^sub>\ \) T) = arity f" using ComposeH.hyps(1) by auto show ?case using intruder_deduct.Compose[OF ** ComposeH.hyps(2)] ComposeH.IH(1) * by auto next case (DecomposeH t K T' t\<^sub>i) have *: "Ana (t \\<^sub>\ \) = (K \\<^sub>\\<^sub>l\<^sub>i\<^sub>s\<^sub>t \, T' \\<^sub>\\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using Ana_abs[OF DecomposeH.hyps(2)] by metis have **: "t\<^sub>i \\<^sub>\ \ \ set (T' \\<^sub>\\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using DecomposeH.hyps(4) abs_in abs_list_set_is_set_abs_set[of T'] by auto have ***: "FP \ k" when k: "k \ set (K \\<^sub>\\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" for k proof - obtain k' where k': "k' \ set K" "k = k' \\<^sub>\ \" by (metis (no_types) k abs_apply_terms_def imageE abs_list_set_is_set_abs_set) show "FP \ k" using DecomposeH.IH k' by blast qed show ?case using intruder_deduct.Decompose[OF _ * _ **] DecomposeH.IH(1) ***(1) by blast qed qed end subsection \Computing and Checking Term Implications and Messages\ context stateful_protocol_model begin abbreviation (input) "absc s \ (Fun (Abs s) []::('fun,'atom,'sets,'lbl) prot_term)" fun absdbupd where "absdbupd [] _ a = a" | "absdbupd (insert\Var y, Fun (Set s) T\#D) x a = ( if x = y then absdbupd D x (insert s a) else absdbupd D x a)" | "absdbupd (delete\Var y, Fun (Set s) T\#D) x a = ( if x = y then absdbupd D x (a - {s}) else absdbupd D x a)" | "absdbupd (_#D) x a = absdbupd D x a" lemma absdbupd_cons_cases: "absdbupd (insert\Var x, Fun (Set s) T\#D) x d = absdbupd D x (insert s d)" "absdbupd (delete\Var x, Fun (Set s) T\#D) x d = absdbupd D x (d - {s})" "t \ Var x \ (\s T. u = Fun (Set s) T) \ absdbupd (insert\t,u\#D) x d = absdbupd D x d" "t \ Var x \ (\s T. u = Fun (Set s) T) \ absdbupd (delete\t,u\#D) x d = absdbupd D x d" proof - assume *: "t \ Var x \ (\s T. u = Fun (Set s) T)" let ?P = "absdbupd (insert\t,u\#D) x d = absdbupd D x d" let ?Q = "absdbupd (delete\t,u\#D) x d = absdbupd D x d" { fix y f T assume "t = Fun f T \ u = Var y" hence ?P ?Q by auto } moreover { fix y f T assume "t = Var y" "u = Fun f T" hence ?P using * by (cases f) auto } moreover { fix y f T assume "t = Var y" "u = Fun f T" hence ?Q using * by (cases f) auto } ultimately show ?P ?Q by (metis term.exhaust)+ qed simp_all lemma absdbupd_filter: "absdbupd S x d = absdbupd (filter is_Update S) x d" by (induction S x d rule: absdbupd.induct) simp_all lemma absdbupd_append: "absdbupd (A@B) x d = absdbupd B x (absdbupd A x d)" proof (induction A arbitrary: d) case (Cons a A) thus ?case proof (cases a) case (Insert t u) thus ?thesis proof (cases "t \ Var x \ (\s T. u = Fun (Set s) T)") case False - then obtain s T where "t = Var x" "u = Fun (Set s) T" by moura + then obtain s T where "t = Var x" "u = Fun (Set s) T" by force thus ?thesis by (simp add: Insert Cons.IH absdbupd_cons_cases(1)) qed (simp_all add: Cons.IH absdbupd_cons_cases(3)) next case (Delete t u) thus ?thesis proof (cases "t \ Var x \ (\s T. u = Fun (Set s) T)") case False - then obtain s T where "t = Var x" "u = Fun (Set s) T" by moura + then obtain s T where "t = Var x" "u = Fun (Set s) T" by force thus ?thesis by (simp add: Delete Cons.IH absdbupd_cons_cases(2)) qed (simp_all add: Cons.IH absdbupd_cons_cases(4)) qed simp_all qed simp lemma absdbupd_wellformed_transaction: assumes T: "wellformed_transaction T" shows "absdbupd (unlabel (transaction_strand T)) = absdbupd (unlabel (transaction_updates T))" proof - define S0 where "S0 \ unlabel (transaction_strand T)" define S1 where "S1 \ unlabel (transaction_receive T)" define S2 where "S2 \ unlabel (transaction_checks T)" define S3 where "S3 \ unlabel (transaction_updates T)" define S4 where "S4 \ unlabel (transaction_send T)" note S_defs = S0_def S1_def S2_def S3_def S4_def have 0: "list_all is_Receive S1" "list_all is_Check_or_Assignment S2" "list_all is_Update S3" "list_all is_Send S4" using T unfolding wellformed_transaction_def S_defs by metis+ have "filter is_Update S1 = []" "filter is_Update S2 = []" "filter is_Update S3 = S3" "filter is_Update S4 = []" using list_all_filter_nil[OF 0(1), of is_Update] list_all_filter_nil[OF 0(2), of is_Update] list_all_filter_eq[OF 0(3)] list_all_filter_nil[OF 0(4), of is_Update] by blast+ moreover have "S0 = S1@S2@S3@S4" unfolding S_defs transaction_strand_def unlabel_def by auto ultimately have "filter is_Update S0 = S3" using filter_append[of is_Update] list_all_append[of is_Update] by simp thus ?thesis using absdbupd_filter[of S0] unfolding S_defs by presburger qed fun abs_substs_set:: "[('fun,'atom,'sets,'lbl) prot_var list, 'sets set list, ('fun,'atom,'sets,'lbl) prot_var \ 'sets set, ('fun,'atom,'sets,'lbl) prot_var \ 'sets set, ('fun,'atom,'sets,'lbl) prot_var \ 'sets set \ bool] \ ((('fun,'atom,'sets,'lbl) prot_var \ 'sets set) list) list" where "abs_substs_set [] _ _ _ _ = [[]]" | "abs_substs_set (x#xs) as posconstrs negconstrs msgconstrs = ( let bs = filter (\a. posconstrs x \ a \ a \ negconstrs x = {} \ msgconstrs x a) as; \ = abs_substs_set xs as posconstrs negconstrs msgconstrs in concat (map (\b. map (\\. (x, b)#\) \) bs))" definition abs_substs_fun:: "[(('fun,'atom,'sets,'lbl) prot_var \ 'sets set) list, ('fun,'atom,'sets,'lbl) prot_var] \ 'sets set" where "abs_substs_fun \ x = (case find (\b. fst b = x) \ of Some (_,a) \ a | None \ {})" lemmas abs_substs_set_induct = abs_substs_set.induct[case_names Nil Cons] fun transaction_poschecks_comp:: "(('fun,'atom,'sets,'lbl) prot_fun, ('fun,'atom,'sets,'lbl) prot_var) stateful_strand \ (('fun,'atom,'sets,'lbl) prot_var \ 'sets set)" where "transaction_poschecks_comp [] = (\_. {})" | "transaction_poschecks_comp (\_: Var x \ Fun (Set s) []\#T) = ( let f = transaction_poschecks_comp T in f(x := insert s (f x)))" | "transaction_poschecks_comp (_#T) = transaction_poschecks_comp T" fun transaction_negchecks_comp:: "(('fun,'atom,'sets,'lbl) prot_fun, ('fun,'atom,'sets,'lbl) prot_var) stateful_strand \ (('fun,'atom,'sets,'lbl) prot_var \ 'sets set)" where "transaction_negchecks_comp [] = (\_. {})" | "transaction_negchecks_comp (\Var x not in Fun (Set s) []\#T) = ( let f = transaction_negchecks_comp T in f(x := insert s (f x)))" | "transaction_negchecks_comp (_#T) = transaction_negchecks_comp T" definition transaction_check_pre where "transaction_check_pre FPT T \ \ let (FP, _, TI) = FPT; C = set (unlabel (transaction_checks T)); xs = fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T)); \ = \\ x. if fst x = TAtom Value then (absc \ \) x else Var x in (\x \ set (transaction_fresh T). \ x = {}) \ (\t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T). intruder_synth_mod_timpls FP TI (t \ \ \)) \ (\u \ C. (is_InSet u \ ( let x = the_elem_term u; s = the_set_term u in (is_Var x \ is_Fun_Set s) \ the_Set (the_Fun s) \ \ (the_Var x))) \ ((is_NegChecks u \ bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p u = [] \ the_eqs u = [] \ length (the_ins u) = 1) \ ( let x = fst (hd (the_ins u)); s = snd (hd (the_ins u)) in (is_Var x \ is_Fun_Set s) \ the_Set (the_Fun s) \ \ (the_Var x))))" definition transaction_check_post where "transaction_check_post FPT T \ \ let (FP, _, TI) = FPT; xs = fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T)); \ = \\ x. if fst x = TAtom Value then (absc \ \) x else Var x; u = \\ x. absdbupd (unlabel (transaction_updates T)) x (\ x) in (\x \ set xs - set (transaction_fresh T). \ x \ u \ x \ List.member TI (\ x, u \ x)) \ (\t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T). intruder_synth_mod_timpls FP TI (t \ \ (u \)))" definition fun_point_inter where "fun_point_inter f g \ \x. f x \ g x" definition fun_point_union where "fun_point_union f g \ \x. f x \ g x" definition fun_point_Inter where "fun_point_Inter fs \ \x. \f \ fs. f x" definition fun_point_Union where "fun_point_Union fs \ \x. \f \ fs. f x" definition fun_point_Inter_list where "fun_point_Inter_list fs \ \x. \(set (map (\f. f x) fs))" definition fun_point_Union_list where "fun_point_Union_list fs \ \x. \(set (map (\f. f x) fs))" definition ticl_abs where "ticl_abs TI a \ set (a#map snd (filter (\p. fst p = a) TI))" definition ticl_abss where "ticl_abss TI as \ \a \ as. ticl_abs TI a" lemma fun_point_Inter_set_eq: "fun_point_Inter (set fs) = fun_point_Inter_list fs" unfolding fun_point_Inter_def fun_point_Inter_list_def by simp lemma fun_point_Union_set_eq: "fun_point_Union (set fs) = fun_point_Union_list fs" unfolding fun_point_Union_def fun_point_Union_list_def by simp lemma ticl_abs_refl_in: "x \ ticl_abs TI x" unfolding ticl_abs_def by simp lemma ticl_abs_iff: assumes TI: "set TI = {(a,b) \ (set TI)\<^sup>+. a \ b}" shows "ticl_abs TI a = {b. (a,b) \ (set TI)\<^sup>*}" proof (intro order_antisym subsetI) fix x assume x: "x \ {b. (a, b) \ (set TI)\<^sup>*}" hence "x = a \ (x \ a \ (a,x) \ (set TI)\<^sup>+)" by (metis mem_Collect_eq rtranclD) moreover have "ticl_abs TI a = {a} \ {b. (a,b) \ set TI}" unfolding ticl_abs_def by force ultimately show "x \ ticl_abs TI a" using TI by blast qed (fastforce simp add: ticl_abs_def) lemma ticl_abs_Inter: assumes xs: "\(ticl_abs TI ` xs) \ {}" and TI: "set TI = {(a,b) \ (set TI)\<^sup>+. a \ b}" shows "\(ticl_abs TI ` \(ticl_abs TI ` xs)) \ \(ticl_abs TI ` xs)" proof fix x assume x: "x \ \(ticl_abs TI ` \(ticl_abs TI ` xs))" have *: "\(ticl_abs TI ` xs) = {b. \a \ xs. (a,b) \ (set TI)\<^sup>*}" unfolding ticl_abs_iff[OF TI] by blast have "(b,x) \ (set TI)\<^sup>*" when b: "\a \ xs. (a,b) \ (set TI)\<^sup>*" for b using x b unfolding ticl_abs_iff[OF TI] by blast hence "(a,x) \ (set TI)\<^sup>*" when "a \ xs" for a using that xs rtrancl.rtrancl_into_rtrancl[of a _ "(set TI)\<^sup>*" x] unfolding * rtrancl_idemp[of "set TI"] by blast thus "x \ \(ticl_abs TI ` xs)" unfolding * by blast qed function (sequential) match_abss' ::"(('a,'b,'c,'d) prot_fun, 'e) term \ (('a,'b,'c,'d) prot_fun, 'e) term \ ('e \ 'c set set) option" where "match_abss' (Var x) (Fun (Abs a) _) = Some ((\_. {})(x := {a}))" | "match_abss' (Fun f ts) (Fun g ss) = ( if f = g \ length ts = length ss then map_option fun_point_Union_list (those (map2 match_abss' ts ss)) else None)" | "match_abss' _ _ = None" by pat_completeness auto termination proof - let ?m = "measures [size \ fst]" have 0: "wf ?m" by simp show ?thesis apply (standard, use 0 in fast) by (metis (no_types) comp_def fst_conv measures_less Fun_zip_size_lt(1)) qed definition match_abss where "match_abss OCC TI t s \ ( let xs = fv t; OCC' = set OCC; f = \\ x. if x \ xs then \ x else OCC'; g = \\ x. \(ticl_abs TI ` \ x) in case match_abss' t s of Some \ \ let \' = g \ in if \x \ xs. \' x \ {} then Some (f \') else None | None \ None)" lemma match_abss'_Var_inv: assumes \: "match_abss' (Var x) t = Some \" shows "\a ts. t = Fun (Abs a) ts \ \ = (\_. {})(x := {a})" proof - obtain f ts where t: "t = Fun f ts" using \ by (cases t) auto then obtain a where a: "f = Abs a" using \ by (cases f) auto show ?thesis using \ unfolding t a by simp qed lemma match_abss'_Fun_inv: assumes "match_abss' (Fun f ts) (Fun g ss) = Some \" shows "f = g" (is ?A) and "length ts = length ss" (is ?B) and "\\. Some \ = those (map2 match_abss' ts ss) \ \ = fun_point_Union_list \" (is ?C) and "\(t,s) \ set (zip ts ss). \\. match_abss' t s = Some \" (is ?D) proof - note 0 = assms match_abss'.simps(2)[of f ts g ss] option.distinct(1) show ?A by (metis 0) show ?B by (metis 0) show ?C by (metis (no_types, opaque_lifting) 0 map_option_eq_Some) thus ?D using map2_those_Some_case[of match_abss' ts ss] by fastforce qed lemma match_abss'_FunI: assumes \: "\i. i < length T \ match_abss' (U ! i) (T ! i) = Some (\ i)" and T: "length T = length U" shows "match_abss' (Fun f U) (Fun f T) = Some (fun_point_Union_list (map \ [0.. [0.. T those_map2_SomeI by metis ultimately show ?thesis by simp qed lemma match_abss'_Fun_param_subset: assumes "match_abss' (Fun f ts) (Fun g ss) = Some \" and "(t,s) \ set (zip ts ss)" and "match_abss' t s = Some \" shows "\ x \ \ x" proof - obtain \ where \: "those (map2 match_abss' ts ss) = Some \" "\ = fun_point_Union_list \" using match_abss'_Fun_inv[OF assms(1)] by metis have "\ \ set \" using \(1) assms(2-) those_Some_iff[of "map2 match_abss' ts ss" \] by force thus ?thesis using \(2) unfolding fun_point_Union_list_def by auto qed lemma match_abss'_fv_is_nonempty: assumes "match_abss' t s = Some \" and "x \ fv t" shows "\ x \ {}" (is "?P \") using assms proof (induction t s arbitrary: \ rule: match_abss'.induct) case (2 f ts g ss) note prems = "2.prems" note IH = "2.IH" have 0: "\(t,s) \ set (zip ts ss). \\. match_abss' t s = Some \" "f = g" "length ts = length ss" using match_abss'_Fun_inv[OF prems(1)] by simp_all obtain t where t: "t \ set ts" "x \ fv t" using prems(2) by auto then obtain s where s: "s \ set ss" "(t,s) \ set (zip ts ss)" by (meson 0(3) in_set_impl_in_set_zip1 in_set_zipE) then obtain \ where \: "match_abss' t s = Some \" using 0(1) by fast show ?case using IH[OF conjI[OF 0(2,3)] s(2) _ \] t(2) match_abss'_Fun_param_subset[OF prems(1) s(2) \] by auto qed auto lemma match_abss'_nonempty_is_fv: fixes s t::"(('a,'b,'c,'d) prot_fun, 'v) term" assumes "match_abss' s t = Some \" and "\ x \ {}" shows "x \ fv s" using assms proof (induction s t arbitrary: \ rule: match_abss'.induct) case (2 f ts g ss) note prems = "2.prems" note IH = "2.IH" obtain \ where \: "Some \ = those (map2 match_abss' ts ss)" "\ = fun_point_Union_list \" and fg: "f = g" "length ts = length ss" using match_abss'_Fun_inv[OF prems(1)] by fast have "\\ \ set \. \ x \ {}" using fg(2) prems \ unfolding fun_point_Union_list_def by auto then obtain t' s' \ where ts': "(t',s') \ set (zip ts ss)" "match_abss' t' s' = Some \" "\ x \ {}" using those_map2_SomeD[OF \(1)[symmetric]] by blast show ?case using ts'(3) IH[OF conjI[OF fg] ts'(1) _ ts'(2)] set_zip_leftD[OF ts'(1)] by force qed auto lemma match_abss'_Abs_in_funs_term: fixes s t::"(('a,'b,'c,'d) prot_fun, 'v) term" assumes "match_abss' s t = Some \" and "a \ \ x" shows "Abs a \ funs_term t" using assms proof (induction s t arbitrary: a \ rule: match_abss'.induct) case (1 y b ts) show ?case using match_abss'_Var_inv[OF "1.prems"(1)] "1.prems"(2) by (cases "x = y") simp_all next case (2 f ts g ss) note prems = "2.prems" note IH = "2.IH" obtain \ where \: "Some \ = those (map2 match_abss' ts ss)" "\ = fun_point_Union_list \" and fg: "f = g" "length ts = length ss" using match_abss'_Fun_inv[OF prems(1)] by fast obtain t' s' \ where ts': "(t',s') \ set (zip ts ss)" "match_abss' t' s' = Some \" "a \ \ x" using fg(2) prems \ those_map2_SomeD[OF \(1)[symmetric]] unfolding fun_point_Union_list_def by fastforce show ?case using ts'(1) IH[OF conjI[OF fg] ts'(1) _ ts'(2,3)] by (meson set_zip_rightD term.set_intros(2)) qed auto lemma match_abss'_subst_fv_ex_abs: assumes "match_abss' s (s \ \) = Some \" and TI: "set TI = {(a,b) \ (set TI)\<^sup>+. a \ b}" shows "\x \ fv s. \a ts. \ x = Fun (Abs a) ts \ \ x = {a}" (is "?P s \") using assms(1) proof (induction s "s \ \" arbitrary: \ rule: match_abss'.induct) case (2 f ts g ss) note prems = "2.prems" note hyps = "2.hyps" obtain \ where \: "Some \ = those (map2 match_abss' ts ss)" "\ = fun_point_Union_list \" and fg: "f = g" "length ts = length ss" "ss = ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" and ts: "\(t,s) \ set (zip ts ss). \\. match_abss' t s = Some \" using match_abss'_Fun_inv[OF prems(1)[unfolded hyps(2)[symmetric]]] hyps(2) by fastforce have 0: "those (map (\t. match_abss' t (t \ \)) ts) = Some \" using \(1) map2_map_subst unfolding fg(3) by metis have 1: "\t \ set ts. \\. match_abss' t (t \ \) = Some \" using ts zip_map_subst[of ts \] unfolding fg(3) by simp have 2: "\' \ set \" when t: "t \ set ts" "match_abss' t (t \ \) = Some \'" for t \' using t 0 those_Some_iff[of "map (\t. match_abss' t (t \ \)) ts" \] by force have 3: "?P t \'" "\' x \ {}" when t: "t \ set ts" "x \ fv t" "match_abss' t (t \ \) = Some \'" for t \' x using t hyps(1)[OF conjI[OF fg(1,2)], of "(t, t \ \)" t \'] zip_map_subst[of ts \] match_abss'_fv_is_nonempty[of t "t \ \" \' x] unfolding fg(3) by auto have 4: "\' x = {}" when t: "x \ fv t" "match_abss' t (t \ \) = Some \'" for t \' x by (meson t match_abss'_nonempty_is_fv) show ?case proof fix x assume "x \ fv (Fun f ts)" then obtain t \' where t: "t \ set ts" "x \ fv t" and \': "match_abss' t (t \ \) = Some \'" using 1 by auto then obtain a tsa where a: "\ x = Fun (Abs a) tsa" using 3[OF t \'] by fast have "\'' x = {a} \ \'' x = {}" when "\'' \ set \" for \'' using that a 0 3[of _ x] 4[of x] unfolding those_Some_iff by fastforce thus "\a ts. \ x = Fun (Abs a) ts \ \ x = {a}" using a 2[OF t(1) \'] 3[OF t \'] unfolding \(2) fun_point_Union_list_def by auto qed qed auto lemma match_abss'_subst_disj_nonempty: assumes TI: "set TI = {(a,b) \ (set TI)\<^sup>+. a \ b}" and "match_abss' s (s \ \) = Some \" and "x \ fv s" shows "\(ticl_abs TI ` \ x) \ {} \ (\a tsa. \ x = Fun (Abs a) tsa \ \ x = {a})" (is "?P \") using assms(2,3) proof (induction s "s \ \" arbitrary: \ rule: match_abss'.induct) case (1 x a ts) thus ?case unfolding ticl_abs_def by force next case (2 f ts g ss) note prems = "2.prems" note hyps = "2.hyps" obtain \ where \: "Some \ = those (map2 match_abss' ts ss)" "\ = fun_point_Union_list \" and fg: "f = g" "length ts = length ss" "ss = ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" and ts: "\(t,s) \ set (zip ts ss). \\. match_abss' t s = Some \" using match_abss'_Fun_inv[OF prems(1)[unfolded hyps(2)[symmetric]]] hyps(2) by fastforce define ts' where "ts' \ filter (\t. x \ fv t) ts" define \' where "\' \ map (\t. (t, the (match_abss' t (t \ \)))) ts" define \'' where "\'' \ map (\t. the (match_abss' t (t \ \))) ts'" have 0: "those (map (\t. match_abss' t (t \ \)) ts) = Some \" using \(1) map2_map_subst unfolding fg(3) by metis have 1: "\t \ set ts. \\. match_abss' t (t \ \) = Some \" using ts zip_map_subst[of ts \] unfolding fg(3) by simp have ts_not_nil: "ts \ []" using prems(2) by fastforce hence "\t \ set ts. x \ fv t" using prems(2) by simp then obtain a tsa where a: "\ x = Fun (Abs a) tsa" using 1 match_abss'_subst_fv_ex_abs[OF _ TI, of _ \] by metis hence a': "\' x = {a}" when "t \ set ts" "x \ fv t" "match_abss' t (t \ \) = Some \'" for t \' using that match_abss'_subst_fv_ex_abs[OF _ TI, of _ \] by fastforce have "ts' \ []" using prems(2) unfolding ts'_def by (simp add: filter_empty_conv) hence \''_not_nil: "\'' \ []" unfolding \''_def by simp have 2: "\' \ set \" when t: "t \ set ts" "match_abss' t (t \ \) = Some \'" for t \' using t 0 those_Some_iff[of "map (\t. match_abss' t (t \ \)) ts" \] by force have 3: "?P \'" "\' x \ {}" when t: "t \ set ts'" "match_abss' t (t \ \) = Some \'" for t \' using t hyps(1)[OF conjI[OF fg(1,2)], of "(t, t \ \)" t \'] zip_map_subst[of ts \] match_abss'_fv_is_nonempty[of t "t \ \" \' x] unfolding fg(3) ts'_def by (force, force) have 4: "\' x = {}" when t: "x \ fv t" "match_abss' t (t \ \) = Some \'" for t \' by (meson t match_abss'_nonempty_is_fv) have 5: "\ = map snd \'" using 0 1 unfolding \'_def by (induct ts arbitrary: \) auto have "fun_point_Union_list (map snd \') x = fun_point_Union_list (map snd (filter (\(t,_). x \ fv t) \')) x" using 1 4 unfolding \'_def fun_point_Union_list_def by fastforce hence 6: "fun_point_Union_list \ x = fun_point_Union_list \'' x" using 0 1 4 unfolding 5 \'_def \''_def fun_point_Union_list_def ts'_def by auto have 7: "?P \'" "\' x \ {}" when \': "\' \ set \''" for \' using that 1 3 unfolding \''_def ts'_def by auto have "\' x = {a}" when \': "\' \ set \''" for \' using \' a' 1 unfolding \''_def ts'_def by fastforce hence "fun_point_Union_list \'' x = {b | b \'. \' \ set \'' \ b \ {a}}" using \''_not_nil unfolding fun_point_Union_list_def by auto hence 8: "fun_point_Union_list \'' x = {a}" using \''_not_nil by auto show ?case using 8 a unfolding \(2) 6 ticl_abs_iff[OF TI] by auto qed simp_all lemma match_abssD: fixes OCC TI s defines "f \ (\\ x. if x \ fv s then \ x else set OCC)" and "g \ (\\ x. \(ticl_abs TI ` \ x))" assumes \': "match_abss OCC TI s t = Some \'" shows "\\. match_abss' s t = Some \ \ \' = f (g \) \ (\x \ fv s. \ x \ {} \ f (g \) x \ {}) \ (set OCC \ {} \ (\x. f (g \) x \ {}))" proof - obtain \ where \: "match_abss' s t = Some \" using \' unfolding match_abss_def by force hence "Some \' = (if \x \ fv s. g \ x \ {} then Some (f (g \)) else None)" using \' unfolding match_abss_def f_def g_def Let_def by simp hence "\' = f (g \)" "\x \ fv s. \ x \ {} \ f (g \) x \ {}" by (metis (no_types, lifting) option.inject option.distinct(1), metis (no_types, lifting) f_def option.distinct(1) match_abss'_fv_is_nonempty[OF \]) thus ?thesis using \ unfolding f_def by force qed lemma match_abss_ticl_abs_Inter_subset: assumes TI: "set TI = {(a,b). (a,b) \ (set TI)\<^sup>+ \ a \ b}" and \: "match_abss OCC TI s t = Some \" and x: "x \ fv s" shows "\(ticl_abs TI ` \ x) \ \ x" proof - let ?h1 = "\\ x. if x \ fv s then \ x else set OCC" let ?h2 = "\\ x. \(ticl_abs TI ` \ x)" obtain \' where \': "match_abss' s t = Some \'" "\ = ?h1 (?h2 \')" "\x \ fv s. \' x \ {} \ \ x \ {}" using match_abssD[OF \] by blast have "\ x = \(ticl_abs TI ` \' x)" "\' x \ {}" "\ x \ {}" using x \'(2,3) by auto thus ?thesis using ticl_abs_Inter TI by simp qed lemma match_abss_fv_has_abs: assumes "match_abss OCC TI s t = Some \" and "x \ fv s" shows "\ x \ {}" using assms match_abssD by fast lemma match_abss_OCC_if_not_fv: fixes s t::"(('a,'b,'c,'d) prot_fun, 'v) term" assumes \': "match_abss OCC TI s t = Some \'" and x: "x \ fv s" shows "\' x = set OCC" proof - define f where "f \ \s::(('a,'b,'c,'d) prot_fun, 'v) term. \\ x. if x \ fv s then \ x else set OCC" define g where "g \ \\. \x::'v. \(ticl_abs TI ` \ x)" obtain \ where \: "match_abss' s t = Some \" "\' = f s (g \)" using match_abssD[OF \'] unfolding f_def g_def by blast show ?thesis using x \(2) unfolding f_def by presburger qed inductive synth_abs_substs_constrs_rel for FP OCC TI where SolveNil: "synth_abs_substs_constrs_rel FP OCC TI [] (\_. set OCC)" | SolveCons: "ts \ [] \ \t \ set ts. synth_abs_substs_constrs_rel FP OCC TI [t] (\ t) \ synth_abs_substs_constrs_rel FP OCC TI ts (fun_point_Inter (\ ` set ts))" | SolvePubConst: "arity c = 0 \ public c \ synth_abs_substs_constrs_rel FP OCC TI [Fun c []] (\_. set OCC)" | SolvePrivConstIn: "arity c = 0 \ \public c \ Fun c [] \ set FP \ synth_abs_substs_constrs_rel FP OCC TI [Fun c []] (\_. set OCC)" | SolvePrivConstNotin: "arity c = 0 \ \public c \ Fun c [] \ set FP \ synth_abs_substs_constrs_rel FP OCC TI [Fun c []] (\_. {})" | SolveValueVar: "\ = ((\_. set OCC)(x := ticl_abss TI {a \ set OCC. \a\\<^sub>a\<^sub>b\<^sub>s \ set FP})) \ synth_abs_substs_constrs_rel FP OCC TI [Var x] \" | SolveComposed: "arity f > 0 \ length ts = arity f \ \\. \ \ \ \ (\s \ set FP. match_abss OCC TI (Fun f ts) s = Some \) \ \1 = fun_point_Union \ \ synth_abs_substs_constrs_rel FP OCC TI ts \2 \ synth_abs_substs_constrs_rel FP OCC TI [Fun f ts] (fun_point_union \1 \2)" fun synth_abs_substs_constrs_aux where "synth_abs_substs_constrs_aux FP OCC TI (Var x) = ( (\_. set OCC)(x := ticl_abss TI (set (filter (\a. \a\\<^sub>a\<^sub>b\<^sub>s \ set FP) OCC))))" | "synth_abs_substs_constrs_aux FP OCC TI (Fun f ts) = ( if ts = [] then if \public f \ Fun f ts \ set FP then (\_. {}) else (\_. set OCC) else let \ = map the (filter (\\. \ \ None) (map (match_abss OCC TI (Fun f ts)) FP)); \1 = fun_point_Union_list \; \2 = fun_point_Inter_list (map (synth_abs_substs_constrs_aux FP OCC TI) ts) in fun_point_union \1 \2)" definition synth_abs_substs_constrs where "synth_abs_substs_constrs FPT T \ let (FP,OCC,TI) = FPT; ts = trms_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_receive T)); f = fun_point_Inter_list \ map (synth_abs_substs_constrs_aux FP OCC TI) in if ts = [] then (\_. set OCC) else f ts" (* definition synth_abs_substs_constrs where "synth_abs_substs_constrs FPT T \ let (FP,OCC,TI) = FPT; negsy = Not \ intruder_synth_mod_timpls FP TI; \ = \\ x. let as = \ x in if as \ {} then as else set OCC; C = unlabel (transaction_checks T); poss = transaction_poschecks_comp C; negs = transaction_negchecks_comp C; ts = trms_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_receive T)); f = \t. let \ = map the (filter (\\. \ \ None) (map (match_abss OCC TI t) FP)) in fun_point_Union_list (map \ \); g = \t. if is_Fun t \ args t \ [] then let s = hd (args t) in case fv_list s of [] \ if negsy s then Some (f t) else None | [x] \ let bs = filter (\a. poss x \ a \ a \ negs x = {}) OCC in if list_all (\b. negsy (s \ Var(x := \b\\<^sub>a\<^sub>b\<^sub>s))) bs then Some (f t) else None | _ \ None else None; h = \t. case g t of Some d \ d | None \ synth_abs_substs_constrs_aux FP OCC TI t in if ts = [] then (\_. set OCC) else fun_point_Inter_list (map h ts)" *) (* poss = transaction_poschecks_comp (C A); negs = transaction_negchecks_comp (C A); bs = filter (\a. poss PK \ a \ a \ negs PK = {}) OCC in if list_all (Not \ sy \ s A) bs then Some (map the (filter (\\. \ \ None) (map (match_abss OCC TI (t' A)) FP))) else None *) definition transaction_check_comp:: "[('fun,'atom,'sets,'lbl) prot_var \ 'sets set \ bool, ('fun,'atom,'sets,'lbl) prot_term list \ 'sets set list \ ('sets set \ 'sets set) list, ('fun,'atom,'sets,'lbl) prot_transaction] \ ((('fun,'atom,'sets,'lbl) prot_var \ 'sets set) list) list" where "transaction_check_comp msgcs FPT T \ let (_, OCC, _) = FPT; S = unlabel (transaction_strand T); C = unlabel (transaction_checks T); xs = filter (\x. x \ set (transaction_fresh T) \ fst x = TAtom Value) (fv_list\<^sub>s\<^sub>s\<^sub>t S); posconstrs = transaction_poschecks_comp C; negconstrs = transaction_negchecks_comp C; pre_check = transaction_check_pre FPT T; \ = abs_substs_set xs OCC posconstrs negconstrs msgcs in filter (\\. pre_check (abs_substs_fun \)) \" definition transaction_check':: "[('fun,'atom,'sets,'lbl) prot_var \ 'sets set \ bool, ('fun,'atom,'sets,'lbl) prot_term list \ 'sets set list \ ('sets set \ 'sets set) list, ('fun,'atom,'sets,'lbl) prot_transaction] \ bool" where "transaction_check' msgcs FPT T \ list_all (\\. transaction_check_post FPT T (abs_substs_fun \)) (transaction_check_comp msgcs FPT T)" definition transaction_check:: "[('fun,'atom,'sets,'lbl) prot_term list \ 'sets set list \ ('sets set \ 'sets set) list, ('fun,'atom,'sets,'lbl) prot_transaction] \ bool" where "transaction_check \ transaction_check' (\_ _. True)" definition transaction_check_coverage_rcv:: "[('fun,'atom,'sets,'lbl) prot_term list \ 'sets set list \ ('sets set \ 'sets set) list, ('fun,'atom,'sets,'lbl) prot_transaction] \ bool" where "transaction_check_coverage_rcv FPT T \ let msgcs = synth_abs_substs_constrs FPT T in transaction_check' (\x a. a \ msgcs x) FPT T" lemma abs_subst_fun_cons: "abs_substs_fun ((x,b)#\) = (abs_substs_fun \)(x := b)" unfolding abs_substs_fun_def by fastforce lemma abs_substs_cons: assumes "\ \ set (abs_substs_set xs as poss negs msgcs)" "b \ set as" "poss x \ b" "b \ negs x = {}" "msgcs x b" shows "(x,b)#\ \ set (abs_substs_set (x#xs) as poss negs msgcs)" using assms by auto lemma abs_substs_cons': assumes \: "\ \ abs_substs_fun ` set (abs_substs_set xs as poss negs msgcs)" and b: "b \ set as" "poss x \ b" "b \ negs x = {}" "msgcs x b" shows "\(x := b) \ abs_substs_fun ` set (abs_substs_set (x#xs) as poss negs msgcs)" proof - obtain \ where \: "\ = abs_substs_fun \" "\ \ set (abs_substs_set xs as poss negs msgcs)" - using \ by moura + using \ by force have "abs_substs_fun ((x, b)#\) \ abs_substs_fun ` set (abs_substs_set (x#xs) as poss negs msgcs)" using abs_substs_cons[OF \(2) b] by blast thus ?thesis using \(1) abs_subst_fun_cons[of x b \] by argo qed lemma abs_substs_has_abs: assumes "\x. x \ set xs \ \ x \ set as" and "\x. x \ set xs \ poss x \ \ x" and "\x. x \ set xs \ \ x \ negs x = {}" and "\x. x \ set xs \ msgcs x (\ x)" and "\x. x \ set xs \ \ x = {}" shows "\ \ abs_substs_fun ` set (abs_substs_set xs as poss negs msgcs)" using assms proof (induction xs arbitrary: \) case (Cons x xs) define \ where "\ \ \y. if y \ set xs then \ y else {}" have "\ \ abs_substs_fun ` set (abs_substs_set xs as poss negs msgcs)" using Cons.prems Cons.IH by (simp add: \_def) moreover have "\ x \ set as" "poss x \ \ x" "\ x \ negs x = {}" "msgcs x (\ x)" by (simp_all add: Cons.prems(1,2,3,4)) ultimately have 0: "\(x := \ x) \ abs_substs_fun ` set (abs_substs_set (x#xs) as poss negs msgcs)" by (metis abs_substs_cons') have "\ = \(x := \ x)" proof fix y show "\ y = (\(x := \ x)) y" proof (cases "y \ set (x#xs)") case False thus ?thesis using Cons.prems(5) by (fastforce simp add: \_def) qed (auto simp add: \_def) qed thus ?case by (metis 0) qed (auto simp add: abs_substs_fun_def) lemma abs_substs_abss_bounded: assumes "\ \ abs_substs_fun ` set (abs_substs_set xs as poss negs msgcs)" and "x \ set xs" shows "\ x \ set as" and "poss x \ \ x" and "\ x \ negs x = {}" and "msgcs x (\ x)" using assms proof (induct xs as poss negs msgcs arbitrary: \ rule: abs_substs_set_induct) case (Cons y xs as poss negs msgcs) { case 1 thus ?case using Cons.hyps(1) unfolding abs_substs_fun_def by fastforce } { case 2 thus ?case proof (cases "x = y") case False then obtain \' where \': "\' \ abs_substs_fun ` set (abs_substs_set xs as poss negs msgcs)" "\' x = \ x" using 2 unfolding abs_substs_fun_def by force moreover have "x \ set xs" using 2(2) False by simp moreover have "\b. b \ set as \ poss y \ b \ b \ negs y = {}" using 2 False by auto ultimately show ?thesis using Cons.hyps(2) by fastforce qed (auto simp add: abs_substs_fun_def) } { case 3 thus ?case proof (cases "x = y") case False then obtain \' where \': "\' \ abs_substs_fun ` set (abs_substs_set xs as poss negs msgcs)" "\' x = \ x" using 3 unfolding abs_substs_fun_def by force moreover have "x \ set xs" using 3(2) False by simp moreover have "\b. b \ set as \ poss y \ b \ b \ negs y = {}" using 3 False by auto ultimately show ?thesis using Cons.hyps(3) by fastforce qed (auto simp add: abs_substs_fun_def) } { case 4 thus ?case proof (cases "x = y") case False then obtain \' where \': "\' \ abs_substs_fun ` set (abs_substs_set xs as poss negs msgcs)" "\' x = \ x" using 4 unfolding abs_substs_fun_def by force moreover have "x \ set xs" using 4(2) False by simp moreover have "\b. b \ set as \ poss y \ b \ b \ negs y = {}" using 4 False by auto ultimately show ?thesis using Cons.hyps(4) by fastforce qed (auto simp add: abs_substs_fun_def) } qed (simp_all add: abs_substs_fun_def) lemma abs_substs_abss_bounded': assumes "\ \ abs_substs_fun ` set (abs_substs_set xs as poss negs msgcs)" and "x \ set xs" shows "\ x = {}" using assms unfolding abs_substs_fun_def by (induct xs as poss negs msgcs arbitrary: \ rule: abs_substs_set_induct) (force, fastforce) lemma transaction_poschecks_comp_unfold: "transaction_poschecks_comp C x = {s. \a. \a: Var x \ Fun (Set s) []\ \ set C}" proof (induction C) case (Cons c C) thus ?case proof (cases "\a y s. c = \a: Var y \ Fun (Set s) []\") case True - then obtain a y s where c: "c = \a: Var y \ Fun (Set s) []\" by moura + then obtain a y s where c: "c = \a: Var y \ Fun (Set s) []\" by force define f where "f \ transaction_poschecks_comp C" have "transaction_poschecks_comp (c#C) = f(y := insert s (f y))" using c by (simp add: f_def Let_def) moreover have "f x = {s. \a. \a: Var x \ Fun (Set s) []\ \ set C}" using Cons.IH unfolding f_def by blast ultimately show ?thesis using c by auto next case False hence "transaction_poschecks_comp (c#C) = transaction_poschecks_comp C" (is ?P) using transaction_poschecks_comp.cases[of "c#C" ?P] by force thus ?thesis using False Cons.IH by auto qed qed simp lemma transaction_poschecks_comp_notin_fv_empty: assumes "x \ fv\<^sub>s\<^sub>s\<^sub>t C" shows "transaction_poschecks_comp C x = {}" using assms transaction_poschecks_comp_unfold[of C x] by fastforce lemma transaction_negchecks_comp_unfold: "transaction_negchecks_comp C x = {s. \Var x not in Fun (Set s) []\ \ set C}" proof (induction C) case (Cons c C) thus ?case proof (cases "\y s. c = \Var y not in Fun (Set s) []\") case True - then obtain y s where c: "c = \Var y not in Fun (Set s) []\" by moura + then obtain y s where c: "c = \Var y not in Fun (Set s) []\" by force define f where "f \ transaction_negchecks_comp C" have "transaction_negchecks_comp (c#C) = f(y := insert s (f y))" using c by (simp add: f_def Let_def) moreover have "f x = {s. \Var x not in Fun (Set s) []\ \ set C}" using Cons.IH unfolding f_def by blast ultimately show ?thesis using c by auto next case False hence "transaction_negchecks_comp (c#C) = transaction_negchecks_comp C" (is ?P) using transaction_negchecks_comp.cases[of "c#C" ?P] by force thus ?thesis using False Cons.IH by fastforce qed qed simp lemma transaction_negchecks_comp_notin_fv_empty: assumes "x \ fv\<^sub>s\<^sub>s\<^sub>t C" shows "transaction_negchecks_comp C x = {}" using assms transaction_negchecks_comp_unfold[of C x] by fastforce lemma transaction_check_preI[intro]: fixes T defines "\ \ \\ x. if fst x = TAtom Value then (absc \ \) x else Var x" and "C \ set (unlabel (transaction_checks T))" assumes a0: "\x \ set (transaction_fresh T). \ x = {}" and a1: "\x \ fv_transaction T - set (transaction_fresh T). fst x = TAtom Value \ \ x \ set OCC" and a2: "\t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T). intruder_synth_mod_timpls FP TI (t \ \ \)" and a3: "\a x s. \a: Var x \ Fun (Set s) []\ \ C \ s \ \ x" and a4: "\x s. \Var x not in Fun (Set s) []\ \ C \ s \ \ x" shows "transaction_check_pre (FP, OCC, TI) T \" proof - let ?P = "\u. is_InSet u \ ( let x = the_elem_term u; s = the_set_term u in (is_Var x \ is_Fun_Set s) \ the_Set (the_Fun s) \ \ (the_Var x))" let ?Q = "\u. (is_NegChecks u \ bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p u = [] \ the_eqs u = [] \ length (the_ins u) = 1) \ ( let x = fst (hd (the_ins u)); s = snd (hd (the_ins u)) in (is_Var x \ is_Fun_Set s) \ the_Set (the_Fun s) \ \ (the_Var x))" have 1: "?P u" when u: "u \ C" for u apply (unfold Let_def, intro impI, elim conjE) using u a3 Fun_Set_InSet_iff[of u] by metis have 2: "?Q u" when u: "u \ C" for u apply (unfold Let_def, intro impI, elim conjE) using u a4 Fun_Set_NotInSet_iff[of u] by metis show ?thesis using a0 a1 a2 1 2 fv_list\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_strand T)"] unfolding transaction_check_pre_def \_def C_def Let_def by blast qed lemma transaction_check_pre_InSetE: assumes T: "transaction_check_pre FPT T \" and u: "u = \a: Var x \ Fun (Set s) []\" "u \ set (unlabel (transaction_checks T))" shows "s \ \ x" proof - have "is_InSet u \ is_Var (the_elem_term u) \ is_Fun_Set (the_set_term u) \ the_Set (the_Fun (the_set_term u)) \ \ (the_Var (the_elem_term u))" using T u unfolding transaction_check_pre_def Let_def by blast thus ?thesis using Fun_Set_InSet_iff[of u a x s] u by argo qed lemma transaction_check_pre_NotInSetE: assumes T: "transaction_check_pre FPT T \" and u: "u = \Var x not in Fun (Set s) []\" "u \ set (unlabel (transaction_checks T))" shows "s \ \ x" proof - have "is_NegChecks u \ bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p u = [] \ the_eqs u = [] \ length (the_ins u) = 1 \ is_Var (fst (hd (the_ins u))) \ is_Fun_Set (snd (hd (the_ins u))) \ the_Set (the_Fun (snd (hd (the_ins u)))) \ \ (the_Var (fst (hd (the_ins u))))" using T u unfolding transaction_check_pre_def Let_def by blast thus ?thesis using Fun_Set_NotInSet_iff[of u x s] u by argo qed lemma transaction_check_pre_ReceiveE: defines "\ \ \\ x. if fst x = TAtom Value then (absc \ \) x else Var x" assumes T: "transaction_check_pre (FP, OCC, TI) T \" and t: "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" shows "intruder_synth_mod_timpls FP TI (t \ \ \)" using T t unfolding transaction_check_pre_def Let_def \_def by blast lemma transaction_check_compI[intro]: assumes T: "transaction_check_pre (FP, OCC, TI) T \" and T_adm: "admissible_transaction' T" and x1: "\x. (x \ fv_transaction T - set (transaction_fresh T) \ fst x = TAtom Value) \ \ x \ set OCC \ msgcs x (\ x)" and x2: "\x. (x \ fv_transaction T - set (transaction_fresh T) \ fst x \ TAtom Value) \ \ x = {}" shows "\ \ abs_substs_fun ` set (transaction_check_comp msgcs (FP, OCC, TI) T)" proof - define S where "S \ unlabel (transaction_strand T)" define C where "C \ unlabel (transaction_checks T)" let ?xs = "fv_list\<^sub>s\<^sub>s\<^sub>t S" define poss where "poss \ transaction_poschecks_comp C" define negs where "negs \ transaction_negchecks_comp C" define ys where "ys \ filter (\x. x \ set (transaction_fresh T) \ fst x = TAtom Value) ?xs" have ys: "{x \ fv_transaction T - set (transaction_fresh T). fst x = TAtom Value} = set ys" using fv_list\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t[of S] unfolding ys_def S_def by force have "\ x \ set OCC" "msgcs x (\ x)" when x: "x \ set ys" for x using x1 x ys by (blast, blast) moreover have "\ x = {}" when x: "x \ set ys" for x using x2 x ys by blast moreover have "poss x \ \ x" when x: "x \ set ys" for x proof - have "s \ \ x" when u: "u = \a: Var x \ Fun (Set s) []\" "u \ set C" for u a s using T u transaction_check_pre_InSetE[of "(FP, OCC, TI)" T \] unfolding C_def by blast thus ?thesis using transaction_poschecks_comp_unfold[of C x] unfolding poss_def by blast qed moreover have "\ x \ negs x = {}" when x: "x \ set ys" for x proof (cases "x \ fv\<^sub>s\<^sub>s\<^sub>t C") case True hence "s \ \ x" when u: "u = \Var x not in Fun (Set s) []\" "u \ set C" for u s using T u transaction_check_pre_NotInSetE[of "(FP, OCC, TI)" T \] unfolding C_def by blast thus ?thesis using transaction_negchecks_comp_unfold[of C x] unfolding negs_def by blast next case False hence "negs x = {}" using x transaction_negchecks_comp_notin_fv_empty unfolding negs_def by blast thus ?thesis by blast qed ultimately have "\ \ abs_substs_fun ` set (abs_substs_set ys OCC poss negs msgcs)" using abs_substs_has_abs[of ys \ OCC poss negs msgcs] by fast thus ?thesis using T unfolding transaction_check_comp_def Let_def S_def C_def ys_def poss_def negs_def by fastforce qed context begin private lemma transaction_check_comp_in_aux: fixes T defines "C \ set (unlabel (transaction_checks T))" assumes T_adm: "admissible_transaction' T" and a1: "\x \ fv_transaction T - set (transaction_fresh T). fst x = TAtom Value \ (\s. select\Var x, Fun (Set s) []\ \ C \ s \ \ x)" and a2: "\x \ fv_transaction T - set (transaction_fresh T). fst x = TAtom Value \ (\s. \Var x in Fun (Set s) []\ \ C \ s \ \ x)" and a3: "\x \ fv_transaction T - set (transaction_fresh T). fst x = TAtom Value \ (\s. \Var x not in Fun (Set s) []\ \ C \ s \ \ x)" shows "\a x s. \a: Var x \ Fun (Set s) []\ \ C \ s \ \ x" (is ?A) and "\x s. \Var x not in Fun (Set s) []\ \ C \ s \ \ x" (is ?B) proof - note * = admissible_transaction_strand_step_cases(2,3)[OF T_adm] have 1: "fst x = TAtom Value" "x \ fv_transaction T - set (transaction_fresh T)" when x: "\a: Var x \ Fun (Set s) []\ \ C" for a x s using * x unfolding C_def by fast+ have 2: "fst x = TAtom Value" "x \ fv_transaction T - set (transaction_fresh T)" when x: "\Var x not in Fun (Set s) []\ \ C" for x s using * x unfolding C_def by fast+ show ?A proof (intro allI impI) fix a x s assume u: "\a: Var x \ Fun (Set s) []\ \ C" thus "s \ \ x" using 1 a1 a2 by (cases a) metis+ qed show ?B proof (intro allI impI) fix x s assume u: "\Var x not in Fun (Set s) []\ \ C" thus "s \ \ x" using 2 a3 by meson qed qed lemma transaction_check_comp_in: fixes T defines "\ \ \\ x. if fst x = TAtom Value then (absc \ \) x else Var x" and "C \ set (unlabel (transaction_checks T))" assumes T_adm: "admissible_transaction' T" and a1: "\x \ set (transaction_fresh T). \ x = {}" and a2: "\t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T). intruder_synth_mod_timpls FP TI (t \ \ \)" and a3: "\x \ fv_transaction T - set (transaction_fresh T). \s. select\Var x, Fun (Set s) []\ \ C \ s \ \ x" and a4: "\x \ fv_transaction T - set (transaction_fresh T). \s. \Var x in Fun (Set s) []\ \ C \ s \ \ x" and a5: "\x \ fv_transaction T - set (transaction_fresh T). \s. \Var x not in Fun (Set s) []\ \ C \ s \ \ x" and a6: "\x \ fv_transaction T - set (transaction_fresh T). fst x = TAtom Value \ \ x \ set OCC" and a7: "\x \ fv_transaction T - set (transaction_fresh T). fst x = TAtom Value \ msgcs x (\ x)" shows "\\ \ abs_substs_fun ` set (transaction_check_comp msgcs (FP, OCC, TI) T). \x \ fv_transaction T. fst x = TAtom Value \ \ x = \ x" proof - let ?xs = "fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T))" let ?ys = "filter (\x. x \ set (transaction_fresh T)) ?xs" define \' where "\' \ \x. if x \ fv_transaction T - set (transaction_fresh T) \ fst x = TAtom Value then \ x else {}" note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] have \\_Fun: "is_Fun (t \ \ \) \ is_Fun (t \ \ \')" for t unfolding \'_def \_def by (induct t) auto have "\t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T). intruder_synth_mod_timpls FP TI (t \ \ \')" proof (intro ballI impI) fix t assume t: "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" have 1: "intruder_synth_mod_timpls FP TI (t \ \ \)" using t a2 by auto obtain r where r: "r \ set (unlabel (transaction_receive T))" "t \ trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p r" using t by auto hence "\ts. r = receive\ts\ \ t \ set ts" using wellformed_transaction_unlabel_cases(1)[OF T_wf] by fastforce hence 2: "fv t \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" using r by force have "fv t \ fv_transaction T" by (metis (no_types, lifting) 2 transaction_strand_def sst_vars_append_subset(1) unlabel_append subset_Un_eq sup.bounded_iff) moreover have "fv t \ set (transaction_fresh T) = {}" using 2 T_wf vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_receive T)"] unfolding wellformed_transaction_def by fast ultimately have "\ \ x = \ \' x" when "x \ fv t" for x using that unfolding \'_def \_def by fastforce hence 3: "t \ \ \ = t \ \ \'" using term_subst_eq by blast show "intruder_synth_mod_timpls FP TI (t \ \ \')" using 1 3 by simp qed moreover have "\x \ fv_transaction T - set (transaction_fresh T). fst x = TAtom Value \ (\s. select\Var x, Fun (Set s) []\ \ C \ s \ \' x)" "\x \ fv_transaction T - set (transaction_fresh T). fst x = TAtom Value \ (\s. \Var x in Fun (Set s) []\ \ C \ s \ \' x)" "\x \ fv_transaction T - set (transaction_fresh T). fst x = TAtom Value \ (\s. \Var x not in Fun (Set s) []\ \ C \ s \ \' x)" using a3 a4 a5 unfolding \'_def \_def C_def by meson+ hence "\a x s. \a: Var x \ Fun (Set s) []\ \ C \ s \ \' x" "\x s. \Var x not in Fun (Set s) []\ \ C \ s \ \' x" using transaction_check_comp_in_aux[OF T_adm, of \'] unfolding C_def by (fast, fast) ultimately have 4: "transaction_check_pre (FP, OCC, TI) T \'" using a6 transaction_check_preI[of T \' OCC FP TI] unfolding \'_def \_def C_def by simp have 5: "\x \ fv_transaction T. fst x = TAtom Value \ \ x = \' x" using a1 by (auto simp add: \'_def) have 6: "\' \ abs_substs_fun ` set (transaction_check_comp msgcs (FP, OCC, TI) T)" using transaction_check_compI[OF 4 T_adm, of msgcs] a6 a7 unfolding \'_def by auto show ?thesis using 5 6 by blast qed end lemma transaction_check_trivial_case: assumes "transaction_updates T = []" and "transaction_send T = []" shows "transaction_check FPT T" using assms by (simp add: list_all_iff transaction_check_def transaction_check'_def transaction_check_post_def) end subsection \Soundness of the Occurs-Message Transformation\ context stateful_protocol_model begin context begin text \The occurs-message transformation, \add_occurs_msgs\, extends a transaction \T\ with additional message-transmission steps such that the following holds: 1. for each fresh variable \x\ of \T\ the message \occurs (Var x)\ now occurs in a send-step, 2. for each of the remaining free variables \x\ of \T\ the message \occurs (Var x)\ now occurs in a receive-step.\ definition add_occurs_msgs where "add_occurs_msgs T \ let frsh = transaction_fresh T; xs = filter (\x. x \ set frsh) (fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T))); f = map (\x. occurs (Var x)); g = \C. if xs = [] then C else \\, receive\f xs\\#C; h = \F. if frsh = [] then F else if F \ [] \ fst (hd F) = \ \ is_Send (snd (hd F)) then \\,send\f frsh@the_msgs (snd (hd F))\\#tl F else \\,send\f frsh\\#F in case T of Transaction A B C D E F \ Transaction A B (g C) D E (h F)" private fun rm_occurs_msgs_constr where "rm_occurs_msgs_constr [] = []" | "rm_occurs_msgs_constr ((l,receive\ts\)#A) = ( if \t. occurs t \ set ts then if \t \ set ts. \s. t \ occurs s then (l,receive\filter (\t. \s. t \ occurs s) ts\)#rm_occurs_msgs_constr A else rm_occurs_msgs_constr A else (l,receive\ts\)#rm_occurs_msgs_constr A)" | "rm_occurs_msgs_constr ((l,send\ts\)#A) = ( if \t. occurs t \ set ts then if \t \ set ts. \s. t \ occurs s then (l,send\filter (\t. \s. t \ occurs s) ts\)#rm_occurs_msgs_constr A else rm_occurs_msgs_constr A else (l,send\ts\)#rm_occurs_msgs_constr A)" | "rm_occurs_msgs_constr (a#A) = a#rm_occurs_msgs_constr A" private lemma add_occurs_msgs_cases: fixes T C frsh xs f defines "T' \ add_occurs_msgs T" and "frsh \ transaction_fresh T" and "xs \ filter (\x. x \ set frsh) (fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T)))" and "xs' \ fv_transaction T - set frsh" and "f \ map (\x. occurs (Var x))" and "C' \ if xs = [] then C else \\, receive\f xs\\#C" and "ts' \ f frsh" assumes T: "T = Transaction A B C D E F" shows "F = \\,send\ts\\#F' \ T' = Transaction A B C' D E (\\,send\ts'@ts\\#F')" (is "?A ts F' \ ?A' ts F'") and "\ts' F'. F = \\,send\ts'\\#F' \ frsh \ [] \ T' = Transaction A B C' D E (\\,send\ts'\\#F)" (is "?B \ ?B' \ ?B''") and "frsh = [] \ T' = Transaction A B C' D E F" (is "?C \ ?C'") and "transaction_decl T' = transaction_decl T" and "transaction_fresh T' = transaction_fresh T" and "xs = [] \ transaction_receive T' = transaction_receive T" and "xs \ [] \ transaction_receive T' = \\,receive\f xs\\#transaction_receive T" and "transaction_checks T' = transaction_checks T" and "transaction_updates T' = transaction_updates T" and "transaction_send T = \\,send\ts\\#F' \ transaction_send T' = \\,send\ts'@ts\\#F'" (is "?D ts F' \ ?D' ts F'") and "\ts' F'. transaction_send T = \\,send\ts'\\#F' \ frsh \ [] \ transaction_send T' = \\,send\ts'\\#transaction_send T" (is "?E \ ?E' \ ?E''") and "frsh = [] \ transaction_send T' = transaction_send T" (is "?F \ ?F'") and "(xs' \ {} \ transaction_receive T' = \\, receive\f xs\\#transaction_receive T) \ (xs' = {} \ transaction_receive T' = transaction_receive T)" (is ?G) and "(frsh \ [] \ (\ts F'. transaction_send T = \\,send\ts\\#F' \ transaction_send T' = \\,send\ts'@ts\\#F')) \ (frsh \ [] \ transaction_send T' = \\,send\ts'\\#transaction_send T) \ (frsh = [] \ transaction_send T' = transaction_send T)" (is ?H) proof - note defs = T'_def T frsh_def xs_def xs'_def f_def C'_def ts'_def add_occurs_msgs_def Let_def show 0: "?A ts F' \ ?A' ts F'" for ts F' unfolding defs by simp have "F = [] \ fst (hd F) \ \ \ \is_Send (snd (hd F))" when ?B using that unfolding is_Send_def by (cases F) auto thus 1: "?B \ ?B' \ ?B''" unfolding defs by force show "?C \ ?C'" unfolding defs by auto show "transaction_decl T' = transaction_decl T" "transaction_fresh T' = transaction_fresh T" "transaction_checks T' = transaction_checks T" "transaction_updates T' = transaction_updates T" unfolding defs by simp_all show "xs = [] \ transaction_receive T' = transaction_receive T" "xs \ [] \ transaction_receive T' = \\, receive\f xs\\#transaction_receive T" unfolding defs by simp_all moreover have "xs = [] \ xs' = {}" using filter_empty_conv[of "\x. x \ set frsh"] fv_list\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_strand T)"] unfolding xs_def xs'_def by blast ultimately show ?G by blast show 2: "?D ts F' \ ?D' ts F'" for ts F' using 0 unfolding T by simp show 3: "?E \ ?E' \ ?E''" using 1 unfolding T by force show 4: "?F \ ?F'" unfolding defs by simp show ?H proof (cases "frsh = []") case False thus ?thesis using 2 3[OF _ False] by (cases "\ts F'. transaction_send T = \\,send\ts\\#F'") (blast,blast) qed (simp add: 4) qed private lemma add_occurs_msgs_transaction_strand_set: fixes T C frsh xs f defines "frsh \ transaction_fresh T" and "xs \ filter (\x. x \ set frsh) (fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T)))" and "f \ map (\x. occurs (Var x))" assumes T: "T = Transaction A B C D E F" shows "F = \\,send\ts\\#F' \ set (transaction_strand (add_occurs_msgs T)) \ set (transaction_strand T) \ {\\,receive\f xs\\,\\,send\f frsh@ts\\}" (is "?A \ ?A'") and "F = \\,send\ts\\#F' \ set (unlabel (transaction_strand (add_occurs_msgs T))) \ set (unlabel (transaction_strand T)) \ {receive\f xs\,send\f frsh@ts\}" (is "?B \ ?B'") and "\ts' F'. F = \\,send\ts'\\#F' \ set (transaction_strand (add_occurs_msgs T)) \ set (transaction_strand T) \ {\\,receive\f xs\\,\\,send\f frsh\\}" (is "?C \ ?C'") and "\ts' F'. F = \\,send\ts'\\#F' \ set (unlabel (transaction_strand (add_occurs_msgs T))) \ set (unlabel (transaction_strand T)) \ {receive\f xs\,send\f frsh\}" (is "?D \ ?D'") proof - note 0 = add_occurs_msgs_cases[ OF T, unfolded frsh_def[symmetric] xs_def[symmetric] f_def[symmetric]] show "?A \ ?A'" using 0(1,3) unfolding T transaction_strand_def by (cases "frsh = []") auto thus "?B \ ?B'" unfolding unlabel_def by force show "?C \ ?C'" using 0(2,3) unfolding T transaction_strand_def by (cases "frsh = []") auto thus "?D \ ?D'" unfolding unlabel_def by auto qed private lemma add_occurs_msgs_transaction_strand_cases: fixes T T'::"('a,'b,'c,'d) prot_transaction" and C frsh xs f \ defines "T' \ add_occurs_msgs T" and "S \ transaction_strand T" and "S' \ transaction_strand T'" and "frsh \ transaction_fresh T" and "xs \ filter (\x. x \ set frsh) (fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T)))" and "f \ map (\x. occurs (Var x))" and "C \ transaction_receive T" and "D \ transaction_checks T" and "E \ transaction_updates T" and "F \ transaction_send T" and "C' \ if xs = [] then C else \\,receive\f xs\\#C" and "C'' \ if xs = [] then dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t C else \\,send\f xs\\#dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t C" and "C''' \ if xs = [] then dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (C \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) else \\,send\f xs \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\#dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (C \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" shows "frsh = [] \ S' = C'@D@E@F" (is "?A \ ?A'") and "frsh \ [] \ \ts F'. F = \\,send\ts\\#F' \ S' = C'@D@E@(\\,send\f frsh\\#F)" (is "?B \ ?B' \ ?B''") and "frsh \ [] \ \ts F'. F = \\,send\ts\\#F' \ \ts F'. F = \\,send\ts\\#F' \ S' = C'@D@E@(\\,send\f frsh@ts\\#F')" (is "?C \ ?C' \ ?C''") and "frsh = [] \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S' = C''@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t D@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t E@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t F" (is "?D \ ?D'") and "frsh \ [] \ \ts F'. F = \\,send\ts\\#F' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S' = C''@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t D@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t E@(\\,receive\f frsh\\#dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t F)" (is "?E \ ?E' \ ?E''") and "frsh \ [] \ \ts F'. F = \\,send\ts\\#F' \ \ts F'. F = \\,send\ts\\#F' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S' = C''@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t D@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t E@(\\,receive\f frsh@ts\\#dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t F')" (is "?F \ ?F' \ ?F''") and "frsh = [] \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = C'''@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (D \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (E \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (F \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" (is "?G \ ?G'") and "frsh \ [] \ \ts F'. F = \\,send\ts\\#F' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = C'''@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (D \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (E \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@ (\\,receive\f frsh \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\#dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (F \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" (is "?H \ ?H' \ ?H''") and "frsh \ [] \ \ts F'. F = \\,send\ts\\#F' \ \ts F'. F = \\,send\ts\\#F' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = C'''@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (D \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (E \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@ (\\,receive\f frsh@ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\#dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (F' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" (is "?I \ ?I' \ ?I''") proof - obtain A' B' CC' D' E' F' where T: "T = Transaction A' B' CC' D' E' F'" by (cases T) simp note 0 = add_occurs_msgs_cases[ OF T, unfolded frsh_def[symmetric] xs_def[symmetric] f_def[symmetric] T'_def[symmetric]] note defs = S'_def C_def D_def E_def F_def C'_def C''_def T transaction_strand_def show A: "?A \ ?A'" using 0(3) unfolding defs by simp show B: "?B \ ?B' \ ?B''" using 0(2) unfolding defs by simp show C: "?C \ ?C' \ ?C''" using 0(1) unfolding defs by force have 1: "C''' = C'' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using subst_lsst_cons[of "\\, send\f xs\\" "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t C" \] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst[of C \] unfolding C'''_def C''_def by (cases "xs = []") auto have 2: "(\\, receive\ts\\#dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t G) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = \\, receive\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\#dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (G \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" for ts and G::"('a,'b,'c,'d) prot_strand" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst[of G \] subst_lsst_cons[of "\\, receive\ts\\" "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t G" \] by simp note 3 = subst_lsst_append[of _ _ \] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst[of _ \] show "?D \ ?D'" using A unfolding defs by fastforce thus "?G \ ?G'" unfolding 1 by (metis 3) show "?E \ ?E' \ ?E''" using B unfolding defs by fastforce thus "?H \ ?H' \ ?H''" unfolding 1 by (metis 2 3) show "?F \ ?F' \ ?F''" using C unfolding defs by fastforce thus "?I \ ?I' \ ?I''" unfolding 1 by (metis 2 3) qed private lemma add_occurs_msgs_trms_transaction: fixes T::"('a,'b,'c,'d) prot_transaction" shows "trms_transaction (add_occurs_msgs T) = trms_transaction T \ (\x. occurs (Var x))`(fv_transaction T \ set (transaction_fresh T))" (is "?A = ?B") proof let ?occs = "(\x. occurs (Var x)) ` (fv_transaction T \ set (transaction_fresh T))" define frsh where "frsh \ transaction_fresh T" define xs where "xs \ filter (\x. x \ set frsh) (fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T)))" define f where "f \ map (\x. occurs (Var x)::('a,'b,'c,'d) prot_term)" obtain A B C D E F where T: "T = Transaction A B C D E F" by (cases T) simp note 0 = add_occurs_msgs_transaction_strand_set(2,4)[ OF T, unfolded f_def[symmetric] frsh_def[symmetric] xs_def[symmetric]] note 1 = add_occurs_msgs_transaction_strand_cases(1,2,3)[ of T, unfolded f_def[symmetric] frsh_def[symmetric] xs_def[symmetric]] have 2: "set (f xs) \ set (f frsh) = ?occs" proof - define ys where "ys \ fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T))" let ?ys' = "fv_transaction T - set frsh" define g where "g \ filter (\x. x \ set frsh)" have "set (g ys) = ?ys'" using fv_list\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_strand T)"] unfolding ys_def g_def by auto hence "set (f (g ys)) = (\x. occurs (Var x)) ` ?ys'" unfolding f_def by force moreover have "set (f frsh) = (\x. occurs (Var x)) ` set frsh" unfolding f_def by force ultimately show ?thesis unfolding xs_def frsh_def[symmetric] ys_def[symmetric] g_def[symmetric] by blast qed have 3: "set (f []) = {}" unfolding f_def by blast have "trms_transaction (add_occurs_msgs T) \ trms_transaction T \ set (f xs) \ set (f frsh)" proof (cases "\ts F'. F = \\, send\ts\\#F'") case True then obtain ts F' where F: "F = \\, send\ts\\#F'" by blast have "set ts \ trms_transaction T" unfolding T F trms_transaction_unfold by auto thus ?thesis using 0(1)[OF F] by force next case False show ?thesis using 0(2)[OF False] by force qed thus "?A \ ?B" using 2 by blast have "trms_transaction T \ set (f xs) \ set (f frsh) \ trms_transaction (add_occurs_msgs T)" proof (cases "frsh = []") case True show ?thesis using 1(1)[OF True] 3 unfolding True by (cases xs) (fastforce,force) next case False note * = 1(2-)[OF False] show ?thesis proof (cases "\ts F'. transaction_send T = \\, send\ts\\#F'") case True show ?thesis using *(2)[OF True] 3 by force next case False show ?thesis using *(1)[OF False] 3 by force qed qed thus "?B \ ?A" using 2 by blast qed private lemma add_occurs_msgs_vars_eq: fixes T::"('fun,'var,'sets,'lbl) prot_transaction" assumes T_adm: "admissible_transaction' T" shows "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive (add_occurs_msgs T)) = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" (is ?A) and "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send (add_occurs_msgs T)) = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \ set (transaction_fresh T)" (is ?B) and "fv_transaction (add_occurs_msgs T) = fv_transaction T" (is ?C) and "bvars_transaction (add_occurs_msgs T) = bvars_transaction T" (is ?D) and "vars_transaction (add_occurs_msgs T) = vars_transaction T" (is ?E) and "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand (add_occurs_msgs T) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" (is ?F) and "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand (add_occurs_msgs T) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" (is ?G) and "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand (add_occurs_msgs T) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" (is ?H) and "set (transaction_fresh (add_occurs_msgs T)) = set (transaction_fresh T)" (is ?I) proof - obtain A B C D E F where T: "T = Transaction A B C D E F" by (cases T) simp have T_fresh: "set (transaction_fresh T) \ fv_transaction T" using admissible_transactionE(7)[OF T_adm] unfolding fv_transaction_unfold by blast note 0 = add_occurs_msgs_cases[OF T] define xs where "xs \ filter (\x. x \ set (transaction_fresh T)) (fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T)))" show D: ?D proof - have "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive (add_occurs_msgs T)) = bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" using 0(6,7) by (cases "xs = []") (auto simp add: xs_def) moreover have "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send (add_occurs_msgs T)) = bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" proof (cases "\ts' F'. F = \\, send\ts'\\#F'") case True thus ?thesis using 0(1) unfolding T by force next case False show ?thesis using 0(2)[OF False] 0(3) unfolding T by (cases "B = []") auto qed ultimately show ?thesis using 0(8,9) unfolding bvars_transaction_unfold by argo qed have T_no_bvars: "bvars_transaction T = {}" "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) = {}" "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) = {}" "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) = {}" "bvars_transaction (add_occurs_msgs T) = {}" using admissible_transactionE(4)[OF T_adm] D unfolding bvars_transaction_unfold by (blast,blast,blast,blast,blast) have T_fv_subst: "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv_transaction T)" (is ?Q1) "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T))" (is ?Q2) "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T))" (is ?Q3) "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))" (is ?Q4) "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand (add_occurs_msgs T) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand (add_occurs_msgs T)))" (is ?Q5) "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive (add_occurs_msgs T) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive (add_occurs_msgs T)))" (is ?Q6) for \ proof - note * = fv\<^sub>s\<^sub>s\<^sub>t_subst_if_no_bvars have **: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive (add_occurs_msgs T)) = {}" using T_no_bvars(5) unfolding bvars_transaction_unfold by fast show ?Q1 using *[OF T_no_bvars(1)] unfolding unlabel_subst by blast show ?Q2 using *[OF T_no_bvars(2)] unfolding unlabel_subst by blast show ?Q3 using *[OF T_no_bvars(3)] unfolding unlabel_subst by blast show ?Q4 using *[OF T_no_bvars(4)] unfolding unlabel_subst by blast show ?Q5 using *[OF T_no_bvars(5)] unfolding unlabel_subst by blast show ?Q6 using *[OF **] unfolding unlabel_subst by blast qed have A: "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive (add_occurs_msgs T) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" for \ proof - define rcv_trms where "rcv_trms \ map (\x. occurs (Var x)::('fun,'var,'sets,'lbl) prot_term) xs" have "fv\<^sub>s\<^sub>e\<^sub>t (set rcv_trms) = fv_transaction T - set (transaction_fresh T)" "rcv_trms = [] \ xs = []" using fv_list\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_strand T)"] unfolding rcv_trms_def xs_def by auto hence 1: "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive (add_occurs_msgs T)) = (fv_transaction T - set (transaction_fresh T)) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" using 0(6,7)[unfolded rcv_trms_def[symmetric] xs_def[symmetric]] by (cases "xs = []") auto have 2: "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ fv_transaction T - set (transaction_fresh T)" using admissible_transactionE(12)[OF T_adm] unfolding fv_transaction_unfold by fast have 3: "fv_transaction T - set (transaction_fresh T) = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" using admissible_transactionE(7,10,12,13)[OF T_adm] unfolding fv_transaction_unfold by blast show ?thesis using 1 2 3 T_fv_subst(2,3,6)[of \] by force qed show ?A using A[of Var] unfolding subst_lsst_id_subst by blast show B: ?B using 0(14) by fastforce have B': "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send (add_occurs_msgs T) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` set (transaction_fresh T))" for \ proof - note * = fv\<^sub>s\<^sub>s\<^sub>t_subst_if_no_bvars[of _ \] have **: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send (add_occurs_msgs T)) = {}" using T_no_bvars(5) unfolding bvars_transaction_unfold by fast show ?thesis using B *[OF T_no_bvars(4)] *[OF **] unfolding unlabel_subst by simp qed show C: ?C using A[of Var] B T_fresh unfolding fv_transaction_unfold 0(8,9) subst_lsst_id_subst by blast show ?E using C D vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t by metis have "fv\<^sub>s\<^sub>e\<^sub>t (\ ` set (transaction_fresh T)) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using T_fresh unfolding fv\<^sub>s\<^sub>s\<^sub>t_subst_if_no_bvars[OF T_no_bvars(1), of \, unfolded unlabel_subst] by auto thus F: ?F using A[of \] B'[of \] fv\<^sub>s\<^sub>s\<^sub>t_append fv\<^sub>s\<^sub>s\<^sub>t_subst_if_no_bvars[OF T_no_bvars(1), of \, unfolded unlabel_subst] fv\<^sub>s\<^sub>s\<^sub>t_subst_if_no_bvars[OF T_no_bvars(5), of \, unfolded unlabel_subst C] unfolding transaction_strand_def by argo show G: ?G using D bvars\<^sub>s\<^sub>s\<^sub>t_subst unlabel_subst by metis show ?H using F G vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t by metis show ?I using 0(5) by argo qed private lemma add_occurs_msgs_trms: "trms_transaction (add_occurs_msgs T) = trms_transaction T \ (\x. occurs (Var x)) ` (set (transaction_fresh T) \ fv_transaction T)" proof - let ?f = "\x. occurs (Var x)" let ?xs = "filter (\x. x \ set (transaction_fresh T)) (fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T)))" obtain A B C D E F where T: "T = Transaction A B C D E F" by (cases T) simp note 0 = add_occurs_msgs_cases[OF T] have "set ?xs = fv_transaction T - set (transaction_fresh T)" using fv_list\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_strand T)"] by auto hence 1: "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive (add_occurs_msgs T)) = trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ ?f ` (fv_transaction T - set (transaction_fresh T))" using 0(6,7) by (cases "?xs = []") auto have 2: "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send (add_occurs_msgs T)) = trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \ ?f ` set (transaction_fresh T)" using 0(10,11,12) by (cases "transaction_fresh T = []") (simp,fastforce) have 3: "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive (add_occurs_msgs T)) \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send (add_occurs_msgs T)) = trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \ ?f ` (set (transaction_fresh T) \ fv_transaction T)" using 1 2 by blast show ?thesis using 3 unfolding trms_transaction_unfold 0(8,9) by blast qed lemma add_occurs_msgs_admissible_occurs_checks: fixes T::"('fun,'atom,'sets,'lbl) prot_transaction" assumes T_adm: "admissible_transaction' T" shows "admissible_transaction' (add_occurs_msgs T)" (is ?A) and "admissible_transaction_occurs_checks (add_occurs_msgs T)" (is ?B) proof - let ?T' = "add_occurs_msgs T" obtain A B C D E F where T: "T = Transaction A B C D E F" by (cases T) simp note defs = T add_occurs_msgs_def Let_def admissible_transaction'_def admissible_transaction_occurs_checks_def note defs' = admissible_transaction_terms_def wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code[symmetric] note 1 = add_occurs_msgs_cases[OF T] note 2 = add_occurs_msgs_vars_eq[OF T_adm] note 3 = add_occurs_msgs_trms[of T] note 4 = add_occurs_msgs_transaction_strand_set[OF T] have occurs_wf: "wf\<^sub>t\<^sub>r\<^sub>m (occurs (Var x))" for x::"('fun,'atom,'sets,'lbl) prot_var" by fastforce have occurs_funs: "funs_term (occurs (Var x)) = {OccursFact, OccursSec}" for x::"('fun,'atom,'sets,'lbl) prot_var" by force have occurs_funs_not_attack: "\(\f \ \(funs_term ` trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p r). is_Attack f)" when "r = receive\map (\x. occurs (Var x)) xs\ \ r = send\map (\x. occurs (Var x)) ys\" for r:: "(('fun,'atom,'sets,'lbl) prot_fun, ('fun,'atom,'sets,'lbl) prot_var) stateful_strand_step" and xs ys::"('fun,'atom,'sets,'lbl) prot_var list" using that by fastforce have occurs_funs_not_attack': "\(\f \ \(funs_term ` trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p r). is_Attack f)" when "r = send\map (\x. occurs (Var x)) xs@ts\" and "\(\f \ \(funs_term ` trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (send\ts\)). is_Attack f)" for r:: "(('fun,'atom,'sets,'lbl) prot_fun, ('fun,'atom,'sets,'lbl) prot_var) stateful_strand_step" and xs::"('fun,'atom,'sets,'lbl) prot_var list" and ts using that by fastforce let ?P1 = "\T. wellformed_transaction T" let ?P2 = "\T. transaction_decl T () = []" let ?P3 = "\T. list_all (\x. fst x = TAtom Value) (transaction_fresh T)" let ?P4 = "\T. \x \ vars_transaction T. is_Var (fst x) \ (the_Var (fst x) = Value)" let ?P5 = "\T. bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T) = {}" let ?P6 = "\T. set (transaction_fresh T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (filter (is_Insert \ snd) (transaction_updates T)) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" let ?P7 = "\T. \x \ fv_transaction T - set (transaction_fresh T). \y \ fv_transaction T - set (transaction_fresh T). x \ y \ \Var x != Var y\ \ set (unlabel (transaction_checks T)) \ \Var y != Var x\ \ set (unlabel (transaction_checks T))" let ?P8 = "\T. fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) - set (transaction_fresh T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" let ?P9 = "\T. \r \ set (unlabel (transaction_checks T)). is_Equality r \ fv (the_rhs r) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" let ?P10 = "\T. fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (filter (\s. is_InSet (snd s) \ the_check (snd s) = Assign) (transaction_checks T))" let ?P11 = "\T. admissible_transaction_checks T" let ?P12 = "\T. admissible_transaction_updates T" let ?P13 = "\T. admissible_transaction_terms T" let ?P14 = "\T. admissible_transaction_send_occurs_form T" let ?P15 = "\T. list_all (\a. is_Receive (snd a) \ the_msgs (snd a) \ []) (transaction_receive T)" let ?P16 = "\T. list_all (\a. is_Send (snd a) \ the_msgs (snd a) \ []) (transaction_send T)" have T_props: "?P1 T" "?P2 T" "?P3 T" "?P4 T" "?P5 T" "?P6 T" "?P7 T" "?P8 T" "?P9 T" "?P10 T" "?P11 T" "?P12 T" "?P13 T" "?P14 T" "?P15 T" "?P16 T" using T_adm unfolding defs by meson+ have 5: "wf'\<^sub>s\<^sub>s\<^sub>t (X \ Y) (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand (add_occurs_msgs T))))" when X: "X = fst ` set (transaction_decl T ())" and Y: "Y = set (transaction_fresh T)" and T_wf: "wf'\<^sub>s\<^sub>s\<^sub>t (X \ Y) (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)))" for X Y proof - define frsh where "frsh \ transaction_fresh T" define xs where "xs \ fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T))" define ys where "ys \ filter (\x. x \ set frsh) xs" let ?snds = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T))" let ?snds' = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive (add_occurs_msgs T)))" let ?chks = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T))" let ?chks' = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks (add_occurs_msgs T)))" let ?upds = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T))" let ?upds' = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates (add_occurs_msgs T)))" let ?rcvs = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))" let ?rcvs' = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send (add_occurs_msgs T)))" have p0: "set ?snds \ set ?snds'" using 1(13) by auto have p1: "?chks = ?chks'" "?upds = ?upds'" using 1(8,9) by (argo,argo) have p2: "wfvarsoccs\<^sub>s\<^sub>s\<^sub>t ?snds \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t ?snds'" "wfvarsoccs\<^sub>s\<^sub>s\<^sub>t (?snds@?chks@?upds) \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t (?snds'@?chks'@?upds')" "X \ Y \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t (?snds@?chks@?upds) \ X \ Y \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t (?snds'@?chks'@?upds')" using p0 p1 unfolding wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_def by auto have "wf'\<^sub>s\<^sub>s\<^sub>t (X \ Y \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t (?snds@?chks@?upds)) ?rcvs" using T_wf wf\<^sub>s\<^sub>s\<^sub>t_append_exec[of "X \ Y" "?snds@?chks@?upds" ?rcvs] unfolding transaction_strand_unlabel_dual_unfold by simp hence r0: "wf'\<^sub>s\<^sub>s\<^sub>t (X \ Y \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t (?snds'@?chks'@?upds')) ?rcvs" using wf\<^sub>s\<^sub>s\<^sub>t_vars_mono'[OF _ p2(3)] by blast have "list_all is_Send (unlabel (transaction_send T))" using admissible_transaction_is_wellformed_transaction(1)[OF T_adm] unfolding wellformed_transaction_def by blast hence "list_all is_Receive ?rcvs" by (metis dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_list_all(2)) hence r1: "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t ?rcvs \ X \ Y \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t (?snds'@?chks'@?upds')" using wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_receives_only_eq wf\<^sub>s\<^sub>s\<^sub>t_receives_only_fv_subset[OF r0] by blast have "fv\<^sub>s\<^sub>e\<^sub>t ((\x. occurs (Var x)) ` set (transaction_fresh T)) \ Y" unfolding Y by auto hence r2: "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t ?rcvs' \ X \ Y \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t (?snds'@?chks'@?upds')" using 1(14) r1 unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by fastforce (* TODO: find faster proof *) have r3: "wf'\<^sub>s\<^sub>s\<^sub>t (X \ Y) (?snds'@?chks'@?upds')" proof - have *: "wf'\<^sub>s\<^sub>s\<^sub>t (X \ Y) (?snds@?chks'@?upds')" using T_wf wf\<^sub>s\<^sub>s\<^sub>t_prefix[of "X \ Y" "?snds@?chks@?upds" ?rcvs] p1 unfolding transaction_strand_unlabel_dual_unfold by simp have "?snds' = ?snds \ (\ts. ?snds' = send\ts\#?snds)" using 1(13) by auto thus ?thesis proof assume "?snds' = ?snds" thus ?thesis using * by simp next assume "\ts. ?snds' = send\ts\#?snds" then obtain ts where "?snds' = send\ts\#?snds" by blast thus ?thesis using wf\<^sub>s\<^sub>s\<^sub>t_sends_only_prepend[OF *, of "[send\ts\]"] by simp qed qed have "wf'\<^sub>s\<^sub>s\<^sub>t (X \ Y) (?snds'@?chks'@?upds'@?rcvs')" using wf\<^sub>s\<^sub>s\<^sub>t_append_suffix''[OF r3] r2 by auto thus ?thesis using unlabel_append dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_append unfolding transaction_strand_def by auto qed have T'_props_1: "?P1 ?T'" unfolding wellformed_transaction_def apply (intro conjI) subgoal using 1(13) T_props(1) unfolding wellformed_transaction_def by force subgoal using 1(8) T_props(1) unfolding wellformed_transaction_def by simp subgoal using 1(9) T_props(1) unfolding wellformed_transaction_def by simp subgoal using 1(14) T_props(1) unfolding wellformed_transaction_def by force subgoal using 1(4) T_props(1) unfolding wellformed_transaction_def by simp subgoal using 1(5) T_props(1) unfolding wellformed_transaction_def by simp subgoal using 1(4,5) T_props(1) unfolding wellformed_transaction_def by simp subgoal using T_props(1) unfolding 2(1) 1(5) wellformed_transaction_def by blast subgoal using 1(5,8) T_props(1) unfolding wellformed_transaction_def by simp subgoal using 1(5) 2(4) T_props(1) unfolding wellformed_transaction_def by simp subgoal using 2(3,4) T_props(1) unfolding wellformed_transaction_def by simp subgoal using 1(4,5) 5 T_props(1) unfolding wellformed_transaction_def by simp done have T'_props_2_12: "?P2 ?T'" "?P3 ?T'" "?P4 ?T'" "?P5 ?T'" "?P6 ?T'" "?P7 ?T'" "?P8 ?T'" "?P9 ?T'" "?P10 ?T'" "?P11 ?T'" "?P12 ?T'" subgoal using T_props(2) unfolding defs by force subgoal using T_props(3) unfolding defs by force subgoal using T_props(4) 2(5) by argo subgoal using T_props(5) 2(4) by argo subgoal using T_props(6) 1(5,8) 2(2) by auto subgoal using T_props(7) 1(5,8) 2(3) by presburger subgoal using T_props(8) 1(5,9) 2(1,2) by auto subgoal using T_props(9) 1(8) 2(1) by auto subgoal using T_props(10) 1(8) 2(1) by auto subgoal using T_props(11) 1(8) unfolding admissible_transaction_checks_def by argo subgoal using T_props(12) 1(9) unfolding admissible_transaction_updates_def by argo done (* TODO: clean up? *) have T'_props_13_aux: "transaction_fresh ?T' = []" (is ?Q1) "is_Send r" (is ?Q2) "length (the_msgs r) = 1" (is ?Q3) "is_Fun_Attack (hd (the_msgs r))" (is ?Q4) when r: "r \ set (unlabel (transaction_strand (add_occurs_msgs T)))" "\f \ \(funs_term ` trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p r). is_Attack f" (is "?Q' (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p r)") for r proof - note q0 = conjunct2[OF conjunct2[OF T_props(13)[unfolded defs']]] let ?Q'' = "\ts' F'. F = \\, send\ts'\\#F'" let ?f = "map (\x. occurs (Var x))" let ?frsh = "transaction_fresh T" let ?xs = "fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T))" have q1: "r \ send\?f ?frsh\" "r \ receive\?f (filter (\x. x \ set ?frsh) ?xs)\" "\f \ \(funs_term ` set (?f ?frsh)). \is_Attack f" using r(2) by (fastforce,fastforce,simp) have q2: "send\ts'\ \ set (unlabel (transaction_strand T))" "r = send\?f ?frsh@ts'\ \ r \ set (unlabel (transaction_strand T))" when "?Q'' ts' F'" for ts' F' subgoal using that unfolding T transaction_strand_def by force subgoal using that r(1) 4(2)[OF that] q1 unfolding T transaction_strand_def by fast done have q3: "?Q' (set ts')" when r': "?Q'' ts' F'" "r \ set (unlabel (transaction_strand T))" for ts' F' proof - have "r = send\?f ?frsh@ts'\" using q2(2)[OF r'(1)] r'(2) by argo thus ?thesis using r(2) by fastforce qed have q4: "r \ set (unlabel (transaction_strand T))" when "\ts' F'. ?Q'' ts' F'" using 4(4)[OF that] r(1) q1(1,2) by blast have "\r' \ set (unlabel (transaction_strand T)). ?Q' (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p r')" when "?Q'' ts' F'" for ts' F' apply (cases "r \ set (unlabel (transaction_strand T))") subgoal using q2(2)[OF that] r(2) by metis subgoal using q2(1)[OF that] q3[OF that] trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p.simps(1)[of ts'] by metis done hence "?frsh = []" when "?Q'' ts' F'" for ts' F' using q0 that by blast hence "r = send\ts'\ \ r \ set (unlabel (transaction_strand T))" when "?Q'' ts' F'" for ts' F' using q2(2)[OF that] that by blast hence "r \ set (unlabel (transaction_strand T))" using q2(1) q4 by fast thus ?Q1 ?Q2 ?Q3 ?Q4 using r(2) q0 unfolding 1(5) by auto qed have T'_props_13: "?P13 ?T'" unfolding defs' 3 apply (intro conjI) subgoal using conjunct1[OF T_props(13)[unfolded defs']] occurs_wf by fast subgoal using conjunct1[OF conjunct2[OF T_props(13)[unfolded defs']]] occurs_funs by auto subgoal using T'_props_13_aux by meson done have T'_props_14: "?P14 ?T'" proof (cases "\ts' F'. transaction_send T = \\,send\ts'\\#F'") case True then obtain ts' F' where F': "transaction_send T = \\,send\ts'\\#F'" by meson show ?thesis using T_props(14) 1(10)[OF F'] F' 1(5,12) unfolding admissible_transaction_send_occurs_form_def Let_def by (cases "transaction_fresh T = []") auto next case False show ?thesis using T_props(14) 1(11)[OF False] 1(5,12) unfolding admissible_transaction_send_occurs_form_def Let_def by (cases "transaction_fresh T = []") auto qed let ?xs = "fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T))" have T'_props_15: "?P15 ?T'" using T_props(15) 1(6,7) unfolding Let_def by (cases "filter (\x. x \ set (transaction_fresh T)) ?xs = []") (simp,fastforce) have T'_props_16: "?P16 ?T'" proof (cases "\ts' F'. transaction_send T = \\,send\ts'\\#F'") case True then obtain ts' F' where F': "transaction_send T = \\,send\ts'\\#F'" by meson show ?thesis using T_props(16) 1(10)[OF F'] F' 1(5,12) unfolding Let_def by (cases "transaction_fresh T = []") auto next case False show ?thesis using T_props(16) 1(11)[OF False] 1(5,12) unfolding Let_def by (cases "transaction_fresh T = []") auto qed note T'_props = T'_props_1 T'_props_2_12 T'_props_13 T'_props_14 T'_props_15 T'_props_16 show ?A using T'_props unfolding admissible_transaction'_def by meson have 5: "set (filter (\x. x \ set (transaction_fresh T)) (fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T)))) = fv_transaction T - set (transaction_fresh T)" using fv_list\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t by fastforce have "transaction_receive ?T' \ []" and "is_Receive (hd (unlabel (transaction_receive ?T')))" and "\x \ fv_transaction ?T' - set (transaction_fresh ?T'). fst x = TAtom Value \ occurs (Var x) \ set (the_msgs (hd (unlabel (transaction_receive ?T'))))" when x: "x \ fv_transaction ?T' - set (transaction_fresh ?T')" "fst x = TAtom Value" for x using 1(13) 5 x unfolding 1(5) 2(3) by (force,force,force) moreover have "transaction_send ?T' \ []" (is ?C) and "is_Send (hd (unlabel (transaction_send ?T')))" (is ?D) and "\x \ set (transaction_fresh ?T'). occurs (Var x) \ set (the_msgs (hd (unlabel (transaction_send ?T'))))" (is ?E) when T'_frsh: "transaction_fresh ?T' \ []" using 1(14) T'_frsh unfolding 1(5) by auto ultimately show ?B using T'_props_14 unfolding admissible_transaction_occurs_checks_def Let_def by blast qed private lemma add_occurs_msgs_in_trms_subst_cases: fixes T::"('fun,'atom,'sets,'lbl) prot_transaction" assumes T_adm: "admissible_transaction' T" and t: "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand (add_occurs_msgs T) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" shows "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ (\x \ fv_transaction T. t = occurs (\ x))" proof - define frsh where "frsh \ transaction_fresh T" define xs where "xs \ filter (\x. x \ set frsh) (fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T)))" define f where "f \ map (\x. occurs (Var x)::('fun,'atom,'sets,'lbl) prot_term)" obtain A B C D E F where T: "T = Transaction A B C D E F" by (cases T) simp note T'_adm = add_occurs_msgs_admissible_occurs_checks(1)[OF T_adm] have 0: "set (transaction_fresh T) \ fv_transaction T" using admissible_transactionE(7)[OF T_adm] unfolding fv_transaction_unfold by blast hence 00: "set (f xs) \ set (f frsh) = (\x. occurs (Var x)) ` fv_transaction T" using fv_list\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_strand T)"] unfolding f_def xs_def frsh_def by auto note 1 = add_occurs_msgs_transaction_strand_set[OF T] have 2: "set (transaction_strand (add_occurs_msgs T)) \ set (transaction_strand T) \ {\\,receive\f xs\\,\\,send\f frsh\\}" when "\ts F'. F = \\,send\ts\\#F'" using 1(3,4)[OF that] unfolding f_def[symmetric] frsh_def[symmetric] xs_def[symmetric] by blast have 3: "trms_transaction (add_occurs_msgs T) = trms_transaction T \ (\x. occurs (Var x)) ` fv_transaction T" using 0 add_occurs_msgs_trms_transaction[of T] by blast have 4: "bvars_transaction T \ subst_domain \ = {}" "bvars_transaction (add_occurs_msgs T) \ subst_domain \ = {}" using admissible_transactionE(4)[OF T_adm] admissible_transactionE(4)[OF T'_adm] by (blast,blast) note 5 = trms\<^sub>s\<^sub>s\<^sub>t_subst[OF 4(1), unfolded unlabel_subst] trms\<^sub>s\<^sub>s\<^sub>t_subst[OF 4(2), unfolded unlabel_subst] note 6 = fv\<^sub>s\<^sub>s\<^sub>t_is_subterm_trms\<^sub>s\<^sub>s\<^sub>t_subst[ OF _ 4(1), unfolded add_occurs_msgs_admissible_occurs_checks(1)[OF T_adm] unlabel_subst] show ?thesis using t 6 unfolding 3 5 by fastforce qed private lemma add_occurs_msgs_updates_send_filter_iff: fixes f defines "f \ \T. list_ex (\a. is_Send (snd a) \ is_Update (snd a)) (transaction_strand T)" and "g \ \T. transaction_fresh T = [] \ f T" shows "map add_occurs_msgs (filter g P) = filter g (map add_occurs_msgs P)" proof - have "g T \ g (add_occurs_msgs T)" for T proof - obtain A B C D E F where T: "T = Transaction A B C D E F" by (cases T) simp_all note 0 = add_occurs_msgs_cases[OF T] show ?thesis using 0(6,7,12) unfolding g_def f_def transaction_strand_def 0(5,8,9) by fastforce qed thus ?thesis by (induct P) simp_all qed lemma add_occurs_msgs_updates_send_filter_iff': fixes f defines "f \ \T. list_ex (\a. is_Send (snd a) \ is_Update (snd a)) (transaction_strand T)" and "g \ \T. transaction_fresh T = [] \ transaction_updates T \ [] \ transaction_send T \ []" shows "map add_occurs_msgs (filter g P) = filter g (map add_occurs_msgs P)" proof - have "g T \ g (add_occurs_msgs T)" for T proof - obtain A B C D E F where T: "T = Transaction A B C D E F" by (cases T) simp_all note 0 = add_occurs_msgs_cases[OF T] show ?thesis using 0(6,7,12) unfolding g_def f_def transaction_strand_def 0(5,8,9) by argo qed thus ?thesis by (induct P) simp_all qed private lemma rm_occurs_msgs_constr_Cons: defines "f \ rm_occurs_msgs_constr" shows "\is_Receive a \ \is_Send a \ f ((l,a)#A) = (l,a)#f A" "is_Receive a \ \t. occurs t \ set (the_msgs a) \ f ((l,a)#A) = (l,a)#f A" "is_Receive a \ \t. occurs t \ set (the_msgs a) \ \t \ set (the_msgs a). \s. t \ occurs s \ f ((l,a)#A) = (l,receive\filter (\t. \s. t \ occurs s) (the_msgs a)\)#f A" "is_Receive a \ \t. occurs t \ set (the_msgs a) \ \t \ set (the_msgs a). \s. t = occurs s \ f ((l,a)#A) = f A" "is_Send a \ \t. occurs t \ set (the_msgs a) \ f ((l,a)#A) = (l,a)#f A" "is_Send a \ \t. occurs t \ set (the_msgs a) \ \t \ set (the_msgs a). \s. t \ occurs s \ f ((l,a)#A) = (l,send\filter (\t. \s. t \ occurs s) (the_msgs a)\)#f A" "is_Send a \ \t. occurs t \ set (the_msgs a) \ \t \ set (the_msgs a). \s. t = occurs s \ f ((l,a)#A) = f A" unfolding f_def by (cases a; auto)+ private lemma rm_occurs_msgs_constr_Cons': defines "f \ rm_occurs_msgs_constr" and "g \ filter (\t. \s. t \ occurs s)" assumes a: "is_Receive a \ is_Send a" shows "\t. occurs t \ set (the_msgs a) \ f ((l,a)#A) = (l,a)#f A" "\t. occurs t \ set (the_msgs a) \ \t \ set (the_msgs a). \s. t \ occurs s \ is_Send a \ f ((l,a)#A) = (l,send\g (the_msgs a)\)#f A" "\t. occurs t \ set (the_msgs a) \ \t \ set (the_msgs a). \s. t \ occurs s \ is_Receive a \ f ((l,a)#A) = (l,receive\g (the_msgs a)\)#f A" "\t. occurs t \ set (the_msgs a) \ \t \ set (the_msgs a). \s. t = occurs s \ f ((l,a)#A) = f A" using a unfolding f_def g_def by (cases a; auto)+ private lemma rm_occurs_msgs_constr_Cons'': defines "f \ rm_occurs_msgs_constr" and "g \ filter (\t. \s. t \ occurs s)" assumes a: "a = receive\ts\ \ a = send\ts\" shows "f ((l,a)#A) = (l,a)#f A \ f ((l,a)#A) = (l,receive\g ts\)#f A \ f ((l,a)#A) = (l,send\g ts\)#f A \ f ((l,a)#A) = f A" using rm_occurs_msgs_constr_Cons(2-)[of a l A] a unfolding f_def g_def by (cases a) auto private lemma rm_occurs_msgs_constr_ik_subset: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (rm_occurs_msgs_constr A) \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" proof (induction A) case (Cons a A) let ?f = "filter (\t. \s. t \ occurs s)" note IH = Cons.IH obtain l b where a: "a = (l,b)" by (metis surj_pair) have 0: "set (unlabel A) \ set (unlabel (a#A))" by auto note 1 = rm_occurs_msgs_constr_Cons[of b l A] note 2 = in_ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t_iff note 3 = ik\<^sub>s\<^sub>s\<^sub>t_set_subset[OF 0] note 4 = ik\<^sub>s\<^sub>s\<^sub>t_append note 5 = 4[of "unlabel [a]" "unlabel A"] 4[of "unlabel [a]" "unlabel (rm_occurs_msgs_constr A)"] show ?case proof (cases "is_Send b \ is_Receive b") case True note b_cases = this define ts where "ts \ the_msgs b" have ts_cases: "is_Send b \ b = send\ts\" "is_Receive b \ b = receive\ts\" unfolding ts_def by simp_all have 6: "is_Send b \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t [(l,b)] = {}" "is_Send b \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t [(l,send\the_msgs b\)] = {}" "is_Send b \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t [(l,send\?f (the_msgs b)\)] = {}" "is_Receive b \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t [(l,b)] = set ts" "is_Receive b \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t [(l,receive\the_msgs b\)] = set ts" "is_Receive b \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t [(l,receive\?f (the_msgs b)\)] = set (?f ts)" using 2[of _ "[(l, send\the_msgs b\)]"] 2[of _ "[(l, send\?f (the_msgs b)\)]"] 2[of _ "[(l, receive\the_msgs b\)]"] 2[of _ "[(l, receive\?f (the_msgs b)\)]"] b_cases ts_cases by auto have "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (rm_occurs_msgs_constr (a#A)) = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (rm_occurs_msgs_constr A)" when b: "is_Send b" proof (cases "\t. occurs t \ set (the_msgs b)") case True note 7 = 1(6,7)[OF b True] show ?thesis proof (cases "\t \ set (the_msgs b). \s. t \ occurs s") case True show ?thesis using 4[of "unlabel [(l,send\?f (the_msgs b)\)]" "unlabel (rm_occurs_msgs_constr A)"] unfolding a 7(1)[OF True] 6(3)[OF b] by simp next case False hence F: "\t \ set (the_msgs b). \s. t = occurs s" by simp show ?thesis using 4[of "unlabel [(l,send\the_msgs b\)]" "unlabel (rm_occurs_msgs_constr A)"] unfolding a 7(2)[OF F] 6(2)[OF b] by simp qed next case False show ?thesis using 4[of "unlabel [(l,b)]" "unlabel (rm_occurs_msgs_constr A)"] unfolding a 1(5)[OF b False] 6(1)[OF b] by auto qed moreover have "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (rm_occurs_msgs_constr (a#A)) \ set ts \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (rm_occurs_msgs_constr A)" when b: "is_Receive b" proof (cases "\t. occurs t \ set (the_msgs b)") case True note 8 = 1(3,4)[OF b True] show ?thesis proof (cases "\t \ set (the_msgs b). \s. t \ occurs s") case True show ?thesis using 4[of "unlabel [(l,receive\?f (the_msgs b)\)]" "unlabel (rm_occurs_msgs_constr A)"] unfolding a 8(1)[OF True] 6(6)[OF b] by auto next case False hence F: "\t \ set (the_msgs b). \s. t = occurs s" by simp show ?thesis using 4[of "unlabel [(l,receive\the_msgs b\)]" "unlabel (rm_occurs_msgs_constr A)"] unfolding a 8(2)[OF F] 6(5)[OF b] by simp qed next case False show ?thesis using 4[of "unlabel [(l,b)]" "unlabel (rm_occurs_msgs_constr A)"] unfolding a 1(2)[OF b False] 6(4)[OF b] by auto qed moreover have "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A) = set ts \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" when b: "is_Receive b" using ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t_Cons(2)[of l ts A] unfolding a ts_cases(2)[OF b] by blast ultimately show ?thesis using IH 3 b_cases by blast qed (use 1(1) IH 5 a in auto) qed simp private lemma rm_occurs_msgs_constr_append: "rm_occurs_msgs_constr (A@B) = rm_occurs_msgs_constr A@rm_occurs_msgs_constr B" by (induction A rule: rm_occurs_msgs_constr.induct) auto private lemma rm_occurs_msgs_constr_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t: "rm_occurs_msgs_constr (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (rm_occurs_msgs_constr A)" proof (induction A) case (Cons a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) show ?case using Cons.IH unfolding a by (cases b) auto qed simp private lemma rm_occurs_msgs_constr_dbupd\<^sub>s\<^sub>s\<^sub>t_eq: "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (rm_occurs_msgs_constr A)) I D = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I D" proof (induction A arbitrary: I D) case (Cons a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) show ?case proof (cases "is_Receive b \ is_Send b") case True then obtain ts where b: "b = receive\ts\ \ b = send\ts\" by (cases b) simp_all show ?thesis using rm_occurs_msgs_constr_Cons''[OF b, of l A] Cons.IH b unfolding a by fastforce next case False thus ?thesis using Cons.IH unfolding a by (cases b) auto qed qed simp private lemma rm_occurs_msgs_constr_subst: fixes A::"('a,'b,'c,'d) prot_strand" and \::"('a,'b,'c,'d) prot_subst" assumes "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. \t. \ x = occurs t" "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. \ x \ Fun OccursSec []" shows "rm_occurs_msgs_constr (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = (rm_occurs_msgs_constr A) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" (is "?f (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = (?f A) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \") using assms proof (induction A) case (Cons a A) note 0 = rm_occurs_msgs_constr_Cons note 1 = rm_occurs_msgs_constr_Cons' define f where "f \ ?f" define not_occ where "not_occ \ \t::('a,'b,'c,'d) prot_term. \s. t \ occurs s" define flt where "flt \ filter not_occ" obtain l b where a: "a = (l,b)" by (metis surj_pair) have 2: "\t. \ x = occurs t" "\ x \ Fun OccursSec []" when b: "is_Receive b \ is_Send b" and t: "t \ set (the_msgs b)" and x: "x \ fv t" for x t using Cons.prems x t b unfolding a by (cases b; auto)+ have IH: "f (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = (f A) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using Cons.prems Cons.IH unfolding f_def by simp show ?case proof (cases "is_Receive b \ is_Send b") case True note T = this then obtain ts where ts: "b = receive\ts\ \ b = send\ts\" by (cases b) simp_all hence ts': "b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = receive\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\ \ b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = send\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\" by auto have the_msgs_b: "the_msgs b = ts" "the_msgs (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" using ts ts' by auto have 4: "is_Receive (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \ is_Send (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" using T by (cases b) simp_all note 6 = 1[OF T, of l A, unfolded f_def[symmetric]] note 7 = 1[OF 4, of l "A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \", unfolded f_def[symmetric]] note 8 = ts IH subst_lsst_cons[of _ _ \] have 9: "t \ \ \ set (the_msgs (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \))" "not_occ (t \ \)" when t: "t \ set (the_msgs b)" "not_occ t" for t proof - show "t \ \ \ set (the_msgs (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \))" using t ts ts' by auto moreover have "not_occ (t \ \)" when "t = Var x" for x using 2[OF T t(1)] t(2) unfolding that not_occ_def by simp moreover have "not_occ (t \ \)" when "t = Fun g ss" "g \ OccursFact" for g ss using 2[OF T t(1)] t(2) that(2) unfolding that(1) not_occ_def by simp moreover have "not_occ (t \ \)" when "t = Fun OccursFact ss" "\s1 s2. ss = [s1,s2]" for ss using 2[OF T t(1)] t(2) that(2) unfolding that(1) not_occ_def by auto moreover have "not_occ (t \ \)" when "t = Fun OccursFact [s1,s2]" for s1 s2 using 2[OF T t(1)] t(2) unfolding that not_occ_def by (cases s1) auto ultimately show "not_occ (t \ \)" by (cases t) (metis, metis) qed have 10: "not_occ t" when t: "t \ set (the_msgs b)" "not_occ (t \ \)" for t proof - have "t \ \ \ set (the_msgs (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \))" using t ts ts' by auto moreover have "not_occ t" when "t = Var x" for x using 2[OF T t(1)] t(2) unfolding that not_occ_def by simp moreover have "not_occ t" when "t = Fun g ss" "g \ OccursFact" for g ss using 2[OF T t(1)] t(2) that(2) unfolding that(1) not_occ_def by simp moreover have "not_occ t" when "t = Fun OccursFact ss" "\s1 s2. ss = [s1,s2]" for ss using 2[OF T t(1)] t(2) that(2) unfolding that(1) not_occ_def by auto moreover have "not_occ t" when "t = Fun OccursFact [s1,s2]" for s1 s2 using 2[OF T t(1)] t(2) unfolding that not_occ_def by (cases s1) auto ultimately show "not_occ t" unfolding not_occ_def by force qed have 11: "not_occ (t \ \) \ not_occ t" when "t \ set ts" for t using that 9 10 unfolding the_msgs_b by blast have 5: "(\t. occurs t \ set ts) \ (\t. occurs t \ set ts \\<^sub>s\<^sub>e\<^sub>t \)" using 11 image_iff unfolding not_occ_def by fastforce have 12: "flt (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \) = (flt ts) \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" using 11 proof (induction ts) case (Cons t ts) hence "not_occ (t \ \) = not_occ t" "flt (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \) = (flt ts) \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" by auto thus ?case unfolding flt_def by auto qed (metis flt_def filter.simps(1) map_is_Nil_conv) show ?thesis proof (cases "\t. occurs t \ set (the_msgs b)") case True note T1 = this hence T2: "\t. occurs t \ set (the_msgs (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \))" using 5 unfolding the_msgs_b by simp show ?thesis proof (cases "\t \ set (the_msgs b). \s. t \ occurs s") case True note T1' = this have T2': "\t \ set (the_msgs (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)). \s. t \ occurs s" using T1' 11 unfolding the_msgs_b not_occ_def by auto show ?thesis using T proof assume b: "is_Receive b" hence b\: "is_Receive (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" using ts by fastforce show ?thesis using 6(3)[OF T1 T1' b] 7(3)[OF T2 T2' b\] IH 12 unfolding f_def[symmetric] a flt_def[symmetric] not_occ_def[symmetric] the_msgs_b by (simp add: subst_lsst_cons) next assume b: "is_Send b" hence b\: "is_Send (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" using ts by fastforce show ?thesis using 6(2)[OF T1 T1' b] 7(2)[OF T2 T2' b\] IH 12 unfolding f_def[symmetric] a flt_def[symmetric] not_occ_def[symmetric] the_msgs_b by (simp add: subst_lsst_cons) qed next case False hence F: "\t \ set (the_msgs b). \s. t = occurs s" by blast hence F': "\t \ set (the_msgs (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)). \s. t = occurs s" unfolding the_msgs_b by auto have *: "\t. occurs t \ set (the_msgs b)" when "the_msgs b \ []" using that F by (cases "the_msgs b") auto hence **: "\t. occurs t \ set (the_msgs (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \))" when "the_msgs b \ []" using that 5 unfolding the_msgs_b by simp show ?thesis proof (cases "ts = []") case True hence ***: "\t. occurs t \ set (the_msgs b)" "\t. occurs t \ set (the_msgs (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \))" unfolding the_msgs_b by simp_all show ?thesis using IH 6(1)[OF ***(1)] 7(1)[OF ***(2)] unfolding a f_def[symmetric] True by (simp add: subst_lsst_cons) next case False thus ?thesis using IH 6(4)[OF * F] 7(4)[OF ** F'] unfolding f_def[symmetric] not_occ_def[symmetric] a the_msgs_b by (simp add: subst_lsst_cons) qed qed next case False note F = this have F': "\t. occurs t \ set (the_msgs (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \))" using F 11 unfolding not_occ_def the_msgs_b by fastforce show ?thesis using IH 6(1)[OF F] 7(1)[OF F'] unfolding a f_def[symmetric] True by (simp add: subst_lsst_cons) qed next case False hence *: "\is_Receive b" "\is_Send b" "\is_Receive (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "\is_Send (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" by (cases b; auto)+ show ?thesis using IH 0(1)[OF *(1,2), of l A] 0(1)[OF *(3,4), of l "A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] subst_lsst_cons[of a _ \] unfolding a f_def by simp qed qed simp private lemma rm_occurs_msgs_constr_transaction_strand: assumes T_adm: "admissible_transaction' T" shows "rm_occurs_msgs_constr (transaction_checks T) = transaction_checks T" (is ?A) and "rm_occurs_msgs_constr (transaction_updates T) = transaction_updates T" (is ?B) and "admissible_transaction_no_occurs_msgs T \ rm_occurs_msgs_constr (transaction_receive T) = transaction_receive T" (is "?C \ ?C'") and "admissible_transaction_no_occurs_msgs T \ rm_occurs_msgs_constr (transaction_send T) = transaction_send T" (is "?D \ ?D'") proof - note 0 = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] note 1 = wellformed_transaction_cases[OF 0] have 2: "\ts. b = receive\ts\ \ (\t. occurs t \ set ts)" when "admissible_transaction_no_occurs_msgs T" "(l,b) \ set (transaction_receive T)" for l b using that 1(1)[OF that(2)] unfolding admissible_transaction_no_occurs_msgs_def Let_def list_all_iff by fastforce have 3: "\ts. b = send\ts\ \ (\t. occurs t \ set ts)" when "admissible_transaction_no_occurs_msgs T" "(l,b) \ set (transaction_send T)" for l b using that 1(4)[OF that(2)] unfolding admissible_transaction_no_occurs_msgs_def Let_def list_all_iff by fastforce define A where "A \ transaction_receive T" define B where "B \ transaction_checks T" define C where "C \ transaction_updates T" define D where "D \ transaction_send T" show ?A using 1(2) unfolding B_def[symmetric] proof (induction B) case (Cons a A) hence IH: "rm_occurs_msgs_constr A = A" by (meson list.set_intros(2)) obtain l b where a: "a = (l,b)" by (metis surj_pair) show ?case using Cons.prems IH unfolding a by (cases b) auto qed simp show ?B using 1(3) unfolding C_def[symmetric] proof (induction C) case (Cons a A) hence IH: "rm_occurs_msgs_constr A = A" by (meson list.set_intros(2)) obtain l b where a: "a = (l,b)" by (metis surj_pair) show ?case using Cons.prems IH unfolding a by (cases b) auto qed simp show ?C' when ?C using 2[OF that] unfolding A_def[symmetric] proof (induction A) case (Cons a A) hence IH: "rm_occurs_msgs_constr A = A" by (meson list.set_intros(2)) obtain l b where a: "a = (l,b)" by (metis surj_pair) obtain ts where b: "b = receive\ts\" using Cons.prems unfolding a by auto show ?case using Cons.prems IH unfolding a b by fastforce qed simp show ?D' when ?D using 3[OF that] unfolding D_def[symmetric] proof (induction D) case (Cons a A) hence IH: "rm_occurs_msgs_constr A = A" by (meson list.set_intros(2)) obtain l b where a: "a = (l,b)" by (metis surj_pair) obtain ts where b: "b = send\ts\" using Cons.prems unfolding a by auto show ?case using Cons.prems IH unfolding a b by fastforce qed simp qed private lemma rm_occurs_msgs_constr_transaction_strand': fixes T::"('fun,'atom,'sets,'lbl) prot_transaction" assumes T_adm: "admissible_transaction' T" and T_no_occ: "admissible_transaction_no_occurs_msgs T" shows "rm_occurs_msgs_constr (transaction_strand (add_occurs_msgs T)) = transaction_strand T" (is "?f (?g (?h T)) = ?g T") proof - obtain A B C D E F where T: "T = Transaction A B C D E F" by (cases T) simp have B: "B = transaction_fresh T" unfolding T by simp have F: "F = transaction_send T" unfolding T by simp define xs where "xs \ filter (\x. x \ set B) (fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T)))" note 0 = rm_occurs_msgs_constr_transaction_strand note 1 = add_occurs_msgs_admissible_occurs_checks[OF T_adm] note 2 = 0(3,4)[OF T_adm T_no_occ] note 3 = add_occurs_msgs_cases[OF T] note 4 = 0(1,2)[OF 1(1)] have 5: "?f (transaction_checks (?h T)) = transaction_checks T" "?f (transaction_updates (?h T)) = transaction_updates T" using 4 3(8,9) by (argo, argo) have 6: "?f (transaction_receive (?h T)) = transaction_receive T" proof (cases "xs = []") case True show ?thesis using 3(6)[OF True[unfolded xs_def B]] 2(1) by simp next case False show ?thesis using False 3(7)[OF False[unfolded xs_def B]] 2(1) rm_occurs_msgs_constr_Cons(4)[ of "receive\map (\x. occurs (Var x)) xs\" \ "transaction_receive T"] unfolding B[symmetric] xs_def[symmetric] by (cases xs) (blast, auto) qed have 7: "?f (transaction_send (?h T)) = transaction_send T" proof (cases "\ts' F'. F = \\, send\ts'\\#F'") case True then obtain ts' F' where F': "F = \\, send\ts'\\#F'" by blast have *: "transaction_send (?h T) = \\, send\map (\x. occurs (Var x)) B@ts'\\#F'" using 3(1)[OF F'] unfolding T by fastforce have **: "ts' \ []" using admissible_transactionE(17)[OF T_adm] unfolding T F' by auto have ***: "\s. t \ occurs s" when t: "t \ set ts'" for t using that T_no_occ unfolding T F' admissible_transaction_no_occurs_msgs_def Let_def list_all_iff by auto let ?ts = "map (\x. occurs (Var x)) B@ts'" have "\t \ set ?ts. \s. t \ occurs s" using ** *** by (cases ts') auto moreover have "filter (\t. \s. t \ occurs s) ?ts = ts'" using *** by simp moreover have "\t. occurs t \ set ?ts" when "B \ []" using that by (cases B) auto moreover have "?f [\\, send\ts'\\] = [\\, send\ts'\\]" using 2(2) ** *** unfolding F[symmetric] F' by force hence "?f F' = F'" using 2(2) rm_occurs_msgs_constr_append[of "[\\, send\ts'\\]" F'] unfolding F[symmetric] F' by fastforce ultimately have "?f (\\, send\?ts\\#F') = \\, send\ts'\\#F'" using 2(2) 3(10)[OF F'[unfolded F]] 3(12) rm_occurs_msgs_constr.simps(3)[of \ ts' F'] rm_occurs_msgs_constr_append[of "[\\, send\ts'\\]" F'] unfolding F[symmetric] F' B[symmetric] by auto thus ?thesis using F * unfolding F' by argo next case False show ?thesis using 3(2)[OF False] 3(3) 2(2) unfolding B[symmetric] xs_def[symmetric] F[symmetric] by (cases B) auto qed show ?thesis using 5 6 7 rm_occurs_msgs_constr_append unfolding transaction_strand_def by metis qed private lemma rm_occurs_msgs_constr_transaction_strand'': fixes T::"('fun,'atom,'sets,'lbl) prot_transaction" assumes T_adm: "admissible_transaction' T" and T_no_occ: "admissible_transaction_no_occurs_msgs T" and \: "\x \ fv_transaction (add_occurs_msgs T). \t. \ x = occurs t" "\x \ fv_transaction (add_occurs_msgs T). \ x \ Fun OccursSec []" shows "rm_occurs_msgs_constr (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand (add_occurs_msgs T) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using rm_occurs_msgs_constr_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t[of "transaction_strand (add_occurs_msgs T) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] rm_occurs_msgs_constr_subst[OF \] rm_occurs_msgs_constr_transaction_strand'[OF T_adm T_no_occ] by argo private lemma rm_occurs_msgs_constr_bvars_subst_eq: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (rm_occurs_msgs_constr A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" proof - have "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (rm_occurs_msgs_constr A) = bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" proof (induction A) case (Cons a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) show ?case using Cons.IH unfolding a by (cases b) auto qed simp thus ?thesis by (metis bvars\<^sub>s\<^sub>s\<^sub>t_subst unlabel_subst) qed private lemma rm_occurs_msgs_constr_reachable_constraints_fv_eq: assumes P: "\T \ set P. admissible_transaction' T" "\T \ set P. admissible_transaction_no_occurs_msgs T" and A: "A \ reachable_constraints (map add_occurs_msgs P)" shows "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (rm_occurs_msgs_constr A) = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" using A proof (induction A rule: reachable_constraints.induct) case (step \ T \ \ \) let ?f = rm_occurs_msgs_constr let ?B = "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" define \ where "\ \ \ \\<^sub>s \ \\<^sub>s \" obtain T' where T': "T' \ set P" "T = add_occurs_msgs T'" using step.hyps(2) by auto have T_adm: "admissible_transaction' T" using add_occurs_msgs_admissible_occurs_checks(1) step.hyps(2) P by auto have T'_adm: "admissible_transaction' T'" and T'_no_occ: "admissible_transaction_no_occurs_msgs T'" using T'(1) P by (blast,blast) have "?f (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using rm_occurs_msgs_constr_transaction_strand''[OF T'_adm T'_no_occ, of \] admissible_transaction_decl_fresh_renaming_subst_not_occurs[OF T_adm step.hyps(3,4,5)] unfolding T'(2) \_def[symmetric] by blast moreover have "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using add_occurs_msgs_vars_eq(6)[OF T'_adm, of \] unfolding T'(2) by blast ultimately have "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?f ?B) = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?B" using fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq unfolding T'(2) \_def[symmetric] by metis thus ?case using step.IH fv\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel ?B"] rm_occurs_msgs_constr_append[of \ ?B] by force qed simp private lemma rm_occurs_msgs_constr_reachable_constraints_vars_eq: assumes P: "\T \ set P. admissible_transaction' T" "\T \ set P. admissible_transaction_no_occurs_msgs T" and A: "A \ reachable_constraints (map add_occurs_msgs P)" shows "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (rm_occurs_msgs_constr A) = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" using rm_occurs_msgs_constr_bvars_subst_eq[of _ Var] rm_occurs_msgs_constr_reachable_constraints_fv_eq[OF P A] by (metis vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t subst_lsst_id_subst) private lemma rm_occurs_msgs_constr_reachable_constraints_trms_cases_aux: assumes A: "x \ fv\<^sub>s\<^sub>s\<^sub>t A" "bvars\<^sub>s\<^sub>s\<^sub>t A = {}" and t: "t = occurs (\ x)" and \: "(\y. \ x = Var y) \ (\c. \ x = Fun c [])" shows "(\x \ fv\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>s\<^sub>s\<^sub>t \). t = occurs (Var x)) \ (\c. Fun c [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>s\<^sub>s\<^sub>t \) \ t = occurs (Fun c []))" using A proof (induction A) case (Cons a A) have 0: "bvars\<^sub>s\<^sub>s\<^sub>t A = {}" "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a) = {}" "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a) \ subst_domain \ = {}" using Cons.prems(2) by auto note 1 = fv\<^sub>s\<^sub>s\<^sub>t_Cons[of a A] trms\<^sub>s\<^sub>s\<^sub>t_cons[of a A] subst_sst_cons[of a A \] show ?case proof (cases "x \ fv\<^sub>s\<^sub>s\<^sub>t A") case False hence x: "x \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" using Cons.prems(1) by simp note 2 = x t \ have 3: "\ x \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" using subst_subterms[OF fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_subterm_trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p[OF x]] trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst[OF 0(3)] by auto have "Fun c [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" when "\ x = Fun c []" for c using that 3 t by argo moreover have "y \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" when "\ x = Var y" for y using that 3 var_subterm_trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p[of y "a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \"] 0(2) unfolding vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst by simp ultimately have "(\x \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \). t = occurs (Var x)) \ (\c. Fun c [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \ t = occurs (Fun c []))" using t \ by fast thus ?thesis using 1 by auto qed (use Cons.IH[OF _ 0(1)] 1 in force) qed simp private lemma rm_occurs_msgs_constr_reachable_constraints_trms_cases: assumes P: "\T \ set P. admissible_transaction' T" "\T \ set P. admissible_transaction_no_occurs_msgs T" and A: "A = rm_occurs_msgs_constr B" and B: "B \ reachable_constraints (map add_occurs_msgs P)" and t: "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B" shows "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ (\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. t = occurs (Var x)) \ (\c. Fun c [] \\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \ t = occurs (Fun c []))" (is "?A A \ ?B A \ ?C A") proof - define rm_occs where "rm_occs \ \A::('fun,'atom,'sets,'lbl) prot_strand. rm_occurs_msgs_constr A" define Q where "Q \ \A. ?A A \ ?B A \ ?C A" have 0: "Q B" when "Q A" "set A \ set B" for A B using that unfolding Q_def fv\<^sub>s\<^sub>s\<^sub>t_def trms\<^sub>s\<^sub>s\<^sub>t_def unlabel_def by auto have "Q A" using B t unfolding A proof (induction rule: reachable_constraints.induct) case (step \ T \ \ \) define \ where "\ \ \ \\<^sub>s \ \\<^sub>s \" define \ where "\ \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" obtain T' where T': "T' \ set P" "T = add_occurs_msgs T'" using step.hyps(2) by auto note T'_adm = bspec[OF P(1) T'(1)] bspec[OF P(2) T'(1)] note T_adm = add_occurs_msgs_admissible_occurs_checks[OF T'_adm(1), unfolded T'(2)[symmetric]] note 1 = \_def[symmetric] \_def[symmetric] rm_occs_def[symmetric] note 2 = rm_occurs_msgs_constr_append[of \ \, unfolded rm_occs_def[symmetric]] note 3 = admissible_transaction_decl_fresh_renaming_subst_not_occurs[ OF T_adm(1) step.hyps(3,4,5)] have 4: "rm_occs (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using 3 rm_occurs_msgs_constr_transaction_strand''[OF T'_adm] unfolding T'(2) 1 by blast have 5: "(\y. \ x = Var y) \ (\c. \ x = Fun c [])" for x using transaction_decl_fresh_renaming_substs_range'(1)[OF step.hyps(3,4,5)] unfolding \_def[symmetric] by blast show ?case proof (cases "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \") case True show ?thesis using 0[OF step.IH[OF True]] unfolding 1 2 by simp next case False hence "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using step.prems unfolding \_def \_def by simp hence "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ (\x \ fv_transaction T'. t = occurs (\ x))" using add_occurs_msgs_in_trms_subst_cases[OF T'_adm(1), of t \] unfolding \_def trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq T'(2) by blast moreover have "(\y. \ x = Var y) \ (\c. \ x = Fun c [])" for x using transaction_decl_fresh_renaming_substs_range'(1)[OF step.hyps(3,4,5)] unfolding \_def[symmetric] by blast ultimately have "Q (rm_occs \)" using rm_occurs_msgs_constr_reachable_constraints_trms_cases_aux[ of _ "unlabel (transaction_strand T')" t \] admissible_transactionE(4)[OF T'_adm(1)] unfolding Q_def \_def 4 trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq unlabel_subst by fast thus ?thesis using 0[of "rm_occs \"] unfolding 1 2 by auto qed qed simp thus ?thesis unfolding Q_def by blast qed private lemma rm_occurs_msgs_constr_receive_attack_iff: fixes A::"('a,'b,'c,'d) prot_strand" shows "(\ts. attack\n\ \ set ts \ receive\ts\ \ set (unlabel A)) \ (\ts. attack\n\ \ set ts \ receive\ts\ \ set (unlabel (rm_occurs_msgs_constr A)))" (is "(\ts. attack\n\ \ set ts \ ?A A ts) \ (\ts. attack\n\ \ set ts \ ?B A ts)") proof let ?att = "\ts. attack\n\ \ set ts" define f where "f \ \ts::('a,'b,'c,'d) prot_term list. filter (\t. \s. t \ occurs s) ts" have 0: "?att ts \ ?att (f ts)" "?att ts \ \t. occurs t \ set ts \ \t \ set ts. \s. t \ occurs s" "\t. occurs t \ set ts \ f ts = ts" for ts::"('a,'b,'c,'d) prot_term list" unfolding f_def subgoal by simp subgoal by auto subgoal by (induct ts) auto done have "?B A (f ts)" when A: "?A A ts" and ts: "?att ts" for ts using A proof (induction A) case (Cons a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) show ?case proof (cases "?A A ts") case True thus ?thesis using Cons.IH unfolding a by (cases b) simp_all next case False hence b: "b = receive\ts\" using Cons.prems unfolding a by simp show ?thesis using 0(2)[OF ts] 0(3) unfolding a b f_def by simp qed qed simp thus "(\ts. ?att ts \ ?A A ts) \ (\ts. ?att ts \ ?B A ts)" using 0(1) by fast have "\ts'. ts = f ts' \ ?A A ts'" when B: "?B A ts" and ts: "?att ts" for ts using B proof (induction A) case (Cons a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) note 1 = rm_occurs_msgs_constr_Cons have 2: "receive\ts\ \ set (unlabel (rm_occurs_msgs_constr A))" when rcv_ts: "receive\ts\ \ set (unlabel (rm_occurs_msgs_constr ((l,send\ts'\)#A)))" for l ts ts' and A::"('a,'b,'c,'d) prot_strand" proof - have *: "is_Send (send\ts'\)" by simp have "set (unlabel (rm_occurs_msgs_constr [(l, send\ts'\)])) \ {send\ts'\, send\f ts'\}" using 1(5-7)[OF *, of l "[]"] unfolding f_def by auto thus ?thesis using rcv_ts rm_occurs_msgs_constr_append[of "[(l,send\ts'\)]" A] by auto qed show ?case proof (cases "?B A ts") case True thus ?thesis using Cons.IH by auto next case False hence 3: "receive\ts\ \ set (unlabel (rm_occurs_msgs_constr [a]))" using rm_occurs_msgs_constr_append[of "[a]" A] Cons.prems by simp obtain ts' where b: "b = receive\ts'\" and b': "is_Receive b" using 2[of ts l _ A] Cons.prems False unfolding a by (cases b) auto have ts': "the_msgs (receive\ts'\) = ts'" by simp have "\t \ set (the_msgs b). \s. t \ occurs s" when "\t. occurs t \ set (the_msgs b)" using that 3 1(4)[OF b' that, of l "[]"] unfolding a by force hence "ts = f ts'" using 3 0(3)[of ts'] 1(3)[OF b', of l "[]", unfolded rm_occurs_msgs_constr.simps(1)] unfolding a b ts' f_def[symmetric] by fastforce thus ?thesis unfolding a b by auto qed qed simp thus "(\ts. ?att ts \ ?B A ts) \ (\ts. ?att ts \ ?A A ts)" using 0 by fast qed private lemma add_occurs_msgs_soundness_aux1: fixes P::"('fun,'atom,'sets,'lbl) prot" defines "wt_attack \ \\ \ l n. welltyped_constraint_model \ (\@[(l, send\[attack\n\]\)])" assumes P: "\T \ set P. admissible_transaction' T" and P_val: "has_initial_value_producing_transaction P" and A: "\ \ reachable_constraints P" "wt_attack \ \ l n" shows "\\ \ reachable_constraints P. \\. wt_attack \ \ l n \ (\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. \n. \ x = Fun (Val n) [])" proof - let ?f = "\(T,\,\,\). dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" let ?g = "concat \ map ?f" let ?rcv_att = "\A n. receive\[attack\n\]\ \ set (unlabel A)" let ?wt_model = welltyped_constraint_model define valconst_cases where "valconst_cases \ \I::('fun,'atom,'sets,'lbl) prot_subst. \x. (\n. I x = Fun (Val n) []) \ (\n. I x = Fun (PubConst Value n) [])" define valconsts_only where "valconsts_only \ \X. \I::('fun,'atom,'sets,'lbl) prot_subst. \x \ X. \n. I x = Fun (Val n) []" define db_eq where "db_eq \ \A B::('fun,'atom,'sets,'lbl) prot_constr. \s. \upds::('fun,'atom,'sets,'lbl) prot_constr. let f = filter is_Update \ unlabel; g = filter (\a. \l t ts. a = (l,insert\t,Fun (Set s) ts\)) in (upds = [] \ f A = f B) \ (upds \ [] \ f (g A) = f (g B))" define db_upds_consts_fresh where "db_upds_consts_fresh \ \A::('fun,'atom,'sets,'lbl) prot_constr. \X. \J::('fun,'atom,'sets,'lbl) prot_subst. \x \ X. (\n. \ x = Fun (PubConst Value n) []) \ (\n s. insert\Fun (Val n) [],s\ \ set (unlabel A) \ delete\Fun (Val n) [],s\ \ set (unlabel A) \ J x \ Fun (Val n) [])" define subst_eq_on_privvals where "subst_eq_on_privvals \ \X \. \x \ X. (\n. \ x = Fun (Val n) []) \ \ x = \ x" define subst_in_ik_if_subst_pubval where "subst_in_ik_if_subst_pubval \ \X \. \\::('fun,'atom,'sets,'lbl) prot_constr. \x \ X. (\n. \ x = Fun (PubConst Value n) []) \ \ x \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" define subst_eq_iff where "subst_eq_iff \ \X. \\::('fun,'atom,'sets,'lbl) prot_subst. \x \ X. \y \ X. \ x = \ y \ \ x = \ y" obtain x_val T_val T_upds s_val ts_val l1_val l2_val where x_val: "T_val \ set P" "Var x_val \ set ts_val" "\\<^sub>v x_val = TAtom Value" "fv\<^sub>s\<^sub>e\<^sub>t (set ts_val) = {x_val}" "\n. \(Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t set ts_val)" "T_val = Transaction (\(). []) [x_val] [] [] T_upds [(l1_val,send\ts_val\)]" "T_upds = [] \ (T_upds = [(l2_val,insert\Var x_val,\s_val\\<^sub>s\)] \ (\T \ set P. \(l,a) \ set (transaction_strand T). \t. a \ select\t,\s_val\\<^sub>s\ \ a \ \t in \s_val\\<^sub>s\ \ a \ \t not in \s_val\\<^sub>s\ \ a \ delete\t,\s_val\\<^sub>s\))" using has_initial_value_producing_transactionE[OF P_val P, of thesis] by (auto simp add: disj_commute) have 0: "\n. Fun (PubConst Value n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" when "\ \ reachable_constraints P" for \ using reachable_constraints_val_funs_private(1)[OF that P(1)] funs_term_Fun_subterm' unfolding is_PubConstValue_def by fastforce have I: "?wt_model \ \" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using welltyped_constraint_model_prefix[OF A(2)[unfolded wt_attack_def]] A(2)[unfolded wt_attack_def welltyped_constraint_model_def constraint_model_def] by blast+ have 1: "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. valconst_cases \ x" using reachable_constraints_fv_Value_const_cases[OF P(1) A(1) I(1)] unfolding valconst_cases_def by blast have 2: "?rcv_att \ n" using A(2) strand_sem_append_stateful[of "{}" "{}" "unlabel \" "[send\[attack\n\]\]" \] reachable_constraints_receive_attack_if_attack'(2)[OF A(1) P(1) I(1)] unfolding wt_attack_def welltyped_constraint_model_def constraint_model_def by simp note \_empty = admissible_transaction_decl_subst_empty[OF bspec[OF P(1)]] have lmm: "\\ \ reachable_constraints P. \\. ?wt_model \ \ \ valconsts_only (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X) \ \ (?rcv_att \ n \ ?rcv_att \ n) \ subst_eq_on_privvals (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X) \ \ subst_in_ik_if_subst_pubval (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X) \ \ \ subst_eq_iff (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X) \ \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ (\n \ N. \(Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ db_eq \ \ s_val T_upds \ db_upds_consts_fresh \ (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X) \" when "finite N" "\n \ N. \(Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "X \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = {}" "finite X" "\x \ X. valconst_cases \ x" "\x \ X. \\<^sub>v x = TAtom Value" for N X using A(1) I(1) 1 that proof (induction arbitrary: N X rule: reachable_constraints.induct) case init define pubvals where "pubvals \ {n | n x. x \ X \ \ x = Fun (PubConst Value n) []}" define X_vals where "X_vals \ {n | n x. x \ X \ \ x = Fun (Val n) []}" have X_vals_finite: "finite X_vals" using finite_surj[OF init.prems(6), of X_vals "\x. THE n. \ x = Fun (Val n) []"] unfolding X_vals_def by force have pubvals_finite: "finite pubvals" using finite_surj[OF init.prems(6), of pubvals "\x. THE n. \ x = Fun (PubConst Value n) []"] unfolding pubvals_def by force obtain T_val_fresh_vals and \::"nat \ nat" where T_val_fresh_vals: "T_val_fresh_vals \ (N \ X_vals) = {}" and \: "inj \" "\ ` pubvals = T_val_fresh_vals" using ex_finite_disj_nat_inj[OF pubvals_finite finite_UnI[OF init.prems(3) X_vals_finite]] by blast have T_val_fresh_vals_finite: "finite T_val_fresh_vals" using pubvals_finite \(2) by blast obtain \::"('fun,'atom,'sets,'lbl) prot_constr" where B: "\ \ reachable_constraints P" "T_upds = [] \ list_all is_Receive (unlabel \)" "T_upds \ [] \ list_all (\a. is_Insert a \ is_Receive a) (unlabel \)" "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = {}" "\n. Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ Fun (Val n) [] \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" "T_val_fresh_vals = {n. Fun (Val n) [] \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \}" "\l a. (l,a) \ set \ \ is_Insert a \ (l = l2_val \ (\n. a = insert\Fun (Val n) [],\s_val\\<^sub>s\))" using reachable_constraints_initial_value_transaction[ OF P reachable_constraints.init T_val_fresh_vals_finite _ x_val] by auto define \ where "\ \ \x. if x \ X \ (\n. \ x = Fun (PubConst Value n) []) then Fun (Val (\ (THE n. \ x = Fun (PubConst Value n) []))) [] else \ x" have 0: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t [] \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t [] \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" "?rcv_att [] n \ ?rcv_att \ n" "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t [] = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t [] = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using B(4) vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel \"] by auto have 1: "db_eq [] \ s_val T_upds" using B(2,3,7) proof (induction \) case (Cons a B) then obtain l b where a: "a = (l,b)" by (metis surj_pair) have IH: "db_eq [] B s_val T_upds" using Cons.prems Cons.IH by auto show ?case proof (cases "T_upds = []") case True hence "is_Receive b" using a Cons.prems(1) by simp thus ?thesis using IH unfolding a db_eq_def Let_def by auto next case False hence "is_Insert b \ is_Receive b" using a Cons.prems(2) by simp hence "\t. a = (l2_val,insert\t,\s_val\\<^sub>s\)" when b: "\is_Receive b" using b Cons.prems(3) unfolding a by (metis list.set_intros(1)) thus ?thesis using IH False unfolding a db_eq_def Let_def by auto qed qed (simp add: db_eq_def) have 2: "?wt_model \ \" unfolding welltyped_constraint_model_def constraint_model_def proof (intro conjI) show "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using I(4) init.prems(8) unfolding \_def wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by fastforce show "strand_sem_stateful {} {} (unlabel \) \" using B(2,3) strand_sem_stateful_if_no_send_or_check[of "unlabel \" "{}" "{}" \] unfolding list_all_iff by blast show "subst_domain \ = UNIV" "ground (subst_range \)" using I(2) unfolding \_def subst_domain_def by auto show "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using I(3) unfolding \_def by fastforce qed have 3: "valconsts_only (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t [] \ X) \" using init.prems(7) unfolding \_def valconsts_only_def valconst_cases_def by fastforce have 4: "subst_eq_on_privvals (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t [] \ X) \" unfolding subst_eq_on_privvals_def \_def by force have 5: "subst_in_ik_if_subst_pubval (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t [] \ X) \ \" proof (unfold subst_in_ik_if_subst_pubval_def; intro ballI impI) fix x assume x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t [] \ X" and "\n. \ x = Fun (PubConst Value n) []" then obtain n where n: "\ x = Fun (PubConst Value n) []" by blast have "n \ pubvals" using x n unfolding pubvals_def by fastforce hence "\ n \ T_val_fresh_vals" using \(2) by fast hence "Fun (Val (\ n)) [] \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using B(6) by fast thus "\ x \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using x n unfolding \_def by simp qed have 6: "subst_eq_iff (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t [] \ X) \" proof (unfold subst_eq_iff_def; intro ballI) fix x y assume "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t [] \ X" "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t [] \ X" hence x: "x \ X" and y: "y \ X" by auto show "\ x = \ y \ \ x = \ y" proof show "\ x = \ y \ \ x = \ y" using x y unfolding \_def by presburger next assume J_eq: "\ x = \ y" show "\ x = \ y" proof (cases "\n. \ x = Fun (PubConst Value n) []") case True then obtain n where n: "\ x = Fun (PubConst Value n) []" by blast hence J_x: "\ x = Fun (Val (\ n)) []" using x unfolding \_def by simp show ?thesis proof (cases "\m. \ y = Fun (PubConst Value m) []") case True then obtain m where m: "\ y = Fun (PubConst Value m) []" by blast have J_y: "\ y = Fun (Val (\ m)) []" using y m unfolding \_def by simp show ?thesis using J_eq J_x J_y injD[OF \(1), of n m] n m by auto next case False then obtain m where m: "\ y = Fun (Val m) []" using init.prems(7) y unfolding valconst_cases_def by blast moreover have "\ n \ T_val_fresh_vals" using \(2) x n unfolding pubvals_def by blast moreover have "m \ X_vals" using y m unfolding X_vals_def by blast ultimately have "\ x \ \ y" using m J_x T_val_fresh_vals by auto moreover have "\ y = \ y" using m unfolding \_def by simp ultimately show ?thesis using J_eq by argo qed next case False then obtain n where n: "\ x = Fun (Val n) []" using init.prems(7) x unfolding valconst_cases_def by blast hence J_x: "\ x = \ x" unfolding \_def by auto show ?thesis proof (cases "\m. \ y = Fun (PubConst Value m) []") case False then obtain m where m: "\ y = Fun (Val m) []" using init.prems(7) y unfolding valconst_cases_def by blast have J_y: "\ y = \ y" using y m unfolding \_def by simp show ?thesis using J_x J_y J_eq by presburger next case True then obtain m where m: "\ y = Fun (PubConst Value m) []" by blast hence "\ y = Fun (Val (\ m)) []" using y unfolding \_def by fastforce moreover have "\ m \ T_val_fresh_vals" using \(2) y m unfolding pubvals_def by blast moreover have "n \ X_vals" using x n unfolding X_vals_def by blast ultimately have "\ y \ \ x" using n J_x T_val_fresh_vals by auto thus ?thesis using J_x J_eq by argo qed qed qed qed have 7: "\n \ N. Fun (Val n) [] \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using B(5,6) T_val_fresh_vals by blast have 8: "db_upds_consts_fresh [] (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t [] \ X) \" unfolding db_upds_consts_fresh_def by simp show ?case using B(1) 0 1 2 3 4 5 6 7 8 by blast next case (step \ T \ \ \ N X') define \ where "\ \ \ \\<^sub>s \ \\<^sub>s \" define T' where "T' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" define T'_pubvals where "T'_pubvals \ {n. \x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'. \ x = Fun (PubConst Value n) []}" define \_vals where "\_vals \ {n. Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \}" define \_vals where "\_vals \ {n. \x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X' \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'. \ x = Fun (Val n) []}" define \_vals where "\_vals \ {n. Fun (Val n) [] \ subst_range \}" have 3: "welltyped_constraint_model \ \" "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. valconst_cases \ x" "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'. valconst_cases \ x" "strand_sem_stateful (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \) (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ {}) (unlabel T') \" "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. valconst_cases \ x" using step.prems(2) welltyped_constraint_model_prefix[OF step.prems(1)] step.prems(1)[unfolded welltyped_constraint_model_def constraint_model_def] strand_sem_append_stateful[of "{}" "{}" "unlabel \" "unlabel T'" \] step.prems(7) unfolding \_def[symmetric] T'_def[symmetric] unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by (blast,blast,blast,simp,blast) note T_adm = bspec[OF P step.hyps(2)] note T_wf = admissible_transaction_is_wellformed_transaction[OF T_adm] note \_empty = admissible_transaction_decl_subst_empty[OF T_adm step.hyps(3)] note 4 = admissible_transaction_sem_iff note 5 = iffD1[OF 4[OF T_adm I(2,3), of "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ {}" \, unfolded T'_def[symmetric]] 3(4)] note \_dom = transaction_fresh_subst_domain[OF step.hyps(4)] have \_ran: "\n. t = Fun (Val n) []" when t: "t \ subst_range \" for t proof - obtain x where x: "x \ set (transaction_fresh T)" "t = \ x" using \_dom t by auto show ?thesis using x(1) admissible_transactionE(2)[OF T_adm] transaction_fresh_subst_sends_to_val[OF step.hyps(4) x(1)] unfolding x(2) by meson qed have T'_vals_in_\: "Fun (Val k) [] \ subst_range \" when k: "Fun (Val k) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" for k proof - have "Fun (Val k) [] \ (subterms\<^sub>s\<^sub>e\<^sub>t (trms_transaction T)) \\<^sub>s\<^sub>e\<^sub>t \" using k admissible_transactionE(4)[OF T_adm] transaction_decl_fresh_renaming_substs_trms[ OF step.hyps(3,4,5), of "transaction_strand T"] unfolding T'_def \_def[symmetric] trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq by fast then obtain t where t: "t \\<^sub>s\<^sub>e\<^sub>t trms_transaction T" "t \ \ = Fun (Val k) []" by force hence "Fun (Val k) [] \ subst_range \" using admissible_transactions_no_Value_consts(1)[OF T_adm] by (cases t) force+ thus ?thesis using transaction_decl_fresh_renaming_substs_range'(4)[OF step.hyps(3,4,5)] \_empty unfolding \_def[symmetric] by blast qed have \_vals_is_T'_vals: "k \ \_vals \ Fun (Val k) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" for k proof show "k \ \_vals" when "Fun (Val k) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" using that T'_vals_in_\ unfolding \_vals_def by blast show "Fun (Val k) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" when k: "k \ \_vals" proof - have "Fun (Val k) [] \ subst_range \" using k unfolding \_vals_def by fast then obtain x where x: "x \ fv_transaction T" "\ x = Fun (Val k) []" using admissible_transactionE(7)[OF T_adm] transaction_fresh_subst_domain[OF step.hyps(4)] unfolding fv_transaction_unfold by fastforce have "\ x = Fun (Val k) []" using x(2) unfolding \_def \_empty subst_compose_def by auto thus ?thesis using fv\<^sub>s\<^sub>s\<^sub>t_is_subterm_trms\<^sub>s\<^sub>s\<^sub>t_subst[OF x(1), of \] admissible_transactionE(4)[OF T_adm] unfolding T'_def trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq unlabel_subst by simp qed qed have \_vals_N_disj: "N \ \_vals = {}" using step.prems(4) \_vals_is_T'_vals unfolding \_def[symmetric] T'_def[symmetric] unlabel_append trms\<^sub>s\<^sub>s\<^sub>t_append by blast have T'_pubvals_finite: "finite T'_pubvals" using finite_surj[OF fv\<^sub>s\<^sub>s\<^sub>t_finite[of "unlabel T'"], of T'_pubvals "\x. THE n. \ x = Fun (PubConst Value n) []"] unfolding T'_pubvals_def by force have \_vals_finite: "finite \_vals" proof - have *: "finite (subst_range \)" using transaction_fresh_subst_domain[OF step.hyps(4)] by simp show ?thesis using finite_surj[OF *, of \_vals "\t. THE n. t = Fun (Val n) []"] unfolding \_vals_def by force qed have \_vals_finite: "finite \_vals" proof - have *: "\_vals \ (\t. THE n. t = Fun (Val n) []) ` subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" unfolding \_vals_def by force show ?thesis by (rule finite_surj[OF subterms_union_finite[OF trms\<^sub>s\<^sub>s\<^sub>t_finite] *]) qed have \_vals_finite: "finite \_vals" proof - define X where "X \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X' \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" have *: "finite X" using fv\<^sub>s\<^sub>s\<^sub>t_finite step.prems(6) unfolding X_def by blast show ?thesis using finite_surj[OF *, of \_vals "\x. THE n. \ x = Fun (Val n) []"] unfolding \_vals_def X_def[symmetric] by force qed obtain T_val_fresh_vals and \::"nat \ nat" where T_val_fresh_vals: "T_val_fresh_vals \ (N \ \_vals \ \_vals \ \_vals) = {}" and \: "inj \" "\ ` T'_pubvals = T_val_fresh_vals" using step.prems(3) T'_pubvals_finite \_vals_finite \_vals_finite \_vals_finite by (metis finite_UnI ex_finite_disj_nat_inj) define N' where "N' \ N \ \_vals \ T_val_fresh_vals" have T_val_fresh_vals_finite: "finite T_val_fresh_vals" using T'_pubvals_finite \(2) by blast have N'_finite: "finite N'" using step.prems(3) T'_pubvals_finite T_val_fresh_vals_finite \_vals_finite unfolding N'_def by auto have \_vals_trms_in: "n \ \_vals" when "Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" for n using that unfolding \_vals_def by blast have N'_notin_\: "\(Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" when n: "n \ N'" for n proof - have ?thesis when n': "n \ N" using n' step.prems(4) unfolding N'_def unlabel_append trms\<^sub>s\<^sub>s\<^sub>t_append by blast moreover have ?thesis when n': "n \ \_vals" using n' step.hyps(4) unfolding \_vals_def transaction_fresh_subst_def by blast moreover have ?thesis when n': "n \ T_val_fresh_vals" using n' T_val_fresh_vals \_vals_trms_in by blast ultimately show ?thesis using n unfolding N'_def by blast qed have T'_fv_\_disj: "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' = {}" using transaction_decl_fresh_renaming_substs_vars_disj(8)[OF step.hyps(3,4,5)] transaction_decl_fresh_renaming_substs_vars_subset(4)[OF step.hyps(3,4,5,2)] unfolding \_def[symmetric] T'_def fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq by blast have X'_disj: "X' \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = {}" "X' \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' = {}" using step.prems(5) unfolding \_def[symmetric] T'_def[symmetric] unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by (blast, blast) have X'_disj': "(X' \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T') \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = {}" using X'_disj(1) T'_fv_\_disj by blast have X'_finite: "finite (X' \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" using step.prems(6) fv\<^sub>s\<^sub>s\<^sub>t_finite by blast have \_X'_valconstcases: "\x \ X' \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'. valconst_cases \ x" using 3(3,5) by blast have T'_value_vars: "\\<^sub>v x = TAtom Value" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" for x using x reachable_constraints_fv_Value_typed[ OF P reachable_constraints.step[OF step.hyps]] unfolding \_def[symmetric] T'_def[symmetric] unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by blast have X'_T'_value_vars: "\x \ X' \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'. \\<^sub>v x = TAtom Value" using step.prems(8) T'_value_vars by blast have N'_not_subterms_\: "\n \ N'. \(Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using N'_notin_\ by blast obtain \ \ where B: "\ \ reachable_constraints P" "?wt_model \ \" "valconsts_only (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X' \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T') \" "?rcv_att \ n \ ?rcv_att \ n" "subst_eq_on_privvals (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X' \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T') \" "subst_in_ik_if_subst_pubval (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X' \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T') \ \" "subst_eq_iff (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X' \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T') \" "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" "\n \ N'. \(Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "db_eq \ \ s_val T_upds" "db_upds_consts_fresh \ (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X' \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T') \" using step.IH[OF 3(1,2) N'_finite N'_not_subterms_\ X'_disj' X'_finite \_X'_valconstcases X'_T'_value_vars] unfolding Un_assoc by fast have J: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "constr_sem_stateful \ (unlabel \)" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using B(2) unfolding welltyped_constraint_model_def constraint_model_def by blast+ have T_val_fresh_vals_notin_\: "\(Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" when "n \ T_val_fresh_vals" for n using that B(12) unfolding N'_def by blast hence "\n \ T_val_fresh_vals. \(Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by blast then obtain T_val_constr::"('fun,'atom,'sets,'lbl) prot_constr" where T_val_constr: "\@T_val_constr \ reachable_constraints P" "T_val_constr \ reachable_constraints P" "T_upds = [] \ list_all is_Receive (unlabel T_val_constr)" "T_upds \ [] \ list_all (\a. is_Insert a \ is_Receive a) (unlabel T_val_constr)" "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T_val_constr = {}" "\n. Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ Fun (Val n) [] \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t T_val_constr" "\n. Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T_val_constr \ Fun (Val n) [] \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t T_val_constr" "T_val_fresh_vals = {n. Fun (Val n) [] \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t T_val_constr}" "\l a. (l,a) \ set T_val_constr \ is_Insert a \ (l = l2_val \ (\n. a = insert\Fun (Val n) [],\s_val\\<^sub>s\))" using reachable_constraints_initial_value_transaction[ OF P B(1) T_val_fresh_vals_finite _ x_val] by blast have T_val_constr_no_upds_if_no_T_upds: "filter is_Update (unlabel T_val_constr) = []" when "T_upds = []" using T_val_constr(3)[OF that] by (induct T_val_constr) auto have T_val_fresh_vals_is_T_val_constr_vals: "k \ T_val_fresh_vals \ Fun (Val k) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T_val_constr" for k using that T_val_constr(7,8) ik\<^sub>s\<^sub>s\<^sub>t_trms\<^sub>s\<^sub>s\<^sub>t_subset by fast have T_val_constr_no_fv: "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T_val_constr = {}" using T_val_constr(5) vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t by fast have T_val_\: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T_val_constr))" proof - have "\(t \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T_val_constr))" when t: "t \ subst_range \" for t proof - obtain k where k: "t = Fun (Val k) []" using t \_ran by fast have "k \ \_vals" using t unfolding k \_vals_def by blast thus ?thesis using B(12) T_val_fresh_vals T_val_constr(7,8) unfolding N'_def k unlabel_append trms\<^sub>s\<^sub>s\<^sub>t_append by blast qed thus ?thesis using step.hyps(4) unfolding transaction_fresh_subst_def by fast qed have T_val_\: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T_val_constr))" using step.hyps B(8) T_val_constr(5) by auto define \' where "\' \ \@T_val_constr@T'" define \ where "\ \ \x. if x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' then if \n. \ x = Fun (PubConst Value n) [] then if \y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. \ y = \ x then \ (SOME y. y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X' \ \ y = \ x) else Fun (Val (\ (THE n. \ x = Fun (PubConst Value n) []))) [] else \ x else \ x" have \_ground_ran: "ground (subst_range \)" "range_vars \ = {}" and \_ran_bvars_disj: "range_vars \ \ bvars_transaction T = {}" using transaction_fresh_subst_domain[OF step.hyps(4)] transaction_fresh_subst_range_vars_empty[OF step.hyps(4)] transaction_decl_subst_range_vars_empty[OF step.hyps(3)] by (metis range_vars_alt_def, argo, blast) have \_T'_fv_disj: "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' = {}" using T'_fv_\_disj unfolding B(9) by argo hence \_\_fv_\_eq: "\ x = \ x" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'" for x using x X'_disj unfolding \_def by auto have B'1: "\' \ reachable_constraints P" using reachable_constraints.step[OF T_val_constr(1) step.hyps(2,3) T_val_\ T_val_\] unfolding \'_def T'_def \_def by simp have "\n. \ x = Fun (Val n) []" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \ X'" for x proof (cases "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'") case True note T = this show ?thesis proof (cases "\n. \ x = Fun (PubConst Value n) []") case True thus ?thesis using T B(3,9) someI_ex[of "\y. y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X' \ \ y = \ x"] unfolding \_def valconsts_only_def by (cases "\y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. \ y = \ x") (meson, auto) next case False thus ?thesis using T 3(3) unfolding \_def valconst_cases_def by fastforce qed next case False thus ?thesis using x B(3) unfolding \_def valconsts_only_def by auto qed hence B'3: "valconsts_only (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \ X') \" unfolding valconsts_only_def by blast have B'4: "?rcv_att \' n" when "?rcv_att (\@T') n" using that B(4) unfolding \'_def by auto have "\ x = \ x" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \ X'" "\ x = Fun (Val n) []" for x n proof - have "\ x = \ x" when "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" using that unfolding \_def by meson moreover have "\ x = \ x" when "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" using that x unfolding \_def by simp ultimately show ?thesis using B(5) x unfolding subst_eq_on_privvals_def unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by (cases "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'") auto qed hence B'5: "subst_eq_on_privvals (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \ X') \" unfolding subst_eq_on_privvals_def by blast have \_fv_\_eq_\: "\ x = \ x" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'" for x proof - have "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" using x T'_fv_\_disj X'_disj by blast thus ?thesis unfolding \_def by argo qed have T'_fv_\_val_\_eq_\: "\ x = \ x" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" "\n. \ x = Fun (PubConst Value n) []" for x using x B'5 3(3) unfolding unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append valconst_cases_def \_def by meson have T'_fv_\_pubval_\_eq_\_fresh_val: "\ x = Fun (Val (\ n)) []" "\ n \ T_val_fresh_vals" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" "\ x = Fun (PubConst Value n) []" "\y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. \ y \ \ x" for x n proof - show "\ x = Fun (Val (\ n)) []" using x unfolding \_def by auto show "\ n \ T_val_fresh_vals" using \(2) x unfolding T'_pubvals_def by blast qed have T'_fv_\_pubval_\_eq_\_val: "\y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. \m. \ y = \ x \ \ x = \ y \ \ x = Fun (Val m) []" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" "\ x = Fun (PubConst Value n) []" "\y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. \ y = \ x" for x n proof - have "\ x = \ (SOME y. y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X' \ \ y = \ x)" using x unfolding \_def by meson then obtain y where y: "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'" "\ y = \ x" "\ x = \ y" using x(3) someI_ex[of "\y. y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X' \ \ y = \ x"] by blast thus ?thesis using B(3,9) unfolding valconsts_only_def by auto qed have T'_fv_\_pubval_\_eq_val: "\n. \ x = Fun (Val n) []" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" "\ x = Fun (PubConst Value n) []" for x n using T'_fv_\_pubval_\_eq_\_fresh_val[OF x] T'_fv_\_pubval_\_eq_\_val[OF x] by auto have B'6': "\ x \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'" "\ x = Fun (PubConst Value n) []" for x n using x B(6) \_fv_\_eq_\ x(2) unfolding B(8) subst_in_ik_if_subst_pubval_def by simp have B'6'': "\ x \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t T_val_constr" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" "\ x = Fun (PubConst Value n) []" for x n proof (cases "\y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. \ y = \ x") case True thus ?thesis using B(6) x(2) T'_fv_\_pubval_\_eq_\_val[OF x True] unfolding B(9) subst_in_ik_if_subst_pubval_def by force next case False thus ?thesis using T_val_constr(8) T'_fv_\_pubval_\_eq_\_fresh_val[OF x] by force qed have "\ x \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \'" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \ X'" "\ x = Fun (PubConst Value n) []" for x n proof (cases "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'") case True thus ?thesis using B'6''[OF _ x(2)] unfolding \'_def by auto next case False hence "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'" using x(1) unfolding unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by blast thus ?thesis using B'6' x(2) unfolding \'_def by simp qed hence B'6: "subst_in_ik_if_subst_pubval (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \ X') \ \'" unfolding subst_in_ik_if_subst_pubval_def by blast have B'7: "subst_eq_iff (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \ X') \" proof (unfold subst_eq_iff_def; intro ballI) fix x y assume xy: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \ X'" "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \ X'" let ?Q = "\x y. \ x = \ y \ \ x = \ y" have *: "?Q x y" when xy: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'" "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'" "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" for x y using B(7) xy unfolding \_def subst_eq_iff_def by force have **: "?Q x y" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'" and y: "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" for x y proof - have xy_neq: "x \ y" using x y T'_fv_\_disj X'_disj by blast have x_eq: "\ x = \ x" using \_fv_\_eq_\ x by blast have x_eq_if_val: "\ x = \ x" when "\ x = Fun (Val n) []" for n using that x B(5) unfolding subst_eq_on_privvals_def by blast have x_neq_if_neq_val: "\ x \ \ x" when "\ x = Fun (PubConst Value n) []" for n by (metis that B(3) x UnI1 prot_fun.distinct(37) term.inject(2) valconsts_only_def) have y_eq_if_val: "\ y = \ y" when "\ y = Fun (Val n) []" for n using that y B'5 unfolding subst_eq_on_privvals_def by simp have y_eq: "\ y = Fun (Val (\ n)) []" when "\ y = Fun (PubConst Value n) []" "\z \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. \ z \ \ y" for n by (rule T'_fv_\_pubval_\_eq_\_fresh_val(1)[OF y that]) have y_eq': "\z \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. \m. \ z = \ y \ \ y = \ z \ \ y = Fun (Val m) []" when "\ y = Fun (PubConst Value n) []" "\z \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. \ z = \ y" for n by (rule T'_fv_\_pubval_\_eq_\_val[OF y that]) have K_eq_if_I_eq: "\ x = \ y \ \ x = \ y" apply (cases "\n. \ x = Fun (PubConst Value n) []") subgoal using B(7,9) unfolding subst_eq_iff_def by (metis UnI1 x x_eq y_eq') subgoal by (metis x_eq x T'_fv_\_val_\_eq_\[OF y] 3(5) valconst_cases_def x_eq_if_val) done have K_neq_if_I_neq_val: "\ x \ \ y" when n: "\ y = Fun (Val n) []" and m: "\ x = Fun (PubConst Value m) []" for n m proof - have I_neq: "\ x \ \ y" using n m by simp note y_eq'' = y_eq_if_val[OF n] note x_neq = x_neq_if_neq_val[OF m] have x_ex: "\z \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. \ z = \ x" using x unfolding B(9) by blast have J1: "\ x \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using B(6) x m unfolding subst_in_ik_if_subst_pubval_def by fast have J2: "\ x \ \ x" by (metis m B(3) x UnI1 prot_fun.distinct(37) term.inject(2) valconsts_only_def) have J3: "\ y = \ y" using B(5) n y unfolding subst_eq_on_privvals_def by blast have K_x: "\ x \ \ x" using J2 x_eq by presburger have x_notin: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" using x T'_fv_\_disj X'_disj by blast have K_x': "\ x = \ x" using x_notin unfolding \_def by argo have K_y: "\ y = \ y" using y_eq'' J3 by argo have "\ x \ \ y" using I_neq x y B(7) unfolding subst_eq_iff_def by blast thus ?thesis using K_x' K_y by argo qed show ?thesis proof show "\ x = \ y \ \ x = \ y" by (rule K_eq_if_I_eq) next assume xy_eq: "\ x = \ y" show "\ x = \ y" proof (cases "\n. \ y = Fun (PubConst Value n) []") case True then obtain n where n: "\ y = Fun (PubConst Value n) []" by blast show ?thesis proof (cases "\z \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. \ z = \ y") case True thus ?thesis using B(7,9) unfolding subst_eq_iff_def by (metis xy_eq UnI1 x x_eq y_eq'[OF n]) next case False hence F: "\z \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. \ z \ \ y" by blast note y_eq'' = y_eq[OF n F] have n_in: "\ n \ T_val_fresh_vals" using \(2) x_eq xy_eq T_val_fresh_vals_notin_\ y n unfolding T'_pubvals_def by blast hence y_notin: "\(\ y \\<^sub>s\<^sub>e\<^sub>t ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using T_val_fresh_vals_notin_\ y_eq'' ik\<^sub>s\<^sub>s\<^sub>t_trms\<^sub>s\<^sub>s\<^sub>t_subset[of "unlabel \"] by auto show ?thesis proof (cases "\m. \ x = Fun (PubConst Value m) []") case True thus ?thesis using y_notin B(6) x xy_eq x_eq unfolding B(9) subst_in_ik_if_subst_pubval_def by fastforce next case False then obtain m where m: "\ x = Fun (Val m) []" using 3(5) x unfolding valconst_cases_def by fast hence "\ x = \ x" using x B(5) unfolding subst_eq_on_privvals_def by blast hence "\ x = Fun (Val m) []" using m x_eq by argo moreover have "m \ T_val_fresh_vals" using m T_val_fresh_vals x unfolding \_vals_def by blast hence "m \ \ n" using n_in by blast ultimately have False using xy_eq y_eq'' by force thus ?thesis by simp qed qed next case False then obtain n where n: "\ y = Fun (Val n) []" using 3(3) y unfolding valconst_cases_def by fast note y_eq'' = y_eq_if_val[OF n] show ?thesis proof (cases "\m. \ x = Fun (Val m) []") case True thus ?thesis by (metis xy_eq x_eq y_eq'' x_eq_if_val) next case False then obtain m where m: "\ x = Fun (PubConst Value m) []" using 3(5) x unfolding valconst_cases_def by blast show ?thesis using K_neq_if_I_neq_val[OF n m] xy_eq by blast qed qed qed qed have ***: "?Q x y" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" and y: "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" for x y proof assume xy_eq: "\ x = \ y" show "\ x = \ y" proof (cases "\n. \ x = Fun (PubConst Value n) []") case True thus ?thesis using xy_eq x y B(7) T'_fv_\_pubval_\_eq_\_fresh_val(1) T'_fv_\_pubval_\_eq_\_val unfolding B(9) subst_eq_iff_def by (metis (no_types) UnI1) qed (metis xy_eq x y T'_fv_\_val_\_eq_\) next assume xy_eq: "\ x = \ y" have case1: False when x': "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" and y': "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" and xy_eq': "\ x = \ y" and m: "\ x = Fun (PubConst Value m) []" and n: "\ y = Fun (Val n) []" for x y n m proof - have F: "\n. \ y = Fun (PubConst Value n) []" using n by auto note x_eq = T'_fv_\_pubval_\_eq_\_fresh_val[OF x' m] note y_eq = T'_fv_\_val_\_eq_\[OF y' F] have "\z \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. \ z = \ x" proof (rule ccontr) assume no_z: "\(\z \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. \ z = \ x)" hence "n \ \ m" using n y' x_eq(2) T_val_fresh_vals unfolding \_vals_def by blast thus False using xy_eq' x_eq y_eq(1) n no_z by auto qed then obtain z k where z: "z \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'" "\ z = \ x" "\ x = \ z" "\ x = Fun (Val k) []" using T'_fv_\_pubval_\_eq_\_val[OF x' m] by blast have "\ y = \ z" using z(2,3) y_eq xy_eq' by presburger hence "\ x = \ y" using z(1,2) ** B(9) \_\_fv_\_eq y' y_eq by metis thus False using n m by simp qed have case2: "m = n" when x': "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" and y': "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" and xy_eq': "\ x = \ y" and m: "\ x = Fun (PubConst Value m) []" and n: "\ y = Fun (PubConst Value n) []" for x y n m proof (cases "\z \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. \ z \ \ x") case True show ?thesis apply (cases "\z \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. \ z \ \ y") subgoal using xy_eq' m n T'_fv_\_pubval_\_eq_\_fresh_val[OF x' m True] T'_fv_\_pubval_\_eq_\_fresh_val[OF y' n] injD[OF \(1), of m n] by fastforce subgoal by (metis x' y' xy_eq' ** B(9) True) done qed (metis x' y' xy_eq' m n ** B(9) prot_fun.inject(6) term.inject(2)) have case3: "m = n" when x': "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" and y': "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" and xy_eq': "\ x = \ y" and m: "\ x = Fun (Val m) []" and n: "\ y = Fun (Val n) []" for x y n m using x' y' xy_eq' m n T'_fv_\_val_\_eq_\ by fastforce show "\ x = \ y" using x y xy_eq case1 case2 case3 3(3) unfolding valconst_cases_def by metis qed have ****: "?Q x y" when xy: "x \ X'" "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" for x y proof - have xy': "y \ X'" "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using xy X'_disj T'_fv_\_disj by (blast, blast, blast) have K_x: "\ x = \ x" using xy(2) unfolding \_def by argo have I_iff_J: "\ x = \ y \ \ x = \ y" using xy B(7) unfolding subst_eq_iff_def by fast show ?thesis using K_x I_iff_J K_x injD[OF \(1)] xy B(7,9) ** by (meson UnCI) qed show "?Q x y" using xy * **[of x y] **[of y x] ***[of x y] ****[of x y] ****[of y x] unfolding unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by (metis Un_iff) qed have B'8_9: "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \'" "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \'" using B(8,9) T_val_constr(5) vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel T_val_constr"] unfolding \'_def unlabel_append vars\<^sub>s\<^sub>s\<^sub>t_append fv\<^sub>s\<^sub>s\<^sub>t_append by simp_all have I': "\n. \ x = Fun (PubConst Value n) []" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" "\n. \ x = Fun (Val n) []" for x using x 3(3) unfolding valconst_cases_def by fast have B'10_11: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \'" "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \'" using B(10,11) unfolding N'_def \'_def unlabel_append trms\<^sub>s\<^sub>s\<^sub>t_append ik\<^sub>s\<^sub>s\<^sub>t_append by (blast, blast) have B'12: "\n \ N. \(Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \')" using B(12) \_vals_is_T'_vals \_vals_N_disj T_val_fresh_vals T_val_fresh_vals_is_T_val_constr_vals unfolding N'_def \'_def unfolding unlabel_append trms\<^sub>s\<^sub>s\<^sub>t_append ik\<^sub>s\<^sub>s\<^sub>t_append by fast have B'13: "db_eq (\@T') \' s_val T_upds" proof - let ?f = "filter is_Update \ unlabel" let ?g = "filter (\a. \l t ts. a = (l, insert\t,Fun (Set s_val) ts\))" have "?f (?g T_val_constr) = []" using T_val_constr(3,4,9) proof (induction T_val_constr) case (Cons a B) then obtain l b where a: "a = (l,b)" by (metis surj_pair) have IH: "?f (?g B) = []" using Cons.prems Cons.IH by auto show ?case proof (cases "T_upds = []") case True hence "is_Receive b" using a Cons.prems(1,2) by simp thus ?thesis using IH unfolding a Let_def by auto next case False hence "is_Insert b \ is_Receive b" using a Cons.prems(1,2) by simp hence "\t. a = (l2_val, insert\t,\s_val\\<^sub>s\)" when b: "\is_Receive b" using b Cons.prems(3) unfolding a by (metis list.set_intros(1)) thus ?thesis using IH unfolding a Let_def by auto qed qed simp hence "?f (?g (\@T')) = ?f (?g \)@?f (?g T')" "?f (?g (\@T_val_constr@T')) = ?f (?g \)@?f (?g T')" when "T_upds \ []" by simp_all moreover have "?f T_val_constr = []" when "T_upds = []" using T_val_constr_no_upds_if_no_T_upds[OF that] by force hence "?f (\@T') = ?f \@?f T'" "?f (\@T_val_constr@T') = ?f \@?f T'" when "T_upds = []" using that by auto ultimately show ?thesis using B(13) unfolding \'_def db_eq_def Let_def by presburger qed have B'14: "db_upds_consts_fresh (\@T') (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \ X') \" proof (unfold db_upds_consts_fresh_def; intro ballI allI impI; elim exE) fix x s n m assume x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \ X'" and n: "insert\Fun (Val n) [],s\ \ set (unlabel (\@T')) \ delete\Fun (Val n) [],s\ \ set (unlabel (\@T'))" (is "?A (\@T')") and m: "\ x = Fun (PubConst Value m) []" have A_cases: "?A \ \ ?A T'" using n by force have n_in_case: "n \ \_vals" when A: "?A T'" proof - obtain t s' where t: "insert\t,s'\ \ set (unlabel (transaction_strand T)) \ delete\t,s'\ \ set (unlabel (transaction_strand T))" "Fun (Val n) [] = t \ \" using A dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(4,5) stateful_strand_step_mem_substD(4,5)[of _ _ _ \] subst_lsst_unlabel[of _ \] unfolding T'_def by (metis (no_types, opaque_lifting)) have "Fun (Val n) [] \ subst_range \" using t transaction_inserts_are_Value_vars(1)[OF T_wf(1,3), of t s'] transaction_deletes_are_Value_vars(1)[OF T_wf(1,3), of t s'] by force hence "Fun (Val n) [] \ subst_range \" using transaction_decl_fresh_renaming_substs_range'(4)[ OF step.hyps(3,4,5) _ \_empty] unfolding \_def by blast thus ?thesis unfolding \_vals_def by fast qed have in_A_case: "\ x \ Fun (Val n) []" when y: "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'" "\ x = \ y" "\ x = \ y" for y using A_cases proof assume "?A \" thus ?thesis using B(14) m y(1,3) unfolding db_upds_consts_fresh_def y(2) by auto next assume "?A T'" hence "n \ N'" using n_in_case unfolding N'_def by blast moreover have "\ y \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using B(6) y(1) m ik\<^sub>s\<^sub>s\<^sub>t_trms\<^sub>s\<^sub>s\<^sub>t_subset unfolding y(2) subst_in_ik_if_subst_pubval_def by blast ultimately show ?thesis using B(12) y(3) by fastforce qed show "\ x \ Fun (Val n) []" proof (cases "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'") case True note 0 = T'_fv_\_pubval_\_eq_\_fresh_val[OF True m, unfolded B(9)[symmetric]] note 1 = T'_fv_\_pubval_\_eq_\_val[OF True m, unfolded B(9)[symmetric]] show ?thesis proof (cases "\y\ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. \ y \ \ x") case True show ?thesis using A_cases 0[OF True] T_val_fresh_vals n_in_case unfolding \_vals_def by force next case False then obtain y where "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'" "\ y = \ x" "\ x = \ y" using 1 by blast thus ?thesis using in_A_case by auto qed next case False hence x_in: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'" using x unfolding unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by fast hence x_eq: "\ x = \ x" using \_fv_\_eq_\ by blast show ?thesis using in_A_case[OF x_in _ x_eq] by blast qed qed have B'2: "?wt_model \ \'" proof (unfold welltyped_constraint_model_def; intro conjI) have "\ (\ x) = \\<^sub>v x" for x proof - have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using B(2) 3(1) unfolding welltyped_constraint_model_def by (blast,blast) hence *: "\y. \ (\ y) = \\<^sub>v y" "\y. \ (\ y) = \\<^sub>v y" unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by auto show ?thesis proof (cases "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'") case True note x = this show ?thesis proof (cases "\n. \ x = Fun (PubConst Value n) []") case True thus ?thesis using T'_fv_\_pubval_\_eq_val[OF x] T'_value_vars[OF x] by force next case False thus ?thesis using x * unfolding \_def by presburger qed next case False thus ?thesis using *(1) unfolding \_def by presburger qed qed thus "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" unfolding \_def wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by force show "constraint_model \ \'" proof (unfold constraint_model_def; intro conjI) have *: "strand_sem_stateful {} {} (unlabel \) \" "strand_sem_stateful {} {} (unlabel \) \" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using B(2) 3(1) unfolding welltyped_constraint_model_def constraint_model_def by fast+ show K0: "subst_domain \ = UNIV" proof - have "x \ subst_domain \" for x proof (cases "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'") case True thus ?thesis using T'_fv_\_pubval_\_eq_val[OF True] T'_fv_\_val_\_eq_\[OF True] *(3) unfolding subst_domain_def by (cases "\n. \ x = Fun (PubConst Value n) []") auto next case False thus ?thesis using *(4) unfolding \_def subst_domain_def by auto qed thus ?thesis by blast qed have "fv (\ x) = {}" for x using interpretation_grounds_all[OF *(3)] interpretation_grounds_all[OF *(4)] unfolding \_def by simp thus K1: "ground (subst_range \)" by simp have "wf\<^sub>t\<^sub>r\<^sub>m (Fun (Val n) [])" for n by fastforce moreover have "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" for x using *(5,6) by (fastforce,fastforce) ultimately have "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" for x unfolding \_def by auto thus K2: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by simp show "strand_sem_stateful {} {} (unlabel \') \" proof (unfold \'_def unlabel_append strand_sem_append_stateful Un_empty_left; intro conjI) let ?sem = "\M D A. strand_sem_stateful M D (unlabel A) \" let ?M1 = "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" let ?M2 = "?M1 \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t T_val_constr \\<^sub>s\<^sub>e\<^sub>t \)" let ?D1 = "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ {}" let ?D2 = "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel T_val_constr) \ ?D1" show "?sem {} {} \" using \_\_fv_\_eq strand_sem_model_swap[OF _ *(2)] by blast show "?sem ?M1 ?D1 T_val_constr" using T_val_constr(3,4) strand_sem_stateful_if_no_send_or_check unfolding list_all_iff by blast have D2: "?D2 = ?D1 \ {(t \ \, s \ \) | t s. insert\t,s\ \ set (unlabel T_val_constr)}" using T_val_constr(3,4) dbupd\<^sub>s\<^sub>s\<^sub>t_no_deletes unfolding list_all_iff by blast have K3: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using K0 K1 by argo have rcv_\_is_\: "t \ \ = t \ \" when t: "(l,receive\ts\) \ set (transaction_receive T)" "t \ set ts" for l ts t proof - have "fv t \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" using t(2) stateful_strand_step_fv_subset_cases(2)[OF unlabel_in[OF t(1)]] by auto hence "t \ \ = t" using t \_dom \_ran admissible_transactionE(12,13)[OF T_adm] by blast thus ?thesis unfolding \_def \_empty by simp qed have eq_\_is_\: "t \ \ = t \ \" "s \ \ = s \ \" when t: "(l,\ac: t \ s\) \ set (transaction_checks T)" for l ac t s proof - have "fv t \ fv s \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" using stateful_strand_step_fv_subset_cases(3)[OF unlabel_in[OF t]] by auto hence "t \ \ = t" "s \ \ = s" using t \_dom \_ran admissible_transactionE(12,13)[OF T_adm] by (blast, blast) thus "t \ \ = t \ \" "s \ \ = s \ \" unfolding \_def \_empty by simp_all qed have noteq_\_is_\: "t \ \ = t \ \" "s \ \ = s \ \" when t: "(l,\t != s\) \ set (transaction_checks T)" for l t s proof - have "fv t \ fv s \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" using stateful_strand_step_fv_subset_cases(8)[OF unlabel_in[OF t]] by auto hence "t \ \ = t" "s \ \ = s" using t \_dom \_ran admissible_transactionE(12,13)[OF T_adm] by (blast, blast) thus "t \ \ = t \ \" "s \ \ = s \ \" unfolding \_def \_empty by simp_all qed have in_\_is_\: "t \ \ = t \ \" "s \ \ = s \ \" when t: "(l,\ac: t \ s\) \ set (transaction_checks T)" for l ac t s proof - have "fv t \ fv s \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" using stateful_strand_step_fv_subset_cases(6)[OF unlabel_in[OF t]] by auto hence "t \ \ = t" "s \ \ = s" using t \_dom \_ran admissible_transactionE(12,13)[OF T_adm] by (blast, blast) thus "t \ \ = t \ \" "s \ \ = s \ \" unfolding \_def \_empty by simp_all qed have notin_\_is_\: "t \ \ = t \ \" "s \ \ = s \ \" when t: "(l,\t not in s\) \ set (transaction_checks T)" for l t s proof - have "fv t \ fv s \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" using stateful_strand_step_fv_subset_cases(9)[OF unlabel_in[OF t]] by auto hence "t \ \ = t" "s \ \ = s" using t \_dom \_ran admissible_transactionE(12,13)[OF T_adm] by (blast, blast) thus "t \ \ = t \ \" "s \ \ = s \ \" unfolding \_def \_empty by simp_all qed have T'_trm_no_val: "\n. s = Fun (Val n) [] \ s = Fun (PubConst Value n) []" when t: "t \ trms_transaction T" "s \ t \ \" for t s proof - have ?thesis when "s \ t" using that t admissible_transactions_no_Value_consts'[OF T_adm] admissible_transactions_no_PubConsts[OF T_adm] by blast moreover have "Fun k [] \ u" when "Fun k [] \ u \ \" for k u using that proof (induction u) case (Var x) thus ?case using transaction_renaming_subst_is_renaming(2)[OF step.hyps(5), of x] by fastforce qed auto ultimately show ?thesis using t by blast qed define flt1 where "flt1 \ \A::('fun,'atom,'sets,'lbl) prot_constr. filter is_Update (unlabel A)" define flt2 where "flt2 \ \A::('fun,'atom,'sets,'lbl) prot_constr. filter (\a. \l t ts. a = (l, insert\t,\s_val\ts\\\<^sub>s\)) A" define flt3 where "flt3 \ \A::(('fun,'atom,'sets,'lbl) prot_fun, ('fun,'atom,'sets,'lbl) prot_var) stateful_strand. filter (\a. \t ts. a = insert\t,\s_val\ts\\\<^sub>s\) A" have flt2_subset: "set (unlabel (flt2 A)) \ set (unlabel A)" for A unfolding flt2_def unlabel_def by auto have flt2_unlabel: "unlabel (flt2 A) = flt3 (unlabel A)" for A unfolding flt2_def flt3_def by (induct A) auto have flt2_suffix: "suffix (filter (\a. \t ts. a = insert\t,\s_val\ts\\\<^sub>s\) A) (unlabel (flt2 B))" when "suffix A (unlabel B)" for A B using that unfolding flt2_def by (induct B arbitrary: A rule: List.rev_induct) auto have flt_AB: "flt1 (flt2 \) = flt1 (flt2 \)" proof - have *: "flt1 (flt2 \) = filter is_Update (flt3 (unlabel \))" "flt1 (flt2 \) = filter is_Update (flt3 (unlabel \))" using flt2_unlabel unfolding flt1_def by presburger+ have **: "filter is_Update (flt3 C) = flt3 (filter is_Update C)" for C proof (induction C) case Nil thus ?case unfolding flt3_def by force next case (Cons c C) thus ?case unfolding flt3_def by (cases c) auto qed show ?thesis proof (cases "T_upds = []") case True hence "filter is_Update (unlabel \) = filter is_Update (unlabel \)" using B(13) unfolding db_eq_def by fastforce thus ?thesis using ** unfolding * by presburger next case False thus ?thesis using B(13) unfolding flt1_def flt2_def db_eq_def Let_def by force qed qed have A_setops_Fun: "\t s. insert\t,s\ \ set (unlabel \) \ (\g ts. s = Fun g ts)" using reachable_constraints_setops_form[OF step.hyps(1) P] unfolding setops\<^sub>s\<^sub>s\<^sub>t_def by fastforce have A_insert_delete_not_subterm: "\ x = \ x \ (\(\ x \ t) \ \(\ x \ s) \ \(\ x \ t) \ \(\ x \ s))" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ fv t \ fv s" and x_neq: "\ x \ \ x" and ts: "insert\t,s\ \ set (unlabel \) \ delete\t,s\ \ set (unlabel \)" for x t s proof - have x_in: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" using ts x stateful_strand_step_fv_subset_cases(4,5) by blast note ts' = reachable_constraints_insert_delete_form[OF step.hyps(1) P ts] have *: "\ x = \ x" when n: "\ x = Fun (Val n) []" for n using n B'5 x_in unfolding subst_eq_on_privvals_def unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by blast have **: "\(\ x \ t)" "\(\ x \ s)" "\(\ x \ t)" "\(\ x \ s)" when n: "\ x = Fun (PubConst Value n) []" for n proof - show "\(\ x \ s)" using ts'(1) x_in 3(2,3) unfolding valconst_cases_def by fastforce show "\(\ x \ s)" using ts'(1) x_in B'3 unfolding valconsts_only_def unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by force show "\(\ x \ t)" using n ts'(3) by fastforce from ts'(3) have "\ x \ t" proof assume "\y. t = Var y" thus ?thesis using B'3 x_in unfolding valconsts_only_def by force next assume "\k. t = Fun (Val k) []" thus ?thesis using B'14 n x_in ts unfolding db_upds_consts_fresh_def by auto qed thus "\(\ x \ t)" using ts'(3) by auto qed show ?thesis using * ** 3(2,3) x_in unfolding valconst_cases_def by fast qed have flt2_insert_in_iff: "insert\u,v\ \ set (unlabel A) \ insert\u,v\ \ set (unlabel (flt2 A))" (is "?A A \ ?B A") when h: "s = \h\\<^sub>s" "h \ s_val" and t: "(t \ I,s \ I) = (u,v) \\<^sub>p I" for t s h u v A and I::"('fun,'atom,'sets,'lbl) prot_subst" proof show "?B A \ ?A A" using flt2_subset by fast show "?A A \ ?B A" proof (induction A) case (Cons a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) show ?case proof (cases "b = insert\u,v\") case True thus ?thesis using h t unfolding a flt2_def by force next case False thus ?thesis using Cons.prems Cons.IH unfolding a flt2_def by auto qed qed simp qed have flt2_inset_iff: "(t \ \, s \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (flt2 \)) \ {} \ (t \ \, s \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ {}" (is "?A \ ?B") when h: "s = \h\\<^sub>s" "h \ s_val" for t s h proof let ?C1 = "\u v B C. suffix (delete\u,v\#B) (unlabel C)" let ?C2 = "\t s u v. (t,s) = (u,v) \\<^sub>p \" let ?C3 = "\t s C. \u v. ?C2 t s u v \ insert\u,v\ \ set C" let ?D = "\t s C. \u v B. ?C1 u v B C \ ?C2 t s u v \ ?C3 t s B" let ?db = "\C D. dbupd\<^sub>s\<^sub>s\<^sub>t C \ D" have "?C3 t s B" when "?D t s (flt2 \)" "?C1 u v B \" "?C2 t s u v" for u v B t s using that flt2_suffix flt2_subset by fastforce thus "?A \ ?B" using flt2_subset unfolding dbupd\<^sub>s\<^sub>s\<^sub>t_in_iff by blast show ?A when ?B using that proof (induction \ rule: List.rev_induct) case (snoc a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) have *: "?db (unlabel (A@[a])) {} = ?db [b] (?db (unlabel A) {})" "?db (unlabel (flt2 (A@[a]))) {} = ?db (unlabel (flt2 [a])) (?db (unlabel (flt2 A)) {})" using dbupd\<^sub>s\<^sub>s\<^sub>t_append[of _ _ \ "{}"] unfolding a flt2_def by auto show ?case proof (cases "\u v. b = insert\u,v\ \ (t \ \, s \ \) = (u,v) \\<^sub>p \") case True then obtain u v where "b = insert\u,v\" "(t \ \, s \ \) = (u, v) \\<^sub>p \" by force thus ?thesis using h *(2) unfolding a flt2_def by auto next case False hence IH: "(t \ \, s \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (flt2 A)) \ {}" using snoc.prems snoc.IH unfolding *(1) by (cases b) auto show ?thesis proof (cases "is_Delete b") case True then obtain u v where b: "b = delete\u,v\" by (cases b) auto have b': "unlabel (flt2 [a]) = [b]" "unlabel (flt2 (A@[a])) = unlabel (flt2 A)@[b]" unfolding a flt2_def b by (fastforce,fastforce) have "(t \ \, s \ \) \ (u,v) \\<^sub>p \" using *(1) snoc.prems unfolding b' b by simp thus ?thesis using *(2) IH unfolding b' b by simp next case False thus ?thesis using *(2) IH unfolding a flt2_def by (cases b) auto qed qed qed simp qed have inset_model_swap: "(t \ \, s \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ {} \ (t \ \, s \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ {}" (is "?in \ (unlabel \) \ ?in \ (unlabel \)") when h: "s = \h\\<^sub>s" "h \ s_val \ filter is_Update (unlabel \) = filter is_Update (unlabel \)" and t: "t = Var tx" and t_s_fv: "fv t \ fv s \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" and q: "\x \ fv t \ fv s. \ x = \ x \ (\(\ x \ t) \ \(\ x \ s) \ \(\ x \ t) \ \(\ x \ s))" "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv t \ fv s. \c. \ x = Fun c []" "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv t \ fv s. \c. \ x = Fun c []" "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv t \ fv s. \y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv t \ fv s. \ x = \ y \ \ x = \ y" for t s h tx proof - let ?upds = "\A. filter is_Update (unlabel A)" have flt2_fv: "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (flt2 \) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using fv\<^sub>s\<^sub>s\<^sub>t_mono[OF flt2_subset[of \]] by blast have upds_fv: "fv\<^sub>s\<^sub>s\<^sub>t (?upds \) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" by auto have flt2_upds_fv: "fv\<^sub>s\<^sub>s\<^sub>t (?upds (flt2 \)) \ fv\<^sub>s\<^sub>s\<^sub>t (?upds \)" using flt2_subset[of \] by auto have h_neq: "Set h \ (Set s_val::('fun,'atom,'sets,'lbl) prot_fun)" when "h \ s_val" using that by simp have *: "\(fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r ` {}) = {}" "{} \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ = {}" "{} \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ = {}" by blast+ have "?in \ (?upds (flt2 \)) \ ?in \ (?upds (flt2 \))" proof let ?X = "fv\<^sub>s\<^sub>s\<^sub>t (?upds (flt2 \)) \ fv t \ fv s \ \(fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r ` ({}::(('fun,'atom,'sets,'lbl) prot_term \ ('fun,'atom,'sets,'lbl) prot_term) set))" let ?q0 = "\\ \. \x \ ?X. \ x = \ x \ (\(\ x \ t) \ \(\ x \ s) \\(\ x \ t) \ \(\ x \ s) \ (\(u,v) \ {}. \(\ x \ u) \ \(\ x \ v) \ \(\ x \ u) \ \(\ x \ v)) \ (\u v. insert\u,v\ \ set (?upds (flt2 \)) \ delete\u,v\ \ set (?upds (flt2 \)) \ \(\ x \ u) \ \(\ x \ v) \ \(\ x \ u) \ \(\ x \ v)))" let ?q1 = "\\. \x \ ?X. \c. \ x = Fun c []" let ?q2 = "\\ \. \x \ ?X. \y \ ?X. \ x = \ y \ \ x = \ y" have q0: "?q0 \ \" "?q0 \ \" proof - have upd_ex: "\u v. x \ fv u \ fv v \ (insert\u,v\ \ set (?upds A) \ delete\u,v\ \ set (?upds A))" when "x \ fv\<^sub>s\<^sub>s\<^sub>t (?upds A)" for x and A::"('fun,'atom,'sets,'lbl) prot_constr" using that proof (induction A) case (Cons a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) show ?case using Cons.IH Cons.prems unfolding a by (cases b) auto qed simp have "\(\ x \ t)" "\(\ x \ s)" "\(\ x \ t)" "\(\ x \ s)" when x: "x \ fv\<^sub>s\<^sub>s\<^sub>t (?upds (flt2 \)) \ fv t \ fv s" and x_neq: "\ x \ \ x" for x proof - have "\(\ x \ t) \ \(\ x \ s) \ \(\ x \ t) \ \(\ x \ s)" proof (cases "x \ fv t \ fv s") case True thus ?thesis using q(1) x_neq by blast next case False hence "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using x flt2_upds_fv upds_fv by blast hence "\n. \ x = Fun (Val n) []" "\n. \ x = Fun (Val n) [] \ \ x = Fun (PubConst Value n) []" using B'3 3(2) unfolding valconst_cases_def valconsts_only_def unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by (blast, blast) thus ?thesis unfolding t h(1) by auto qed thus "\(\ x \ t)" "\(\ x \ s)" "\(\ x \ t)" "\(\ x \ s)" by simp_all qed moreover have "\(\ x \ u)" "\(\ x \ v)" "\(\ x \ u)" "\(\ x \ v)" when x: "x \ fv\<^sub>s\<^sub>s\<^sub>t (?upds (flt2 \)) \ fv t \ fv s" and x_neq: "\ x \ \ x" and uv: "insert\u,v\ \ set (?upds (flt2 \)) \ delete\u,v\ \ set (?upds (flt2 \))" for x u v proof - have uv': "insert\u,v\ \ set (unlabel \) \ delete\u,v\ \ set (unlabel \)" using uv flt2_subset by auto have x_in: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ fv u \ fv v" using t_s_fv x flt2_upds_fv upds_fv by blast show "\(\ x \ u)" "\(\ x \ v)" "\(\ x \ u)" "\(\ x \ v)" using x_neq A_insert_delete_not_subterm[OF x_in x_neq uv'] by simp_all qed ultimately show "?q0 \ \" unfolding upd_ex unfolding * by (metis (no_types, lifting) empty_iff sup_bot_right) thus "?q0 \ \" by (metis (lifting) empty_iff) qed have q1: "?q1 \" "?q1 \" using q(2,3) flt2_upds_fv upds_fv by (blast,blast) have q2: "?q2 \ \" "?q2 \ \" using q(4) flt2_upds_fv upds_fv unfolding * by (blast,blast) show "?in \ (?upds (flt2 \)) \ ?in \ (?upds (flt2 \))" using dbupd\<^sub>s\<^sub>s\<^sub>t_subst_const_swap[OF _ q0(1) q1(1,2) q2(1)] by force show "?in \ (?upds (flt2 \)) \ ?in \ (?upds (flt2 \))" using dbupd\<^sub>s\<^sub>s\<^sub>t_subst_const_swap[OF _ q0(2) q1(2,1) q2(2)] by force qed hence flt2_subst_swap: "?in \ (unlabel (flt2 \)) \ ?in \ (unlabel (flt2 \))" using dbupd\<^sub>s\<^sub>s\<^sub>t_filter by blast (* TODO: merge with similar proof above? *) have "?in \ (?upds \) \ ?in \ (?upds \)" proof let ?X = "fv\<^sub>s\<^sub>s\<^sub>t (?upds \) \ fv t \ fv s \ \(fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r ` ({}::(('fun,'atom,'sets,'lbl) prot_term \ ('fun,'atom,'sets,'lbl) prot_term) set))" let ?q0 = "\\ \. \x \ ?X. \ x = \ x \ (\(\ x \ t) \ \(\ x \ s) \\(\ x \ t) \ \(\ x \ s) \ (\(u,v) \ {}. \(\ x \ u) \ \(\ x \ v) \ \(\ x \ u) \ \(\ x \ v)) \ (\u v. insert\u,v\ \ set (?upds \) \ delete\u,v\ \ set (?upds \) \ \(\ x \ u) \ \(\ x \ v) \ \(\ x \ u) \ \(\ x \ v)))" let ?q1 = "\\. \x \ ?X. \c. \ x = Fun c []" let ?q2 = "\\ \. \x \ ?X. \y \ ?X. \ x = \ y \ \ x = \ y" have q0: "?q0 \ \" "?q0 \ \" proof - have upd_ex: "\u v. x \ fv u \ fv v \ (insert\u,v\ \ set (?upds A) \ delete\u,v\ \ set (?upds A))" when "x \ fv\<^sub>s\<^sub>s\<^sub>t (?upds A)" for x and A::"('fun,'atom,'sets,'lbl) prot_constr" using that proof (induction A) case (Cons a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) show ?case using Cons.IH Cons.prems unfolding a by (cases b) auto qed simp have "\(\ x \ t)" "\(\ x \ s)" "\(\ x \ t)" "\(\ x \ s)" when x: "x \ fv\<^sub>s\<^sub>s\<^sub>t (?upds \) \ fv t \ fv s" and x_neq: "\ x \ \ x" for x proof - have "\(\ x \ t) \ \(\ x \ s) \ \(\ x \ t) \ \(\ x \ s)" proof (cases "x \ fv t \ fv s") case True thus ?thesis using q(1) x_neq by blast next case False hence "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using x flt2_upds_fv upds_fv by blast hence "\n. \ x = Fun (Val n) []" "\n. \ x = Fun (Val n) [] \ \ x = Fun (PubConst Value n) []" using B'3 3(2) unfolding valconst_cases_def valconsts_only_def unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by (blast, blast) thus ?thesis unfolding t h(1) by auto qed thus "\(\ x \ t)" "\(\ x \ s)" "\(\ x \ t)" "\(\ x \ s)" by simp_all qed moreover have "\(\ x \ u)" "\(\ x \ v)" "\(\ x \ u)" "\(\ x \ v)" when x: "x \ fv\<^sub>s\<^sub>s\<^sub>t (?upds \) \ fv t \ fv s" and x_neq: "\ x \ \ x" and uv: "insert\u,v\ \ set (?upds \) \ delete\u,v\ \ set (?upds \)" for x u v proof - have uv': "insert\u,v\ \ set (unlabel \) \ delete\u,v\ \ set (unlabel \)" using uv flt2_subset by auto have x_in: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ fv u \ fv v" using t_s_fv x flt2_upds_fv upds_fv by blast show "\(\ x \ u)" "\(\ x \ v)" "\(\ x \ u)" "\(\ x \ v)" using x_neq A_insert_delete_not_subterm[OF x_in x_neq uv'] by simp_all qed ultimately show "?q0 \ \" unfolding upd_ex unfolding * by (metis (no_types, lifting) empty_iff sup_bot_right) thus "?q0 \ \" by (metis (lifting) empty_iff) qed have q1: "?q1 \" "?q1 \" using q(2,3) flt2_upds_fv upds_fv by (blast,blast) have q2: "?q2 \ \" "?q2 \ \" using q(4) flt2_upds_fv upds_fv unfolding * by (blast,blast) show "?in \ (?upds \) \ ?in \ (?upds \)" using dbupd\<^sub>s\<^sub>s\<^sub>t_subst_const_swap[OF _ q0(1) q1(1,2) q2(1)] by force show "?in \ (?upds \) \ ?in \ (?upds \)" using dbupd\<^sub>s\<^sub>s\<^sub>t_subst_const_swap[OF _ q0(2) q1(2,1) q2(2)] by force qed hence db_subst_swap: "?in \ (unlabel \) \ ?in \ (unlabel \)" using dbupd\<^sub>s\<^sub>s\<^sub>t_filter by blast have "?in \ (unlabel \)" when A: "?in \ (unlabel \)" using h(2) proof assume h': "h \ s_val" have "?in \ (unlabel (flt2 \))" using A flt2_unlabel dbupd\<^sub>s\<^sub>s\<^sub>t_set_term_neq_in_iff[OF h_neq[OF h'] A_setops_Fun] unfolding h(1) flt3_def by simp hence "?in \ (unlabel (flt2 \))" using flt2_subst_swap by blast hence "?in \ (flt1 (flt2 \))" using dbupd\<^sub>s\<^sub>s\<^sub>t_filter unfolding flt1_def by blast hence "?in \ (flt1 (flt2 \))" using flt_AB by simp hence "?in \ (unlabel (flt2 \))" using dbupd\<^sub>s\<^sub>s\<^sub>t_filter unfolding flt1_def by blast thus ?thesis using flt2_inset_iff[OF h(1) h'] by fast next assume h': "filter is_Update (unlabel \) = filter is_Update (unlabel \)" have "?in \ (unlabel \)" using A db_subst_swap by blast hence "?in \ (flt1 \)" using dbupd\<^sub>s\<^sub>s\<^sub>t_filter unfolding flt1_def by blast hence "?in \ (flt1 \)" using h' unfolding flt1_def by simp thus ?thesis using dbupd\<^sub>s\<^sub>s\<^sub>t_filter unfolding flt1_def by blast qed moreover have "\?in \ (unlabel \)" when A: "\?in \ (unlabel \)" using h(2) proof assume h': "h \ s_val" have "\?in \ (unlabel (flt2 \))" using A flt2_unlabel dbupd\<^sub>s\<^sub>s\<^sub>t_set_term_neq_in_iff[OF h_neq[OF h'] A_setops_Fun] unfolding h(1) flt3_def by simp hence "\?in \ (unlabel (flt2 \))" using flt2_subst_swap by blast hence "\?in \ (flt1 (flt2 \))" using dbupd\<^sub>s\<^sub>s\<^sub>t_filter unfolding flt1_def by blast hence "\?in \ (flt1 (flt2 \))" using flt_AB by simp hence "\?in \ (unlabel (flt2 \))" using dbupd\<^sub>s\<^sub>s\<^sub>t_filter unfolding flt1_def by blast thus ?thesis using flt2_inset_iff[OF h(1) h'] by fast next assume h': "filter is_Update (unlabel \) = filter is_Update (unlabel \)" have "\?in \ (unlabel \)" using A db_subst_swap by blast hence "\?in \ (flt1 \)" using dbupd\<^sub>s\<^sub>s\<^sub>t_filter unfolding flt1_def by blast hence "\?in \ (flt1 \)" using h' unfolding flt1_def by simp thus ?thesis using dbupd\<^sub>s\<^sub>s\<^sub>t_filter unfolding flt1_def by blast qed ultimately show ?thesis by blast qed have "?M2 \ t \ \ \ \" when ts: "(l, receive\ts\) \ set (transaction_receive T)" "t \ set ts" for l t ts proof - have *: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \ \ \" using 5 ts by blast note t\\ = rcv_\_is_\[OF ts] have t_T'_trm: "t \ trms_transaction T" using trms\<^sub>s\<^sub>s\<^sub>t_memI(2)[OF unlabel_in[OF ts(1)] ts(2)] unfolding trms_transaction_unfold by blast have t_T'_trm': "t \ \ \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" using trms\<^sub>s\<^sub>s\<^sub>t_memI(2)[ OF stateful_strand_step_subst_inI(2)[ OF unlabel_in[OF ts(1)], unfolded unlabel_subst]] ts(2) unfolding T'_def trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq trms_transaction_subst_unfold by auto note t_no_val = T'_trm_no_val[OF t_T'_trm, unfolded t\\[symmetric]] have t_fv_T': "fv (t \ \) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" using ts(2) stateful_strand_step_fv_subset_cases(2)[ OF stateful_strand_step_subst_inI(2)[OF unlabel_in[OF ts(1)], of \]] unfolding T'_def unlabel_subst fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq fv_transaction_subst_unfold by auto have ik_B_fv_subset: "fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" by (meson UnE fv_ik\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t subset_iff) let ?fresh_vals = "(\n. Fun (Val n) []) ` T_val_fresh_vals" have q0: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \ \ \" using * B(10) by (blast intro: ideduct_mono) have q1: "\x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ fv (t \ \). valconst_cases \ x" using 3(2,3) t_fv_T' ik_B_fv_subset unfolding B(9) by blast have q2: "\x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ fv (t \ \). \n. \ x = Fun (Val n) []" using B'3 t_fv_T' ik_B_fv_subset unfolding valconsts_only_def unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append B(9) by blast have T_val_constr_ik: "\M. ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t T_val_constr = M \ ?fresh_vals" "\M. ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t T_val_constr \\<^sub>s\<^sub>e\<^sub>t \ = (M \\<^sub>s\<^sub>e\<^sub>t \) \ ?fresh_vals" proof - obtain M where M: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t T_val_constr = M \ ?fresh_vals" using T_val_constr(8) by blast have "?fresh_vals \\<^sub>s\<^sub>e\<^sub>t \ = ?fresh_vals" by fastforce thus "\M. ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t T_val_constr = M \ ?fresh_vals" "\M. ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t T_val_constr \\<^sub>s\<^sub>e\<^sub>t \ = (M \\<^sub>s\<^sub>e\<^sub>t \) \ ?fresh_vals" using M by (fastforce, fastforce) qed have "\ x \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t T_val_constr" when x: "x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ fv (t \ \)" "\ x = Fun (PubConst Value n) []" for x n using x(1) B'6'[OF _ x(2)] B'6''[OF _ x(2)] t_fv_T' ik_B_fv_subset unfolding B(9) unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by blast hence q3: "\x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ fv (t \ \). (\n. \ x = Fun (PubConst Value n) []) \ \ x \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ ?fresh_vals" using T_val_constr_ik(1) T_val_constr(8) q2 unfolding B(9) \'_def unlabel_append ik\<^sub>s\<^sub>s\<^sub>t_append fv\<^sub>s\<^sub>s\<^sub>t_append by (metis (no_types, lifting) UnE UnI1 UnI2 image_iff mem_Collect_eq) have q4: "\x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ fv (t \ \). (\n. \ x = Fun (Val n) []) \ \ x = \ x" using B'5 t_fv_T' ik_B_fv_subset unfolding subst_eq_on_privvals_def B(9) unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by blast have q5: "\x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ fv (t \ \). \y \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ fv (t \ \). \ x = \ y \ \ x = \ y" using B'7 t_fv_T' ik_B_fv_subset unfolding subst_eq_iff_def B(9) unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by blast have q6: "\n. \(Fun (PubConst Value n) [] \\<^sub>s\<^sub>e\<^sub>t insert (t \ \) (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" proof - have "\n. s = Fun (PubConst Value n) []" when s: "s \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \'" for s proof - have "f \ PubConst Value n" when f: "f \ funs_term s" for f n using f s reachable_constraints_val_funs_private(1)[OF B'1 P, of f] unfolding is_PubConstValue_def is_PubConst_def the_PubConst_type_def by (metis (mono_tags, lifting) UN_I funs_term_subterms_eq(2) prot_fun.simps(85)) thus ?thesis by fastforce qed moreover have "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \'" using ik\<^sub>s\<^sub>s\<^sub>t_trms\<^sub>s\<^sub>s\<^sub>t_subset unfolding \'_def unlabel_append trms\<^sub>s\<^sub>s\<^sub>t_append by blast ultimately show ?thesis using t_no_val by blast qed show ?thesis using deduct_val_const_swap[OF q0 q1[unfolded valconst_cases_def] q2 q3 q4 q5 q6] T_val_constr_ik(2) by (blast intro: ideduct_mono) qed moreover have "t \ \ \ \ = s \ \ \ \" when ts: "(l, \ac: t \ s\) \ set (transaction_checks T)" for l ac t s proof - have q0: "t \ \ \ \ = s \ \ \ \" using 5 ts by blast have "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\ac: (t \ \) \ (s \ \)\) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using stateful_strand_step_fv_subset_cases(3)[ OF stateful_strand_step_subst_inI(3)[OF unlabel_in[OF ts], of \]] unfolding unlabel_subst by simp hence t_s_fv: "fv (t \ \) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" "fv (s \ \) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" unfolding T'_def fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq fv_transaction_subst_unfold[of T \] by (fastforce, fastforce) have "t \ trms_transaction T" "s \ trms_transaction T" using trms\<^sub>s\<^sub>s\<^sub>t_memI(3,4)[OF unlabel_in[OF ts]] unfolding trms_transaction_unfold by (blast, blast) hence "\n. u = Fun (Val n) [] \ u = Fun (PubConst Value n) []" when u: "u \ t \ \ \ u \ s \ \" for u using u T'_trm_no_val unfolding eq_\_is_\[OF ts] by blast hence "\(\ x \ t \ \)" "\(\ x \ s \ \)" when x: "x \ fv (t \ \) \ fv (s \ \)" for x using x t_s_fv I' by (fast, fast) hence q1: "\x \ fv (t \ \) \ fv (s \ \). \ x = \ x \ (\(\ x \ t \ \) \ \(\ x \ s \ \))" by blast have q2: "\x \ fv (t \ \) \ fv (s \ \). \c. \ x = Fun c []" using t_s_fv 3(3) unfolding valconst_cases_def by blast have q3: "\x \ fv (t \ \) \ fv (s \ \). \c. \ x = Fun c []" using t_s_fv B'3 unfolding valconsts_only_def unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by blast have q4: "\x \ fv (t \ \) \ fv (s \ \). \y \ fv (t \ \) \ fv (s \ \). \ x = \ y \ \ x = \ y" using B'7 t_s_fv unfolding subst_eq_iff_def B(9) unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by blast show ?thesis by (rule subst_const_swap_eq'[OF q0 q1 q2 q3 q4]) qed moreover have "t \ \ \ \ \ s \ \ \ \" when ts: "(l, \t != s\) \ set (transaction_checks T)" for l t s proof - have q0: "t \ \ \ \ \ s \ \ \ \" using 5 ts by blast have "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\(t \ \) != (s \ \)\) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using stateful_strand_step_fv_subset_cases(8)[ OF stateful_strand_step_subst_inI(8)[OF unlabel_in[OF ts], of \]] unfolding unlabel_subst by simp hence t_s_fv: "fv (t \ \) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" "fv (s \ \) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" unfolding T'_def fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq fv_transaction_subst_unfold[of T \] by (fastforce, fastforce) have "t \ trms_transaction T" "s \ trms_transaction T" using trms\<^sub>s\<^sub>s\<^sub>t_memI(9)[OF unlabel_in[OF ts]] unfolding trms_transaction_unfold by auto hence "\n. u = Fun (Val n) []" when u: "u \ t \ \ \ u \ s \ \" for u using u T'_trm_no_val unfolding noteq_\_is_\[OF ts] by blast hence "\(\ x \ t \ \)" "\(\ x \ s \ \)" when x: "x \ fv (t \ \) \ fv (s \ \)" for x using x t_s_fv B'3 unfolding valconsts_only_def unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by (fast, fast) hence q1: "\x \ fv (t \ \) \ fv (s \ \). \ x = \ x \ (\(\ x \ t \ \) \ \(\ x \ s \ \))" by blast have q2: "\x \ fv (t \ \) \ fv (s \ \). \c. \ x = Fun c []" using t_s_fv B'3 unfolding valconsts_only_def unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by blast have q3: "\x \ fv (t \ \) \ fv (s \ \). \c. \ x = Fun c []" using t_s_fv 3(3) unfolding valconst_cases_def by blast have q4: "\x \ fv (t \ \) \ fv (s \ \). \y \ fv (t \ \) \ fv (s \ \). \ x = \ y \ \ x = \ y" using B'7 t_s_fv unfolding subst_eq_iff_def B(9) unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by blast show ?thesis using q0 subst_const_swap_eq'[OF _ q1 q2 q3 q4] by fast qed moreover have "(t \ \ \ \, s \ \ \ \) \ ?D2" when ts: "(l, \ac: t \ s\) \ set (transaction_checks T)" for l ac t s proof - have s_neq_s_val: "s \ \s_val\\<^sub>s \ filter is_Update (unlabel \) = filter is_Update (unlabel \)" proof (cases "T_upds = []") case False thus ?thesis using step.hyps(2) ts x_val(7) unfolding transaction_strand_def by (cases ac) fastforce+ qed (use B(13)[unfolded db_eq_def] in simp) have ts': "\ac: t \ s\ \ set (unlabel (transaction_strand T))" using ts unlabel_in[OF ts] unfolding transaction_strand_def by fastforce have "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\ac: (t \ \) \ (s \ \)\) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using stateful_strand_step_fv_subset_cases(6)[ OF stateful_strand_step_subst_inI(6)[OF unlabel_in[OF ts], of \]] unfolding unlabel_subst by simp hence t_s_fv: "fv (t \ \) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" "fv (s \ \) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" unfolding T'_def fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq fv_transaction_subst_unfold[of T \] by (fastforce, fastforce) have "t \ trms_transaction T" "s \ trms_transaction T" using ts' unfolding trms\<^sub>s\<^sub>s\<^sub>t_def by (force, force) hence "\n. u = Fun (Val n) [] \ u = Fun (PubConst Value n) []" when u: "u \ t \ \ \ u \ s \ \" for u using u T'_trm_no_val unfolding in_\_is_\[OF ts] by blast hence "\(\ x \ t \ \)" "\(\ x \ s \ \)" "\(\ x \ t \ \)" "\(\ x \ s \ \)" when x: "x \ fv (t \ \) \ fv (s \ \)" for x using x t_s_fv B'3 I' unfolding valconsts_only_def unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by (fast,fast,fast,fast) hence q1: "\x \ fv (t \ \) \ fv (s \ \). \ x = \ x \ (\(\ x \ t \ \) \ \(\ x \ s \ \) \ \(\ x \ t \ \) \ \(\ x \ s \ \))" by blast have q2: "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv (t \ \) \ fv (s \ \). \c. \ x = Fun c []" using t_s_fv 3(2,3) unfolding valconst_cases_def by blast have q3: "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv (t \ \) \ fv (s \ \). \c. \ x = Fun c []" using t_s_fv B'3 unfolding valconsts_only_def unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by blast have q4: "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv (t \ \) \ fv (s \ \). \y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv (t \ \) \ fv (s \ \). \ x = \ y \ \ x = \ y" using B'7 t_s_fv unfolding subst_eq_iff_def B(9) unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by blast obtain h tx where s: "s = \h\\<^sub>s" and tx: "t = Var tx" using ts' transaction_selects_are_Value_vars[OF T_wf(1,2), of t s] transaction_inset_checks_are_Value_vars[OF T_adm, of t s] by (cases ac) auto have h: "s \ \ = \h\\<^sub>s" "h \ s_val \ filter is_Update (unlabel \) = filter is_Update (unlabel \)" using s s_neq_s_val by (simp,blast) obtain ty where ty: "t \ \ = Var ty" using tx transaction_renaming_subst_is_renaming(2)[OF step.hyps(5), of tx] unfolding in_\_is_\[OF ts] by force have "(t \ \ \ \, s \ \ \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ {}" using 5 ts by blast hence "(t \ \ \ \, s \ \ \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ {}" using inset_model_swap[OF h ty _ q1 q2 q3 q4] t_s_fv by simp thus ?thesis unfolding D2 by blast qed moreover have "(t \ \ \ \, s \ \ \ \) \ ?D2" when ts: "(l, \t not in s\) \ set (transaction_checks T)" for l t s proof - have s_neq_s_val: "(T_upds \ [] \ s \ \s_val\\<^sub>s) \ (T_upds = [] \ filter is_Update (unlabel \) = filter is_Update (unlabel \))" proof (cases "T_upds = []") case False thus ?thesis using step.hyps(2) ts x_val(7) unfolding transaction_strand_def by force qed (use B(13)[unfolded db_eq_def] in simp) have ts': "\t not in s\ \ set (unlabel (transaction_strand T))" using ts unlabel_in[OF ts] unfolding transaction_strand_def by fastforce have "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\(t \ \) not in (s \ \)\) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using stateful_strand_step_fv_subset_cases(9)[ OF stateful_strand_step_subst_inI(9)[OF unlabel_in[OF ts], of \]] unfolding unlabel_subst by simp hence t_s_fv: "fv (t \ \) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" "fv (s \ \) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" unfolding T'_def fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq fv_transaction_subst_unfold[of T \] by (fastforce, fastforce) have "t \ trms_transaction T" "s \ trms_transaction T" using ts' unfolding trms\<^sub>s\<^sub>s\<^sub>t_def by (force, force) hence "\n. u = Fun (Val n) [] \ u = Fun (PubConst Value n) []" when u: "u \ t \ \ \ u \ s \ \" for u using u T'_trm_no_val unfolding notin_\_is_\[OF ts] by blast hence "\(\ x \ t \ \)" "\(\ x \ s \ \)" "\(\ x \ t \ \)" "\(\ x \ s \ \)" when x: "x \ fv (t \ \) \ fv (s \ \)" for x using x t_s_fv B'3 I' unfolding valconsts_only_def unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by (fast,fast,fast,fast) hence q1: "\x \ fv (t \ \) \ fv (s \ \). \ x = \ x \ (\(\ x \ t \ \) \ \(\ x \ s \ \) \ \(\ x \ t \ \) \ \(\ x \ s \ \))" by blast have q2: "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv (t \ \) \ fv (s \ \). \c. \ x = Fun c []" using t_s_fv 3(2,3) unfolding valconst_cases_def by blast have q3: "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv (t \ \) \ fv (s \ \). \c. \ x = Fun c []" using t_s_fv B'3 unfolding valconsts_only_def unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by blast have q4: "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv (t \ \) \ fv (s \ \). \y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv (t \ \) \ fv (s \ \). \ x = \ y \ \ x = \ y" using B'7 t_s_fv unfolding subst_eq_iff_def B(9) unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by blast obtain h tx where s: "s = \h\\<^sub>s" and tx: "t = Var tx" using transaction_notinset_checks_are_Value_vars(1,2)[OF T_adm ts', of t s] by auto have h: "s \ \ = \h\\<^sub>s" "h \ s_val \ filter is_Update (unlabel \) = filter is_Update (unlabel \)" "T_upds \ [] \ h \ s_val" using s s_neq_s_val by (simp,blast,blast) obtain ty where ty: "t \ \ = Var ty" using tx transaction_renaming_subst_is_renaming(2)[OF step.hyps(5), of tx] unfolding notin_\_is_\[OF ts] by force have *: "(t \ \ \ \, s \ \ \ \) \ (u \ \, v \ \)" when u: "insert\u,v\ \ set (unlabel T_val_constr)" and h': "h \ s_val" for u v proof - have "v = \s_val\\<^sub>s" using T_val_constr(9) unlabel_mem_has_label[OF u] by force thus ?thesis using h(1) h' by simp qed have "(t \ \ \ \, s \ \ \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ {}" using 5 ts by blast hence **: "(t \ \ \ \, s \ \ \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ {}" using inset_model_swap[OF h(1,2) ty _ q1 q2 q3 q4] t_s_fv by simp show ?thesis proof (cases "T_upds = []") case True have "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel T_val_constr) I D = D" for I D using T_val_constr_no_upds_if_no_T_upds[OF True] dbupd\<^sub>s\<^sub>s\<^sub>t_filter[of "unlabel T_val_constr"] by force thus ?thesis using ** by simp next case False thus ?thesis using ** * h(3) T_val_constr_no_upds_if_no_T_upds unfolding D2 by blast qed qed ultimately show "?sem ?M2 ?D2 T'" unfolding T'_def 4[OF T_adm K3 K2] by blast qed qed qed show ?case using B'1 B'2 B'3 B'4 B'5 B'6 B'7 B'8_9 B'10_11 B'12 B'13 B'14 unfolding \_def[symmetric] T'_def[symmetric] by blast qed obtain \ \ where B: "\ \ reachable_constraints P" "?wt_model \ \" "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. \n. \ x = Fun (Val n) []" "?rcv_att \ n \ ?rcv_att \ n" "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using lmm[OF finite.emptyI _ _ finite.emptyI] unfolding valconsts_only_def by auto show ?thesis using B(1,3) welltyped_constraint_model_attack_if_receive_attack[OF B(2)] B(4) 2 unfolding wt_attack_def B(5) by (meson list.set_intros(1)) qed private lemma add_occurs_msgs_soundness_aux2: assumes P: "\T \ set P. admissible_transaction T" and A: "\ \ reachable_constraints P" shows "\\ \ reachable_constraints (map add_occurs_msgs P). \ = rm_occurs_msgs_constr \" using A proof (induction rule: reachable_constraints.induct) case (step \ T \ \ \) define \ where "\ \ \ \\<^sub>s \ \\<^sub>s \" let ?A' = "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" let ?B' = "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand (add_occurs_msgs T) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" obtain A B C D E F where T: "T = Transaction A B C D E F" by (cases T) simp have P': "\T \ set P. admissible_transaction' T" "\T \ set P. admissible_transaction_no_occurs_msgs T" using P admissible_transactionE'(1,2) by (blast,blast) note T_adm' = bspec[OF P step.hyps(2)] note T_adm = bspec[OF P'(1) step.hyps(2)] note \_empty = admissible_transaction_decl_subst_empty[OF T_adm step.hyps(3)] note T_fresh_val = admissible_transactionE(2)[OF T_adm] note T_no_occ = admissible_transactionE'(2)[OF T_adm'] obtain \ where B: "\ \ reachable_constraints (map add_occurs_msgs P)" "\ = rm_occurs_msgs_constr \" using step.IH by blast note 0 = add_occurs_msgs_cases[OF T] note 1 = add_occurs_msgs_vars_eq[OF bspec[OF P'(1)]] note 2 = add_occurs_msgs_trms[of T] note 3 = add_occurs_msgs_transaction_strand_set[OF T] have 4: "add_occurs_msgs T \ set (map add_occurs_msgs P)" using step.hyps(2) by simp have 5: "transaction_decl_subst \ (add_occurs_msgs T)" using step.hyps(3) 0(4) unfolding transaction_decl_subst_def by argo have "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms_transaction (add_occurs_msgs T))" when t: "t \ subst_range \" for t proof - obtain c where c: "t = Fun (Val c) []" using t T_fresh_val transaction_fresh_subst_domain[OF step.hyps(4)] transaction_fresh_subst_sends_to_val[OF step.hyps(4), of _ thesis] by fastforce have *: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms_transaction T)" using t step.hyps(4) unfolding transaction_fresh_subst_def by (fast,fast) have "t \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ (\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. t \ occurs (Var x)) \ (\c. Fun c [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ t \ occurs (Fun c []))" when t: "t \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using t rm_occurs_msgs_constr_reachable_constraints_trms_cases[OF P' B(2,1)] by fast thus "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using *(1) unfolding c by fastforce show "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms_transaction (add_occurs_msgs T))" using *(2) unfolding 2 c by force qed hence 6: "transaction_fresh_subst \ (add_occurs_msgs T) (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using step.hyps(4) unfolding transaction_fresh_subst_def 0(5) 2 by fast have 7: "transaction_renaming_subst \ (map add_occurs_msgs P) (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using step.hyps(5) rm_occurs_msgs_constr_reachable_constraints_vars_eq[OF P' B(1)] B(2) 1(5) unfolding transaction_renaming_subst_def by simp have "?A' = rm_occurs_msgs_constr ?B'" using admissible_transaction_decl_fresh_renaming_subst_not_occurs[OF T_adm step.hyps(3,4,5)] rm_occurs_msgs_constr_transaction_strand''[OF T_adm T_no_occ] unfolding \_def[symmetric] by metis hence 8: "\@?A' = rm_occurs_msgs_constr (\@?B')" by (metis rm_occurs_msgs_constr_append B(2)) show ?case using reachable_constraints.step[OF B(1) 4 5 6 7] 8 unfolding \_def by blast qed (metis reachable_constraints.init rm_occurs_msgs_constr.simps(1)) private lemma add_occurs_msgs_soundness_aux3: assumes P: "\T \ set P. admissible_transaction T" and A: "\ \ reachable_constraints (map add_occurs_msgs P)" "welltyped_constraint_model \ (rm_occurs_msgs_constr \)" and I: "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. \n. \ x = Fun (Val n) []" (is "?I \") shows "welltyped_constraint_model \ \" (is "?Q \ \") using A I proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \ \) let ?f = rm_occurs_msgs_constr let ?sem = "\B. strand_sem_stateful (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \) (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ {}) (unlabel B) \" define \ where "\ \ \ \\<^sub>s \ \\<^sub>s \" define \ where "\ \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" obtain T' where T': "T' \ set P" "T = add_occurs_msgs T'" using step.hyps(2) by fastforce then obtain A' B' C' D' E' F' where T'': "T' = Transaction A' B' C' D' E' F'" using prot_transaction.exhaust by blast have P': "\T \ set (map add_occurs_msgs P). admissible_transaction' T" "\T \ set (map add_occurs_msgs P). admissible_transaction_occurs_checks T" "\T \ set P. admissible_transaction' T" "\T \ set P. admissible_transaction_no_occurs_msgs T" using P admissible_transactionE' add_occurs_msgs_admissible_occurs_checks by (fastforce,fastforce,fastforce,fastforce) note T_adm = bspec[OF P'(1) step.hyps(2)] note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] note T'_adm = bspec[OF P'(3) T'(1)] note T'_no_occ = bspec[OF P'(4) T'(1)] note T'_wf = admissible_transaction_is_wellformed_transaction(1)[OF T'_adm] note \_empty = admissible_transaction_decl_subst_empty[OF T_adm step.hyps(3)] note T_fresh_val = admissible_transactionE(2)[OF T_adm] have 0: "?Q \ (?f \)" "?I \" "?I \" by (metis step.prems(1) welltyped_constraint_model_prefix rm_occurs_msgs_constr_append, simp_all add: step.prems(2) \_def \_def) note IH = step.IH[OF 0(1,2)] have I': "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using step.prems(1) unfolding welltyped_constraint_model_def constraint_model_def by blast+ have 1: "\x \ fv_transaction T. \t. \ x = occurs t" "\x \ fv_transaction T. \ x \ Fun OccursSec []" using admissible_transaction_decl_fresh_renaming_subst_not_occurs[OF T_adm step.hyps(3,4,5)] unfolding \_def[symmetric] by simp_all have "(ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?f \) \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \) = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" using rm_occurs_msgs_constr_ik_subset by fast hence 2: "?sem (?f \)" using step.prems(1) strand_sem_append_stateful[of "{}" "{}" "unlabel (?f \)" "unlabel (?f \)"] rm_occurs_msgs_constr_dbupd\<^sub>s\<^sub>s\<^sub>t_eq[of \ \ "{}"] rm_occurs_msgs_constr_append[of \ \] strand_sem_ik_mono_stateful[of "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?f \) \\<^sub>s\<^sub>e\<^sub>t \" _ _ _ "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \"] unfolding welltyped_constraint_model_def constraint_model_def \_def[symmetric] \_def[symmetric] by auto note 3 = rm_occurs_msgs_constr_transaction_strand''[ OF T'_adm T'_no_occ 1[unfolded T'(2)], unfolded \_def[symmetric] T'(2)[symmetric]] note 4 = add_occurs_msgs_cases[OF T'', unfolded T'(2)[symmetric]] define xs where "xs \ fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T'))" define flt where "flt \ filter (\x. x \ set (transaction_fresh T'))" define occs where "occs \ map (\x. occurs (Var x)::('fun,'atom,'sets,'lbl) prot_term)" note 6 = add_occurs_msgs_transaction_strand_cases(7,8,9)[ of T' \, unfolded xs_def[symmetric] flt_def[symmetric] occs_def[symmetric] T'(2)[symmetric]] have 7: "x \ fv_transaction T - set (transaction_fresh T)" when x: "x \ set (flt xs)" for x using that fv_list\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t add_occurs_msgs_vars_eq(3,9)[OF T'_adm] unfolding xs_def flt_def T'(2) by force have 9: "\y. \ x = Var y" when x: "x \ set (flt xs)" for x proof - have *: "x \ fst ` set (transaction_decl T ())" using admissible_transactionE(1)[OF T_adm] by simp have **: "x \ set (transaction_fresh T)" using 7[OF x] by simp show ?thesis using transaction_decl_fresh_renaming_substs_range(4)[OF step.hyps(3,4,5) * **] unfolding \_def by blast qed have 8: "\y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. \ x = Var y" when x: "x \ set (flt xs)" for x proof - note * = 7[OF x] obtain y where y: "\ x = Var y" using 9[OF x] by blast have "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T))" by (metis * Diff_iff fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq) have "\ x \ \ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)" using * by fast hence "fv (\ x) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv_transaction T)" by force hence "fv (\ x) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using fv\<^sub>s\<^sub>s\<^sub>t_subst_if_no_bvars[OF admissible_transactionE(4)[OF T_adm], of \] by (metis unlabel_subst) hence "fv (\ x) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" by (metis fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq \_def) thus ?thesis using y by simp qed have \_var_is_\_val: "\n. \ x = Fun (Val n) []" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" for x using step.prems(2) x unfolding \_def[symmetric] \_def[symmetric] by auto have T'_var_is_\\_val: "\n. \ x \ \ = Fun (Val n) []" when x: "x \ set (flt xs)" for x using 8[OF x] \_var_is_\_val by force (* TODO: extract lemma *) have poschecks_has_occ: "occurs (Fun (Val n) []) \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" when x: "\ac: t \ s\ \ set (unlabel \)" and n: "t \ \ = Fun (Val n) []" for ac t s n proof - have *: "(t \ \, s \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ {}" proof - obtain t' s' where t': "\ac: t' \ s'\ \ set (unlabel (transaction_checks T'))" "t = t' \ \" "s = s' \ \" using 4(8) x stateful_strand_step_mem_substD(6) wellformed_transaction_strand_unlabel_memberD(10)[OF T_wf(1)] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(6) unfolding \_def by (metis (no_types) unlabel_subst dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst) have "(t' \ \ \ \, s' \ \ \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ {}" using t'(1) 2 wellformed_transaction_unlabel_sem_iff[ OF T'_wf(1) I'(2,3), of "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ {}" \] unfolding 3 by blast thus ?thesis using t'(2,3) by simp qed have "t' \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" when "insert\t',s'\ \ set (unlabel \)" for t' s' using that by force have "t \ \ \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" proof - obtain t' s' where t': "insert\t',s'\ \ set (unlabel \)" "t \ \ = t' \ \" "s \ \ = s' \ \" using * db\<^sub>s\<^sub>s\<^sub>t_in_cases[of "t \ \" "s \ \" "unlabel \" \ "[]"] db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t[of "unlabel \" \ "[]"] by auto have t'': "t' = t \ \ \ (\y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. t' = Var y \ \ y = t \ \)" using t'(1,2) stateful_strand_step_fv_subset_cases(4) unfolding n by (cases t') (force,force) thus ?thesis proof assume "t' = t \ \" thus ?thesis using t'(1) by force next assume "\y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. t' = Var y \ \ y = t \ \" then obtain y where y: "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" "\ y = t \ \" by blast have "\\<^sub>v y = TAtom Value" using y(2) wt_subst_trm''[OF I'(1), of "Var y"] unfolding n by simp hence "\B. prefix B \ \ t \ \ \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B" by (metis y constraint_model_Value_var_in_constr_prefix[OF step.hyps(1) IH P'(1,2)]) thus ?thesis unfolding prefix_def by auto qed qed thus ?thesis using reachable_constraints_occurs_fact_ik_case'[OF step.hyps(1) P'(1,2)] unfolding n by blast qed (* TODO: extract lemma *) have snds_has_occ: "occurs (Fun (Val n) []) \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" when ts: "send\ts\ \ set (unlabel \)" and n: "Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t set ts \\<^sub>s\<^sub>e\<^sub>t \" for ts n proof - have "receive\ts\ \ set (unlabel (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using ts dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(2) unfolding \_def by metis then obtain ts' where ts': "receive\ts'\ \ set (unlabel (transaction_strand T))" "ts = ts' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" by (metis subst_lsst_memD(1) unlabel_in unlabel_mem_has_label) have "?sem (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using 2 strand_sem_append_stateful[of "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ {}"] unfolding 3 transaction_dual_subst_unlabel_unfold by blast moreover have "list_all is_Receive (unlabel (transaction_receive T'))" using T'_wf unfolding wellformed_transaction_def by blast hence "list_all is_Send (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" by (metis subst_lsst_unlabel subst_sst_list_all(2) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_list_all(1)) hence "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) = {}" using in_ik\<^sub>s\<^sub>s\<^sub>t_iff unfolding list_all_iff is_Send_def by fast ultimately have *: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" when "send\ts\ \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" "t \ set ts" for t ts using strand_sem_stateful_sends_deduct[OF _ that] by simp hence *: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \ \ \" when ts: "receive\ts\ \ set (unlabel (transaction_receive T'))" "t \ set ts" for t ts using ts(2) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(2)[of "ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" "transaction_receive T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] stateful_strand_step_subst_inI(2)[OF ts(1), of \, unfolded unlabel_subst] by auto have **: "set (flt xs) = fv_transaction T' - set (transaction_fresh T')" using fv_list\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t unfolding flt_def xs_def by fastforce have rcv_case: ?thesis when "ts = ts' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" "Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t set ts \\<^sub>s\<^sub>e\<^sub>t \" "receive\ts'\ \ set (unlabel (transaction_receive T'))" for ts ts' using that * reachable_constraints_occurs_fact_ik_case''[OF step.hyps(1) IH P'(1,2)] by auto have "receive\ts'\ \ set (unlabel (transaction_receive T))" using wellformed_transaction_strand_unlabel_memberD(1)[OF T_wf] ts'(1) by blast hence "(ts' = map (\x. occurs (Var x)) (flt xs) \ ts' \ []) \ receive\ts'\ \ set (unlabel (transaction_receive T'))" (is "?A \ ?B") using ** ts'(1) add_occurs_msgs_cases(13)[OF T''] unfolding T'(2)[symmetric] xs_def[symmetric] flt_def[symmetric] by force thus ?thesis proof assume ?A then obtain x where x: "x \ set (flt xs)" "Fun (Val n) [] \ \ x \ \" using ts' n by fastforce have x': "\ x \ \ = Fun (Val n) []" "x \ fv_transaction T" "x \ set (transaction_fresh T)" "x \ fv_transaction T'" "x \ set (transaction_fresh T')" using x(2) T'_var_is_\\_val[OF x(1)] 7[OF x(1)] ** x(1) by fastforce+ let ?snds = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" let ?chks = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" have B_subsets: "set ?chks \ set (unlabel \)" unfolding \_def transaction_dual_subst_unlabel_unfold 4(8) by fastforce from admissible_transaction_fv_in_receives_or_selects_dual_subst[OF T'_adm x'(4,5), of \] show ?thesis proof assume "\ts. send\ts\ \ set ?snds \ \ x \\<^sub>s\<^sub>e\<^sub>t set ts" then obtain ss where ss: "send\ss\ \ set ?snds" "\ x \\<^sub>s\<^sub>e\<^sub>t set ss" by blast obtain ss' where ss': "ss = ss' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" "receive\ss'\ \ set (unlabel (transaction_receive T'))" by (metis ss(1) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(2) subst_lsst_memD(1) unlabel_in unlabel_mem_has_label) show ?thesis using rcv_case[OF ss'(1) _ ss'(2)] subst_subterms[OF ss(2), of \] x'(1) by argo qed (use B_subsets poschecks_has_occ[OF _ x'(1)] in blast) qed (metis rcv_case[OF ts'(2) n]) qed (* TODO: extract lemma *) have "occurs (\ x \ \) \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" when x: "x \ set (flt xs)" for x proof - have "(\ac s. \ac: \ x \ s\ \ set (unlabel \)) \ (\ts. send\ts\ \ set (unlabel \) \ \ x \\<^sub>s\<^sub>e\<^sub>t set ts)" (is "(\ac s. ?A ac s) \ (\ts. ?B1 ts \ ?B2 ts)") using 7[OF x] admissible_transaction_fv_in_receives_or_selects_dual_subst[OF T_adm, of x \] unfolding \_def transaction_dual_subst_unlabel_unfold by auto thus ?thesis proof assume "\ac s. ?A ac s" then obtain ac s where s: "?A ac s" by blast show ?thesis using poschecks_has_occ[OF s] T'_var_is_\\_val[OF x] by force next assume "\ts. ?B1 ts \ ?B2 ts" then obtain ts where ts: "?B1 ts" "?B2 ts" by meson have ts': "\ x \ \ \\<^sub>s\<^sub>e\<^sub>t set ts \\<^sub>s\<^sub>e\<^sub>t \" by (metis ts(2) subst_subterms) show ?thesis using snds_has_occ[OF ts(1)] ts' T'_var_is_\\_val[OF x] by force qed qed hence "occurs (\ x \ \) \ \ \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" when "x \ set (flt xs)" for x using that by fast moreover have "occurs (\ x \ \) \ \ = occurs (\ x \ \)" for x using subst_ground_ident[OF interpretation_grounds[OF I'(2), of "\ x"], of \] by simp ultimately have "occurs (\ x \ \) \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" when "x \ set (flt xs)" for x using that by auto hence "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" when "t \ set (occs (flt xs) \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" for t using that unfolding occs_def by auto hence occs_sem: "?sem [\\, send\occs (flt xs) \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\]" by auto (* TODO: extract lemma *) have "?sem \" proof - let ?IK = "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" let ?DB = "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ {}" let ?snds = "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" let ?snds_occs = "(\\, send\occs (flt xs) \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\)#?snds" let ?chks = "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" let ?upds = "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" let ?rcvs = "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" note * = strand_sem_append_stateful[of _ _ _ _ \] note ** = transaction_dual_subst_unlabel_unfold have ***: "\M. M \ (ik\<^sub>s\<^sub>s\<^sub>t [] \\<^sub>s\<^sub>e\<^sub>t \) = M" "\D. dbupd\<^sub>s\<^sub>s\<^sub>t [] \ D = D" by simp_all have snds_sem: "?sem ?snds" "?sem ?snds_occs" using 2 occs_sem *[of ?IK ?DB] unfolding 3 ** by (blast, fastforce) have "list_all is_Receive (unlabel (transaction_receive T'))" using T'_wf unfolding wellformed_transaction_def by blast hence "list_all is_Send (unlabel ?snds)" "list_all is_Send (unlabel ?snds_occs)" using subst_sst_list_all(2) unlabel_subst dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_list_all(1) by (metis, metis (no_types) list.pred_inject(2) stateful_strand_step.disc(1) unlabel_Cons(1)) hence "\a \ set (unlabel ?snds). \is_Receive a \ \is_Insert a \ \is_Delete a" "\a \ set (unlabel ?snds_occs). \is_Receive a \ \is_Insert a \ \is_Delete a" unfolding list_all_iff by (blast,blast) hence snds_no_upds: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?snds \\<^sub>s\<^sub>e\<^sub>t \ = {}" "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel ?snds) \ ?DB = ?DB" "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?snds_occs) \\<^sub>s\<^sub>e\<^sub>t \ = {}" "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel ?snds_occs) \ ?DB = ?DB" by (metis ik\<^sub>s\<^sub>s\<^sub>t_snoc_no_receive_empty, metis dbupd\<^sub>s\<^sub>s\<^sub>t_no_upd, metis ik\<^sub>s\<^sub>s\<^sub>t_snoc_no_receive_empty, metis dbupd\<^sub>s\<^sub>s\<^sub>t_no_upd) have chks_sem: "?sem ?chks" using 2 snds_no_upds * unfolding 3 ** by auto have "list_all is_Check_or_Assignment (unlabel (transaction_checks T'))" using T'_wf unfolding wellformed_transaction_def by blast hence "list_all is_Check_or_Assignment (unlabel ?chks)" by (metis (no_types) subst_sst_list_all(11) unlabel_subst dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_list_all(11)) hence "\a \ set (unlabel ?chks). \is_Receive a \ \is_Insert a \ \is_Delete a" unfolding list_all_iff by blast hence chks_no_upds: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?chks \\<^sub>s\<^sub>e\<^sub>t \ = {}" "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel ?chks) \ ?DB = ?DB" by (metis ik\<^sub>s\<^sub>s\<^sub>t_snoc_no_receive_empty, metis dbupd\<^sub>s\<^sub>s\<^sub>t_no_upd) have upds_sem: "?sem ?upds" using 2 snds_no_upds chks_no_upds * unfolding 3 ** by auto have "list_all is_Send (unlabel (transaction_send T'))" using T'_wf unfolding wellformed_transaction_def by fast hence "list_all is_Send (unlabel (transaction_send T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" by (metis (no_types, opaque_lifting) subst_sst_list_all(1) unlabel_subst) hence rcvs_is_rcvs: "list_all is_Receive (unlabel ?rcvs)" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_list_all(2) by blast have rcvs_sem: "strand_sem_stateful M D (unlabel rcvs) \" when "list_all is_Receive (unlabel rcvs)" for M D and rcvs::"('fun, 'atom, 'sets, 'lbl) prot_strand" using rcvs_is_rcvs strand_sem_receive_prepend_stateful[of M D "[]" \, OF _ that] by auto have B_sem: "?sem (?snds@?chks@?upds@rcvs)" "?sem (?snds_occs@?chks@?upds@rcvs)" when "list_all is_Receive (unlabel rcvs)" for rcvs using strand_sem_append_stateful[of _ _ _ _ \] snds_sem snds_no_upds chks_sem chks_no_upds upds_sem rcvs_sem[OF that] by (force, force) show ?thesis proof (cases "transaction_fresh T' = []") case True show ?thesis using B_sem[OF rcvs_is_rcvs] unfolding \_def 6(1)[OF True] by force next case False note F = this show ?thesis proof (cases "\ts F'. transaction_send T' = \\, send\ts\\#F'") case True obtain ts F' rcvs' where F': "transaction_send T' = \\, send\ts\\#F'" "\ = (if flt xs = [] then ?snds else ?snds_occs)@?chks@?upds@rcvs'" "rcvs' = \\, receive\occs (transaction_fresh T')@ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\#dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (F' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using 6(3)[OF F True] unfolding \_def by blast have *: "list_all is_Receive (unlabel rcvs')" using rcvs_is_rcvs dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_Cons(1)[of \ ts "F' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] subst_lsst_cons[of "\\, send\ts\\" F' \] unfolding F'(1,3) list_all_iff by auto show ?thesis using B_sem[OF *] unfolding F'(2) by fastforce next case False have *: "list_all is_Receive (unlabel (\\, receive\occs (transaction_fresh T') \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\#?rcvs))" using rcvs_is_rcvs by auto show ?thesis using B_sem[OF *] unfolding \_def 6(2)[OF F False] by fastforce qed qed qed thus ?case using IH strand_sem_append_stateful[of "{}" "{}" "unlabel \" "unlabel \" \] unfolding welltyped_constraint_model_def constraint_model_def \_def[symmetric] \_def[symmetric] by simp qed simp theorem add_occurs_msgs_soundness: defines "wt_attack \ \\ \ l n. welltyped_constraint_model \ (\@[(l, send\[attack\n\]\)])" assumes P: "\T \ set P. admissible_transaction T" "has_initial_value_producing_transaction P" and A: "\ \ reachable_constraints P" "wt_attack \ \ l n" shows "\\ \ reachable_constraints (map add_occurs_msgs P). \\. wt_attack \ \ l n" proof - have P': "\T \ set (map add_occurs_msgs P). admissible_transaction' T" "\T \ set (map add_occurs_msgs P). admissible_transaction_occurs_checks T" "\T \ set P. admissible_transaction' T" "\T \ set P. admissible_transaction_no_occurs_msgs T" using P admissible_transactionE' add_occurs_msgs_admissible_occurs_checks by (fastforce,fastforce,fastforce,fastforce) obtain \' \ where A': "\' \ reachable_constraints P" "wt_attack \ \' l n" "\x\fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \'. \n. \ x = Fun (Val n) []" using add_occurs_msgs_soundness_aux1[OF P'(3) P(2) A[unfolded wt_attack_def]] unfolding wt_attack_def by blast have J: "welltyped_constraint_model \ \'" using A'(2) welltyped_constraint_model_prefix unfolding wt_attack_def by blast obtain \ where B: "\ \ reachable_constraints (map add_occurs_msgs P)" "\' = rm_occurs_msgs_constr \" using add_occurs_msgs_soundness_aux2[OF P(1) A'(1)] by blast have J': "welltyped_constraint_model \ \" using add_occurs_msgs_soundness_aux3[OF P(1) B(1) J[unfolded B(2)]] A'(3) rm_occurs_msgs_constr_reachable_constraints_fv_eq[OF P'(3,4) B(1)] unfolding wt_attack_def B(2) by blast obtain ts where ts: "receive\ts\ \ set (unlabel \)" "attack\n\ \ set ts" using reachable_constraints_receive_attack_if_attack''[OF P'(3) A'(1,2)[unfolded wt_attack_def]] rm_occurs_msgs_constr_receive_attack_iff[of n \] unfolding B(2)[symmetric] by auto have J'': "wt_attack \ \ l n" using welltyped_constraint_model_attack_if_receive_attack[OF J' ts] unfolding wt_attack_def by fast show ?thesis using B(1) J'' by blast qed end end subsection \Automatically Checking Protocol Security in a Typed Model\ context stateful_protocol_model begin definition abs_intruder_knowledge ("\\<^sub>i\<^sub>k") where "\\<^sub>i\<^sub>k S \ \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \)" definition abs_value_constants ("\\<^sub>v\<^sub>a\<^sub>l\<^sub>s") where "\\<^sub>v\<^sub>a\<^sub>l\<^sub>s S \ \ {t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t S) \\<^sub>s\<^sub>e\<^sub>t \. \n. t = Fun (Val n) []} \\<^sub>\\<^sub>s\<^sub>e\<^sub>t \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \)" definition abs_term_implications ("\\<^sub>t\<^sub>i") where "\\<^sub>t\<^sub>i \ T \ \ \ {(s,t) | s t x. s \ t \ x \ fv_transaction T \ x \ set (transaction_fresh T) \ Fun (Abs s) [] = \ x \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) \ Fun (Abs t) [] = \ x \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \)}" lemma abs_intruder_knowledge_append: "\\<^sub>i\<^sub>k (A@B) \ = (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@B) \) \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t B \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@B) \)" by (metis unlabel_append abs_set_union image_Un ik\<^sub>s\<^sub>s\<^sub>t_append abs_intruder_knowledge_def) lemma abs_value_constants_append: fixes A B::"('a,'b,'c,'d) prot_strand" shows "\\<^sub>v\<^sub>a\<^sub>l\<^sub>s (A@B) \ = {t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \\<^sub>s\<^sub>e\<^sub>t \. \n. t = Fun (Val n) []} \\<^sub>\\<^sub>s\<^sub>e\<^sub>t \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@B) \) \ {t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B) \\<^sub>s\<^sub>e\<^sub>t \. \n. t = Fun (Val n) []} \\<^sub>\\<^sub>s\<^sub>e\<^sub>t \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@B) \)" proof - define a0 where "a0 \ \\<^sub>0 (db\<^sub>s\<^sub>s\<^sub>t (unlabel (A@B)) \)" define M where "M \ \a::('a,'b,'c,'d) prot_strand. {t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t a) \\<^sub>s\<^sub>e\<^sub>t \. \n. t = Fun (Val n) []}" have "M (A@B) = M A \ M B" using image_Un[of "\x. x \ \" "subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" "subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)"] unfolding M_def unlabel_append[of A B] trms\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel A" "unlabel B"] by blast hence "M (A@B) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a0 = (M A \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a0) \ (M B \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a0)" by (simp add: abs_set_union) thus ?thesis unfolding abs_value_constants_def a0_def M_def by force qed lemma transaction_renaming_subst_has_no_pubconsts_abss: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_renaming_subst \ P A" shows "subst_range \ \ pubval_terms = {}" (is ?A) and "subst_range \ \ abs_terms = {}" (is ?B) proof - { fix t assume "t \ subst_range \" then obtain x where "t = Var x" using transaction_renaming_subst_is_renaming(1)[OF assms] by force hence "t \ pubval_terms" "t \ abs_terms" by simp_all } thus ?A ?B by auto qed lemma transaction_fresh_subst_has_no_pubconsts_abss: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_fresh_subst \ T \" "\x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" shows "subst_range \ \ pubval_terms = {}" (is ?A) and "subst_range \ \ abs_terms = {}" (is ?B) proof - { fix t assume "t \ subst_range \" then obtain x where "x \ set (transaction_fresh T)" "\ x = t" using assms(1) unfolding transaction_fresh_subst_def by auto then obtain n where "t = Fun (Val n) []" using transaction_fresh_subst_sends_to_val[OF assms(1)] assms(2) by meson hence "t \ pubval_terms" "t \ abs_terms" unfolding is_PubConstValue_def by simp_all } thus ?A ?B by auto qed lemma reachable_constraints_GSMP_no_pubvals_abss: assumes "\ \ reachable_constraints P" and P: "\T \ set P. admissible_transaction' T" and \: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "\n. PubConst Value n \ \(funs_term ` (\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" "\n. Abs n \ \(funs_term ` (\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" shows "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ GSMP (\T \ set P. trms_transaction T) - (pubval_terms \ abs_terms)" (is "?A \ ?B") using assms(1) \(4,5) proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \ \) define trms_P where "trms_P \ (\T \ set P. trms_transaction T)" define T' where "T' \ transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \" have \_elim: "\ \\<^sub>s \ \\<^sub>s \ = \ \\<^sub>s \" using admissible_transaction_decl_subst_empty[OF bspec[OF P step.hyps(2)] step.hyps(3)] by simp note T_fresh = admissible_transactionE(2)[OF bspec[OF P step.hyps(2)]] note T_no_bvars = admissible_transactionE(4)[OF bspec[OF P step.hyps(2)]] have T_no_PubVal: "\T \ set P. \n. PubConst Value n \ \(funs_term ` trms_transaction T)" and T_no_Abs: "\T \ set P. \n. Abs n \ \(funs_term ` trms_transaction T)" using admissible_transactions_no_Value_consts''[OF bspec[OF P]] by metis+ have \': "\n. PubConst Value n \ \ (funs_term ` (\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" "\n. Abs n \ \ (funs_term ` (\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using step.prems fv\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \"] unlabel_append[of \] by auto have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (rm_vars (set X) (\ \\<^sub>s \ \\<^sub>s \))" for X using wt_subst_rm_vars[of "\ \\<^sub>s \ \\<^sub>s \" "set X"] transaction_decl_fresh_renaming_substs_wt[OF step.hyps(3-5)] by metis hence wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t ((rm_vars (set X) (\ \\<^sub>s \ \\<^sub>s \)) \\<^sub>s \)" for X using \(2) wt_subst_compose by fast have wftrms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range ((rm_vars (set X) (\ \\<^sub>s \ \\<^sub>s \)) \\<^sub>s \))" for X using wf_trms_subst_compose[OF wf_trms_subst_rm_vars' \(3)] transaction_decl_fresh_renaming_substs_range_wf_trms[OF step.hyps(3-5)] by fast have "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T') \\<^sub>s\<^sub>e\<^sub>t \ \ ?B" proof fix t assume "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T') \\<^sub>s\<^sub>e\<^sub>t \" hence "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \\<^sub>s\<^sub>e\<^sub>t \" using trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq by blast then obtain s X where s: "s \ trms_transaction T" "t = s \ rm_vars (set X) (\ \\<^sub>s \ \\<^sub>s \) \\<^sub>s \" "set X \ bvars_transaction T" using trms\<^sub>s\<^sub>s\<^sub>t_unlabel_subst'' unfolding T'_def by blast define \ where "\ \ rm_vars (set X) (\ \\<^sub>s \ \\<^sub>s \)" note 0 = pubval_terms_subst_range_comp[OF transaction_fresh_subst_has_no_pubconsts_abss(1)[OF step.hyps(4) T_fresh] transaction_renaming_subst_has_no_pubconsts_abss(1)[OF step.hyps(5)]] abs_terms_subst_range_comp[OF transaction_fresh_subst_has_no_pubconsts_abss(2)[OF step.hyps(4) T_fresh] transaction_renaming_subst_has_no_pubconsts_abss(2)[OF step.hyps(5)]] have 1: "s \ trms_P" using step.hyps(2) s(1) unfolding trms_P_def by auto have s_nin: "s \ pubval_terms" "s \ abs_terms" using 1 T_no_PubVal T_no_Abs funs_term_Fun_subterm unfolding trms_P_def is_PubConstValue_def is_Abs_def is_PubConst_def by (fastforce, blast) have 2: "(\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')) \ pubval_terms = {}" "(\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')) \ abs_terms = {}" "subst_range (\ \\<^sub>s \ \\<^sub>s \) \ pubval_terms = {}" "subst_range (\ \\<^sub>s \ \\<^sub>s \) \ abs_terms = {}" using 0 step.prems funs_term_Fun_subterm unfolding T'_def \_def \_elim by (fastforce simp add: is_PubConstValue_def is_PubConst_def, fastforce simp add: is_Abs_def, argo, argo) have "subst_range \ \ subst_range (\ \\<^sub>s \ \\<^sub>s \)" using rm_vars_img_subset unfolding \_def \_elim by blast hence 3: "subst_range \ \ pubval_terms = {}" "subst_range \ \ abs_terms = {}" using 2(3,4) step.prems funs_term_Fun_subterm unfolding T'_def \_def \_elim by (blast,blast) have "(\ ` fv (s \ \)) \ pubval_terms = {}" "(\ ` fv (s \ \)) \ abs_terms = {}" proof - have "\ = \ \\<^sub>s \ \\<^sub>s \" "bvars_transaction T = {}" "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" using s(3) T_no_bvars step.hyps(2) rm_vars_empty vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel T'"] bvars\<^sub>s\<^sub>s\<^sub>t_subst[of "unlabel (transaction_strand T)" "\ \\<^sub>s \ \\<^sub>s \"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] unfolding \_def T'_def by simp_all hence "fv (s \ \) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" using trms\<^sub>s\<^sub>s\<^sub>t_fv_subst_subset[OF s(1), of \] unlabel_subst[of "transaction_strand T" \] unfolding T'_def by auto moreover have "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" using fv\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')"] unlabel_append[of \ "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'"] fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of T'] by simp_all hence "\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ pubval_terms = {}" "\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ abs_terms = {}" using 2(1,2) by blast+ ultimately show "(\ ` fv (s \ \)) \ pubval_terms = {}" "(\ ` fv (s \ \)) \ abs_terms = {}" by blast+ qed hence \\\_disj: "((\ \\<^sub>s \) ` fv s) \ pubval_terms = {}" "((\ \\<^sub>s \) ` fv s) \ abs_terms = {}" using pubval_terms_subst_range_comp'[of \ "fv s" \] abs_terms_subst_range_comp'[of \ "fv s" \] pubval_terms_subst_range_disj[OF 3(1), of s] abs_terms_subst_range_disj[OF 3(2), of s] by (simp_all add: subst_apply_fv_unfold) have 4: "t \ pubval_terms" "t \ abs_terms" using s(2) s_nin \\\_disj pubval_terms_subst[of s "rm_vars (set X) (\ \\<^sub>s \ \\<^sub>s \) \\<^sub>s \"] pubval_terms_subst_range_disj[of "rm_vars (set X) (\ \\<^sub>s \ \\<^sub>s \) \\<^sub>s \" s] abs_terms_subst[of s "rm_vars (set X) (\ \\<^sub>s \ \\<^sub>s \) \\<^sub>s \"] abs_terms_subst_range_disj[of "rm_vars (set X) (\ \\<^sub>s \ \\<^sub>s \) \\<^sub>s \" s] unfolding \_def by blast+ have "t \ SMP trms_P" "fv t = {}" by (metis s(2) SMP.Substitution[OF SMP.MP[OF 1] wt wftrms, of X], metis s(2) subst_subst_compose[of s "rm_vars (set X) (\ \\<^sub>s \ \\<^sub>s \)" \] interpretation_grounds[OF \(1), of "s \ rm_vars (set X) (\ \\<^sub>s \ \\<^sub>s \)"]) hence 5: "t \ GSMP trms_P" unfolding GSMP_def by simp show "t \ ?B" using 4 5 by (auto simp add: trms_P_def) qed thus ?case using step.IH[OF \'] trms\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \"] unlabel_append[of \] image_Un[of "\x. x \ \" "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] by (simp add: T'_def) qed simp lemma \\<^sub>t\<^sub>i_covers_\\<^sub>0_aux: assumes \_reach: "\ \ reachable_constraints P" and T: "T \ set P" and \: "welltyped_constraint_model \ (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" and \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and t: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "t = Fun (Val n) [] \ t = Var x" and neq: "t \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) \ t \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) \)" shows "\y \ fv_transaction T - set (transaction_fresh T). t \ \ = (\ \\<^sub>s \ \\<^sub>s \) y \ \ \ \\<^sub>v y = TAtom Value" proof - let ?\' = "\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" let ?\ = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T))" let ?\' = "?\ \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \" let ?\'' = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" have \_interp: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and \_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (metis \ welltyped_constraint_model_def constraint_model_def, metis \ welltyped_constraint_model_def, metis \ welltyped_constraint_model_def constraint_model_def) note T_adm = bspec[OF P(1) T] note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] note T_adm_upds = admissible_transaction_is_wellformed_transaction(3)[OF T_adm] have T_fresh_vars_value_typed: "\x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" using T P(1) protocol_transaction_vars_TAtom_typed(3)[of T] P(1) by simp note \_empty = admissible_transaction_decl_subst_empty[OF T_adm \] note \\\_wt = transaction_decl_fresh_renaming_substs_wt[OF \ \ \] have \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by (metis reachable_constraints_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s admissible_transactions_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s P(1) \_reach) hence t_wf: "wf\<^sub>t\<^sub>r\<^sub>m t" using t by auto have \_no_val_bvars: "\TAtom Value \ \\<^sub>v x" when "x \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" for x using P(1) reachable_constraints_no_bvars \_reach vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel \"] that admissible_transactionE(4) by fast have x': "x \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" when "t = Var x" using that t by (simp add: var_subterm_trms\<^sub>s\<^sub>s\<^sub>t_is_vars\<^sub>s\<^sub>s\<^sub>t) have "\f \ funs_term (t \ \). is_Val f" using abs_eq_if_no_Val neq by metis hence "\n T. Fun (Val n) T \ t \ \" using funs_term_Fun_subterm unfolding is_Val_def by fast hence "TAtom Value \ \ (Var x)" when "t = Var x" using wt_subst_trm''[OF \_wt, of "Var x"] that subtermeq_imp_subtermtypeeq[of "t \ \"] wf_trm_subst[OF \_wf, of t] t_wf by fastforce hence x_val: "\\<^sub>v x = TAtom Value" when "t = Var x" using reachable_constraints_vars_TAtom_typed[OF \_reach P x'] that by fastforce hence x_fv: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" when "t = Var x" using x' using reachable_constraints_Value_vars_are_fv[OF \_reach P x'] that by blast then obtain m where m: "t \ \ = Fun (Val m) []" using constraint_model_Value_term_is_Val[ OF \_reach welltyped_constraint_model_prefix[OF \] P P_occ, of x] t(2) x_val by force hence 0: "\\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) m \ \\<^sub>0 (db\<^sub>s\<^sub>s\<^sub>t (unlabel \@?\'') \) m" using neq by (simp add: unlabel_def) have t_val: "\ t = TAtom Value" using x_val t by force obtain u s where s: "t \ \ = u \ \" "insert\u,s\ \ set ?\' \ delete\u,s\ \ set ?\'" using to_abs_neq_imp_db_update[OF 0] m by (metis (no_types, lifting) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst subst_lsst_unlabel) then obtain u' s' where s': "u = u' \ \ \\<^sub>s \ \\<^sub>s \" "s = s' \ \ \\<^sub>s \ \\<^sub>s \" "insert\u',s'\ \ set ?\ \ delete\u',s'\ \ set ?\" using stateful_strand_step_mem_substD(4,5) by blast hence s'': "insert\u',s'\ \ set (unlabel (transaction_strand T)) \ delete\u',s'\ \ set (unlabel (transaction_strand T))" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(4,5)[of u' s' "transaction_strand T"] by simp_all then obtain y where y: "y \ fv_transaction T" "u' = Var y" using transaction_inserts_are_Value_vars[OF T_wf T_adm_upds, of u' s'] transaction_deletes_are_Value_vars[OF T_wf T_adm_upds, of u' s'] stateful_strand_step_fv_subset_cases(4,5)[of u' s' "unlabel (transaction_strand T)"] by auto hence 1: "t \ \ = (\ \\<^sub>s \ \\<^sub>s \) y \ \" using y s(1) s'(1) by (metis eval_term.simps(1)) have 2: "y \ set (transaction_fresh T)" when "(\ \\<^sub>s \ \\<^sub>s \) y \ \ \ \ y" using transaction_fresh_subst_grounds_domain[OF \, of y] subst_compose[of \ \ y] that \_empty by (auto simp add: subst_ground_ident) have 3: "y \ set (transaction_fresh T)" when "(\ \\<^sub>s \ \\<^sub>s \) y \ \ \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using 2 that \ unfolding transaction_fresh_subst_def by fastforce have 4: "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. \\<^sub>v x = TAtom Value \ (\B. prefix B \ \ x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t B \ \ x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B))" by (metis welltyped_constraint_model_prefix[OF \] constraint_model_Value_var_in_constr_prefix[OF \_reach _ P P_occ]) have 5: "\\<^sub>v y = TAtom Value" using 1 t_val wt_subst_trm''[OF \\\_wt, of "Var y"] wt_subst_trm''[OF \_wt, of t] wt_subst_trm''[OF \_wt, of "(\ \\<^sub>s \ \\<^sub>s \) y"] by (auto simp del: subst_subst_compose) have "y \ set (transaction_fresh T)" proof (cases "t = Var x") case True (* \ x occurs in \ but not in subst_range \, so y cannot be fresh *) hence *: "\ x = Fun (Val m) []" "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" "\ x = (\ \\<^sub>s \ \\<^sub>s \) y \ \" using m t(1) 1 x_fv x' by (force, blast, force) obtain B where B: "prefix B \" "\ x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" using *(2) 4 x_val[OF True] by fastforce hence "\t \ subst_range \. t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" using transaction_fresh_subst_range_fresh(1)[OF \] trms\<^sub>s\<^sub>s\<^sub>t_unlabel_prefix_subset(1)[of B] unfolding prefix_def by fast thus ?thesis using *(1,3) B(2) 2 by (metis subst_imgI term.distinct(1)) next case False hence "t \ \ \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using t by simp thus ?thesis using 1 3 by argo qed thus ?thesis using 1 5 y(1) by fast qed lemma \\<^sub>t\<^sub>i_covers_\\<^sub>0_Var: assumes \_reach: "\ \ reachable_constraints P" and T: "T \ set P" and \: "welltyped_constraint_model \ (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" and \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" shows "\ x \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) \) \ timpl_closure_set {\ x \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)} (\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \)" proof - define a0 where "a0 \ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" define a0' where "a0' \ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) \)" define a3 where "a3 \ \\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \" have \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by (metis reachable_constraints_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s admissible_transactions_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s P(1) \_reach) note T_adm = bspec[OF P(1) T] note \_empty = admissible_transaction_decl_subst_empty[OF T_adm \] note \\\_wt = transaction_decl_fresh_renaming_substs_wt[OF \ \ \] have \_interp: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (metis \ welltyped_constraint_model_def constraint_model_def, metis \ welltyped_constraint_model_def, metis \ welltyped_constraint_model_def constraint_model_def) have "\\<^sub>v x = Var Value \ (\a. \\<^sub>v x = Var (prot_atom.Atom a))" using reachable_constraints_vars_TAtom_typed[OF \_reach P, of x] x vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel \"] by auto hence "\ x \\<^sub>\ a0' \ timpl_closure_set {\ x \\<^sub>\ a0} a3" proof assume x_val: "\\<^sub>v x = TAtom Value" show "\ x \\<^sub>\ a0' \ timpl_closure_set {\ x \\<^sub>\ a0} a3" proof (cases "\ x \\<^sub>\ a0 = \ x \\<^sub>\ a0'") case False hence "\y \ fv_transaction T - set (transaction_fresh T). \ x = (\ \\<^sub>s \ \\<^sub>s \) y \ \ \ \\<^sub>v y = TAtom Value" using \\<^sub>t\<^sub>i_covers_\\<^sub>0_aux[OF \_reach T \ \ \ \ P P_occ fv\<^sub>s\<^sub>s\<^sub>t_is_subterm_trms\<^sub>s\<^sub>s\<^sub>t[OF x], of _ x] unfolding a0_def a0'_def by fastforce then obtain y where y: "y \ fv_transaction T - set (transaction_fresh T)" "\ x = (\ \\<^sub>s \ \\<^sub>s \) y \ \" "\ x \\<^sub>\ a0 = (\ \\<^sub>s \ \\<^sub>s \) y \ \ \\<^sub>\ a0" "\ x \\<^sub>\ a0' = (\ \\<^sub>s \ \\<^sub>s \) y \ \ \\<^sub>\ a0'" "\\<^sub>v y = TAtom Value" by metis then obtain n where n: "(\ \\<^sub>s \ \\<^sub>s \) y \ \ = Fun (Val n) []" using \\<^sub>v_TAtom''(2)[of y] x x_val transaction_var_becomes_Val[ OF reachable_constraints.step[OF \_reach T \ \ \] \ \ \ \ P P_occ T, of y] by force have "a0 n \ a0' n" "y \ fv_transaction T" "y \ set (transaction_fresh T)" "absc (a0 n) = (\ \\<^sub>s \ \\<^sub>s \) y \ \ \\<^sub>\ a0" "absc (a0' n) = (\ \\<^sub>s \ \\<^sub>s \) y \ \ \\<^sub>\ a0'" using y n False by force+ hence 1: "(a0 n, a0' n) \ a3" unfolding a0_def a0'_def a3_def abs_term_implications_def by blast have 2: "\ x \\<^sub>\ a0' \ set \a0 n --\ a0' n\\\ x \\<^sub>\ a0\" using y n timpl_apply_const by auto show ?thesis using timpl_closure.TI[OF timpl_closure.FP 1] 2 term_variants_pred_iff_in_term_variants[ of "(\_. [])(Abs (a0 n) := [Abs (a0' n)])"] unfolding timpl_closure_set_def timpl_apply_term_def by auto qed (auto intro: timpl_closure_setI) next assume "\a. \\<^sub>v x = TAtom (Atom a)" - then obtain a where x_atom: "\\<^sub>v x = TAtom (Atom a)" by moura + then obtain a where x_atom: "\\<^sub>v x = TAtom (Atom a)" by force obtain f T where fT: "\ x = Fun f T" using interpretation_grounds[OF \_interp, of "Var x"] by (cases "\ x") auto have fT_atom: "\ (Fun f T) = TAtom (Atom a)" using wt_subst_trm''[OF \_wt, of "Var x"] x_atom fT by simp have T: "T = []" using fT wf_trm_subst[OF \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s, of "Var x"] const_type_inv_wf[OF fT_atom] by fastforce have f: "\is_Val f" using fT_atom unfolding is_Val_def by auto have "\ x \\<^sub>\ b = \ x" for b using T fT abs_term_apply_const(2)[OF f] by auto thus "\ x \\<^sub>\ a0' \ timpl_closure_set {\ x \\<^sub>\ a0} a3" by (auto intro: timpl_closure_setI) qed thus ?thesis by (metis a0_def a0'_def a3_def) qed lemma \\<^sub>t\<^sub>i_covers_\\<^sub>0_Val: assumes \_reach: "\ \ reachable_constraints P" and T: "T \ set P" and \: "welltyped_constraint_model \ (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" and \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and n: "Fun (Val n) [] \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" shows "Fun (Val n) [] \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) \) \ timpl_closure_set {Fun (Val n) [] \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)} (\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \)" proof - define T' where "T' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" define a0 where "a0 \ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" define a0' where "a0' \ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \)" define a3 where "a3 \ \\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \" have \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by (metis reachable_constraints_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s admissible_transactions_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s P(1) \_reach) note T_adm = bspec[OF P(1) T] have "Fun (Abs (a0' n)) [] \ timpl_closure_set {Fun (Abs (a0 n)) []} a3" proof (cases "a0 n = a0' n") case False then obtain x where x: "x \ fv_transaction T - set (transaction_fresh T)" "Fun (Val n) [] = (\ \\<^sub>s \ \\<^sub>s \) x \ \" using \\<^sub>t\<^sub>i_covers_\\<^sub>0_aux[OF \_reach T \ \ \ \ P P_occ n] by (fastforce simp add: a0_def a0'_def T'_def) hence "absc (a0 n) = (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ a0" "absc (a0' n) = (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ a0'" by simp_all hence 1: "(a0 n, a0' n) \ a3" using False x(1) unfolding a0_def a0'_def a3_def abs_term_implications_def T'_def by blast show ?thesis using timpl_apply_Abs[of "[]" "[]" "a0 n" "a0' n"] timpl_closure.TI[OF timpl_closure.FP[of "Fun (Abs (a0 n)) []" a3] 1] term_variants_pred_iff_in_term_variants[of "(\_. [])(Abs (a0 n) := [Abs (a0' n)])"] unfolding timpl_closure_set_def timpl_apply_term_def by force qed (auto intro: timpl_closure_setI) thus ?thesis by (simp add: a0_def a0'_def a3_def T'_def) qed lemma \\<^sub>t\<^sub>i_covers_\\<^sub>0_ik: assumes \_reach: "\ \ reachable_constraints P" and T: "T \ set P" and \: "welltyped_constraint_model \ (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" and \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and t: "t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" shows "t \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) \) \ timpl_closure_set {t \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)} (\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \)" proof - define a0 where "a0 \ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" define a0' where "a0' \ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) \)" define a3 where "a3 \ \\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \" let ?U = "\T a. map (\s. s \ \ \\<^sub>\ a) T" have \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by (metis reachable_constraints_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s admissible_transactions_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s P(1) \_reach) note T_adm = bspec[OF P(1) T] have "t \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "wf\<^sub>t\<^sub>r\<^sub>m t" using \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s t ik\<^sub>s\<^sub>s\<^sub>t_trms\<^sub>s\<^sub>s\<^sub>t_subset by force+ hence "\t0 \ subterms t. t0 \ \ \\<^sub>\ a0' \ timpl_closure_set {t0 \ \ \\<^sub>\ a0} a3" proof (induction t) case (Var x) thus ?case using \\<^sub>t\<^sub>i_covers_\\<^sub>0_Var[OF \_reach T \ \ \ \ P P_occ, of x] ik\<^sub>s\<^sub>s\<^sub>t_var_is_fv[of x "unlabel \"] vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel \"] by (simp add: a0_def a0'_def a3_def) next case (Fun f S) have IH: "\t0 \ subterms t. t0 \ \ \\<^sub>\ a0' \ timpl_closure_set {t0 \ \ \\<^sub>\ a0} a3" when "t \ set S" for t using that Fun.prems(1) wf_trm_param[OF Fun.prems(2)] Fun.IH by (meson in_subterms_subset_Union params_subterms subsetCE) hence "t \\<^sub>\ a0' \ timpl_closure_set {t \\<^sub>\ a0} a3" when "t \ set (map (\s. s \ \) S)" for t using that by auto hence "t \\<^sub>\ a0' \ timpl_closure (t \\<^sub>\ a0) a3" when "t \ set (map (\s. s \ \) S)" for t using that timpl_closureton_is_timpl_closure by auto hence "(t \\<^sub>\ a0, t \\<^sub>\ a0') \ timpl_closure' a3" when "t \ set (map (\s. s \ \) S)" for t using that timpl_closure_is_timpl_closure' by auto hence IH': "((?U S a0) ! i, (?U S a0') ! i) \ timpl_closure' a3" when "i < length (map (\s. s \ \ \\<^sub>\ a0) S)" for i using that by auto show ?case proof (cases "\n. f = Val n") case True then obtain n where "Fun f S = Fun (Val n) []" using Fun.prems(2) unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by force moreover have "Fun f S \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using ik\<^sub>s\<^sub>s\<^sub>t_trms\<^sub>s\<^sub>s\<^sub>t_subset Fun.prems(1) by blast ultimately show ?thesis using \\<^sub>t\<^sub>i_covers_\\<^sub>0_Val[OF \_reach T \ \ \ \ P P_occ] by (simp add: a0_def a0'_def a3_def) next case False hence "Fun f S \ \ \\<^sub>\ a = Fun f (map (\t. t \ \ \\<^sub>\ a) S)" for a by (cases f) simp_all hence "(Fun f S \ \ \\<^sub>\ a0, Fun f S \ \ \\<^sub>\ a0') \ timpl_closure' a3" using timpl_closure_FunI[OF IH'] by simp hence "Fun f S \ \ \\<^sub>\ a0' \ timpl_closure_set {Fun f S \ \ \\<^sub>\ a0} a3" using timpl_closureton_is_timpl_closure timpl_closure_is_timpl_closure' by metis thus ?thesis using IH by simp qed qed thus ?thesis by (simp add: a0_def a0'_def a3_def) qed lemma transaction_prop1: assumes "\ \ abs_substs_fun ` set (transaction_check_comp msgcs (FP, OCC, TI) T)" and "x \ fv_transaction T" and "x \ set (transaction_fresh T)" and "\ x \ absdbupd (unlabel (transaction_updates T)) x (\ x)" and "transaction_check' msgcs (FP, OCC, TI) T" and TI: "set TI = {(a,b) \ (set TI)\<^sup>+. a \ b}" shows "(\ x, absdbupd (unlabel (transaction_updates T)) x (\ x)) \ (set TI)\<^sup>+" proof - let ?upd = "\x. absdbupd (unlabel (transaction_updates T)) x (\ x)" have 0: "fv_transaction T = set (fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T)))" by (metis fv_list\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_strand T)"]) have 1: "transaction_check_post (FP, OCC, TI) T \" using assms(1,5) unfolding transaction_check_def transaction_check'_def list_all_iff by blast have "(\ x, ?upd x) \ set TI \ (\ x, ?upd x) \ (set TI)\<^sup>+" using TI using assms(4) by blast thus ?thesis using assms(2,3,4) 0 1 in_trancl_closure_iff_in_trancl_fun[of _ _ TI] unfolding transaction_check_post_def List.member_def Let_def by blast qed lemma transaction_prop2: assumes \: "\ \ abs_substs_fun ` set (transaction_check_comp msgcs (FP, OCC, TI) T)" and x: "x \ fv_transaction T" "fst x = TAtom Value" and T_check: "transaction_check' msgcs (FP, OCC, TI) T" and T_adm: "admissible_transaction' T" and T_occ: "admissible_transaction_occurs_checks T" and FP: "analyzed (timpl_closure_set (set FP) (set TI))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set FP)" and OCC: "\t \ timpl_closure_set (set FP) (set TI). \f \ funs_term t. is_Abs f \ f \ Abs ` set OCC" "timpl_closure_set (absc ` set OCC) (set TI) \ absc ` set OCC" and TI: "set TI = {(a,b) \ (set TI)\<^sup>+. a \ b}" shows "x \ set (transaction_fresh T) \ \ x \ set OCC" (is "?A' \ ?A") and "absdbupd (unlabel (transaction_updates T)) x (\ x) \ set OCC" (is ?B) proof - let ?xs = "fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T))" let ?ys = "filter (\x. x \ set (transaction_fresh T) \ fst x = TAtom Value) ?xs" let ?C = "unlabel (transaction_checks T)" let ?poss = "transaction_poschecks_comp ?C" let ?negs = "transaction_negchecks_comp ?C" let ?\upd = "\y. absdbupd (unlabel (transaction_updates T)) y (\ y)" note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] have 0: "{x \ fv_transaction T - set (transaction_fresh T). fst x = TAtom Value} = set ?ys" using fv_list\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_strand T)"] by force have 1: "transaction_check_pre (FP, OCC, TI) T \" using \ unfolding transaction_check_comp_def Let_def by fastforce have 2: "transaction_check_post (FP, OCC, TI) T \" using \ T_check unfolding transaction_check_def transaction_check'_def list_all_iff by auto have 3: "\ \ abs_substs_fun ` set (abs_substs_set ?ys OCC ?poss ?negs msgcs)" using \ unfolding transaction_check_comp_def Let_def by force show A: ?A when ?A' using that 0 3 x abs_substs_abss_bounded by blast have 4: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" when x': "x \ set (transaction_fresh T)" using x' admissible_transactionE(7)[OF T_adm] by blast have "intruder_synth_mod_timpls FP TI (occurs (absc (?\upd x)))" when x': "x \ set (transaction_fresh T)" proof - obtain l ts S where ts: "transaction_send T = (l,send\ts\)#S" "occurs (Var x) \ set ts" using admissible_transaction_occurs_checksE2[OF T_occ x'] by blast hence "occurs (Var x) \ set ts" "send\ts\ \ set (unlabel (transaction_send T))" using x' unfolding suffix_def by (fastforce, fastforce) thus ?thesis using 2 x unfolding transaction_check_post_def by fastforce qed hence "timpl_closure_set (set FP) (set TI) \\<^sub>c occurs (absc (?\upd x))" when x': "x \ set (transaction_fresh T)" using x' intruder_synth_mod_timpls_is_synth_timpl_closure_set[ OF TI, of FP "occurs (absc (?\upd x))"] by argo hence "Abs (?\upd x) \ \(funs_term ` timpl_closure_set (set FP) (set TI))" when x': "x \ set (transaction_fresh T)" using x' ideduct_synth_priv_fun_in_ik[ of "timpl_closure_set (set FP) (set TI)" "occurs (absc (?\upd x))"] by simp hence "\t \ timpl_closure_set (set FP) (set TI). Abs (?\upd x) \ funs_term t" when x': "x \ set (transaction_fresh T)" using x' by force hence 5: "?\upd x \ set OCC" when x': "x \ set (transaction_fresh T)" using x' OCC by fastforce have 6: "?\upd x \ set OCC" when x': "x \ set (transaction_fresh T)" proof (cases "\ x = ?\upd x") case False hence "(\ x, ?\upd x) \ (set TI)\<^sup>+" "\ x \ set OCC" using A 2 x' x TI unfolding transaction_check_post_def fv_list\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t Let_def in_trancl_closure_iff_in_trancl_fun[symmetric] List.member_def by blast+ thus ?thesis using timpl_closure_set_absc_subset_in[OF OCC(2)] by blast qed (simp add: A x' x(1)) show ?B by (metis 5 6) qed lemma transaction_prop3: assumes \_reach: "\ \ reachable_constraints P" and T: "T \ set P" and \: "welltyped_constraint_model \ (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" and \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and FP: "analyzed (timpl_closure_set (set FP) (set TI))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set FP)" "\t \ \\<^sub>i\<^sub>k \ \. timpl_closure_set (set FP) (set TI) \\<^sub>c t" and OCC: "\t \ timpl_closure_set (set FP) (set TI). \f \ funs_term t. is_Abs f \ f \ Abs ` set OCC" "timpl_closure_set (absc ` set OCC) (set TI) \ absc ` set OCC" "\\<^sub>v\<^sub>a\<^sub>l\<^sub>s \ \ \ absc ` set OCC" and TI: "set TI = {(a,b) \ (set TI)\<^sup>+. a \ b}" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" shows "\x \ set (transaction_fresh T). (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) = absc {}" (is ?A) and "\t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T). intruder_synth_mod_timpls FP TI (t \ (\ \\<^sub>s \ \\<^sub>s \) \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \))" (is ?B) and "\x \ fv_transaction T - set (transaction_fresh T). \s. select\Var x,Fun (Set s) []\ \ set (unlabel (transaction_checks T)) \ (\ss. (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) = absc ss \ s \ ss)" (is ?C) and "\x \ fv_transaction T - set (transaction_fresh T). \s. \Var x in Fun (Set s) []\ \ set (unlabel (transaction_checks T)) \ (\ss. (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) = absc ss \ s \ ss)" (is ?D) and "\x \ fv_transaction T - set (transaction_fresh T). \s. \Var x not in Fun (Set s) []\ \ set (unlabel (transaction_checks T)) \ (\ss. (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) = absc ss \ s \ ss)" (is ?E) and "\x \ fv_transaction T - set (transaction_fresh T). \\<^sub>v x = TAtom Value \ (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) \ absc ` set OCC" (is ?F) proof - let ?T' = "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" let ?\ = "\ \\<^sub>s \ \\<^sub>s \" define a0 where "a0 \ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" define a0' where "a0' \ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@?T') \)" define fv_AT' where "fv_AT' \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@?T')" note T_adm = bspec[OF P(1) T] note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] note T_adm' = admissible_transaction_is_wellformed_transaction(2-4)[OF T_adm] note \_empty = admissible_transaction_decl_subst_empty[OF T_adm \] hence \_elim: "?\ = \ \\<^sub>s \" by simp have \': "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "\n. PubConst Value n \ \(funs_term ` (\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" "\n. Abs n \ \(funs_term ` (\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" "\n. PubConst Value n \ \(funs_term ` (\ ` fv_AT'))" "\n. Abs n \ \(funs_term ` (\ ` fv_AT'))" using \ admissible_transaction_occurs_checks_prop'[ OF \_reach welltyped_constraint_model_prefix[OF \] P P_occ] admissible_transaction_occurs_checks_prop'[ OF reachable_constraints.step[OF \_reach T \ \ \] \ P P_occ] unfolding welltyped_constraint_model_def constraint_model_def is_Val_def is_Abs_def fv_AT'_def by (meson,meson,meson,metis,metis,metis,metis) have T_no_pubconsts: "\n. PubConst Value n \ \(funs_term ` trms_transaction T)" and T_no_abss: "\n. Abs n \ \(funs_term ` trms_transaction T)" and T_fresh_vars_value_typed: "\x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" and T_fv_const_typed: "\x \ fv_transaction T. \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" using protocol_transaction_vars_TAtom_typed protocol_transactions_no_pubconsts protocol_transactions_no_abss funs_term_Fun_subterm P T by (fast, fast, fast, fast) have wt_\\\: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \ \\<^sub>s \)" using \'(2) wt_subst_compose transaction_fresh_subst_wt[OF \] transaction_renaming_subst_wt[OF \] by blast have 1: "?\ y \ \ = \ y" when "y \ set (transaction_fresh T)" for y using transaction_fresh_subst_grounds_domain[OF \ that] subst_compose[of \ \ y] unfolding \_elim by (simp add: subst_ground_ident) have 2: "?\ y \ \ \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" when "y \ set (transaction_fresh T)" for y using 1[OF that] that \ unfolding transaction_fresh_subst_def by auto have 3: "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. \\<^sub>v x = TAtom Value \ (\B. prefix B \ \ x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t B \ \ x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B))" by (metis welltyped_constraint_model_prefix[OF \] constraint_model_Value_var_in_constr_prefix[OF \_reach _ P P_occ]) have 4: "\n. ?\ y \ \ = Fun (Val n) []" when "y \ fv_transaction T" "\\<^sub>v y = TAtom Value" for y using transaction_var_becomes_Val[ OF reachable_constraints.step[OF \_reach T \ \ \] \ \ \ \ P P_occ T] that T_fv_const_typed \\<^sub>v_TAtom''[of y] by metis have \_is_T_model: "strand_sem_stateful (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \) (set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)) (unlabel ?T') \" using \ unlabel_append[of \ ?T'] db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t[of "unlabel \" \ "[]"] strand_sem_append_stateful[of "{}" "{}" "unlabel \" "unlabel ?T'" \] by (simp add: welltyped_constraint_model_def constraint_model_def db\<^sub>s\<^sub>s\<^sub>t_def) have T_rcv_no_val_bvars: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ subst_domain ?\ = {}" using admissible_transaction_no_bvars[OF T_adm] bvars_transaction_unfold[of T] by blast show ?A proof fix y assume y: "y \ set (transaction_fresh T)" then obtain yn where yn: "(\ \\<^sub>s \ \\<^sub>s \) y \ \ = Fun (Val yn) []" "Fun (Val yn) [] \ subst_range \" by (metis \_elim T_fresh_vars_value_typed transaction_fresh_subst_sends_to_val'[OF \]) { \ \since \y\ is fresh \(\ \\<^sub>s \ \\<^sub>s \) y \ \\ cannot be part of the database state of \\ \\\ fix t' s assume t': "insert\t',s\ \ set (unlabel \)" "t' \ \ = Fun (Val yn) []" then obtain z where t'_z: "t' = Var z" using 2[OF y] yn(1) by (cases t') auto hence z: "z \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" "\ z = (\ \\<^sub>s \ \\<^sub>s \) y \ \" using t' yn(1) by force+ hence z': "\\<^sub>v z = TAtom Value" by (metis \.simps(1) \_consts_simps(2) t'(2) t'_z wt_subst_trm'' \'(2)) obtain B where B: "prefix B \" "\ z \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" using z z' 3 by fastforce hence "\t \ subst_range \. t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" using transaction_fresh_subst_range_fresh(1)[OF \] trms\<^sub>s\<^sub>s\<^sub>t_unlabel_prefix_subset(1)[of B] unfolding prefix_def by fast hence False using B(2) 1[OF y] z yn(1) by (metis subst_imgI term.distinct(1)) } hence "\s. (?\ y \ \, s) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" using db\<^sub>s\<^sub>s\<^sub>t_in_cases[of "?\ y \ \" _ "unlabel \" \ "[]"] yn(1) by (force simp add: db\<^sub>s\<^sub>s\<^sub>t_def) thus "?\ y \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) = absc {}" using to_abs_empty_iff_notin_db[of yn "db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ []"] yn(1) by (simp add: db\<^sub>s\<^sub>s\<^sub>t_def) qed show receives_covered: ?B proof fix t assume t: "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" hence t_in_T: "t \ trms_transaction T" using trms\<^sub>s\<^sub>s\<^sub>t_unlabel_prefix_subset(1)[of "transaction_receive T"] unfolding transaction_strand_def by fast obtain ts where ts: "t \ set ts" "receive\ts\ \ set (unlabel (transaction_receive T))" using t wellformed_transaction_send_receive_trm_cases(1)[OF T_wf] by blast have t_rcv: "receive\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t ?\\ \ set (unlabel (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?\))" using subst_lsst_unlabel_member[of "receive\ts\" "transaction_receive T" ?\] trms\<^sub>s\<^sub>s\<^sub>t_in[OF t] ts by fastforce have "list_all (\t. ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ t \ ?\ \ \) ts" using wellformed_transaction_sem_receives[OF T_wf \_is_T_model t_rcv] unfolding list_all_iff by fastforce hence *: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ t \ ?\ \ \" using ts(1) unfolding list_all_iff by fast have t_fv: "fv (t \ ?\) \ fv_AT'" using fv\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \"] unlabel_append[of \] fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?\"] ts(1) t_rcv fv_transaction_subst_unfold[of T ?\] unfolding fv_AT'_def by force have **: "\t \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a0. timpl_closure_set (set FP) (set TI) \\<^sub>c t" using FP(3) by (auto simp add: a0_def abs_intruder_knowledge_def) note lms1 = pubval_terms_subst[OF _ pubval_terms_subst_range_disj[ OF transaction_fresh_subst_has_no_pubconsts_abss(1)[OF \], of t]] pubval_terms_subst[OF _ pubval_terms_subst_range_disj[ OF transaction_renaming_subst_has_no_pubconsts_abss(1)[OF \], of "t \ \"]] note lms2 = abs_terms_subst[OF _ abs_terms_subst_range_disj[ OF transaction_fresh_subst_has_no_pubconsts_abss(2)[OF \], of t]] abs_terms_subst[OF _ abs_terms_subst_range_disj[ OF transaction_renaming_subst_has_no_pubconsts_abss(2)[OF \], of "t \ \"]] have "t \ (\T\set P. trms_transaction T)" "fv (t \ \ \\<^sub>s \ \ \) = {}" using t_in_T T interpretation_grounds[OF \'(1)] by fast+ moreover have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \ \\<^sub>s \))" using wf_trm_subst_rangeI[of \, OF transaction_fresh_subst_is_wf_trm[OF \]] wf_trm_subst_rangeI[of \, OF transaction_renaming_subst_is_wf_trm[OF \]] wf_trms_subst_compose[of \ \, THEN wf_trms_subst_compose[OF _ \'(3)]] by blast moreover have "t \ pubval_terms" using t_in_T T_no_pubconsts funs_term_Fun_subterm unfolding is_PubConstValue_def is_PubConst_def by fastforce hence "t \ ?\ \ pubval_terms" using lms1 T_fresh_vars_value_typed unfolding \_elim by auto hence "t \ ?\ \ \ \ pubval_terms" using \'(6) t_fv pubval_terms_subst'[of "t \ ?\" \] by auto moreover have "t \ abs_terms" using t_in_T T_no_abss funs_term_Fun_subterm unfolding is_Abs_def by force hence "t \ ?\ \ abs_terms" using lms2 T_fresh_vars_value_typed unfolding \_elim by auto hence "t \ ?\ \ \ \ abs_terms" using \'(7) t_fv abs_terms_subst'[of "t \ ?\" \] by auto ultimately have ***: "t \ \ \\<^sub>s \ \\<^sub>s \ \ \ \ GSMP (\T\set P. trms_transaction T) - (pubval_terms \ abs_terms)" using SMP.Substitution[OF SMP.MP[of t "\T\set P. trms_transaction T"], of "\ \\<^sub>s \ \\<^sub>s \"] subst_subst_compose[of t ?\ \] wt_\\\ unfolding GSMP_def \_elim by fastforce have ****: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ GSMP (\T\set P. trms_transaction T) - (pubval_terms \ abs_terms)" using reachable_constraints_GSMP_no_pubvals_abss[OF \_reach P \'(1-5)] ik\<^sub>s\<^sub>s\<^sub>t_trms\<^sub>s\<^sub>s\<^sub>t_subset[of "unlabel \"] by blast show "intruder_synth_mod_timpls FP TI (t \ ?\ \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \))" using deduct_FP_if_deduct[OF **** ** * ***] deducts_eq_if_analyzed[OF FP(1)] intruder_synth_mod_timpls_is_synth_timpl_closure_set[OF TI, of FP] unfolding a0_def by force qed show ?C proof (intro ballI allI impI) fix y s assume y: "y \ fv_transaction T - set (transaction_fresh T)" and s: "select\Var y, Fun (Set s) []\ \ set (unlabel (transaction_checks T))" hence "select\Var y, Fun (Set s) []\ \ set (unlabel (transaction_strand T))" unfolding transaction_strand_def unlabel_def by auto hence y_val: "\\<^sub>v y = TAtom Value" using transaction_selects_are_Value_vars[OF T_wf T_adm'(1)] by fastforce have "select\?\ y, Fun (Set s) []\ \ set (unlabel (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?\))" using subst_lsst_unlabel_member[OF s] by fastforce hence "((\ \\<^sub>s \ \\<^sub>s \) y \ \, Fun (Set s) []) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" using wellformed_transaction_sem_pos_checks[ OF T_wf \_is_T_model, of Assign "(\ \\<^sub>s \ \\<^sub>s \) y" "Fun (Set s) []"] by simp thus "\ss. (\ \\<^sub>s \ \\<^sub>s \) y \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) = absc ss \ s \ ss" using to_abs_alt_def[of "db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \"] 4[of y] y y_val by auto qed show ?D proof (intro ballI allI impI) fix y s assume y: "y \ fv_transaction T - set (transaction_fresh T)" and s: "\Var y in Fun (Set s) []\ \ set (unlabel (transaction_checks T))" hence "\Var y in Fun (Set s) []\ \ set (unlabel (transaction_strand T))" unfolding transaction_strand_def unlabel_def by auto hence y_val: "\\<^sub>v y = TAtom Value" using transaction_inset_checks_are_Value_vars[OF T_adm] by fastforce have "\?\ y in Fun (Set s) []\ \ set (unlabel (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?\))" using subst_lsst_unlabel_member[OF s] by fastforce hence "(?\ y \ \, Fun (Set s) []) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" using wellformed_transaction_sem_pos_checks[ OF T_wf \_is_T_model, of Check "?\ y" "Fun (Set s) []"] by simp thus "\ss. ?\ y \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) = absc ss \ s \ ss" using to_abs_alt_def[of "db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \"] 4[of y] y y_val by auto qed show ?E proof (intro ballI allI impI) fix y s assume y: "y \ fv_transaction T - set (transaction_fresh T)" and s: "\Var y not in Fun (Set s) []\ \ set (unlabel (transaction_checks T))" hence "\Var y not in Fun (Set s) []\ \ set (unlabel (transaction_strand T))" unfolding transaction_strand_def unlabel_def by auto hence y_val: "\\<^sub>v y = TAtom Value" using transaction_notinset_checks_are_Value_vars(1)[ OF T_adm, of "[]" "[]" "[(Var y, Fun (Set s) [])]" "Var y" "Fun (Set s) []"] by force have "\?\ y not in Fun (Set s) []\ \ set (unlabel (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?\))" using subst_lsst_unlabel_member[OF s] by fastforce hence "(?\ y \ \, Fun (Set s) []) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" using wellformed_transaction_sem_neg_checks'(2)[ OF T_wf \_is_T_model, of "[]" "?\ y" "Fun (Set s) []"] by simp moreover have "list_all admissible_transaction_updates P" using Ball_set[of P "admissible_transaction"] P(1) Ball_set[of P admissible_transaction_updates] admissible_transaction_is_wellformed_transaction(3) by fast moreover have "list_all wellformed_transaction P" using P(1) Ball_set[of P "admissible_transaction"] Ball_set[of P wellformed_transaction] admissible_transaction_is_wellformed_transaction(1) by blast ultimately have "((\ \\<^sub>s \ \\<^sub>s \) y \ \, Fun (Set s) S) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" for S using reachable_constraints_db\<^sub>l\<^sub>s\<^sub>s\<^sub>t_set_args_empty[OF \_reach] unfolding admissible_transaction_updates_def by auto thus "\ss. (\ \\<^sub>s \ \\<^sub>s \) y \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) = absc ss \ s \ ss" using to_abs_alt_def[of "db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \"] 4[of y] y y_val by auto qed show ?F proof (intro ballI impI) fix y assume y: "y \ fv_transaction T - set (transaction_fresh T)" "\\<^sub>v y = TAtom Value" - then obtain yn where yn: "?\ y \ \ = Fun (Val yn) []" using 4 by moura + then obtain yn where yn: "?\ y \ \ = Fun (Val yn) []" using 4 by blast hence y_abs: "?\ y \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) = Fun (Abs (\\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) yn)) []" by simp have "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ (y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) \ (\t s. select\t,s\ \ set (unlabel (transaction_checks T)) \ y \ fv t \ fv s))" using admissible_transaction_fv_in_receives_or_selects[OF T_adm] y by blast thus "?\ y \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) \ absc ` set OCC" proof assume "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" then obtain ts where ts: "receive\ts\ \ set (unlabel (transaction_receive T))" "y \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" using wellformed_transaction_unlabel_cases(1)[OF T_wf] by (force simp add: unlabel_def) have *: "?\ y \ \ \ subterms\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t ?\ \\<^sub>s \)" "list_all (\t. timpl_closure_set (set FP) (set TI) \\<^sub>c t \ ?\ \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)) ts" using ts fv_subterms_substI[of y _ "?\ \\<^sub>s \"] subst_compose[of ?\ \ y] subterms_subst_subset[of "?\ \\<^sub>s \"] receives_covered unfolding intruder_synth_mod_timpls_is_synth_timpl_closure_set[OF TI, symmetric] list_all_iff by fastforce+ have "Abs (\\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) yn) \ \(funs_term ` (timpl_closure_set (set FP) (set TI)))" using * y_abs abs_subterms_in[of "?\ y \ \" _ "\\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)"] ideduct_synth_priv_fun_in_ik[ OF _ funs_term_Fun_subterm'[of "Abs (\\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) yn)" "[]"]] unfolding list_all_iff by fastforce hence "?\ y \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) \ subterms\<^sub>s\<^sub>e\<^sub>t (timpl_closure_set (set FP) (set TI))" using y_abs wf_trms_subterms[OF timpl_closure_set_wf_trms[OF FP(2), of "set TI"]] funs_term_Fun_subterm[of "Abs (\\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) yn)"] unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by fastforce hence "funs_term (?\ y \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)) \ (\t \ timpl_closure_set (set FP) (set TI). funs_term t)" using funs_term_subterms_eq(2)[of "timpl_closure_set (set FP) (set TI)"] by blast thus ?thesis using y_abs OCC(1) by fastforce next assume "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) \ (\t s. select\t,s\ \ set (unlabel (transaction_checks T)) \ y \ fv t \ fv s)" then obtain t t' where "select\t,t'\ \ set (unlabel (transaction_checks T))" "y \ fv t \ fv t'" by blast then obtain l s where "(l,select\Var y, Fun (Set s) []\) \ set (transaction_checks T)" using admissible_transaction_strand_step_cases(2)[OF T_adm] unfolding unlabel_def by fastforce then obtain U where U: "prefix (U@[(l,select\Var y, Fun (Set s) []\)]) (transaction_checks T)" using in_set_conv_decomp[of "(l, select\Var y,Fun (Set s) []\)" "transaction_checks T"] by (auto simp add: prefix_def) hence "select\Var y, Fun (Set s) []\ \ set (unlabel (transaction_checks T))" by (force simp add: prefix_def unlabel_def) hence "select\?\ y, Fun (Set s) []\ \ set (unlabel (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?\))" using subst_lsst_unlabel_member by fastforce hence "(Fun (Val yn) [], Fun (Set s) []) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" using yn wellformed_transaction_sem_pos_checks[ OF T_wf \_is_T_model, of Assign "?\ y" "Fun (Set s) []"] by fastforce hence "Fun (Val yn) [] \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>e\<^sub>t \" using db\<^sub>s\<^sub>s\<^sub>t_in_cases[of "Fun (Val yn) []"] by (fastforce simp add: db\<^sub>s\<^sub>s\<^sub>t_def) thus ?thesis using OCC(3) yn abs_in[of "Fun (Val yn) []" _ "\\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)"] unfolding abs_value_constants_def by (metis (mono_tags, lifting) mem_Collect_eq subsetCE) qed qed qed lemma transaction_prop4: assumes \_reach: "\ \ reachable_constraints P" and T: "T \ set P" and \: "welltyped_constraint_model \ (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" and \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and x: "x \ set (transaction_fresh T)" and y: "y \ fv_transaction T - set (transaction_fresh T)" "\\<^sub>v y = TAtom Value" shows "(\ \\<^sub>s \ \\<^sub>s \) x \ \ \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\ \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" (is ?A) and "(\ \\<^sub>s \ \\<^sub>s \) y \ \ \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\ \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" (is ?B) proof - let ?T' = "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" from \ have \': "welltyped_constraint_model \ \" using welltyped_constraint_model_prefix by auto note T_P_adm = bspec[OF P(1)] note T_adm = bspec[OF P(1) T] note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] have be: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = {}" using T_P_adm \_reach reachable_constraints_no_bvars admissible_transaction_no_bvars(2) by blast have T_no_bvars: "fv_transaction T = vars_transaction T" using admissible_transaction_no_bvars[OF T_adm] by simp note \_empty = admissible_transaction_decl_subst_empty[OF T_adm \] have \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" by (metis \ welltyped_constraint_model_def) have x_val: "fst x = TAtom Value" using x admissible_transactionE(14)[OF T_adm] unfolding list_all_iff by blast obtain xn where xn: "\ x = Fun (Val xn) []" using x_val transaction_fresh_subst_sends_to_val[OF \ x] \\<^sub>v_TAtom''(2)[of x] by meson hence xnxn: "(\ \\<^sub>s \ \\<^sub>s \) x = Fun (Val xn) []" unfolding subst_compose_def \_empty by auto from xn xnxn have a0: "(\ \\<^sub>s \ \\<^sub>s \) x \ \ = Fun (Val xn) []" by auto have b0: "\\<^sub>v x = TAtom Value" using P x T protocol_transaction_vars_TAtom_typed(3) by metis note 0 = a0 b0 have \_x_nin_A: "\ x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" proof - have "\ x \ subst_range \" by (metis transaction_fresh_subst_sends_to_val[OF \ x b0]) moreover have "\t \ subst_range \. t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using \ transaction_fresh_subst_def[of \ T "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] by blast ultimately show ?thesis by auto qed have *: "y \ set (transaction_fresh T)" using assms by auto have **: "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ (y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) \ (\t s. select\t,s\ \ set (unlabel (transaction_checks T)) \ y \ fv t \ fv s))" using * y admissible_transaction_fv_in_receives_or_selects[OF T_adm] by blast have y_fv: "y \ fv_transaction T" using y fv_transaction_unfold by blast have y_val: "fst y = TAtom Value" using y(2) \\<^sub>v_TAtom''(2) by blast have "\ x \ \ \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\ \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" proof (rule ccontr) assume "\\ x \ \ \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\ \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" then have a: "\ x \ \ \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\ \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" by auto then have \_x_I_in_A: "\ x \ \ \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>e\<^sub>t \" using reachable_constraints_subterms_subst[OF \_reach \' P] by blast have "\u. u \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ \ u = \ x" proof - from \_x_I_in_A have "\tu. tu \ \ (subterms ` (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ tu \ \ = \ x \ \" by force then obtain tu where tu: "tu \ \ (subterms ` (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ tu \ \ = \ x \ \" by auto then have "tu \ \ x" using \_x_nin_A by auto moreover have "tu \ \ = \ x" using tu by (simp add: xn) ultimately have "\u. tu = Var u" unfolding xn by (cases tu) auto then obtain u where "tu = Var u" by auto have "u \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ \ u = \ x" proof - have "u \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using \tu = Var u\ tu var_subterm_trms\<^sub>s\<^sub>s\<^sub>t_is_vars\<^sub>s\<^sub>s\<^sub>t by fastforce then have "u \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using be vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel \"] by blast moreover have "\ u = \ x" using \tu = Var u\ \tu \ \ = \ x\ by auto ultimately show ?thesis by auto qed then show "\u. u \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ \ u = \ x" by metis qed then obtain u where u: "u \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" "\ u = \ x" by auto then have u_TA: "\\<^sub>v u = TAtom Value" using P(1) T x_val \\<^sub>v_TAtom''(2)[of x] wt_subst_trm''[OF \_wt, of "Var u"] wt_subst_trm''[of \ "Var x"] transaction_fresh_subst_wt[OF \] protocol_transaction_vars_TAtom_typed(3) by force have "\B. prefix B \ \ u \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t B \ \ u \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" using u u_TA by (metis welltyped_constraint_model_prefix[OF \] constraint_model_Value_var_in_constr_prefix[OF \_reach _ P P_occ]) then obtain B where "prefix B \ \ u \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t B \ \ u \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" by blast moreover have "\(subterms ` trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t xs) \ \(subterms ` trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t ys)" when "prefix xs ys" for xs ys::"('fun,'atom,'sets,'lbl) prot_strand" using that subterms\<^sub>s\<^sub>e\<^sub>t_mono trms\<^sub>s\<^sub>s\<^sub>t_mono unlabel_mono set_mono_prefix by metis ultimately have "\ u \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by blast then have "\ x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using u by auto then show "False" using \_x_nin_A by auto qed then show ?A using eval_term.simps(1)[of _ x \] unfolding subst_compose_def xn \_empty by auto from ** show ?B proof define T' where "T' \ transaction_receive T" define \ where "\ \ \ \\<^sub>s \ \\<^sub>s \" assume y: "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" hence "Var y \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" by (metis T'_def fv\<^sub>s\<^sub>s\<^sub>t_is_subterm_trms\<^sub>s\<^sub>s\<^sub>t) then obtain z where z: "z \ set (unlabel T')" "Var y \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p z)" by (induct T') auto have "is_Receive z" using Ball_set[of "unlabel T'" is_Receive] z(1) admissible_transaction_is_wellformed_transaction(1)[OF T_adm] unfolding wellformed_transaction_def T'_def by blast then obtain tys where "z = receive\tys\" by (cases z) auto hence tys: "receive\tys \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\ \ set (unlabel (T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" "\ y \ subterms\<^sub>s\<^sub>e\<^sub>t (set tys \\<^sub>s\<^sub>e\<^sub>t \)" using z subst_mono unfolding subst_apply_labeled_stateful_strand_def unlabel_def by force+ hence y_deduct: "list_all (\t. ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \ \ \) tys" using transaction_receive_deduct[OF T_wf _ \ \ \] \ unfolding T'_def \_def welltyped_constraint_model_def list_all_iff by auto obtain ty where ty: "ty \ set tys" "\ y \ ty \ \" "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ ty \ \ \ \" using tys y_deduct unfolding list_all_iff by blast obtain zn where zn: "(\ \\<^sub>s \ \\<^sub>s \) y \ \ = Fun (Val zn) []" using transaction_var_becomes_Val[ OF reachable_constraints.step[OF \_reach T \ \ \] \ \ \ \ P P_occ T y_fv y_val] by metis have "(\ \\<^sub>s \ \\<^sub>s \) y \ \ \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \)" using ty tys(2) y_deduct private_fun_deduct_in_ik[of _ _ "Val zn"] by (metis \_def zn subst_mono public.simps(3)) thus ?B using ik\<^sub>s\<^sub>s\<^sub>t_subst[of "unlabel \" \] unlabel_subst[of \ \] subterms\<^sub>s\<^sub>e\<^sub>t_mono[OF ik\<^sub>s\<^sub>s\<^sub>t_trms\<^sub>s\<^sub>s\<^sub>t_subset[of "unlabel (\ \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)"]] by fastforce next assume y': "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) \ (\t s. select\t,s\ \ set (unlabel (transaction_checks T)) \ y \ fv t \ fv s)" then obtain s where s: "select\Var y,s\ \ set (unlabel (transaction_checks T))" "fst y = TAtom Value" using admissible_transaction_strand_step_cases(1,2)[OF T_adm] by fastforce obtain z zn where zn: "(\ \\<^sub>s \ \\<^sub>s \) y = Var z" "\ z = Fun (Val zn) []" using transaction_var_becomes_Val[ OF reachable_constraints.step[OF \_reach T \ \ \] \ \ \ \ P P_occ T y_fv s(2)] transaction_decl_fresh_renaming_substs_range(4)[OF \ \ \ _ *] transaction_decl_subst_empty_inv[OF \[unfolded \_empty]] by auto have transaction_selects_db_here: "\n s. select\Var (TAtom Value, n), Fun (Set s) []\ \ set (unlabel (transaction_checks T)) \ (\ (TAtom Value, n) \ \, Fun (Set s) []) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" using transaction_selects_db[OF T_adm _ \ \ \] \ unfolding welltyped_constraint_model_def by auto have "\n. y = (Var Value, n)" using T \\<^sub>v_TAtom_inv(2) y_fv y(2) by blast moreover have "admissible_transaction_checks T" using admissible_transaction_is_wellformed_transaction(2)[OF T_adm] by blast then have "is_Fun_Set (the_set_term (select\Var y,s\))" using s unfolding admissible_transaction_checks_def by auto then have "\ss. s = Fun (Set ss) []" using is_Fun_Set_exi by auto ultimately obtain n ss where nss: "y = (TAtom Value, n)" "s = Fun (Set ss) []" by auto then have "select\Var (TAtom Value, n), Fun (Set ss) []\ \ set (unlabel (transaction_checks T))" using s by auto then have in_db: "(\ (TAtom Value, n) \ \, Fun (Set ss) []) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" using transaction_selects_db_here[of n ss] by auto have "(\ z, s) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" proof - have "(\ y \ \, s) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" using in_db nss by auto moreover have "\ y = Var z" using zn \_empty * \ by (metis (no_types, opaque_lifting) subst_compose_def subst_imgI subst_to_var_is_var term.distinct(1) transaction_fresh_subst_def var_comp(2)) then have "\ y \ \ = \ z" by auto ultimately show "(\ z, s) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" by auto qed then have "\t' s'. insert\t',s'\ \ set (unlabel \) \ \ z = t' \ \ \ s = s' \ \" using db\<^sub>s\<^sub>s\<^sub>t_in_cases[of "\ z" s "unlabel \" \ "[]"] unfolding db\<^sub>s\<^sub>s\<^sub>t_def by auto then obtain t' s' where t's': "insert\t',s'\ \ set (unlabel \) \ \ z = t' \ \ \ s = s' \ \" by auto then have "t' \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by force then have "t' \ \ \ (subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \\<^sub>s\<^sub>e\<^sub>t \" by auto then have "\ z \ (subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \\<^sub>s\<^sub>e\<^sub>t \" using t's' by auto then have "\ z \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\ \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using reachable_constraints_subterms_subst[ OF \_reach welltyped_constraint_model_prefix[OF \] P] by auto then show ?B using zn(1) by simp qed qed lemma transaction_prop5: fixes T \ \ \ \ \ T' a0 a0' \ defines "T' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" and "a0 \ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" and "a0' \ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \)" and "\ \ \\ x. if fst x = TAtom Value then (absc \ \) x else Var x" assumes \_reach: "\ \ reachable_constraints P" and T: "T \ set P" and \: "welltyped_constraint_model \ (\@T')" and \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and FP: "analyzed (timpl_closure_set (set FP) (set TI))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set FP)" "\t \ \\<^sub>i\<^sub>k \ \. timpl_closure_set (set FP) (set TI) \\<^sub>c t" and OCC: "\t \ timpl_closure_set (set FP) (set TI). \f \ funs_term t. is_Abs f \ f \ Abs ` set OCC" "timpl_closure_set (absc ` set OCC) (set TI) \ absc ` set OCC" "\\<^sub>v\<^sub>a\<^sub>l\<^sub>s \ \ \ absc ` set OCC" and TI: "set TI = {(a,b) \ (set TI)\<^sup>+. a \ b}" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and step: "list_all (transaction_check (FP, OCC, TI)) P" shows "\\ \ abs_substs_fun ` set (transaction_check_comp (\_ _. True) (FP, OCC, TI) T). \x \ fv_transaction T. \\<^sub>v x = TAtom Value \ (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ a0 = absc (\ x) \ (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ a0' = absc (absdbupd (unlabel (transaction_updates T)) x (\ x))" proof - define comp0 where "comp0 \ abs_substs_fun ` set (transaction_check_comp (\_ _. True) (FP, OCC, TI) T)" define check0 where "check0 \ transaction_check (FP, OCC, TI) T" define upd where "upd \ \\ x. absdbupd (unlabel (transaction_updates T)) x (\ x)" define b0 where "b0 \ \x. THE b. absc b = (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ a0" note all_defs = comp0_def check0_def a0_def a0'_def upd_def b0_def \_def T'_def have \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by (metis reachable_constraints_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s admissible_transactions_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s P(1) \_reach) have \_interp: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and \_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (metis \ welltyped_constraint_model_def constraint_model_def, metis \ welltyped_constraint_model_def, metis \ welltyped_constraint_model_def constraint_model_def) have \_is_T_model: "strand_sem_stateful (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \) (set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)) (unlabel T') \" using \ unlabel_append[of \ T'] db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t[of "unlabel \" \ "[]"] strand_sem_append_stateful[of "{}" "{}" "unlabel \" "unlabel T'" \] by (simp add: welltyped_constraint_model_def constraint_model_def db\<^sub>s\<^sub>s\<^sub>t_def) note T_adm = bspec[OF P(1) T] note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] have T_no_bvars: "fv_transaction T = vars_transaction T" "bvars_transaction T = {}" using admissible_transaction_no_bvars[OF T_adm] by simp_all have T_vars_const_typed: "\x \ fv_transaction T. \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" and T_fresh_vars_value_typed: "\x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" using T P protocol_transaction_vars_TAtom_typed(2,3)[of T] by simp_all note \_empty = admissible_transaction_decl_subst_empty[OF T_adm \] have wt_\\\: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \ \\<^sub>s \ \\<^sub>s \)" and wt_\\: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \ \\<^sub>s \)" using \_wt wt_subst_compose transaction_decl_fresh_renaming_substs_wt[OF \ \ \] by (blast, blast) have T_vars_vals: "\x \ fv_transaction T. \n. (\ \\<^sub>s \ \\<^sub>s \) x \ \ = Fun (Val n) []" proof fix x assume x: "x \ fv_transaction T" have "\n. (\ \\<^sub>s \) x \ \ = Fun (Val n) []" proof (cases "x \ subst_domain \") case True then obtain n where "\ x = Fun (Val n) []" using transaction_fresh_subst_sends_to_val[OF \] transaction_fresh_subst_domain[OF \] T_fresh_vars_value_typed by metis thus ?thesis by (simp add: subst_compose_def) next case False hence *: "(\ \\<^sub>s \) x = \ x" by (auto simp add: subst_compose_def) obtain y where y: "\\<^sub>v x = \\<^sub>v y" "\ x = Var y" using transaction_renaming_subst_wt[OF \] transaction_renaming_subst_is_renaming(1)[OF \] by (metis \.simps(1) prod.exhaust wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) hence "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)" using x * T_no_bvars(2) unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \"] fv\<^sub>s\<^sub>s\<^sub>t_subst_fv_subset[of x "unlabel (transaction_strand T)" "\ \\<^sub>s \"] by auto hence "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \))" using fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \"] fv\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \"] unlabel_append[of \] by auto thus ?thesis using x y * T P (* T_vars_const_typed *) constraint_model_Value_term_is_Val[ OF reachable_constraints.step[OF \_reach T \ \ \] \[unfolded T'_def] P P_occ, of y] admissible_transaction_Value_vars[of T] \_empty by simp qed thus "\n. (\ \\<^sub>s \ \\<^sub>s \) x \ \ = Fun (Val n) []" using \_empty by simp qed have T_vars_absc: "\x \ fv_transaction T. \!n. (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ a0 = absc n" using T_vars_vals by fastforce hence "(absc \ b0) x = (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ a0" when "x \ fv_transaction T" for x using that unfolding b0_def by fastforce hence T_vars_absc': "t \ (absc \ b0) = t \ (\ \\<^sub>s \ \\<^sub>s \) \ \ \\<^sub>\ a0" when "fv t \ fv_transaction T" "\n T. Fun (Val n) T \ subterms t" for t using that(1) abs_term_subst_eq'[OF _ that(2), of "\ \\<^sub>s \ \\<^sub>s \ \\<^sub>s \" a0 "absc \ b0"] subst_compose[of "\ \\<^sub>s \ \\<^sub>s \" \] subst_subst_compose[of t "\ \\<^sub>s \ \\<^sub>s \" \] by fastforce have "\\ \ comp0. \x \ fv_transaction T. fst x = TAtom Value \ b0 x = \ x" proof - let ?C = "set (unlabel (transaction_checks T))" let ?xs = "fv_transaction T - set (transaction_fresh T)" note * = transaction_prop3[OF \_reach T \[unfolded T'_def] \ \ \ FP OCC TI P P_occ] have **: "\x \ set (transaction_fresh T). b0 x = {}" "\t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T). intruder_synth_mod_timpls FP TI (t \ \ b0)" (is ?B) proof - show ?B proof (intro ballI impI) fix t assume t: "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" hence t': "fv t \ fv_transaction T" "\n T. Fun (Val n) T \ subterms t" using trms_transaction_unfold[of T] vars_transaction_unfold[of T] trms\<^sub>s\<^sub>s\<^sub>t_fv_vars\<^sub>s\<^sub>s\<^sub>t_subset[of t "unlabel (transaction_strand T)"] admissible_transactions_no_Value_consts'[OF T_adm] wellformed_transaction_send_receive_fv_subset(1)[OF T_wf t(1)] by blast+ have "intruder_synth_mod_timpls FP TI (t \ (absc \ b0))" using t(1) t' *(2) T_vars_absc' by (metis a0_def) moreover have "(absc \ b0) x = (\ b0) x" when "x \ fv t" for x using that T P admissible_transaction_Value_vars[of T] \fv t \ fv_transaction T\ \\<^sub>v_TAtom''(2)[of x] unfolding \_def by fastforce hence "t \ (absc \ b0) = t \ \ b0" using term_subst_eq[of t "absc \ b0" "\ b0"] by argo ultimately show "intruder_synth_mod_timpls FP TI (t \ \ b0)" using intruder_synth.simps[of "set FP"] by (cases "t \ \ b0") metis+ qed qed (simp add: *(1) a0_def b0_def) have ***: "\x \ ?xs. \s. select\Var x,Fun (Set s) []\ \ ?C \ s \ b0 x" "\x \ ?xs. \s. \Var x in Fun (Set s) []\ \ ?C \ s \ b0 x" "\x \ ?xs. \s. \Var x not in Fun (Set s) []\ \ ?C \ s \ b0 x" "\x \ ?xs. fst x = TAtom Value \ b0 x \ set OCC" unfolding a0_def b0_def using *(3,4) apply (force, force) using *(5) apply force using *(6) admissible_transaction_Value_vars[OF bspec[OF P T]] by force show ?thesis using transaction_check_comp_in[OF T_adm **[unfolded \_def] ***] unfolding comp0_def by metis qed hence 1: "\\ \ comp0. \x \ fv_transaction T. fst x = TAtom Value \ (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ a0 = absc (\ x)" using T_vars_absc unfolding b0_def a0_def by fastforce obtain \ where \: "\ \ comp0" "\x \ fv_transaction T. fst x = TAtom Value \ (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ a0 = absc (\ x)" - using 1 by moura + using 1 by force have 2: "\ x \ \ \\<^sub>\ \\<^sub>0 (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ D) = absc (absdbupd (unlabel A) x d)" when "\ x \ \ \\<^sub>\ \\<^sub>0 D = absc d" and "\t u. insert\t,u\ \ set (unlabel A) \ (\y s. t = Var y \ u = Fun (Set s) [])" and "\t u. delete\t,u\ \ set (unlabel A) \ (\y s. t = Var y \ u = Fun (Set s) [])" and "\y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. \ x \ \ = \ y \ \ \ x = y" and "\y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. \n. \ y \ \ = Fun (Val n) []" and x: "\ x \ \ = Fun (Val n) []" and D: "\d \ set D. \s. snd d = Fun (Set s) []" for A::"('fun,'atom,'sets,'lbl) prot_strand" and x \ D n d using that(2,3,4,5) proof (induction A rule: List.rev_induct) case (snoc a A) then obtain l b where a: "a = (l,b)" by (metis surj_pair) have IH: "\\<^sub>0 (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ D) n = absdbupd (unlabel A) x d" using snoc unlabel_append[of A "[a]"] a x by (simp del: unlabel_append) have b_prems: "\y \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p b. \ x \ \ = \ y \ \ \ x = y" "\y \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p b. \n. \ y \ \ = Fun (Val n) []" using snoc.prems(3,4) a by (simp_all add: unlabel_def) have *: "filter is_Update (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@[a] \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) = filter is_Update (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" "filter is_Update (unlabel (A@[a])) = filter is_Update (unlabel A)" when "\is_Update b" using that a by (cases b, simp_all add: dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def unlabel_def subst_apply_labeled_stateful_strand_def)+ note ** = IH a dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst_append[of A "[a]" \] note *** = * absdbupd_filter[of "unlabel (A@[a])"] absdbupd_filter[of "unlabel A"] db\<^sub>s\<^sub>s\<^sub>t_filter[of "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@[a] \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))"] db\<^sub>s\<^sub>s\<^sub>t_filter[of "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))"] note **** = **(2,3) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst_snoc[of A a \] unlabel_append[of "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" "[dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \]"] db\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "unlabel [dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \]" \ D] have "\\<^sub>0 (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@[a] \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ D) n = absdbupd (unlabel (A@[a])) x d" using ** *** proof (cases b) case (Insert t t') then obtain y s m where y: "t = Var y" "t' = Fun (Set s) []" "\ y \ \ = Fun (Val m) []" using snoc.prems(1) b_prems(2) a by (fastforce simp add: unlabel_def) hence a': "db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@[a] \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ D = List.insert ((Fun (Val m) [], Fun (Set s) [])) (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ D)" "unlabel [dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \] = [insert\\ y, Fun (Set s) []\]" "unlabel [a] = [insert\Var y, Fun (Set s) []\]" using **** Insert by simp_all show ?thesis proof (cases "x = y") case True hence "\ x \ \ = \ y \ \" by simp hence "\\<^sub>0 (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@[a] \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ D) n = insert s (\\<^sub>0 (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ D) n)" by (metis (no_types, lifting) y(3) a'(1) x dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst to_abs_list_insert') thus ?thesis using True IH a'(3) absdbupd_append[of "unlabel A"] by (simp add: unlabel_def) next case False hence "\ x \ \ \ \ y \ \" using b_prems(1) y Insert by simp hence "\\<^sub>0 (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@[a] \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ D) n = \\<^sub>0 (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ D) n" by (metis (no_types, lifting) y(3) a'(1) x dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst to_abs_list_insert) thus ?thesis using False IH a'(3) absdbupd_append[of "unlabel A"] by (simp add: unlabel_def) qed next case (Delete t t') then obtain y s m where y: "t = Var y" "t' = Fun (Set s) []" "\ y \ \ = Fun (Val m) []" using snoc.prems(2) b_prems(2) a by (fastforce simp add: unlabel_def) hence a': "db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@[a] \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ D = List.removeAll ((Fun (Val m) [], Fun (Set s) [])) (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ D)" "unlabel [dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \] = [delete\\ y, Fun (Set s) []\]" "unlabel [a] = [delete\Var y, Fun (Set s) []\]" using **** Delete by simp_all have "\s S. snd d = Fun (Set s) []" when "d \ set (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ D)" for d using snoc.prems(1,2) db\<^sub>l\<^sub>s\<^sub>s\<^sub>t_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_set_ex[OF that _ _ D] by (simp add: unlabel_def) moreover { fix t::"('fun,'atom,'sets,'lbl) prot_term" and D::"(('fun,'atom,'sets,'lbl) prot_term \ ('fun,'atom,'sets,'lbl) prot_term) list" assume "\d \ set D. \s. snd d = Fun (Set s) []" hence "removeAll (t, Fun (Set s) []) D = filter (\d. \S. d = (t, Fun (Set s) S)) D" by (induct D) auto } ultimately have a'': "List.removeAll ((Fun (Val m) [], Fun (Set s) [])) (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ D) = filter (\d. \S. d = (Fun (Val m) [], Fun (Set s) S)) (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ D)" by simp show ?thesis proof (cases "x = y") case True hence "\ x \ \ = \ y \ \" by simp hence "\\<^sub>0 (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@[a] \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ D) n = (\\<^sub>0 (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ D) n) - {s}" using y(3) a'' a'(1) x by (simp add: dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst to_abs_list_remove_all') thus ?thesis using True IH a'(3) absdbupd_append[of "unlabel A"] by (simp add: unlabel_def) next case False hence "\ x \ \ \ \ y \ \" using b_prems(1) y Delete by simp hence "\\<^sub>0 (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@[a] \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ D) n = \\<^sub>0 (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ D) n" by (metis (no_types, lifting) y(3) a'(1) x dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst to_abs_list_remove_all) thus ?thesis using False IH a'(3) absdbupd_append[of "unlabel A"] by (simp add: unlabel_def) qed qed simp_all thus ?case by (simp add: x) qed (simp add: that(1)) have 3: "x = y" when xy: "(\ \\<^sub>s \ \\<^sub>s \) x \ \ = (\ \\<^sub>s \ \\<^sub>s \) y \ \" "x \ fv_transaction T" "y \ fv_transaction T" for x y proof - have "x \ set (transaction_fresh T) \ y \ set (transaction_fresh T) \ ?thesis" using xy admissible_transaction_strand_sem_fv_ineq[OF T_adm \_is_T_model[unfolded T'_def]] by fast moreover { assume *: "x \ set (transaction_fresh T)" "y \ set (transaction_fresh T)" hence "\\<^sub>v x = TAtom Value" "\\<^sub>v y = TAtom Value" using T_fresh_vars_value_typed by (blast, blast) then obtain xn yn where "\ x = Fun (Val xn) []" "\ y = Fun (Val yn) []" using * transaction_fresh_subst_sends_to_val[OF \] by meson hence "\ x = \ y" using that(1) \_empty by (simp add: subst_compose) moreover have "inj_on \ (subst_domain \)" "x \ subst_domain \" "y \ subst_domain \" using * \ unfolding transaction_fresh_subst_def by auto ultimately have ?thesis unfolding inj_on_def by blast } moreover have False when "x \ set (transaction_fresh T)" "y \ set (transaction_fresh T)" using that(2) xy T_no_bvars admissible_transaction_Value_vars[OF bspec[OF P T], of y] transaction_prop4[OF \_reach T \[unfolded T'_def] \ \ \ P P_occ that(1), of y] by auto moreover have False when "x \ set (transaction_fresh T)" "y \ set (transaction_fresh T)" using that(1) xy T_no_bvars admissible_transaction_Value_vars[OF bspec[OF P T], of x] transaction_prop4[OF \_reach T \[unfolded T'_def] \ \ \ P P_occ that(2), of x] by fastforce ultimately show ?thesis by metis qed have 4: "\y s. t = Var y \ u = Fun (Set s) []" when "insert\t,u\ \ set (unlabel (transaction_strand T))" for t u using that admissible_transaction_strand_step_cases(3)[OF T_adm] T_wf by blast have 5: "\y s. t = Var y \ u = Fun (Set s) []" when "delete\t,u\ \ set (unlabel (transaction_strand T))" for t u using that admissible_transaction_strand_step_cases(3)[OF T_adm] T_wf by blast have 6: "\n. (\ \\<^sub>s \ \\<^sub>s \) y \ \ = Fun (Val n) []" when "y \ fv_transaction T" for y using that by (simp add: T_vars_vals) have "list_all wellformed_transaction P" "list_all admissible_transaction_updates P" using P(1) Ball_set[of P admissible_transaction'] Ball_set[of P wellformed_transaction] Ball_set[of P admissible_transaction_updates] admissible_transaction_is_wellformed_transaction(1,3) by fastforce+ hence 7: "\s. snd d = Fun (Set s) []" when "d \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" for d using that reachable_constraints_db\<^sub>l\<^sub>s\<^sub>s\<^sub>t_set_args_empty[OF \_reach] unfolding admissible_transaction_updates_def by (cases d) simp have "(\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ a0' = absc (upd \ x)" when x: "x \ fv_transaction T" "fst x = TAtom Value" for x proof - have "(\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ \\<^sub>0 (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) \ (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)) = absc (absdbupd (unlabel (transaction_strand T)) x (\ x))" using 2[of "\ \\<^sub>s \ \\<^sub>s \" x "db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \" "\ x" "transaction_strand T"] 3[OF _ x(1)] 4 5 6[OF that(1)] 6 7 x \(2) unfolding all_defs by blast thus ?thesis using x db\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \"] absdbupd_wellformed_transaction[OF T_wf] unfolding all_defs db\<^sub>s\<^sub>s\<^sub>t_def by force qed thus ?thesis using \ \\<^sub>v_TAtom''(2) unfolding all_defs by blast qed lemma transaction_prop6: fixes T \ \ \ \ \ T' a0 a0' defines "T' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" and "a0 \ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" and "a0' \ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \)" assumes \_reach: "\ \ reachable_constraints P" and T: "T \ set P" and \: "welltyped_constraint_model \ (\@T')" and \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and FP: "analyzed (timpl_closure_set (set FP) (set TI))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set FP)" "\t \ \\<^sub>i\<^sub>k \ \. timpl_closure_set (set FP) (set TI) \\<^sub>c t" and OCC: "\t \ timpl_closure_set (set FP) (set TI). \f \ funs_term t. is_Abs f \ f \ Abs ` set OCC" "timpl_closure_set (absc ` set OCC) (set TI) \ absc ` set OCC" "\\<^sub>v\<^sub>a\<^sub>l\<^sub>s \ \ \ absc ` set OCC" and TI: "set TI = {(a,b) \ (set TI)\<^sup>+. a \ b}" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and step: "list_all (transaction_check (FP, OCC, TI)) P" shows "\t \ timpl_closure_set (\\<^sub>i\<^sub>k \ \) (\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \). timpl_closure_set (set FP) (set TI) \\<^sub>c t" (is ?A) and "timpl_closure_set (\\<^sub>v\<^sub>a\<^sub>l\<^sub>s \ \) (\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \) \ absc ` set OCC" (is ?B) and "\t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T). is_Fun (t \ (\ \\<^sub>s \ \\<^sub>s \) \ \ \\<^sub>\ a0') \ timpl_closure_set (set FP) (set TI) \\<^sub>c t \ (\ \\<^sub>s \ \\<^sub>s \) \ \ \\<^sub>\ a0'" (is ?C) and "\x \ fv_transaction T. \\<^sub>v x = TAtom Value \ (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ a0' \ absc ` set OCC" (is ?D) proof - define comp0 where "comp0 \ abs_substs_fun ` set (transaction_check_comp (\_ _. True) (FP, OCC, TI) T)" define check0 where "check0 \ transaction_check (FP, OCC, TI) T" define upd where "upd \ \\ x. absdbupd (unlabel (transaction_updates T)) x (\ x)" define \ where "\ \ \\ x. if fst x = TAtom Value then (absc \ \) x else Var x" note T_adm = bspec[OF P T] note T_occ = bspec[OF P_occ T] note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] have \_prop: "\ \ x = absc (\ x)" when "\\<^sub>v x = TAtom Value" for \ x using that \\<^sub>v_TAtom''(2)[of x] unfolding \_def by simp \ \The set-membership status of all value constants in T under \\\, \\\, \\\ are covered by the check\ have 0: "\\ \ comp0. \x \ fv_transaction T. \\<^sub>v x = TAtom Value \ (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ a0 = absc (\ x) \ (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ a0' = absc (upd \ x)" using transaction_prop5[OF \_reach T \[unfolded T'_def] \ \ \ FP OCC TI P P_occ step] unfolding a0_def a0'_def T'_def upd_def comp0_def by blast \ \All set-membership changes are covered by the term implication graph\ have 1: "(\ x, upd \ x) \ (set TI)\<^sup>+" when "\ \ comp0" "\ x \ upd \ x" "x \ fv_transaction T" "x \ set (transaction_fresh T)" for x \ using T that step Ball_set[of P "transaction_check (FP, OCC, TI)"] transaction_prop1[of \ "\_ _. True" FP OCC TI T x] TI unfolding upd_def comp0_def transaction_check_def by blast \ \All set-membership changes are covered by the fixed point\ have 2: (* "\ x \ set OCC" *) "upd \ x \ set OCC" when "\ \ comp0" "x \ fv_transaction T" "fst x = TAtom Value" for x \ using T that step Ball_set[of P "transaction_check (FP, OCC, TI)"] T_adm T_occ FP OCC TI transaction_prop2[of \ "\_ _. True" FP OCC TI T x] unfolding upd_def comp0_def transaction_check_def by blast obtain \ where \: "\ \ comp0" "\x \ fv_transaction T. \\<^sub>v x = TAtom Value \ (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ a0 = absc (\ x) \ (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ a0' = absc (upd \ x)" - using 0 by moura + using 0 by force have "\x. ab = (\ x, upd \ x) \ x \ fv_transaction T - set (transaction_fresh T) \ \ x \ upd \ x" when ab: "ab \ \\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \" for ab proof - obtain a b where ab': "ab = (a,b)" by (metis surj_pair) then obtain x where x: "a \ b" "x \ fv_transaction T" "x \ set (transaction_fresh T)" "absc a = (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ a0" "absc b = (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ a0'" using ab unfolding abs_term_implications_def a0_def a0'_def T'_def by blast hence "absc a = absc (\ x)" "absc b = absc (upd \ x)" using \(2) admissible_transaction_Value_vars[OF bspec[OF P T] x(2)] by metis+ thus ?thesis using x ab' by blast qed hence \\<^sub>t\<^sub>i_TI_subset: "\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \ \ {(a,b) \ (set TI)\<^sup>+. a \ b}" using 1[OF \(1)] by blast have "timpl_closure_set (timpl_closure_set (set FP) (set TI)) (\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \) \\<^sub>c t" when t: "t \ timpl_closure_set (\\<^sub>i\<^sub>k \ \) (\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \)" for t using timpl_closure_set_is_timpl_closure_union[of "\\<^sub>i\<^sub>k \ \" "\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \"] intruder_synth_timpl_closure_set FP(3) t by blast thus ?A using ideduct_synth_mono[OF _ timpl_closure_set_mono[OF subset_refl[of "timpl_closure_set (set FP) (set TI)"] \\<^sub>t\<^sub>i_TI_subset]] timpl_closure_set_timpls_trancl_eq'[of "timpl_closure_set (set FP) (set TI)" "set TI"] unfolding timpl_closure_set_idem by force have "timpl_closure_set (\\<^sub>v\<^sub>a\<^sub>l\<^sub>s \ \) (\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \) \ timpl_closure_set (absc ` set OCC) {(a,b) \ (set TI)\<^sup>+. a \ b}" using timpl_closure_set_mono[OF _ \\<^sub>t\<^sub>i_TI_subset] OCC(3) by blast thus ?B using OCC(2) timpl_closure_set_timpls_trancl_subset' by blast have "transaction_check_post (FP, OCC, TI) T \" using T \(1) step unfolding transaction_check_def transaction_check'_def comp0_def list_all_iff by fastforce hence 3: "timpl_closure_set (set FP) (set TI) \\<^sub>c t \ \ (upd \)" when "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "is_Fun (t \ \ (upd \))" for t using that unfolding transaction_check_post_def upd_def \_def intruder_synth_mod_timpls_is_synth_timpl_closure_set[OF TI, symmetric] by fastforce have 4: "\x \ fv t. (\ \\<^sub>s \ \\<^sub>s \ \\<^sub>s \) x \\<^sub>\ a0' = \ (upd \) x" when "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" for t using wellformed_transaction_send_receive_fv_subset(2)[OF T_wf that] \(2) subst_compose[of "\ \\<^sub>s \ \\<^sub>s \" \] \_prop admissible_transaction_Value_vars[OF bspec[OF P T]] by fastforce have 5: "\n T. Fun (Val n) T \ subterms t" when "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" for t using that admissible_transactions_no_Value_consts'[OF T_adm] trms_transaction_unfold[of T] by blast show ?D using 2[OF \(1)] \(2) \\<^sub>v_TAtom''(2) unfolding a0'_def T'_def by blast show ?C using 3 abs_term_subst_eq'[OF 4 5] by simp qed lemma reachable_constraints_covered_step: fixes \::"('fun,'atom,'sets,'lbl) prot_constr" assumes \_reach: "\ \ reachable_constraints P" and T: "T \ set P" and \: "welltyped_constraint_model \ (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" and \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and FP: "analyzed (timpl_closure_set (set FP) (set TI))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set FP)" "\t \ \\<^sub>i\<^sub>k \ \. timpl_closure_set (set FP) (set TI) \\<^sub>c t" "ground (set FP)" and OCC: "\t \ timpl_closure_set (set FP) (set TI). \f \ funs_term t. is_Abs f \ f \ Abs ` set OCC" "timpl_closure_set (absc ` set OCC) (set TI) \ absc ` set OCC" "\\<^sub>v\<^sub>a\<^sub>l\<^sub>s \ \ \ absc ` set OCC" and TI: "set TI = {(a,b) \ (set TI)\<^sup>+. a \ b}" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and transactions_covered: "list_all (transaction_check (FP, OCC, TI)) P" shows "\t \ \\<^sub>i\<^sub>k (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) \. timpl_closure_set (set FP) (set TI) \\<^sub>c t" (is ?A) and "\\<^sub>v\<^sub>a\<^sub>l\<^sub>s (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) \ \ absc ` set OCC" (is ?B) proof - note step_props = transaction_prop6[OF \_reach T \ \ \ \ FP(1,2,3) OCC TI P P_occ transactions_covered] define T' where "T' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" define a0 where "a0 \ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" define a0' where "a0' \ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \)" define vals where "vals \ \S::('fun,'atom,'sets,'lbl) prot_constr. {t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t S) \\<^sub>s\<^sub>e\<^sub>t \. \n. t = Fun (Val n) []}" define vals_sym where "vals_sym \ \S::('fun,'atom,'sets,'lbl) prot_constr. {t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t S). (\n. t = Fun (Val n) []) \ (\m. t = Var (TAtom Value,m))}" have \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" by (metis \ welltyped_constraint_model_def) have \_grounds: "fv (t \ \) = {}" for t using \ interpretation_grounds[of \] unfolding welltyped_constraint_model_def constraint_model_def by auto have wt_\\\: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \ \\<^sub>s \ \\<^sub>s \)" and wt_\\: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \ \\<^sub>s \)" using \_wt wt_subst_compose transaction_decl_fresh_renaming_substs_wt[OF \ \ \] by (blast, blast) have "\T\set P. bvars_transaction T = {}" using P admissible_transactionE(4) by metis hence \_no_bvars: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = {}" using reachable_constraints_no_bvars[OF \_reach] by metis have \_vals: "\n. \ (TAtom Value, m) = Fun (Val n) []" when "(TAtom Value, m) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" for m using constraint_model_Value_term_is_Val'[ OF \_reach welltyped_constraint_model_prefix[OF \] P P_occ] \_no_bvars vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel \"] that by blast have vals_sym_vals: "t \ \ \ vals \" when t: "t \ vals_sym \" for t proof (cases t) case (Var x) then obtain m where *: "x = (TAtom Value,m)" using t unfolding vals_sym_def by blast moreover have "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using t unfolding vals_sym_def by blast hence "t \ \ \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>e\<^sub>t \" "\n. \ (Var Value, m) = Fun (Val n) []" using Var * \_vals[of m] var_subterm_trms\<^sub>s\<^sub>s\<^sub>t_is_vars\<^sub>s\<^sub>s\<^sub>t[of x "unlabel \"] \\<^sub>v_TAtom[of Value m] reachable_constraints_Value_vars_are_fv[OF \_reach P(1), of x] by blast+ ultimately show ?thesis using Var unfolding vals_def by auto next case (Fun f T) then obtain n where "f = Val n" "T = []" using t unfolding vals_sym_def by blast moreover have "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using t unfolding vals_sym_def by blast hence "t \ \ \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>e\<^sub>t \" using Fun by blast ultimately show ?thesis using Fun unfolding vals_def by auto qed have vals_vals_sym: "\s. s \ vals_sym \ \ t = s \ \" when "t \ vals \" for t using that constraint_model_Val_is_Value_term[OF \] unfolding vals_def vals_sym_def by fast note T_adm = bspec[OF P T] note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] have 0: "\\<^sub>i\<^sub>k (\@T') \ = (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a0' \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a0'" "\\<^sub>v\<^sub>a\<^sub>l\<^sub>s (\@T') \ = vals \ \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a0' \ vals T' \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a0'" by (metis abs_intruder_knowledge_append a0'_def, metis abs_value_constants_append[of \ T' \] a0'_def vals_def) have 1: "(ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a0' = (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t (\ \\<^sub>s \ \\<^sub>s \) \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a0'" by (metis T'_def dual_transaction_ik_is_transaction_send''[OF T_wf]) have 2: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T) \ subst_domain \ = {}" "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T) \ subst_domain \ = {}" "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T) \ subst_domain \ = {}" using admissible_transactionE(4)[OF T_adm] by blast+ have "vals T' \ (\ \\<^sub>s \ \\<^sub>s \) ` fv_transaction T \\<^sub>s\<^sub>e\<^sub>t \" proof fix t assume "t \ vals T'" then obtain s n where s: "s \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" "t = s \ \" "t = Fun (Val n) []" unfolding vals_def by fast then obtain u where u: "u \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T))" "s = u \ (\ \\<^sub>s \ \\<^sub>s \)" using transaction_decl_fresh_renaming_substs_trms[OF \ \ \ 2] trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \"] unfolding T'_def by blast have *: "t = u \ (\ \\<^sub>s \ \\<^sub>s \ \\<^sub>s \)" using s(2) u(2) subst_subst_compose by simp then obtain x where x: "u = Var x" using s(3) admissible_transactions_no_Value_consts(1)[OF T_adm u(1)] by (cases u) force+ hence **: "x \ vars_transaction T" by (metis u(1) var_subterm_trms\<^sub>s\<^sub>s\<^sub>t_is_vars\<^sub>s\<^sub>s\<^sub>t) have "\\<^sub>v x = TAtom Value" using * x s(3) wt_subst_trm''[OF wt_\\\, of u] by simp thus "t \ (\ \\<^sub>s \ \\<^sub>s \) ` fv_transaction T \\<^sub>s\<^sub>e\<^sub>t \" using admissible_transaction_Value_vars_are_fv[OF T_adm **] x * eval_term.simps(1)[of _ x "\ \\<^sub>s \ \\<^sub>s \ \\<^sub>s \"] subst_comp_set_image[of "\ \\<^sub>s \ \\<^sub>s \" \ "fv_transaction T"] by blast qed hence 3: "vals T' \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a0' \ ((\ \\<^sub>s \ \\<^sub>s \) ` fv_transaction T \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a0'" by (simp add: abs_apply_terms_def image_mono) have "t \ \ \\<^sub>\ a0' \ timpl_closure_set (\\<^sub>i\<^sub>k \ \) (\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \)" when "t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" for t using that abs_in[OF imageI[OF that]] \\<^sub>t\<^sub>i_covers_\\<^sub>0_ik[OF \_reach T \ \ \ \ P P_occ] timpl_closure_set_mono[of "{t \ \ \\<^sub>\ a0}" "\\<^sub>i\<^sub>k \ \" "\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \" "\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \"] unfolding a0_def a0'_def T'_def abs_intruder_knowledge_def by fast hence A: "\\<^sub>i\<^sub>k (\@T') \ \ timpl_closure_set (\\<^sub>i\<^sub>k \ \) (\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \) \ (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t (\ \\<^sub>s \ \\<^sub>s \) \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a0'" using 0(1) 1 by (auto simp add: abs_apply_terms_def) have "t \ \ \\<^sub>\ a0' \ timpl_closure_set {t \ \ \\<^sub>\ a0} (\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \)" when t: "t \ vals_sym \" for t proof - have "(\n. t = Fun (Val n) [] \ t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ (\n. t = Var (TAtom Value,n) \ (TAtom Value,n) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" (is "?P \ ?Q") using t var_subterm_trms\<^sub>s\<^sub>s\<^sub>t_is_vars\<^sub>s\<^sub>s\<^sub>t[of _ "unlabel \"] \\<^sub>v_TAtom[of Value] reachable_constraints_Value_vars_are_fv[OF \_reach P(1)] unfolding vals_sym_def by fast thus ?thesis proof assume ?P - then obtain n where n: "t = Fun (Val n) []" "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by moura + then obtain n where n: "t = Fun (Val n) []" "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by force thus ?thesis using \\<^sub>t\<^sub>i_covers_\\<^sub>0_Val[OF \_reach T \ \ \ \ P P_occ, of n] unfolding a0_def a0'_def T'_def by fastforce next assume ?Q thus ?thesis using \\<^sub>t\<^sub>i_covers_\\<^sub>0_Var[OF \_reach T \ \ \ \ P P_occ] unfolding a0_def a0'_def T'_def by fastforce qed qed moreover have "t \ \ \\<^sub>\ a0 \ \\<^sub>v\<^sub>a\<^sub>l\<^sub>s \ \" when "t \ vals_sym \" for t using that abs_in vals_sym_vals unfolding a0_def abs_value_constants_def vals_sym_def vals_def by (metis (mono_tags, lifting)) ultimately have "t \ \ \\<^sub>\ a0' \ timpl_closure_set (\\<^sub>v\<^sub>a\<^sub>l\<^sub>s \ \) (\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \)" when t: "t \ vals_sym \" for t using t timpl_closure_set_mono[of "{t \ \ \\<^sub>\ a0}" "\\<^sub>v\<^sub>a\<^sub>l\<^sub>s \ \" "\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \" "\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \"] by blast hence "t \\<^sub>\ a0' \ timpl_closure_set (\\<^sub>v\<^sub>a\<^sub>l\<^sub>s \ \) (\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \)" when t: "t \ vals \" for t using vals_vals_sym[OF t] by blast hence B: "\\<^sub>v\<^sub>a\<^sub>l\<^sub>s (\@T') \ \ timpl_closure_set (\\<^sub>v\<^sub>a\<^sub>l\<^sub>s \ \) (\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \) \ ((\ \\<^sub>s \ \\<^sub>s \) ` fv_transaction T \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a0'" using 0(2) 3 by (simp add: abs_apply_terms_def image_subset_iff) have 4: "fv (t \ \ \\<^sub>s \ \\<^sub>s \ \ \ \\<^sub>\ a) = {}" for t a using \_grounds[of "t \ \ \\<^sub>s \ \\<^sub>s \"] abs_fv[of "t \ \ \\<^sub>s \ \\<^sub>s \ \ \" a] by argo have "is_Fun (t \ \ \\<^sub>s \ \\<^sub>s \ \ \ \\<^sub>\ a0')" for t using 4[of t a0'] by force thus ?A using A step_props(1,3) unfolding T'_def a0_def a0'_def abs_apply_terms_def by blast show ?B using B step_props(2,4) admissible_transaction_Value_vars[OF bspec[OF P T]] by (auto simp add: T'_def a0_def a0'_def abs_apply_terms_def) qed lemma reachable_constraints_covered: assumes \_reach: "\ \ reachable_constraints P" and \: "welltyped_constraint_model \ \" and FP: "analyzed (timpl_closure_set (set FP) (set TI))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set FP)" "ground (set FP)" and OCC: "\t \ timpl_closure_set (set FP) (set TI). \f \ funs_term t. is_Abs f \ f \ Abs ` set OCC" "timpl_closure_set (absc ` set OCC) (set TI) \ absc ` set OCC" and TI: "set TI = {(a,b) \ (set TI)\<^sup>+. a \ b}" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and transactions_covered: "list_all (transaction_check (FP, OCC, TI)) P" shows "\t \ \\<^sub>i\<^sub>k \ \. timpl_closure_set (set FP) (set TI) \\<^sub>c t" and "\\<^sub>v\<^sub>a\<^sub>l\<^sub>s \ \ \ absc ` set OCC" using \_reach \ proof (induction rule: reachable_constraints.induct) case init { case 1 show ?case by (simp add: abs_intruder_knowledge_def) } { case 2 show ?case by (simp add: abs_value_constants_def) } next case (step \ T \ \ \) { case 1 hence "welltyped_constraint_model \ \" by (metis welltyped_constraint_model_prefix) hence IH: "\t \ \\<^sub>i\<^sub>k \ \. timpl_closure_set (set FP) (set TI) \\<^sub>c t" "\\<^sub>v\<^sub>a\<^sub>l\<^sub>s \ \ \ absc ` set OCC" using step.IH by metis+ show ?case using reachable_constraints_covered_step[ OF step.hyps(1,2) "1.prems" step.hyps(3-5) FP(1,2) IH(1) FP(3) OCC IH(2) TI P P_occ transactions_covered] by metis } { case 2 hence "welltyped_constraint_model \ \" by (metis welltyped_constraint_model_prefix) hence IH: "\t \ \\<^sub>i\<^sub>k \ \. timpl_closure_set (set FP) (set TI) \\<^sub>c t" "\\<^sub>v\<^sub>a\<^sub>l\<^sub>s \ \ \ absc ` set OCC" using step.IH by metis+ show ?case using reachable_constraints_covered_step[ OF step.hyps(1,2) "2.prems" step.hyps(3-5) FP(1,2) IH(1) FP(3) OCC IH(2) TI P P_occ transactions_covered] by metis } qed lemma attack_in_fixpoint_if_attack_in_ik: fixes FP::"('fun,'atom,'sets,'lbl) prot_terms" assumes "\t \ IK \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a. FP \\<^sub>c t" and "attack\n\ \ IK" shows "attack\n\ \ FP" proof - have "attack\n\ \\<^sub>\ a \ IK \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a" by (rule abs_in[OF assms(2)]) hence "FP \\<^sub>c attack\n\ \\<^sub>\ a" using assms(1) by blast moreover have "attack\n\ \\<^sub>\ a = attack\n\" by simp ultimately have "FP \\<^sub>c attack\n\" by metis thus ?thesis using ideduct_synth_priv_const_in_ik[of FP "Attack n"] by simp qed lemma attack_in_fixpoint_if_attack_in_timpl_closure_set: fixes FP::"('fun,'atom,'sets,'lbl) prot_terms" assumes "attack\n\ \ timpl_closure_set FP TI" shows "attack\n\ \ FP" proof - have "\f \ funs_term (attack\n\). \is_Abs f" by auto thus ?thesis using timpl_closure_set_no_Abs_in_set[OF assms] by blast qed theorem prot_secure_if_fixpoint_covered_typed: assumes FP: "analyzed (timpl_closure_set (set FP) (set TI))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set FP)" "ground (set FP)" and OCC: "\t \ timpl_closure_set (set FP) (set TI). \f \ funs_term t. is_Abs f \ f \ Abs ` set OCC" "timpl_closure_set (absc ` set OCC) (set TI) \ absc ` set OCC" and TI: "set TI = {(a,b) \ (set TI)\<^sup>+. a \ b}" and P: "\T \ set P. admissible_transaction T" "has_initial_value_producing_transaction P" and transactions_covered: "list_all (transaction_check (FP, OCC, TI)) (map add_occurs_msgs P)" and attack_notin_FP: "attack\n\ \ set FP" and \: "\ \ reachable_constraints P" shows "\\. welltyped_constraint_model \ (\@[(l, send\[attack\n\]\)])" (is "\\. ?Q \") proof assume "\\. ?Q \" then obtain \ \ where \: "\ \ reachable_constraints (map add_occurs_msgs P)" and \: "welltyped_constraint_model \ (\@[(l, send\[attack\n\]\)])" using add_occurs_msgs_soundness[OF P \] unfolding list_all_iff by blast have P': "\T \ set (map add_occurs_msgs P). admissible_transaction' T" "\T \ set (map add_occurs_msgs P). admissible_transaction_occurs_checks T" using P add_occurs_msgs_admissible_occurs_checks[OF admissible_transactionE'(1)] by auto have 0: "attack\n\ \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" using welltyped_constraint_model_prefix[OF \] reachable_constraints_covered(1)[OF \ _ FP OCC TI P' transactions_covered] attack_in_fixpoint_if_attack_in_ik[ of "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" "\\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" "timpl_closure_set (set FP) (set TI)" n] attack_in_fixpoint_if_attack_in_timpl_closure_set attack_notin_FP unfolding abs_intruder_knowledge_def by blast have 1: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ attack\n\" using \ strand_sem_append_stateful[of "{}" "{}" "unlabel \" _ \] unfolding welltyped_constraint_model_def constraint_model_def by force show False using 0 private_const_deduct[OF _ 1] reachable_constraints_receive_attack_if_attack'(1)[ OF \ P'(1) welltyped_constraint_model_prefix[OF \] 1] by simp qed end subsection \Theorem: A Protocol is Secure if it is Covered by a Fixed-Point\ context stateful_protocol_model begin theorem prot_secure_if_fixpoint_covered: fixes P assumes FP: "analyzed (timpl_closure_set (set FP) (set TI))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set FP)" "ground (set FP)" and OCC: "\t \ timpl_closure_set (set FP) (set TI). \f \ funs_term t. is_Abs f \ f \ Abs ` set OCC" "timpl_closure_set (absc ` set OCC) (set TI) \ absc ` set OCC" and TI: "set TI = {(a,b) \ (set TI)\<^sup>+. a \ b}" and M: "has_all_wt_instances_of \ (\T \ set P. trms_transaction T) N" "finite N" "tfr\<^sub>s\<^sub>e\<^sub>t N" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s N" and P: "\T \ set P. admissible_transaction T" "\T \ set P. list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (transaction_strand T))" "has_initial_value_producing_transaction P" and transactions_covered: "list_all (transaction_check (FP, OCC, TI)) (map add_occurs_msgs P)" and attack_notin_FP: "attack\n\ \ set FP" and A: "\ \ reachable_constraints P" shows "\\. constraint_model \ (\@[(l, send\[attack\n\]\)])" (is "\\. constraint_model \ ?A") proof assume "\\. constraint_model \ ?A" - then obtain \ where "constraint_model \ ?A" by moura + then obtain \ where "constraint_model \ ?A" by force then obtain \\<^sub>\ where I: "welltyped_constraint_model \\<^sub>\ ?A" using reachable_constraints_typing_result[OF M P(1,2) A] by blast note a = FP OCC TI P(1,3) transactions_covered attack_notin_FP A show False using prot_secure_if_fixpoint_covered_typed[OF a] I by force qed end subsection \Alternative Protocol-Coverage Check\ context stateful_protocol_model begin context begin private lemma transaction_check_variant_soundness_aux0: assumes S: "S \ unlabel (transaction_strand T)" and xs: "xs \ filter (\x. x \ set (transaction_fresh T) \ fst x = TAtom Value) (fv_list\<^sub>s\<^sub>s\<^sub>t S)" and x: "fst x = Var Value" "x \ fv_transaction T" "x \ set (transaction_fresh T)" shows "x \ set xs" using x fv_list\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_strand T)"] unfolding xs S by auto private lemma transaction_check_variant_soundness_aux1: fixes T FP S C xs OCC negs poss as assumes C: "C \ unlabel (transaction_checks T)" and S: "S \ unlabel (transaction_strand T)" and xs: "xs \ filter (\x. x \ set (transaction_fresh T) \ fst x = TAtom Value) (fv_list\<^sub>s\<^sub>s\<^sub>t S)" and poss: "poss \ transaction_poschecks_comp C" and negs: "negs \ transaction_negchecks_comp C" and as: "as \ map (\x. (x, set (filter (\ab. poss x \ ab \ negs x \ ab = {}) OCC))) xs" and f: "f \ \x. case List.find (\p. fst p = x) as of Some p \ snd p | None \ {}" and x: "x \ set xs" shows "f x = set (filter (\ab. poss x \ ab \ negs x \ ab = {}) OCC)" proof - define g where "g \ \x. set (filter (\ab. poss x \ ab \ negs x \ ab = {}) OCC)" define gs where "gs \ map (\x. (x, g x)) xs" have 1: "(x, g x) \ set gs" using x unfolding gs_def by simp have 2: "distinct xs" unfolding xs fv_list\<^sub>s\<^sub>s\<^sub>t_def by auto have "\i < length xs. xs ! i = x \ (\j < i. xs ! j \ x)" when x: "x \ set xs" for x proof (rule ex1E[OF distinct_Ex1[OF 2 x]]) fix i assume i: "i < length xs \ xs ! i = x" and "\j. j < length xs \ xs ! j = x \ j = i" hence "\j < length xs. xs ! j = x \ j = i" by blast hence "\j < i. xs ! j = x \ j = i" using i by auto hence "\j < i. xs ! j \ x" by blast thus ?thesis using i by blast qed hence "\i < length gs. gs ! i = (x, g x) \ (\j < i. gs ! j \ (x, g x))" using 1 unfolding gs_def by fastforce hence "\i < length gs. fst (gs ! i) = x \ (x, g x) = gs ! i \ (\j < i. fst (gs ! j) \ x)" using nth_map[of _ xs "\x. (x, g x)"] length_map[of "\x. (x, g x)" xs] unfolding gs_def by (metis (no_types, lifting) fstI min.strict_order_iff min_less_iff_conj) hence "List.find (\p. fst p = x) gs = Some (x, g x)" using find_Some_iff[of "\p. fst p = x" gs "(x, g x)"] by blast thus ?thesis unfolding f as gs_def g_def by force qed private lemma transaction_check_variant_soundness_aux2: fixes T FP S C xs OCC negs poss as assumes C: "C \ unlabel (transaction_checks T)" and S: "S \ unlabel (transaction_strand T)" and xs: "xs \ filter (\x. x \ set (transaction_fresh T) \ fst x = TAtom Value) (fv_list\<^sub>s\<^sub>s\<^sub>t S)" and poss: "poss \ transaction_poschecks_comp C" and negs: "negs \ transaction_negchecks_comp C" and as: "as \ map (\x. (x, set (filter (\ab. poss x \ ab \ negs x \ ab = {}) OCC))) xs" and f: "f \ \x. case List.find (\p. fst p = x) as of Some p \ snd p | None \ {}" and x: "x \ set xs" shows "f x = {}" proof - define g where "g \ \x. set (filter (\ab. poss x \ ab \ negs x \ ab = {}) OCC)" define gs where "gs \ map (\x. (x, g x)) xs" have "(x, y) \ set gs" for y using x unfolding gs_def by force thus ?thesis using find_None_iff[of "\p. fst p = x" gs] unfolding f as gs_def g_def by fastforce qed private lemma synth_abs_substs_constrs_rel_if_synth_abs_substs_constrs: fixes T OCC negs poss defines "\ \ \\ x. if fst x = TAtom Value then (absc \ \) x else Var x" and "ts \ trms_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_receive T))" assumes ts_wf: "\t \ set ts. wf\<^sub>t\<^sub>r\<^sub>m t" and FP_ground: "ground (set FP)" and FP_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set FP)" shows "synth_abs_substs_constrs_rel FP OCC TI ts (synth_abs_substs_constrs (FP,OCC,TI) T)" proof - let ?R = "synth_abs_substs_constrs_rel FP OCC TI" let ?D = "synth_abs_substs_constrs_aux FP OCC TI" have *: "?R [t] (?D t)" when "wf\<^sub>t\<^sub>r\<^sub>m t" for t using that proof (induction t) case (Var x) thus ?case using synth_abs_substs_constrs_rel.SolveValueVar[of "?D (Var x)" OCC x TI FP] by fastforce next case (Fun f ss) let ?xs = "fv (Fun f ss)" let ?lst = "map (match_abss OCC TI (Fun f ss)) FP" define flt where "flt = (\\::(('fun,'atom,'sets,'lbl) prot_var \ 'sets set set) option. \ \ None)" define \ where "\ = map the (filter flt (map (match_abss OCC TI (Fun f ss)) FP))" define \1 where "\1 = fun_point_Union (set \)" define \2 where "\2 = fun_point_Inter (?D ` set ss)" have f: "arity f = length ss" using wf_trm_arity[OF Fun.prems] by simp have IH: "?R [s] (?D s)" when s: "s \ set ss" for s using Fun.IH[OF s wf_trm_subterm[OF Fun.prems Fun_param_is_subterm[OF s]]] s by force have \3: "\\. \ \ set \ \ (\s \ set FP. match_abss OCC TI (Fun f ss) s = Some \)" (is "\\. \ \ set \ \ ?P \") proof (intro allI iffI) fix \ assume "\ \ set \" then obtain u where "u \ set FP" "match_abss OCC TI (Fun f ss) u = Some \" unfolding \_def flt_def by fastforce thus "?P \" by blast next fix \ assume "?P \" then obtain u where u: "u \ set FP" "match_abss OCC TI (Fun f ss) u = Some \" by blast have "Some \ \ set ?lst" using u unfolding flt_def by force hence "Some \ \ set (filter flt ?lst)" unfolding flt_def by force moreover have "\\. d = Some \" when d: "d \ set (filter flt ?lst)" for d using d unfolding flt_def by simp ultimately have "\ \ set (map the (filter flt ?lst))" by force thus "\ \ set \" unfolding \_def by blast qed show ?case proof (cases "ss = []") case True note ss = this show ?thesis proof (cases "\public f \ Fun f ss \ set FP") case True thus ?thesis using synth_abs_substs_constrs_rel.SolvePrivConstNotin[of f FP OCC TI] unfolding ss by force next case False thus ?thesis using f synth_abs_substs_constrs_rel.SolvePubConst[of f FP OCC TI] synth_abs_substs_constrs_rel.SolvePrivConstIn[of f FP OCC TI] unfolding ss by auto qed next case False note ss = this hence f': "arity f > 0" using f by auto have IH': "?R ss (fun_point_Inter (?D ` set ss))" using IH synth_abs_substs_constrs_rel.SolveCons[OF ss, of FP OCC TI ?D] by blast have "?D (Fun f ss) = ( fun_point_union (fun_point_Union_list \) (fun_point_Inter_list (map ?D ss)))" using ss synth_abs_substs_constrs_aux.simps(2)[of FP OCC TI f ss] unfolding Let_def \_def flt_def by argo hence "?D (Fun f ss) = fun_point_union \1 \2" using fun_point_Inter_set_eq[of "map ?D ss"] fun_point_Union_set_eq[of \] unfolding \1_def \2_def by simp thus ?thesis using synth_abs_substs_constrs_rel.SolveComposed[ OF f' f[symmetric] \3 \1_def IH'] unfolding \2_def by argo qed qed note l0 = synth_abs_substs_constrs_rel.SolveNil[of FP OCC TI] note d0 = synth_abs_substs_constrs_def ts_def note l1 = * ts_wf synth_abs_substs_constrs_rel.SolveCons[of ts FP OCC TI ?D] note d1 = d0 Let_def fun_point_Inter_set_eq[symmetric] fun_point_Inter_def show ?thesis proof (cases "ts = []") case True thus ?thesis using l0 unfolding d0 by simp next case False thus ?thesis using l1 unfolding d1 by auto qed qed private function (sequential) match_abss'_timpls_transform ::"('c set \ 'c set) list \ ('a,'b,'c,'d) prot_subst \ ('a,'b,'c,'d) prot_term \ ('a,'b,'c,'d) prot_term \ (('a,'b,'c,'d) prot_var \ 'c set set) option" where "match_abss'_timpls_transform TI \ (Var x) (Fun (Abs a) _) = ( if \b ts. \ x = Fun (Abs b) ts \ (a = b \ (a,b) \ set TI) then Some ((\_. {})(x := {a})) else None)" | "match_abss'_timpls_transform TI \ (Fun f ts) (Fun g ss) = ( if f = g \ length ts = length ss then map_option fun_point_Union_list (those (map2 (match_abss'_timpls_transform TI \) ts ss)) else None)" | "match_abss'_timpls_transform _ _ _ _ = None" by pat_completeness auto termination proof - let ?m = "measures [size \ fst \ snd \ snd]" have 0: "wf ?m" by simp show ?thesis apply (standard, use 0 in fast) by (metis (no_types) comp_def fst_conv snd_conv measures_less Fun_zip_size_lt(1)) qed private lemma match_abss'_timpls_transform_Var_inv: assumes "match_abss'_timpls_transform TI \ (Var x) (Fun (Abs a) ts) = Some \" shows "\b ts. \ x = Fun (Abs b) ts \ (a = b \ (a, b) \ set TI)" and "\ = ((\_. {})(x := {a}))" using assms match_abss'_timpls_transform.simps(1)[of TI \ x a ts] by (metis option.distinct(1), metis option.distinct(1) option.inject) private lemma match_abss'_timpls_transform_Fun_inv: assumes "match_abss'_timpls_transform TI \ (Fun f ts) (Fun g ss) = Some \" shows "f = g" (is ?A) and "length ts = length ss" (is ?B) and "\\. Some \ = those (map2 (match_abss'_timpls_transform TI \) ts ss) \ \ = fun_point_Union_list \" (is ?C) and "\(t,s) \ set (zip ts ss). \\'. match_abss'_timpls_transform TI \ t s = Some \'" (is ?D) proof - note 0 = assms match_abss'_timpls_transform.simps(2)[of TI \ f ts g ss] option.distinct(1) show ?A by (metis 0) show ?B by (metis 0) show ?C by (metis (no_types, opaque_lifting) 0 map_option_eq_Some) thus ?D using map2_those_Some_case[of "match_abss'_timpls_transform TI \" ts ss] by fastforce qed private lemma match_abss'_timpl_transform_nonempty_is_fv: assumes "match_abss'_timpls_transform TI \ s t = Some \" and "\ x \ {}" shows "x \ fv s" using assms proof (induction s t arbitrary: TI \ \ rule: match_abss'_timpls_transform.induct) case (1 TI \ y a ts) show ?case using match_abss'_timpls_transform_Var_inv[OF "1.prems"(1)] "1.prems"(2) by fastforce next case (2 TI \ f ts g ss) note prems = "2.prems" note IH = "2.IH" obtain \ where \: "Some \ = those (map2 (match_abss'_timpls_transform TI \) ts ss)" "\ = fun_point_Union_list \" and fg: "f = g" "length ts = length ss" using match_abss'_timpls_transform_Fun_inv[OF prems(1)] by fast have "\\ \ set \. \ x \ {}" using fg(2) prems \ unfolding fun_point_Union_list_def by auto then obtain t' s' \ where ts': "(t',s') \ set (zip ts ss)" "match_abss'_timpls_transform TI \ t' s' = Some \" "\ x \ {}" using those_map2_SomeD[OF \(1)[symmetric]] by blast show ?case using ts'(3) IH[OF conjI[OF fg] ts'(1) _ ts'(2)] set_zip_leftD[OF ts'(1)] by force qed auto private lemma match_abss'_timpls_transformI: fixes s t::"('a,'b,'c,'d) prot_term" and \::"('a,'b,'c,'d) prot_subst" and \::"('a,'b,'c,'d) prot_var \ 'c set set" assumes TI: "set TI = {(a,b) \ (set TI)\<^sup>+. a \ b}" and \: "timpls_transformable_to TI t (s \ \)" and \: "match_abss' s t = Some \" and t: "fv t = {}" and s: "\f \ funs_term s. \is_Abs f" "\x \ fv s. \a. \ x = \a\\<^sub>a\<^sub>b\<^sub>s" shows "match_abss'_timpls_transform TI \ s t = Some \" using \ \ t s proof (induction t arbitrary: s \ \) case (Fun f ts) note prems = Fun.prems note IH = Fun.IH show ?case proof (cases s) case (Var x) obtain a b where a: "f = Abs a" "\ = (\_. {})(x := {a})" and b: "\ x = \b\\<^sub>a\<^sub>b\<^sub>s" using match_abss'_Var_inv[OF prems(2)[unfolded Var]] prems(5)[unfolded Var] by auto thus ?thesis using prems(1) timpls_transformable_to_inv(3)[of TI f ts "Abs b" "[]"] unfolding Var by auto next case (Fun g ss) note 0 = timpls_transformable_to_inv[OF prems(1)[unfolded Fun eval_term.simps(2)]] note 1 = match_abss'_Fun_inv[OF prems(2)[unfolded Fun]] obtain \ where \: "those (map2 match_abss' ss ts) = Some \" "\ = fun_point_Union_list \" using 1(3) by force have "timpls_transformable_to TI t' (s' \ \)" "\\'. match_abss' s' t' = Some \'" when "(t',s') \ set (zip ts ss)" for s' t' by (metis 0(2) nth_map[of _ ss] zip_arg_index[OF that], use that 1(4) in_set_zip_swap[of t' s' ts ss] in fast) hence IH': "match_abss'_timpls_transform TI \ s' t' = Some \'" when "(t',s') \ set (zip ts ss)" "match_abss' s' t' = Some \'" for s' t' \' using that IH[of t' s' \ \'] prems(3-) unfolding Fun by (metis (no_types, lifting) set_zip_leftD set_zip_rightD subsetI subset_empty term.set_intros(2) term.set_intros(4)) have "those (map2 (match_abss'_timpls_transform TI \) ss ts) = Some \" using IH' \(1) 1(4) in_set_zip_swap[of _ _ ss ts] those_Some_iff[of "map2 match_abss' ss ts" \] those_Some_iff[of "map2 (match_abss'_timpls_transform TI \) ss ts" \] by auto thus ?thesis using \(2) 1(1,2) Fun by simp qed qed simp lemma timpls_transformable_to_match_abss'_nonempty_disj': fixes s t::"('a,'b,'c,'d) prot_term" and \::"('a,'b,'c,'d) prot_subst" and \::"('a,'b,'c,'d) prot_var \ 'c set set" assumes TI: "set TI = {(a,b) \ (set TI)\<^sup>+. a \ b}" and \: "timpls_transformable_to TI t (s \ \)" and \: "match_abss' s t = Some \" and x: "x \ fv s" and t: "fv t = {}" and s: "\f \ funs_term s. \is_Abs f" "\x \ fv s. \a. \ x = \a\\<^sub>a\<^sub>b\<^sub>s" and a: "\ x = \a\\<^sub>a\<^sub>b\<^sub>s" shows "\b \ \ x. (b,a) \ (set TI)\<^sup>*" (is "?P \ x") proof - note 0 = match_abss'_subst_disj_nonempty[OF TI] have 1: "s \ \ \ timpl_closure t (set TI)" using timpls_transformable_to_iff_in_timpl_closure[OF TI] \ by blast have 2: "match_abss'_timpls_transform TI \ s t = Some \" using match_abss'_timpls_transformI[OF TI \ \ t s] by simp show "?P \ x" using 2 TI x t s a proof (induction TI \ s t arbitrary: \ rule: match_abss'_timpls_transform.induct) case (1 TI \ y c ts) thus ?case using match_abss'_timpls_transform_Var_inv[OF "1.prems"(1)] by auto next case (2 TI \ f ts g ss) obtain \ where fg: "f = g" "length ts = length ss" and \: "Some \ = those (map2 (match_abss'_timpls_transform TI \) ts ss)" "\ = fun_point_Union_list \" "\(t, s)\set (zip ts ss). \\'. match_abss'_timpls_transform TI \ t s = Some \'" using match_abss'_timpls_transform_Fun_inv[OF "2.prems"(1)] by blast have "(b,a) \ (set TI)\<^sup>*" when \': "\' \ set \" "b \ \' x" for \' b proof - obtain t' s' where t': "(t',s') \ set (zip ts ss)" "match_abss'_timpls_transform TI \ t' s' = Some \'" using those_map2_SomeD[OF \(1)[symmetric] \'(1)] by blast have *: "fv s' = {}" "\f \ funs_term t'. \is_Abs f" "\x \ fv t'. \a. \ x = \a\\<^sub>a\<^sub>b\<^sub>s" using "2.prems"(4-6) set_zip_leftD[OF t'(1)] set_zip_rightD[OF t'(1)] by (fastforce, fastforce, fastforce) have **: "x \ fv t'" using \'(2) match_abss'_timpl_transform_nonempty_is_fv[OF t'(2)] by blast show ?thesis using \'(2) "2.IH"[OF conjI[OF fg] t'(1) _ t'(2) "2.prems"(2) ** * "2.prems"(7)] by blast qed thus ?case using \(1) unfolding \(2) fun_point_Union_list_def by simp qed auto qed lemma timpls_transformable_to_match_abss'_nonempty_disj: fixes s t::"('a,'b,'c,'d) prot_term" and \::"('a,'b,'c,'d) prot_subst" and \::"('a,'b,'c,'d) prot_var \ 'c set set" assumes TI: "set TI = {(a,b) \ (set TI)\<^sup>+. a \ b}" and \: "timpls_transformable_to TI t (s \ \)" and \: "match_abss' s t = Some \" and x: "x \ fv s" and t: "fv t = {}" and s: "\f \ funs_term s. \is_Abs f" "\x \ fv s. \a. \ x = \a\\<^sub>a\<^sub>b\<^sub>s" shows "\(ticl_abs TI ` \ x) \ {}" proof - have 0: "(a,b) \ (set TI)\<^sup>*" when y: "y \ fv s" "a \ \ y" "\ y = \b\\<^sub>a\<^sub>b\<^sub>s" for a b y using timpls_transformable_to_match_abss'_nonempty_disj'[OF TI \ \ y(1) t s y(3)] y(2) by blast obtain b where b: "\ x = \b\\<^sub>a\<^sub>b\<^sub>s" using x s(2) by blast have "b \ ticl_abs TI a" when a: "a \ \ x" for a using 0[OF x a b] unfolding ticl_abs_iff[OF TI] by blast thus ?thesis by blast qed lemma timpls_transformable_to_subst_subterm: fixes s t::"(('a,'b,'c,'d) prot_fun, 'v) term" and \ \::"(('a,'b,'c,'d) prot_fun, 'v) subst" assumes "timpls_transformable_to TI (t \ \) (t \ \)" and "s \ t" shows "timpls_transformable_to TI (s \ \) (s \ \)" using assms proof (induction "t \ \" "t \ \" arbitrary: t rule: timpls_transformable_to.induct) case (1 TI x y) thus ?case by (cases t) auto next case (2 TI f T g S) note prems = "2.prems" note hyps = "2.hyps"(2-) note IH = "2.hyps"(1) show ?case proof (cases "s = t") case False then obtain h U u where t: "t = Fun h U" "u \ set U" "s \ u" using prems(2) by (cases t) auto then obtain i where i: "i < length U" "U ! i = u" by (metis in_set_conv_nth) have "timpls_transformable_to TI (u \ \) (u \ \)" using t i prems(1) timpls_transformable_to_inv(2)[of TI h "U \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" h "U \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" i] by simp thus ?thesis using IH hyps t by auto qed (use prems in auto) qed simp_all lemma timpls_transformable_to_subst_match_case: assumes "timpls_transformable_to TI s (t \ \)" and "fv s = {}" and "\f \ funs_term t. \is_Abs f" and "distinct (fv_list t)" and "\x \ fv t. \a. \ x = \a\\<^sub>a\<^sub>b\<^sub>s" shows "\\. s = t \ \" using assms proof (induction s "t \ \" arbitrary: t rule: timpls_transformable_to.induct) case (2 TI f T g S) note prems = "2.prems" note hyps = "2.hyps"(2-) note IH = "2.hyps"(1) show ?case proof (cases t) case (Var x) then obtain a where a: "t \ \ = \a\\<^sub>a\<^sub>b\<^sub>s" using prems(5) by fastforce show ?thesis using hyps timpls_transformable_to_inv'[OF prems(1)[unfolded a]] unfolding Var by force next case (Fun h U) have g: "g = h" and S: "S = U \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" using hyps unfolding Fun by simp_all note 0 = distinct_fv_list_Fun_param[OF prems(4)[unfolded Fun]] have 1: "\f \ funs_term u. \is_Abs f" when u: "u \ set U" for u using prems(3) u unfolding Fun by fastforce have 2: "fv t' = {}" when t': "t' \ set T" for t' using t' prems(2) by simp have 3: "\x \ fv u. \a. \ x = \a\\<^sub>a\<^sub>b\<^sub>s" when u: "u \ set U" for u using u prems(5) unfolding Fun by simp have "\is_Abs f" using prems(3) timpls_transformable_to_inv(3)[OF prems(1)[unfolded hyps[symmetric] S g]] unfolding Fun by auto hence f: "f = h" and T: "length T = length U" using prems(1) timpls_transformable_to_inv(1,3)[of TI f T h "U \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \"] unfolding Fun by fastforce+ define \ where "\ \ \i. if i < length T then SOME \. T ! i = U ! i \ \ else undefined" have "timpls_transformable_to TI (T ! i) (U ! i \ \)" when i: "i < length T" for i using prems(1)[unfolded Fun] T i timpls_transformable_to_inv(2)[of TI f T h "U \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" i] by auto hence "\\. T ! i = U ! i \ \" when i: "i < length T" for i using i T IH[OF _ _ _ 2 1 0 3, of "T ! i" "U ! i"] unfolding Fun g S by simp hence \: "T ! i = U ! i \ \ i" when i: "i < length T" for i using i someI2[of "\\. T ! i = U ! i \ \" _ "\\. T ! i = U ! i \ \"] unfolding \_def by fastforce define \ where "\ \ \x. if \i < length T. x \ fv (U ! i) then \ (SOME i. i < length T \ x \ fv (U ! i)) x else undefined" have "T ! i = U ! i \ \" when i: "i < length T" for i proof - have "j = i" when x: "x \ fv (U ! i)" and j: "j < length T" "x \ fv (U ! j)" for j x using x i j T distinct_fv_list_idx_fv_disjoint[OF prems(4)[unfolded Fun], of h U] by (metis (no_types, lifting) disjoint_iff_not_equal neqE term.dual_order.refl) hence "\ x = \ i x" when x: "x \ fv (U ! i)" for x using x i some_equality[of "\i. i < length T \ x \ fv (U ! i)" i] unfolding \_def by (metis (no_types, lifting)) thus ?thesis by (metis \ i term_subst_eq) qed hence "T = U \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" by (metis (no_types, lifting) T length_map nth_equalityI nth_map) hence "Fun f T = Fun f U \ \" by simp thus ?thesis using Fun f by fast qed qed simp_all lemma timpls_transformable_to_match_abss'_case: assumes "timpls_transformable_to TI s (t \ \)" and "fv s = {}" and "\f \ funs_term t. \is_Abs f" and "\x \ fv t. \a. \ x = \a\\<^sub>a\<^sub>b\<^sub>s" shows "\\. match_abss' t s = Some \" using assms proof (induction s "t \ \" arbitrary: t rule: timpls_transformable_to.induct) case (2 TI f T g S) note prems = "2.prems" note hyps = "2.hyps"(2-) note IH = "2.hyps"(1) show ?case proof (cases t) case (Var x) then obtain a where a: "t \ \ = \a\\<^sub>a\<^sub>b\<^sub>s" using prems(4) by fastforce thus ?thesis using timpls_transformable_to_inv'(4)[OF prems(1)[unfolded a]] by (metis (no_types) Var is_Abs_def term.sel(2) match_abss'.simps(1)) next case (Fun h U) have g: "g = h" and S: "S = U \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" using hyps unfolding Fun by simp_all have 1: "\f \ funs_term u. \is_Abs f" when u: "u \ set U" for u using prems(3) u unfolding Fun by fastforce have 2: "fv t' = {}" when t': "t' \ set T" for t' using t' prems(2) by simp have 3: "\x \ fv u. \a. \ x = \a\\<^sub>a\<^sub>b\<^sub>s" when u: "u \ set U" for u using u prems(4) unfolding Fun by simp have "\is_Abs f" using prems(3) timpls_transformable_to_inv(3)[OF prems(1)[unfolded hyps[symmetric] S g]] unfolding Fun by auto hence f: "f = h" and T: "length T = length U" using prems(1) timpls_transformable_to_inv(1,3)[of TI f T h "U \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \"] unfolding Fun by fastforce+ define \ where "\ \ \i. if i < length T then SOME \. match_abss' (U ! i) (T ! i) = Some \ else undefined" have "timpls_transformable_to TI (T ! i) (U ! i \ \)" when i: "i < length T" for i using prems(1)[unfolded Fun] T i timpls_transformable_to_inv(2)[of TI f T h "U \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" i] by auto hence "\\. match_abss' (U ! i) (T ! i) = Some \" when i: "i < length T" for i using i T IH[OF _ _ _ 2 1 3, of "T ! i" "U ! i"] unfolding Fun g S by simp hence "match_abss' (U ! i) (T ! i) = Some (\ i)" when i: "i < length T" for i using i someI2[of "\\. match_abss' (U ! i) (T ! i) = Some \" _ "\\. match_abss' (U ! i) (T ! i) = Some \"] unfolding \_def by fastforce thus ?thesis using match_abss'_FunI[OF _ T] unfolding Fun f by auto qed qed simp_all lemma timpls_transformable_to_match_abss_case: assumes TI: "set TI = {(a,b) \ (set TI)\<^sup>+. a \ b}" and "timpls_transformable_to TI s (t \ \)" and "fv s = {}" and "\f \ funs_term t. \is_Abs f" and "\x \ fv t. \a. \ x = \a\\<^sub>a\<^sub>b\<^sub>s" shows "\\. match_abss OCC TI t s = Some \" proof - obtain \ where \: "match_abss' t s = Some \" using timpls_transformable_to_match_abss'_case[OF assms(2-)] by blast show ?thesis using \ timpls_transformable_to_match_abss'_nonempty_disj[OF assms(1,2) \ _ assms(3-5)] unfolding match_abss_def by simp qed private lemma transaction_check_variant_soundness_aux3: fixes T FP S C xs OCC negs poss as defines "\ \ \\ x. if fst x = TAtom Value then (absc \ \) x else Var x" and "C \ unlabel (transaction_checks T)" and "S \ unlabel (transaction_strand T)" and "ts \ trms_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_receive T))" and "xs \ filter (\x. x \ set (transaction_fresh T) \ fst x = TAtom Value) (fv_list\<^sub>s\<^sub>s\<^sub>t S)" assumes TI0: "\(a,b) \ set TI. \(c,d) \ set TI. b = c \ a \ d \ (a,d) \ set TI" "\(a,b) \ set TI. a \ b" and OCC: "\t \ set FP. \a. Abs a \ funs_term t \ a \ set OCC" and FP_ground: "ground (set FP)" and x: "x \ set xs" and xs: "\x. x \ set xs \ \ x \ set OCC" "\x. x \ set xs \ poss x \ \ x" "\x. x \ set xs \ \ x \ negs x = {}" "\x. x \ set xs \ \ x = {}" and ts: "\t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T). intruder_synth_mod_timpls FP TI (t \ \ \)" "\t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T). \f \ funs_term t. \is_Abs f" "\x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)). fst x = TAtom Value" and C: "\a x s. \a: Var x \ Fun (Set s) []\ \ set C \ s \ \ x" "\x s. \Var x not in Fun (Set s) []\ \ set C \ s \ \ x" and \: "synth_abs_substs_constrs_rel FP OCC TI ts \" shows "\ x \ \ x" proof - note defs = assms(1-5) note TI = trancl_eqI'[OF TI0] have \x0: "\ x \ set OCC" "poss x \ \ x" "\ x \ negs x = {}" using x xs by (blast,blast,blast) have ts0: "\t \ set ts. intruder_synth_mod_timpls FP TI (t \ \ \)" using ts(1) trms_list\<^sub>s\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>s\<^sub>t unfolding ts_def by blast have ts1: "\Fun (Abs n) S \\<^sub>s\<^sub>e\<^sub>t set ts" for n S using ts(2) funs_term_Fun_subterm' unfolding ts_def trms_transaction_unfold trms_list\<^sub>s\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>s\<^sub>t[symmetric] is_Abs_def by fastforce have ts2: "\x \ fv\<^sub>s\<^sub>e\<^sub>t (set ts). fst x = TAtom Value" using ts(3) unfolding ts_def trms_transaction_unfold trms_list\<^sub>s\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>s\<^sub>t[symmetric] by fastforce show ?thesis using \ ts0 ts1 ts2 proof (induction rule: synth_abs_substs_constrs_rel.induct) case (SolvePrivConstNotin c) hence "intruder_synth_mod_timpls FP TI (Fun c [])" by force hence "list_ex (\t. timpls_transformable_to TI t (Fun c [])) FP" using SolvePrivConstNotin.hyps(1,2) by simp then obtain t where t: "t \ set FP" "timpls_transformable_to TI t (Fun c [])" unfolding list_ex_iff by blast have "\is_Abs c" using SolvePrivConstNotin.prems(2)[of _ "[]"] by (metis in_subterms_Union is_Abs_def list.set_intros(1)) hence "t = Fun c []" using t(2) timpls_transformable_to_inv[of TI] by (cases t) auto thus ?case using t(1) SolvePrivConstNotin.hyps(3) by fast next case (SolveValueVar \1 y) have "list_ex (\t. timpls_transformable_to TI t \\ y\\<^sub>a\<^sub>b\<^sub>s) FP" using SolveValueVar.prems(1-3) unfolding \_def by simp then obtain t where t: "t \ set FP" "timpls_transformable_to TI t \\ y\\<^sub>a\<^sub>b\<^sub>s" unfolding list_ex_iff by blast obtain a where a: "t = \a\\<^sub>a\<^sub>b\<^sub>s" "a = \ y \ (a, \ y) \ set TI" proof - obtain ft tst where ft: "t = Fun ft tst" using t(2) timpls_transformable_to_inv_Var(1)[of TI _ "\\ y\\<^sub>a\<^sub>b\<^sub>s"] by (cases t) auto have "tst = []" "is_Abs ft" "the_Abs ft = \ y \ (the_Abs ft, \ y) \ set TI" using timpls_transformable_to_inv'(2,4,5)[OF t(2)[unfolded ft]] by (simp, force, force) thus thesis using that[of "the_Abs ft"] ft by force qed have "a \ set OCC" using t(1)[unfolded a(1)] OCC by auto thus ?case using \x0(1) t(1)[unfolded a(1)] a(2) unfolding SolveValueVar.hyps(1) ticl_abss_def ticl_abs_def by force next case (SolveComposed g us \ \1 \2) show ?case proof (cases "\t \ set us. intruder_synth_mod_timpls FP TI (t \ \ \)") case True hence "\ x \ \2 x" using SolveComposed.IH SolveComposed.prems(2,3) distinct_fv_list_Fun_param[of g us] by auto thus ?thesis unfolding fun_point_union_def by simp next case False hence "list_ex (\t. timpls_transformable_to TI t (Fun g us \ \ \)) FP" using SolveComposed.prems(1) intruder_synth_mod_timpls.simps(2)[of FP TI g "us \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \"] unfolding list_all_iff by auto then obtain t where t: "t \ set FP" "timpls_transformable_to TI t (Fun g us \ \ \)" unfolding list_ex_iff by blast have t_ground: "fv t = {}" using t(1) FP_ground by simp have g_no_abs: "\is_Abs f" when f: "f \ funs_term (Fun g us)" for f proof - obtain fts where "Fun f fts \ Fun g us" using funs_term_Fun_subterm[OF f] by blast thus ?thesis using SolveComposed.prems(2)[of _ fts] by (cases f) auto qed have g_\_abs: "\a. \ \ y = \a\\<^sub>a\<^sub>b\<^sub>s" when y: "y \ fv (Fun g us)" for y using y SolveComposed.prems(3) unfolding \_def by fastforce obtain \' where \': "match_abss OCC TI (Fun g us) t = Some \'" using g_no_abs g_\_abs timpls_transformable_to_match_abss_case[OF TI t(2) t_ground ] by blast let ?h1 = "\\ x. if x \ fv (Fun g us) then \ x else set OCC" let ?h2 = "\\ x. \(ticl_abs TI ` \ x)" obtain \'' where \'': "match_abss' (Fun g us) t = Some \''" "\' = ?h1 (?h2 \'')" "\x \ fv (Fun g us). \' x \ {} \ \' x \ {}" using match_abssD[OF \'] by blast have \'_\: "\' \ \" using t(1) \' SolveComposed.hyps(3) by metis have "\ x \ ticl_abs TI a" when a: "a \ \'' x" and x_in_g: "x \ fv (Fun g us)" for a proof - have "fst x = TAtom Value" using x_in_g SolveComposed.prems(3) by auto hence "\ \ x = \\ x\\<^sub>a\<^sub>b\<^sub>s" unfolding \_def by simp hence "(a, \ x) \ (set TI)\<^sup>*" using timpls_transformable_to_match_abss'_nonempty_disj'[ OF TI t(2) \''(1) x_in_g t_ground] g_no_abs g_\_abs a by fastforce thus "\ x \ ticl_abs TI a" unfolding ticl_abs_iff[OF TI] by force qed hence "\ x \ \' x" when x_in_g: "x \ fv (Fun g us)" using \''(2,3) x_in_g unfolding \''(1) by simp hence "\ x \ \' x" using match_abss_OCC_if_not_fv[OF \'] \x0(1) by blast hence "\ x \ \1 x" using \'_\ \x0(1) unfolding SolveComposed.hyps(4,5) fun_point_Union_def by auto thus ?thesis unfolding fun_point_union_def by simp qed qed (auto simp add: \x0 fun_point_Inter_def) qed private lemma transaction_check_variant_soundness_aux4: fixes T FP S C xs OCC negs poss as defines "\ \ \\ x. if fst x = TAtom Value then (absc \ \) x else Var x" and "C \ unlabel (transaction_checks T)" and "S \ unlabel (transaction_strand T)" and "xas \ (the_Abs \ the_Fun) ` set (filter (\t. is_Fun t \ is_Abs (the_Fun t)) FP)" and "ts \ trms_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_receive T))" and "xs \ filter (\x. x \ set (transaction_fresh T) \ fst x = TAtom Value) (fv_list\<^sub>s\<^sub>s\<^sub>t S)" and "poss \ transaction_poschecks_comp C" and "negs \ transaction_negchecks_comp C" and "as \ map (\x. (x, set (filter (\ab. poss x \ ab \ negs x \ ab = {}) OCC))) xs" and "f \ \x. case List.find (\p. fst p = x) as of Some p \ snd p | None \ {}" assumes T_adm: "admissible_transaction' T" and TI0: "\(a,b) \ set TI. \(c,d) \ set TI. b = c \ a \ d \ (a,d) \ set TI" "\(a,b) \ set TI. a \ b" and OCC: "\t \ set FP. \a. Abs a \ funs_term t \ a \ set OCC" and FP_ground: "ground (set FP)" and FP_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set FP)" and "x \ set xs" and "\x. x \ set xs \ \ x \ set OCC" and "\x. x \ set xs \ poss x \ \ x" and "\x. x \ set xs \ \ x \ negs x = {}" and "\x. x \ set xs \ \ x = {}" and "\t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T). intruder_synth_mod_timpls FP TI (t \ \ \)" and "\a x s. \a: Var x \ Fun (Set s) []\ \ set C \ s \ \ x" and "\x s. \Var x not in Fun (Set s) []\ \ set C \ s \ \ x" shows "\ x \ synth_abs_substs_constrs (FP,OCC,TI) T x" proof - let ?FPT = "(FP,OCC,TI)" let ?P = "\s u. let \ = mgu s u in \ \ None \ (\x \ fv s. is_Fun (the \ x) \ is_Abs (the_Fun (the \ x)))" define \0 where "\0 \ \x. if fst x = TAtom Value \ x \ fv_transaction T \ x \ set (transaction_fresh T) then {ab \ set OCC. poss x \ ab \ negs x \ ab = {}} else {}" define g where "g \ \x. set (filter (\ab. poss x \ ab \ negs x \ ab = {}) OCC)" define gs where "gs \ map (\x. (x, g x)) xs" note defs = assms(3-10) \0_def note assm = assms(11-)[unfolded defs] have ts0: "\t \ set ts. wf\<^sub>t\<^sub>r\<^sub>m t" using admissible_transaction_is_wellformed_transaction(4)[OF T_adm] unfolding admissible_transaction_terms_def wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code[symmetric] ts_def trms_list\<^sub>s\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>s\<^sub>t[symmetric] trms_transaction_unfold by fast have ts1: "\t \ set ts. \f \ funs_term t. \is_Abs f" using protocol_transactions_no_abss[OF T_adm] funs_term_Fun_subterm trms\<^sub>s\<^sub>s\<^sub>t_unlabel_prefix_subset(1) unfolding ts_def trms_list\<^sub>s\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>s\<^sub>t[symmetric] is_Abs_def transaction_strand_def by (metis (no_types, opaque_lifting) in_subterms_Union in_subterms_subset_Union subset_iff) have ts2: "\x \ fv\<^sub>s\<^sub>e\<^sub>t (set ts). fst x = TAtom Value" using admissible_transaction_Value_vars[OF T_adm] wellformed_transaction_send_receive_fv_subset(1)[ OF admissible_transaction_is_wellformed_transaction(1)[OF T_adm]] unfolding ts_def trms_transaction_unfold trms_list\<^sub>s\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>s\<^sub>t[symmetric] \\<^sub>v_TAtom''(2) by fastforce have "f x = \0 x" for x proof (cases "fst x = Var Value \ x \ fv_transaction T \ x \ set (transaction_fresh T)") case True hence "\0 x = {ab \ set OCC. poss x \ ab \ negs x \ ab = {}}" unfolding \0_def by argo thus ?thesis using True transaction_check_variant_soundness_aux0[OF S_def xs_def, of x] transaction_check_variant_soundness_aux1[ OF C_def S_def xs_def poss_def negs_def as_def f_def, of x] by simp next case False hence 0: "\0 x = {}" unfolding \0_def by argo have "x \ set xs" using False fv_list\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_strand T)"] unfolding xs_def S_def by fastforce hence "List.find (\p. fst p = x) gs = None" using find_None_iff[of "\p. fst p = x" gs] unfolding gs_def by simp hence "f x = {}" unfolding f_def as_def gs_def g_def by force thus ?thesis using 0 by simp qed thus ?thesis using synth_abs_substs_constrs_rel_if_synth_abs_substs_constrs[ OF _ FP_ground FP_wf, of T, unfolded trms_list\<^sub>s\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>s\<^sub>t ts_def[symmetric], OF ts0] transaction_check_variant_soundness_aux3[ OF TI0 OCC FP_ground assm(7-11), of "synth_abs_substs_constrs ?FPT T", unfolded trms_list\<^sub>s\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>s\<^sub>t ts_def[symmetric], OF assm(12)[unfolded \_def trms_list\<^sub>s\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>s\<^sub>t ts_def[symmetric]] ts1 ts2 assm(13-)[unfolded C_def]] unfolding defs synth_abs_substs_constrs_def Let_def by blast qed private lemma transaction_check_variant_soundness_aux5: fixes FP OCC TI T S C defines "msgcs \ \x a. a \ synth_abs_substs_constrs (FP,OCC,TI) T x" and "S \ unlabel (transaction_strand T)" and "C \ unlabel (transaction_checks T)" and "xs \ filter (\x. x \ set (transaction_fresh T) \ fst x = TAtom Value) (fv_list\<^sub>s\<^sub>s\<^sub>t S)" and "poss \ transaction_poschecks_comp C" and "negs \ transaction_negchecks_comp C" assumes T_adm: "admissible_transaction' T" and TI: "\(a,b) \ set TI. \(c,d) \ set TI. b = c \ a \ d \ (a,d) \ set TI" "\(a,b) \ set TI. a \ b" and OCC: "\t \ set FP. \a. Abs a \ funs_term t \ a \ set OCC" and FP: "ground (set FP)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set FP)" and \: "\ \ abs_substs_fun ` set (abs_substs_set xs OCC poss negs (\_ _. True))" "transaction_check_pre (FP,OCC,TI) T \" shows "\ \ abs_substs_fun ` set (abs_substs_set xs OCC poss negs msgcs)" proof - have 0: "\ x \ set OCC" "poss x \ \ x" "\ x \ negs x = {}" when x: "x \ set xs" for x using abs_substs_abss_bounded[OF \(1) x] by simp_all have 1: "\ x = {}" when x: "x \ set xs" for x by (rule abs_substs_abss_bounded'[OF \(1) x]) have 2: "msgcs x (\ x)" when x: "x \ set xs" for x using 0 1 x transaction_check_variant_soundness_aux4[OF T_adm TI OCC FP, of x \] transaction_check_pre_ReceiveE[OF \(2)] transaction_check_pre_InSetE[OF \(2)] transaction_check_pre_NotInSetE[OF \(2)] unfolding msgcs_def xs_def C_def S_def negs_def poss_def by fast show ?thesis using abs_substs_has_abs[of xs \ OCC poss negs msgcs] 0 1 2 by blast qed theorem transaction_check_variant_soundness: assumes P_adm: "\T \ set P. admissible_transaction' T" and TI: "\(a,b) \ set TI. \(c,d) \ set TI. b = c \ a \ d \ (a,d) \ set TI" "\(a,b) \ set TI. a \ b" and OCC: "\t \ set FP. \a. Abs a \ funs_term t \ a \ set OCC" and FP: "ground (set FP)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set FP)" and T_in: "T \ set P" and T_check: "transaction_check_coverage_rcv (FP,OCC,TI) T" shows "transaction_check (FP,OCC,TI) T" proof - have 0: "admissible_transaction' T" using P_adm T_in by blast show ?thesis using T_check transaction_check_variant_soundness_aux5[OF 0 TI OCC FP] unfolding transaction_check_def transaction_check'_def transaction_check_coverage_rcv_def transaction_check_comp_def list_all_iff Let_def by force qed end end subsection \Automatic Fixed-Point Computation\ context stateful_protocol_model begin fun reduce_fixpoint' where "reduce_fixpoint' FP _ [] = FP" | "reduce_fixpoint' FP TI (t#M) = ( let FP' = List.removeAll t FP in if intruder_synth_mod_timpls FP' TI t then FP' else reduce_fixpoint' FP TI M)" definition reduce_fixpoint where "reduce_fixpoint FP TI \ let f = \FP. reduce_fixpoint' FP TI FP in while (\M. set (f M) \ set M) f FP" definition compute_fixpoint_fun' where "compute_fixpoint_fun' P (n::nat option) enable_traces \ S0 \ let P' = map add_occurs_msgs P; sy = intruder_synth_mod_timpls; FP' = \S. fst (fst S); TI' = \S. snd (fst S); OCC' = \S. remdups ( (map (\t. the_Abs (the_Fun (args t ! 1))) (filter (\t. is_Fun t \ the_Fun t = OccursFact) (FP' S)))@ (map snd (TI' S))); equal_states = \S S'. set (FP' S) = set (FP' S') \ set (TI' S) = set (TI' S'); trace' = \S. snd S; close = \M f. let g = remdups \ f in while (\A. set (g A) \ set A) g M; close' = \M f. let g = remdups \ f in while (\A. set (g A) \ set A) g M; trancl_minus_refl = \TI. let aux = \ts p. map (\q. (fst p,snd q)) (filter ((=) (snd p) \ fst) ts) in filter (\p. fst p \ snd p) (close' TI (\ts. concat (map (aux ts) ts)@ts)); snd_Ana = \N M TI. let N' = filter (\t. \k \ set (fst (Ana t)). sy M TI k) N in filter (\t. \sy M TI t) (concat (map (\t. filter (\s. s \ set (snd (Ana t))) (args t)) N')); Ana_cl = \FP TI. close FP (\M. (M@snd_Ana M M TI)); TI_cl = \FP TI. close FP (\M. (M@filter (\t. \sy M TI t) (concat (map (\m. concat (map (\(a,b). \a --\ b\\m\) TI)) M)))); Ana_cl' = \FP TI. let K = \t. set (fst (Ana t)); flt = \M t. (\k \ K t. \sy M TI k) \ (\k \ K t. \f \ funs_term k. is_Abs f); N = \M. comp_timpl_closure_list (filter (flt M) M) TI in close FP (\M. M@snd_Ana (N M) M TI); \' = \S. \ (FP' S, OCC' S, TI' S); result = \S T \. let not_fresh = \x. x \ set (transaction_fresh T); xs = filter not_fresh (fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T))); u = \\ x. absdbupd (unlabel (transaction_strand T)) x (\ x) in (remdups (filter (\t. \sy (FP' S) (TI' S) t) (concat (map (\ts. the_msgs ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (absc \ u \)) (filter is_Send (unlabel (transaction_send T)))))), remdups (filter (\s. fst s \ snd s) (map (\x. (\ x, u \ x)) xs))); result_tuple = \S T \. (result S T (abs_substs_fun \), if enable_traces then \ else []); update_state = \S. if list_ex (\t. is_Fun t \ is_Attack (the_Fun t)) (FP' S) then S else let results = map (\T. map (result_tuple S T) (\' S T)) P'; newtrace_flt = (\n. let x = map fst (results ! n); y = map fst x; z = map snd x in set (concat y) - set (FP' S) \ {} \ set (concat z) - set (TI' S) \ {}); trace = if enable_traces then trace' S@[concat (map (\i. map (\a. (i,snd a)) (results ! i)) (filter newtrace_flt [0..x. fst x \ snd x) (concat (map snd U)@TI' S))), trace); W = ((Ana_cl (TI_cl (FP' V) (TI' V)) (TI' V), trancl_minus_refl (TI' V)), trace' V) in if \equal_states W S then W else ((Ana_cl' (FP' W) (TI' W), TI' W), trace' W); S = ((\h. case n of None \ while (\S. \equal_states S (h S)) h | Some m \ h ^^ m) update_state S0) in ((reduce_fixpoint (FP' S) (TI' S), OCC' S, TI' S), trace' S)" definition compute_fixpoint_fun where "compute_fixpoint_fun P \ let P' = remdups (filter (\T. transaction_updates T \ [] \ transaction_send T \ []) P); f = (\FPT T. let msgcs = synth_abs_substs_constrs FPT T in transaction_check_comp (\x a. a \ msgcs x) FPT T) in fst (compute_fixpoint_fun' P' None False f (([],[]),[]))" definition compute_fixpoint_with_trace where "compute_fixpoint_with_trace P \ compute_fixpoint_fun' P None True (transaction_check_comp (\_ _. True)) (([],[]),[])" definition compute_fixpoint_from_trace where "compute_fixpoint_from_trace P trace \ let P' = map add_occurs_msgs P; \ = \FPT T. let pre_check = transaction_check_pre FPT T; \s = map snd (filter (\(i,as). P' ! i = T) (concat trace)) in filter (\\. pre_check (abs_substs_fun \)) \s; f = compute_fixpoint_fun' \ map (nth P); g = \L FPT. fst ((f L (Some 1) False \ ((fst FPT,snd (snd FPT)),[]))) in fold g (map (map fst) trace) ([],[],[])" definition compute_reduced_attack_trace where "compute_reduced_attack_trace P trace \ let attack_in_fixpoint = list_ex (\t. \f \ funs_term t. is_Attack f) \ fst; is_attack_trace = attack_in_fixpoint \ compute_fixpoint_from_trace P; trace' = let is_attack_transaction = list_ex is_Fun_Attack \ concat \ map the_msgs \ filter is_Send \ unlabel \ transaction_send; trace' = if trace = [] then [] else butlast trace@[filter (is_attack_transaction \ nth P \ fst) (last trace)] in trace'; iter = \trace_prev trace_rest elem (prev,rest). let next = if is_attack_trace (trace_prev@(prev@rest)#trace_rest) then prev else prev@[elem] in (next, tl rest); iter' = \trace_part (trace_prev,trace_rest). let updated = foldr (iter trace_prev (tl trace_rest)) trace_part ([],tl (rev trace_part)) in (trace_prev@[rev (fst updated)], tl trace_rest); reduced_trace = fst (fold iter' trace' ([],trace')) in concat reduced_trace" end subsection \Locales for Protocols Proven Secure through Fixed-Point Coverage\ type_synonym ('f,'a,'s,'l) fixpoint_triple = "('f,'a,'s,'l) prot_term list \ 's set list \ ('s set \ 's set) list" context stateful_protocol_model begin definition "attack_notin_fixpoint (FPT::('fun,'atom,'sets,'lbl) fixpoint_triple) \ list_all (\t. \f \ funs_term t. \is_Attack f) (fst FPT)" definition "protocol_covered_by_fixpoint (FPT::('fun,'atom,'sets,'lbl) fixpoint_triple) P \ list_all (transaction_check FPT) (filter (\T. transaction_updates T \ [] \ transaction_send T \ []) (map add_occurs_msgs P))" definition "protocol_covered_by_fixpoint_coverage_rcv (FPT::('fun,'atom,'sets,'lbl) fixpoint_triple) P \ list_all (transaction_check_coverage_rcv FPT) (filter (\T. transaction_updates T \ [] \ transaction_send T \ []) (map add_occurs_msgs P))" definition "analyzed_fixpoint (FPT::('fun,'atom,'sets,'lbl) fixpoint_triple) \ let (FP, _, TI) = FPT in analyzed_closed_mod_timpls FP TI" definition "wellformed_protocol_SMP_set (P::('fun,'atom,'sets,'lbl) prot) N \ has_all_wt_instances_of \ (\T \ set P. trms_transaction T) (set N) \ comp_tfr\<^sub>s\<^sub>e\<^sub>t arity Ana \ (set N) \ list_all (\T. list_all (comp_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ Pair) (unlabel (transaction_strand T))) P" (* TODO: try to avoid checking the "list_all is_*" conditions here... *) definition "wellformed_protocol'' (P::('fun,'atom,'sets,'lbl) prot) N \ let f = \T. transaction_fresh T = [] \ transaction_updates T \ [] \ transaction_send T \ [] in list_all (\T. list_all is_Receive (unlabel (transaction_receive T)) \ list_all is_Check_or_Assignment (unlabel (transaction_checks T)) \ list_all is_Update (unlabel (transaction_updates T)) \ list_all is_Send (unlabel (transaction_send T))) P \ list_all admissible_transaction (filter f P) \ wellformed_protocol_SMP_set P N" definition "wellformed_protocol' (P::('fun,'atom,'sets,'lbl) prot) N \ wellformed_protocol'' P N \ has_initial_value_producing_transaction P" definition "wellformed_protocol (P::('fun,'atom,'sets,'lbl) prot) \ let f = \M. remdups (concat (map subterms_list M@map (fst \ Ana) M)); N0 = remdups (concat (map (trms_list\<^sub>s\<^sub>s\<^sub>t \ unlabel \ transaction_strand) P)); N = while (\A. set (f A) \ set A) f N0 in wellformed_protocol' P N" definition "wellformed_fixpoint' (FPT::('fun,'atom,'sets,'lbl) fixpoint_triple) \ let (FP, OCC, TI) = FPT; OCC' = set OCC in list_all (\t. wf\<^sub>t\<^sub>r\<^sub>m' arity t \ fv t = {}) FP \ list_all (\a. a \ OCC') (map snd TI) \ list_all (\t. \f \ funs_term t. is_Abs f \ the_Abs f \ OCC') FP" definition "wellformed_term_implication_graph (FPT::('fun,'atom,'sets,'lbl) fixpoint_triple) \ let (_, _, TI) = FPT in list_all (\(a,b). list_all (\(c,d). b = c \ a \ d \ List.member TI (a,d)) TI) TI \ list_all (\p. fst p \ snd p) TI" definition "wellformed_fixpoint (FPT::('fun,'atom,'sets,'lbl) fixpoint_triple) \ wellformed_fixpoint' FPT \ wellformed_term_implication_graph FPT" lemma wellformed_protocol_SMP_set_mono: assumes "wellformed_protocol_SMP_set P S" and "set P' \ set P" shows "wellformed_protocol_SMP_set P' S" using assms unfolding wellformed_protocol_SMP_set_def comp_tfr\<^sub>s\<^sub>e\<^sub>t_def has_all_wt_instances_of_def wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s'_def list_all_iff by fast lemma wellformed_protocol''_mono: assumes "wellformed_protocol'' P S" and "set P' \ set P" shows "wellformed_protocol'' P' S" using assms wellformed_protocol_SMP_set_mono[of P S P'] unfolding wellformed_protocol''_def list_all_iff by auto lemma wellformed_protocol'_mono: assumes "wellformed_protocol' P S" and "set P' \ set P" and "has_initial_value_producing_transaction P'" shows "wellformed_protocol' P' S" using assms wellformed_protocol_SMP_set_mono[of P S P'] wellformed_protocol''_mono unfolding wellformed_protocol'_def by blast lemma protocol_covered_by_fixpoint_if_protocol_covered_by_fixpoint_coverage_rcv: assumes P: "wellformed_protocol'' P P_SMP" and FPT: "wellformed_fixpoint FPT" and covered: "protocol_covered_by_fixpoint_coverage_rcv FPT P" shows "protocol_covered_by_fixpoint FPT P" proof - obtain FP OCC TI where FPT': "FPT = (FP,OCC,TI)" by (metis surj_pair) note defs = FPT' wellformed_protocol''_def wellformed_fixpoint_def wellformed_fixpoint'_def wellformed_term_implication_graph_def Let_def wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code[symmetric] wf\<^sub>t\<^sub>r\<^sub>m_code[symmetric] member_def case_prod_unfold list_all_iff let ?f = "\T. transaction_fresh T = [] \ transaction_updates T \ [] \ transaction_send T \ []" have TI: "\(a,b) \ set TI. \(c,d) \ set TI. b = c \ a \ d \ (a,d) \ set TI" "\(a,b) \ set TI. a \ b" and OCC: "\t \ set FP. \a. Abs a \ funs_term t \ a \ set OCC" and FP: "ground (set FP)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set FP)" using FPT unfolding defs by (simp, simp, fastforce, simp, simp) have P_adm: "\T \ set (filter ?f (map add_occurs_msgs P)). admissible_transaction' T" using P add_occurs_msgs_admissible_occurs_checks(1)[OF admissible_transactionE'(1)] unfolding defs add_occurs_msgs_updates_send_filter_iff'[of P, symmetric] by auto show ?thesis using covered transaction_check_variant_soundness[OF P_adm TI OCC FP] unfolding protocol_covered_by_fixpoint_def protocol_covered_by_fixpoint_coverage_rcv_def FPT' list_all_iff by fastforce qed lemma protocol_covered_by_fixpoint_if_protocol_covered_by_fixpoint_coverage_rcv': assumes P: "wellformed_protocol'' P P_SMP" and P': "set P' \ set P" and FPT: "wellformed_fixpoint FPT" and covered: "protocol_covered_by_fixpoint_coverage_rcv FPT P'" shows "protocol_covered_by_fixpoint FPT P'" using protocol_covered_by_fixpoint_if_protocol_covered_by_fixpoint_coverage_rcv[OF _ FPT covered] wellformed_protocol''_mono[OF P P'] by simp lemma protocol_covered_by_fixpoint_trivial_case: assumes "list_all (\T. transaction_updates T = [] \ transaction_send T = []) (map add_occurs_msgs P)" shows "protocol_covered_by_fixpoint FPT P" using assms by (simp add: list_all_iff transaction_check_trivial_case protocol_covered_by_fixpoint_def) lemma protocol_covered_by_fixpoint_empty[simp]: "protocol_covered_by_fixpoint FPT []" by (simp add: protocol_covered_by_fixpoint_def) lemma protocol_covered_by_fixpoint_Cons[simp]: "protocol_covered_by_fixpoint FPT (T#P) \ transaction_check FPT (add_occurs_msgs T) \ protocol_covered_by_fixpoint FPT P" using transaction_check_trivial_case[of "add_occurs_msgs T"] unfolding protocol_covered_by_fixpoint_def case_prod_unfold by simp lemma protocol_covered_by_fixpoint_append[simp]: "protocol_covered_by_fixpoint FPT (P1@P2) \ protocol_covered_by_fixpoint FPT P1 \ protocol_covered_by_fixpoint FPT P2" by (simp add: protocol_covered_by_fixpoint_def case_prod_unfold) lemma protocol_covered_by_fixpoint_I1[intro]: assumes "list_all (protocol_covered_by_fixpoint FPT) P" shows "protocol_covered_by_fixpoint FPT (concat P)" using assms by (auto simp add: protocol_covered_by_fixpoint_def list_all_iff) lemma protocol_covered_by_fixpoint_I2[intro]: assumes "protocol_covered_by_fixpoint FPT P1" and "protocol_covered_by_fixpoint FPT P2" shows "protocol_covered_by_fixpoint FPT (P1@P2)" using assms by (auto simp add: protocol_covered_by_fixpoint_def) lemma protocol_covered_by_fixpoint_I3: assumes "\T \ set P. \\::('fun,'atom,'sets,'lbl) prot_var \ 'sets set. transaction_check_pre FPT (add_occurs_msgs T) \ \ transaction_check_post FPT (add_occurs_msgs T) \" shows "protocol_covered_by_fixpoint FPT P" using assms unfolding protocol_covered_by_fixpoint_def transaction_check_def transaction_check'_def transaction_check_comp_def list_all_iff Let_def case_prod_unfold Product_Type.fst_conv Product_Type.snd_conv by fastforce lemmas protocol_covered_by_fixpoint_intros = protocol_covered_by_fixpoint_I1 protocol_covered_by_fixpoint_I2 protocol_covered_by_fixpoint_I3 lemma prot_secure_if_prot_checks: fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list" and FP_OCC_TI:: "('fun,'atom,'sets,'lbl) fixpoint_triple" assumes attack_notin_fixpoint: "attack_notin_fixpoint FP_OCC_TI" and transactions_covered: "protocol_covered_by_fixpoint FP_OCC_TI P" and analyzed_fixpoint: "analyzed_fixpoint FP_OCC_TI" and wellformed_protocol: "wellformed_protocol' P N" and wellformed_fixpoint: "wellformed_fixpoint FP_OCC_TI" shows "\\ \ reachable_constraints P. \\. constraint_model \ (\@[(l, send\[attack\n\]\)])" (is "?secure P") proof - define FP where "FP \ let (FP,_,_) = FP_OCC_TI in FP" define OCC where "OCC \ let (_,OCC,_) = FP_OCC_TI in OCC" define TI where "TI \ let (_,_,TI) = FP_OCC_TI in TI" define f where "f \ \T::('fun,'atom,'sets,'lbl) prot_transaction. transaction_fresh T = [] \ transaction_updates T \ [] \ transaction_send T \ []" define g where "g \ \T::('fun,'atom,'sets,'lbl) prot_transaction. transaction_fresh T = [] \ list_ex (\a. is_Update (snd a) \ is_Send (snd a)) (transaction_strand T)" note wellformed_prot_defs = wellformed_protocol'_def wellformed_protocol''_def wellformed_protocol_SMP_set_def have attack_notin_FP: "attack\n\ \ set FP" using attack_notin_fixpoint[unfolded attack_notin_fixpoint_def] unfolding list_all_iff FP_def by force have 1: "\(a,b) \ set TI. \(c,d) \ set TI. b = c \ a \ d \ (a,d) \ set TI" using wellformed_fixpoint unfolding wellformed_fixpoint_def wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code[symmetric] Let_def TI_def list_all_iff member_def case_prod_unfold wellformed_term_implication_graph_def by auto have 0: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set FP)" and 2: "\(a,b) \ set TI. a \ b" and 3: "snd ` set TI \ set OCC" and 4: "\t \ set FP. \f \ funs_term t. is_Abs f \ f \ Abs ` set OCC" and 5: "ground (set FP)" using wellformed_fixpoint unfolding wellformed_fixpoint_def wf\<^sub>t\<^sub>r\<^sub>m_code[symmetric] is_Abs_def the_Abs_def list_all_iff Let_def case_prod_unfold set_map FP_def OCC_def TI_def wellformed_fixpoint'_def wellformed_term_implication_graph_def by (fast, fast, blast, fastforce, simp) have 8: "finite (set N)" and 9: "has_all_wt_instances_of \ (\T \ set (filter g P). trms_transaction T) (set N)" and 10: "tfr\<^sub>s\<^sub>e\<^sub>t (set N)" and 11: "\T \ set (filter f P). list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (transaction_strand T))" and 12: "\T \ set (filter f P). admissible_transaction T" using wellformed_protocol[unfolded wellformed_prot_defs] tfr\<^sub>s\<^sub>e\<^sub>t_if_comp_tfr\<^sub>s\<^sub>e\<^sub>t[of "set N"] unfolding Let_def list_all_iff wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code[symmetric] tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_comp_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p[symmetric] has_all_wt_instances_of_def f_def[symmetric] by (fast, fastforce, fast, fastforce, fast) have 13: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set N)" using wellformed_protocol[unfolded wellformed_prot_defs] finite_SMP_representationD unfolding wf\<^sub>t\<^sub>r\<^sub>m_code[symmetric] wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s'_def comp_tfr\<^sub>s\<^sub>e\<^sub>t_def list_all_iff Let_def by fast have 14: "has_initial_value_producing_transaction (filter g P)" using wellformed_protocol has_initial_value_producing_transaction_update_send_ex_filter unfolding wellformed_protocol'_def Let_def g_def by blast note TI0 = trancl_eqI'[OF 1 2] have "analyzed (timpl_closure_set (set FP) (set TI))" using analyzed_fixpoint[unfolded analyzed_fixpoint_def] analyzed_closed_mod_timpls_is_analyzed_timpl_closure_set[OF TI0 0] unfolding FP_def TI_def by force note FP0 = this 0 5 note OCC0 = funs_term_OCC_TI_subset(1)[OF 4 3] timpl_closure_set_supset'[OF funs_term_OCC_TI_subset(2)[OF 4 3]] note M0 = 9 8 10 13 have "f T \ g T" when T: "T \ set P" for T proof - have *: "list_all stateful_strand_step.is_Receive (unlabel (transaction_receive T))" "list_all is_Check_or_Assignment (unlabel (transaction_checks T))" "list_all is_Update (unlabel (transaction_updates T))" "list_all stateful_strand_step.is_Send (unlabel (transaction_send T))" using T wellformed_protocol unfolding wellformed_protocol_def wellformed_prot_defs Let_def list_all_iff by (fast, fast, fast, fast) show ?thesis using transaction_updates_send_ex_iff[OF *] unfolding f_def g_def by (metis (no_types, lifting) list_ex_cong) qed hence 15: "\T \ set (filter g P). list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (transaction_strand T))" and 16: "\T \ set (filter g P). admissible_transaction T" using 11 12 by auto have "list_all (transaction_check (FP, OCC, TI)) (map add_occurs_msgs (filter g P))" using transactions_covered[unfolded protocol_covered_by_fixpoint_def] transaction_check_trivial_case[of _ FP_OCC_TI] unfolding FP_def OCC_def TI_def list_all_iff Let_def case_prod_unfold by auto note P0 = 16 15 14 this attack_notin_FP show ?thesis using prot_secure_if_fixpoint_covered[OF FP0 OCC0 TI0 M0 P0] reachable_constraints_secure_if_filter_secure_case[unfolded g_def[symmetric]] by fast qed lemma prot_secure_if_prot_checks_coverage_rcv: fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list" and FP_OCC_TI:: "('fun,'atom,'sets,'lbl) fixpoint_triple" assumes attack_notin_fixpoint: "attack_notin_fixpoint FP_OCC_TI" and transactions_covered: "protocol_covered_by_fixpoint_coverage_rcv FP_OCC_TI P" and analyzed_fixpoint: "analyzed_fixpoint FP_OCC_TI" and wellformed_protocol: "wellformed_protocol' P N" and wellformed_fixpoint: "wellformed_fixpoint FP_OCC_TI" shows "\\ \ reachable_constraints P. \\. constraint_model \ (\@[(l, send\[attack\n\]\)])" using prot_secure_if_prot_checks[ OF attack_notin_fixpoint _ analyzed_fixpoint wellformed_protocol wellformed_fixpoint] protocol_covered_by_fixpoint_if_protocol_covered_by_fixpoint_coverage_rcv[ OF _ wellformed_fixpoint transactions_covered] wellformed_protocol[unfolded wellformed_protocol'_def] by blast end locale secure_stateful_protocol = pm: stateful_protocol_model arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 for arity\<^sub>f::"'fun \ nat" and arity\<^sub>s::"'sets \ nat" and public\<^sub>f::"'fun \ bool" and Ana\<^sub>f::"'fun \ ((('fun,'atom::finite,'sets,'lbl) prot_fun, nat) term list \ nat list)" and \\<^sub>f::"'fun \ 'atom option" and label_witness1::"'lbl" and label_witness2::"'lbl" + fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list" and FP_OCC_TI:: "('fun,'atom,'sets,'lbl) fixpoint_triple" and P_SMP::"('fun,'atom,'sets,'lbl) prot_term list" assumes attack_notin_fixpoint: "pm.attack_notin_fixpoint FP_OCC_TI" and transactions_covered: "pm.protocol_covered_by_fixpoint FP_OCC_TI P" and analyzed_fixpoint: "pm.analyzed_fixpoint FP_OCC_TI" and wellformed_protocol: "pm.wellformed_protocol' P P_SMP" and wellformed_fixpoint: "pm.wellformed_fixpoint FP_OCC_TI" begin theorem protocol_secure: "\\ \ pm.reachable_constraints P. \\. pm.constraint_model \ (\@[(l, send\[attack\n\]\)])" by (rule pm.prot_secure_if_prot_checks[OF attack_notin_fixpoint transactions_covered analyzed_fixpoint wellformed_protocol wellformed_fixpoint]) corollary protocol_welltyped_secure: "\\ \ pm.reachable_constraints P. \\. pm.welltyped_constraint_model \ (\@[(l, send\[attack\n\]\)])" using protocol_secure unfolding pm.welltyped_constraint_model_def by fast end locale secure_stateful_protocol' = pm: stateful_protocol_model arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 for arity\<^sub>f::"'fun \ nat" and arity\<^sub>s::"'sets \ nat" and public\<^sub>f::"'fun \ bool" and Ana\<^sub>f::"'fun \ ((('fun,'atom::finite,'sets,'lbl) prot_fun, nat) term list \ nat list)" and \\<^sub>f::"'fun \ 'atom option" and label_witness1::"'lbl" and label_witness2::"'lbl" + fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list" and FP_OCC_TI:: "('fun,'atom,'sets,'lbl) fixpoint_triple" assumes attack_notin_fixpoint': "pm.attack_notin_fixpoint FP_OCC_TI" and transactions_covered': "pm.protocol_covered_by_fixpoint FP_OCC_TI P" and analyzed_fixpoint': "pm.analyzed_fixpoint FP_OCC_TI" and wellformed_protocol': "pm.wellformed_protocol P" and wellformed_fixpoint': "pm.wellformed_fixpoint FP_OCC_TI" begin sublocale secure_stateful_protocol arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 P FP_OCC_TI "let f = \M. remdups (concat (map subterms_list M@map (fst \ pm.Ana) M)); N0 = remdups (concat (map (trms_list\<^sub>s\<^sub>s\<^sub>t \ unlabel \ transaction_strand) P)) in while (\A. set (f A) \ set A) f N0" apply unfold_locales using attack_notin_fixpoint' transactions_covered' analyzed_fixpoint' wellformed_protocol'[unfolded pm.wellformed_protocol_def Let_def] wellformed_fixpoint' unfolding Let_def by blast+ end locale secure_stateful_protocol'' = pm: stateful_protocol_model arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 for arity\<^sub>f::"'fun \ nat" and arity\<^sub>s::"'sets \ nat" and public\<^sub>f::"'fun \ bool" and Ana\<^sub>f::"'fun \ ((('fun,'atom::finite,'sets,'lbl) prot_fun, nat) term list \ nat list)" and \\<^sub>f::"'fun \ 'atom option" and label_witness1::"'lbl" and label_witness2::"'lbl" + fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list" assumes checks: "let FPT = pm.compute_fixpoint_fun P in pm.attack_notin_fixpoint FPT \ pm.protocol_covered_by_fixpoint FPT P \ pm.analyzed_fixpoint FPT \ pm.wellformed_protocol P \ pm.wellformed_fixpoint FPT" begin sublocale secure_stateful_protocol' arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 P "pm.compute_fixpoint_fun P" using checks[unfolded Let_def case_prod_unfold] by unfold_locales meson+ end locale secure_stateful_protocol''' = pm: stateful_protocol_model arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 for arity\<^sub>f::"'fun \ nat" and arity\<^sub>s::"'sets \ nat" and public\<^sub>f::"'fun \ bool" and Ana\<^sub>f::"'fun \ ((('fun,'atom::finite,'sets,'lbl) prot_fun, nat) term list \ nat list)" and \\<^sub>f::"'fun \ 'atom option" and label_witness1::"'lbl" and label_witness2::"'lbl" + fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list" and FP_OCC_TI:: "('fun,'atom,'sets,'lbl) fixpoint_triple" and P_SMP::"('fun,'atom,'sets,'lbl) prot_term list" assumes checks': "let P' = P; FPT = FP_OCC_TI; P'_SMP = P_SMP in pm.attack_notin_fixpoint FPT \ pm.protocol_covered_by_fixpoint FPT P' \ pm.analyzed_fixpoint FPT \ pm.wellformed_protocol' P' P'_SMP \ pm.wellformed_fixpoint FPT" begin sublocale secure_stateful_protocol arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 P FP_OCC_TI P_SMP using checks'[unfolded Let_def case_prod_unfold] by unfold_locales meson+ end locale secure_stateful_protocol'''' = pm: stateful_protocol_model arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 for arity\<^sub>f::"'fun \ nat" and arity\<^sub>s::"'sets \ nat" and public\<^sub>f::"'fun \ bool" and Ana\<^sub>f::"'fun \ ((('fun,'atom::finite,'sets,'lbl) prot_fun, nat) term list \ nat list)" and \\<^sub>f::"'fun \ 'atom option" and label_witness1::"'lbl" and label_witness2::"'lbl" + fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list" and FP_OCC_TI:: "('fun,'atom,'sets,'lbl) fixpoint_triple" assumes checks'': "let P' = P; FPT = FP_OCC_TI in pm.attack_notin_fixpoint FPT \ pm.protocol_covered_by_fixpoint FPT P' \ pm.analyzed_fixpoint FPT \ pm.wellformed_protocol P' \ pm.wellformed_fixpoint FPT" begin sublocale secure_stateful_protocol' arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 P FP_OCC_TI using checks''[unfolded Let_def case_prod_unfold] by unfold_locales meson+ end locale secure_stateful_protocol_coverage_rcv = pm: stateful_protocol_model arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 for arity\<^sub>f::"'fun \ nat" and arity\<^sub>s::"'sets \ nat" and public\<^sub>f::"'fun \ bool" and Ana\<^sub>f::"'fun \ ((('fun,'atom::finite,'sets,'lbl) prot_fun, nat) term list \ nat list)" and \\<^sub>f::"'fun \ 'atom option" and label_witness1::"'lbl" and label_witness2::"'lbl" + fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list" and FP_OCC_TI:: "('fun,'atom,'sets,'lbl) fixpoint_triple" and P_SMP::"('fun,'atom,'sets,'lbl) prot_term list" assumes attack_notin_fixpoint_coverage_rcv: "pm.attack_notin_fixpoint FP_OCC_TI" and transactions_covered_coverage_rcv: "pm.protocol_covered_by_fixpoint_coverage_rcv FP_OCC_TI P" and analyzed_fixpoint_coverage_rcv: "pm.analyzed_fixpoint FP_OCC_TI" and wellformed_protocol_coverage_rcv: "pm.wellformed_protocol' P P_SMP" and wellformed_fixpoint_coverage_rcv: "pm.wellformed_fixpoint FP_OCC_TI" begin sublocale secure_stateful_protocol arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 P FP_OCC_TI P_SMP using pm.protocol_covered_by_fixpoint_if_protocol_covered_by_fixpoint_coverage_rcv[ OF _ wellformed_fixpoint_coverage_rcv transactions_covered_coverage_rcv] attack_notin_fixpoint_coverage_rcv analyzed_fixpoint_coverage_rcv wellformed_protocol_coverage_rcv wellformed_fixpoint_coverage_rcv wellformed_protocol_coverage_rcv[unfolded pm.wellformed_protocol'_def] by unfold_locales meson+ end locale secure_stateful_protocol_coverage_rcv' = pm: stateful_protocol_model arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 for arity\<^sub>f::"'fun \ nat" and arity\<^sub>s::"'sets \ nat" and public\<^sub>f::"'fun \ bool" and Ana\<^sub>f::"'fun \ ((('fun,'atom::finite,'sets,'lbl) prot_fun, nat) term list \ nat list)" and \\<^sub>f::"'fun \ 'atom option" and label_witness1::"'lbl" and label_witness2::"'lbl" + fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list" and FP_OCC_TI:: "('fun,'atom,'sets,'lbl) fixpoint_triple" assumes attack_notin_fixpoint_coverage_rcv': "pm.attack_notin_fixpoint FP_OCC_TI" and transactions_covered_coverage_rcv': "pm.protocol_covered_by_fixpoint_coverage_rcv FP_OCC_TI P" and analyzed_fixpoint_coverage_rcv': "pm.analyzed_fixpoint FP_OCC_TI" and wellformed_protocol_coverage_rcv': "pm.wellformed_protocol P" and wellformed_fixpoint_coverage_rcv': "pm.wellformed_fixpoint FP_OCC_TI" begin sublocale secure_stateful_protocol_coverage_rcv arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 P FP_OCC_TI "let f = \M. remdups (concat (map subterms_list M@map (fst \ pm.Ana) M)); N0 = remdups (concat (map (trms_list\<^sub>s\<^sub>s\<^sub>t \ unlabel \ transaction_strand) P)) in while (\A. set (f A) \ set A) f N0" apply unfold_locales using attack_notin_fixpoint_coverage_rcv' transactions_covered_coverage_rcv' analyzed_fixpoint_coverage_rcv' wellformed_protocol_coverage_rcv'[unfolded pm.wellformed_protocol_def Let_def] wellformed_fixpoint_coverage_rcv' unfolding Let_def by blast+ end locale secure_stateful_protocol_coverage_rcv'' = pm: stateful_protocol_model arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 for arity\<^sub>f::"'fun \ nat" and arity\<^sub>s::"'sets \ nat" and public\<^sub>f::"'fun \ bool" and Ana\<^sub>f::"'fun \ ((('fun,'atom::finite,'sets,'lbl) prot_fun, nat) term list \ nat list)" and \\<^sub>f::"'fun \ 'atom option" and label_witness1::"'lbl" and label_witness2::"'lbl" + fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list" assumes checks_coverage_rcv: "let FPT = pm.compute_fixpoint_fun P in pm.attack_notin_fixpoint FPT \ pm.protocol_covered_by_fixpoint_coverage_rcv FPT P \ pm.analyzed_fixpoint FPT \ pm.wellformed_protocol P \ pm.wellformed_fixpoint FPT" begin sublocale secure_stateful_protocol_coverage_rcv' arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 P "pm.compute_fixpoint_fun P" using checks_coverage_rcv[unfolded Let_def case_prod_unfold] by unfold_locales meson+ end locale secure_stateful_protocol_coverage_rcv''' = pm: stateful_protocol_model arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 for arity\<^sub>f::"'fun \ nat" and arity\<^sub>s::"'sets \ nat" and public\<^sub>f::"'fun \ bool" and Ana\<^sub>f::"'fun \ ((('fun,'atom::finite,'sets,'lbl) prot_fun, nat) term list \ nat list)" and \\<^sub>f::"'fun \ 'atom option" and label_witness1::"'lbl" and label_witness2::"'lbl" + fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list" and FP_OCC_TI:: "('fun,'atom,'sets,'lbl) fixpoint_triple" and P_SMP::"('fun,'atom,'sets,'lbl) prot_term list" assumes checks_coverage_rcv': "let P' = P; FPT = FP_OCC_TI; P'_SMP = P_SMP in pm.attack_notin_fixpoint FPT \ pm.protocol_covered_by_fixpoint_coverage_rcv FPT P' \ pm.analyzed_fixpoint FPT \ pm.wellformed_protocol' P' P'_SMP \ pm.wellformed_fixpoint FPT" begin sublocale secure_stateful_protocol_coverage_rcv arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 P FP_OCC_TI P_SMP using checks_coverage_rcv'[unfolded Let_def case_prod_unfold] by unfold_locales meson+ end locale secure_stateful_protocol_coverage_rcv'''' = pm: stateful_protocol_model arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 for arity\<^sub>f::"'fun \ nat" and arity\<^sub>s::"'sets \ nat" and public\<^sub>f::"'fun \ bool" and Ana\<^sub>f::"'fun \ ((('fun,'atom::finite,'sets,'lbl) prot_fun, nat) term list \ nat list)" and \\<^sub>f::"'fun \ 'atom option" and label_witness1::"'lbl" and label_witness2::"'lbl" + fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list" and FP_OCC_TI:: "('fun,'atom,'sets,'lbl) fixpoint_triple" assumes checks_coverage_rcv'': "let P' = P; FPT = FP_OCC_TI in pm.attack_notin_fixpoint FPT \ pm.protocol_covered_by_fixpoint_coverage_rcv FPT P' \ pm.analyzed_fixpoint FPT \ pm.wellformed_protocol P' \ pm.wellformed_fixpoint FPT" begin sublocale secure_stateful_protocol_coverage_rcv' arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 P FP_OCC_TI using checks_coverage_rcv''[unfolded Let_def case_prod_unfold] by unfold_locales meson+ end subsection \Automatic Protocol Composition\ context stateful_protocol_model begin definition welltyped_leakage_free_protocol where "welltyped_leakage_free_protocol S P \ let f = \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}; Sec = (f (set S)) - {m. {} \\<^sub>c m} in \\ \ reachable_constraints P. \\\<^sub>\ s. (\l ts. suffix [(l, receive\ts\)] \) \ s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>\ \ welltyped_constraint_model \\<^sub>\ (\@[\\, send\[s]\\])" definition wellformed_composable_protocols where "wellformed_composable_protocols (P::('fun,'atom,'sets,'lbl) prot list) N \ let Ts = concat P; steps = remdups (concat (map transaction_strand Ts)); MP0 = \T \ set Ts. trms_transaction T \ pair' Pair ` setops_transaction T in list_all (wf\<^sub>t\<^sub>r\<^sub>m' arity) N \ has_all_wt_instances_of \ MP0 (set N) \ comp_tfr\<^sub>s\<^sub>e\<^sub>t arity Ana \ (set N) \ list_all (comp_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ Pair \ snd) steps \ list_all admissible_transaction_terms Ts \ list_all (list_all (\x. \\<^sub>v x = TAtom Value \ (is_Var (\\<^sub>v x) \ is_Atom (the_Var (\\<^sub>v x)))) \ transaction_fresh) Ts \ list_all (\T. \x \ vars_transaction T. \TAtom AttackType \ \\<^sub>v x) Ts \ list_all (\T. \x \ vars_transaction T. \f \ funs_term (\\<^sub>v x). f \ Pair \ f \ OccursFact) Ts \ list_all (list_all (\s. is_Send (snd s) \ length (the_msgs (snd s)) = 1 \ is_Fun_Attack (hd (the_msgs (snd s))) \ the_Attack_label (the_Fun (hd (the_msgs (snd s)))) = fst s) \ transaction_strand) Ts \ list_all (\r. (\f \ \(funs_term ` (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd r))). f = OccursFact \ f = OccursSec) \ (is_Receive (snd r) \ is_Send (snd r)) \ fst r = \ \ (\t \ set (the_msgs (snd r)). (OccursFact \ funs_term t \ OccursSec \ funs_term t) \ is_Fun t \ length (args t) = 2 \ t = occurs (args t ! 1) \ is_Var (args t ! 1) \ (\ (args t ! 1) = TAtom Value))) steps" definition wellformed_composable_protocols' where "wellformed_composable_protocols' (P::('fun,'atom,'sets,'lbl) prot list) \ let Ts = concat P in list_all wellformed_transaction Ts \ list_all (list_all (\p. let (x,cs) = p in is_Var (\\<^sub>v x) \ is_Atom (the_Var (\\<^sub>v x)) \ (\c \ cs. \\<^sub>v x = \ (Fun (Fu c) []::('fun,'atom,'sets,'lbl) prot_term))) \ (\T. transaction_decl T ())) Ts" definition composable_protocols where "composable_protocols (P::('fun,'atom,'sets,'lbl) prot list) Ms S \ let steps = concat (map transaction_strand (concat P)); M_fun = (\l. case find ((=) l \ fst) Ms of Some M \ set (snd M) | None \ {}) in comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t public arity Ana \ Pair steps M_fun (set S)" lemma composable_protocols_par_comp_constr: fixes S f defines "f \ \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}" and "Sec \ (f (set S)) - {m. intruder_synth {} m}" assumes Ps_pc: "wellformed_composable_protocols Ps N" "wellformed_composable_protocols' Ps" "composable_protocols Ps Ms S" shows "\\ \ reachable_constraints (concat Ps). \\. constraint_model \ \ \ (\\\<^sub>\. welltyped_constraint_model \\<^sub>\ \ \ ((\n. welltyped_constraint_model \\<^sub>\ (proj n \)) \ (\\' l t. prefix \' \ \ suffix [(l, receive\t\)] \' \ strand_leaks\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' Sec \\<^sub>\)))" (is "\\ \ _. \_. _ \ ?Q \ \") proof (intro allI ballI impI) fix \ \ assume \: "\ \ reachable_constraints (concat Ps)" and \: "constraint_model \ \" let ?Ts = "concat Ps" let ?steps = "concat (map transaction_strand ?Ts)" let ?MP0 = "\T \ set ?Ts. trms_transaction T \ pair' Pair ` setops_transaction T" let ?M_fun = "\l. case find ((=) l \ fst) Ms of Some M \ set (snd M) | None \ {}" have M: "has_all_wt_instances_of \ ?MP0 (set N)" "finite (set N)" "tfr\<^sub>s\<^sub>e\<^sub>t (set N)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set N)" using Ps_pc tfr\<^sub>s\<^sub>e\<^sub>t_if_comp_tfr\<^sub>s\<^sub>e\<^sub>t[of "set N"] unfolding composable_protocols_def wellformed_composable_protocols_def Let_def list_all_iff wf\<^sub>t\<^sub>r\<^sub>m_code[symmetric] by fast+ have P: "\T \ set ?Ts. wellformed_transaction T" "\T \ set ?Ts. wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (trms_transaction T)" "\T \ set ?Ts. list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (transaction_strand T))" "comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t public arity Ana \ Pair ?steps ?M_fun (set S)" using Ps_pc tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_comp_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p unfolding wellformed_composable_protocols_def wellformed_composable_protocols'_def composable_protocols_def Let_def list_all_iff unlabel_def wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code[symmetric] admissible_transaction_terms_def by (meson, meson, fastforce, blast) show "?Q \ \" using reachable_constraints_par_comp_constr[OF M P \ \] unfolding Sec_def f_def by fast qed context begin private lemma reachable_constraints_no_leakage_alt_aux: fixes P lbls L defines "lbls \ \T. map (the_LabelN o fst) (filter (Not o has_LabelS) (transaction_strand T))" and "L \ set (remdups (concat (map lbls P)))" assumes l: "l \ L" shows "map (transaction_proj l) P = map transaction_star_proj P" proof - have 0: "\list_ex (has_LabelN l) (transaction_strand T)" when "T \ set P" for T using that l unfolding L_def lbls_def list_ex_iff by force have 1: "\list_ex (has_LabelN l) (transaction_strand T)" when T: "T \ set (map (transaction_proj l) P)" for T proof - obtain T' where T': "T' \ set P" "T = transaction_proj l T'" using T by auto show ?thesis using T'(2) 0[OF T'(1)] proj_set_subset[of l "transaction_strand T'"] transaction_strand_proj[of l T'] unfolding list_ex_iff by fastforce qed have "list_all has_LabelS (transaction_strand T)" when "T \ set (map (transaction_proj l) P)" for T using that 1[OF that] transaction_proj_idem[of l] transaction_strand_proj[of l "transaction_proj l T"] has_LabelS_proj_iff_not_has_LabelN[of l "transaction_strand (transaction_proj l T)"] by (metis (no_types) ex_map_conv) thus ?thesis using transaction_star_proj_ident_iff transaction_proj_member transaction_star_proj_negates_transaction_proj(1) by (metis (mono_tags, lifting) map_eq_conv) qed private lemma reachable_constraints_star_no_leakage: fixes Sec P lbls k defines "no_leakage \ \\. \\\<^sub>\ \' s. prefix \' \ \ (\l ts. suffix [(l, receive\ts\)] \') \ s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\ \ welltyped_constraint_model \\<^sub>\ (\'@[(k,send\[s]\)])" assumes Sec: "\s \ Sec. \{} \\<^sub>c s" "ground Sec" shows "\\ \ reachable_constraints (map transaction_star_proj P). no_leakage \" proof fix A assume A: "A \ reachable_constraints (map transaction_star_proj P)" have A': "\(l,a) \ set A. l = \" using reachable_constraints_preserves_labels[OF A] transaction_star_proj_has_star_labels unfolding list_all_iff by fastforce show "no_leakage A" using constr_sem_stateful_star_proj_no_leakage[OF Sec(2) A'] unlabel_append[of A] singleton_lst_proj(4)[of k] unfolding no_leakage_def welltyped_constraint_model_def constraint_model_def by fastforce qed private lemma reachable_constraints_no_leakage_alt: fixes Sec P lbls k defines "no_leakage \ \\. \\\<^sub>\ \' s. prefix \' \ \ (\l ts. suffix [(l, receive\ts\)] \') \ s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\ \ welltyped_constraint_model \\<^sub>\ (\'@[(k,send\[s]\)])" and "lbls \ \T. map (the_LabelN o fst) (filter (Not o has_LabelS) (transaction_strand T))" and "L \ set (remdups (concat (map lbls P)))" assumes Sec: "\s \ Sec. \{} \\<^sub>c s" "ground Sec" and lbl: "\l \ L. \\ \ reachable_constraints (map (transaction_proj l) P). no_leakage \" shows "\l. \\ \ reachable_constraints (map (transaction_proj l) P). \\\<^sub>\ \'. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\) \ prefix \' \ \ (\l' ts. suffix [(l', receive\ts\)] \') \ strand_leaks\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' Sec \\<^sub>\" proof (intro allI ballI) fix l \ assume \: "\ \ reachable_constraints (map (transaction_proj l) P)" let ?Q = "\\\<^sub>\ \'. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\) \ prefix \' \ \ (\l' t. suffix [(l', receive\t\)] \') \ strand_leaks\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' Sec \\<^sub>\" show "\\\<^sub>\ \'. ?Q \\<^sub>\ \'" proof assume "\\\<^sub>\ \'. ?Q \\<^sub>\ \'" then obtain \\<^sub>\ \' t n l' ts' where \\<^sub>\: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)" and \': "prefix \' \" "suffix [(l', receive\ts'\)] \'" and t: "t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\" and n: "constr_sem_stateful \\<^sub>\ (proj_unl n \'@[send\[t]\])" unfolding strand_leaks\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by blast hence 0: "welltyped_constraint_model \\<^sub>\ (proj n \'@[(m,send\[t]\)])" for m unfolding welltyped_constraint_model_def constraint_model_def by fastforce have t_Sec: "\{} \\<^sub>c t" "t \ \\<^sub>\ = t" using t Sec subst_ground_ident[of t \\<^sub>\] by auto obtain B k' s where B: "constr_sem_stateful \\<^sub>\ (proj_unl n B@[send\[t]\])" "prefix (proj n B) (proj n \)" "suffix [(k', receive\s\)] (proj n B)" "t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n B) \\<^sub>\" using constr_sem_stateful_proj_priv_term_prefix_obtain[OF \'(1) n t t_Sec] by metis hence 1: "welltyped_constraint_model \\<^sub>\ (proj n B@[(m,send\[t]\)])" for m using 0 unfolding welltyped_constraint_model_def constraint_model_def by fastforce note 2 = reachable_constraints_transaction_proj_proj_eq note 3 = reachable_constraints_transaction_proj_star_proj note 4 = reachable_constraints_no_leakage_alt_aux note star_case = 0 t t_Sec(1) reachable_constraints_star_no_leakage[OF Sec] \'(2) 3[OF \] prefix_proj(1)[OF \'(1)] suffix_proj(1)[OF \'(2)] declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_proj_eq note lbl_case = 0 t(1) \ \' lbl 2(2)[OF \ \'(1)] show False proof (cases "l = n") case True thus ?thesis proof (cases "l \ L") case False hence "map (transaction_proj l) P = map transaction_star_proj P" using 4 unfolding L_def lbls_def by fast thus ?thesis using lbl_case(1-4,7) star_case(4,5) True by metis qed (metis lbl_case no_leakage_def) next case False hence "no_leakage (proj n \)" using star_case(4,6) unfolding no_leakage_def by fast thus ?thesis by (metis B(2-4) 1 no_leakage_def) qed qed qed private lemma reachable_constraints_no_leakage_alt'_aux1: fixes P::"('a,'b,'c,'d) prot_transaction list" defines "f \ list_all ((list_all (Not \ has_LabelS)) \ tl \ transaction_send)" assumes P: "f P" shows "f (map (transaction_proj l) P)" and "f (map transaction_star_proj P)" proof - let ?g = "\T. tl (transaction_send T)" have "set (?g (transaction_proj l T)) \ set (?g T)" (is "?A \ ?C") and "set (?g (transaction_star_proj T)) \ set (?g T)" (is "?B \ ?C") for T::"('a,'b,'c,'d) prot_transaction" proof - obtain T1 T2 T3 T4 T5 T6 where T: "T = Transaction T1 T2 T3 T4 T5 T6" by (cases T) simp have "transaction_send (transaction_proj l T) = proj l (transaction_send T)" "transaction_send (transaction_star_proj T) = filter has_LabelS (transaction_send T)" using transaction_proj.simps[of l T1 T2 T3 T4 T5 T6] transaction_star_proj.simps[of T1 T2 T3 T4 T5 T6] unfolding T proj_def Let_def by auto hence "set (?g (transaction_proj l T)) \ set (proj l (?g T))" "set (?g (transaction_star_proj T)) \ set (filter has_LabelS (?g T))" unfolding proj_def by (metis (no_types, lifting) filter.simps(2) list.collapse list.sel(2,3) list.set_sel(2) subsetI)+ thus "?A \ ?C" "?B \ ?C" using T unfolding proj_def by auto qed thus "f (map (transaction_proj l) P)" "f (map transaction_star_proj P)" using P unfolding f_def list_all_iff by fastforce+ qed private lemma reachable_constraints_no_leakage_alt'_aux2: fixes P defines "f \ \T. list_all is_Receive (unlabel (transaction_receive T)) \ list_all is_Check_or_Assignment (unlabel (transaction_checks T)) \ list_all is_Update (unlabel (transaction_updates T)) \ list_all is_Send (unlabel (transaction_send T))" assumes P: "list_all f P" shows "list_all f (map (transaction_proj l) P)" (is ?A) and "list_all f (map transaction_star_proj P)" (is ?B) proof - have "f (transaction_proj l T)" (is ?A') and "f (transaction_star_proj T)" (is ?B') when T_in: "T \ set P" for T proof - obtain T1 T2 T3 T4 T5 T6 where T: "T = Transaction T1 T2 T3 T4 T5 T6" by (cases T) have "f T" using P T_in unfolding list_all_iff by simp thus ?A' ?B' unfolding f_def T unlabel_def proj_def Let_def list_all_iff transaction_proj.simps[of l T1 T2 T3 T4 T5 T6] transaction_star_proj.simps[of T1 T2 T3 T4 T5 T6] by auto qed thus ?A ?B unfolding list_all_iff by auto qed private lemma reachable_constraints_no_leakage_alt': fixes Sec P lbls k defines "no_leakage \ \\. \\\<^sub>\ \' s. prefix \' \ \ (\l ts. suffix [(l, receive\ts\)] \') \ s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\ \ welltyped_constraint_model \\<^sub>\ (\'@[(k,send\[s]\)])" and "no_leakage' \ \\. \\\<^sub>\ s. (\l ts. suffix [(l, receive\ts\)] \) \ s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>\ \ welltyped_constraint_model \\<^sub>\ (\@[(k,send\[s]\)])" assumes P: "list_all wellformed_transaction P" "list_all ((list_all (Not \ has_LabelS)) \ tl \ transaction_send) P" and Sec: "\s \ Sec. \{} \\<^sub>c s" "ground Sec" and lbl: "\l \ L. \\ \ reachable_constraints (map (transaction_proj l) P). no_leakage' \" shows "\l \ L. \\ \ reachable_constraints (map (transaction_proj l) P). no_leakage \" (is ?A) and "\\ \ reachable_constraints (map transaction_star_proj P). no_leakage \" (is ?B) proof - define f where "f \ \T::('fun,'atom,'sets,'lbl) prot_transaction. list_all is_Receive (unlabel (transaction_receive T)) \ list_all is_Check_or_Assignment (unlabel (transaction_checks T)) \ list_all is_Update (unlabel (transaction_updates T)) \ list_all is_Send (unlabel (transaction_send T))" define g where "(g::('fun,'atom,'sets,'lbl) prot_transaction \ bool) \ list_all (Not \ has_LabelS) \ tl \ transaction_send" have P': "list_all f P" using P(1) unfolding wellformed_transaction_def f_def list_all_iff by fastforce note 0 = reachable_constraints_no_leakage_alt'_aux1[OF P(2), unfolded g_def[symmetric]] note 1 = reachable_constraints_no_leakage_alt'_aux2[ OF P'[unfolded f_def], unfolded f_def[symmetric]] note 2 = reachable_constraints_aligned_prefix_ex[unfolded f_def[symmetric] g_def[symmetric]] have 3: "\\ \ reachable_constraints (map transaction_star_proj P). no_leakage' \" using reachable_constraints_star_no_leakage[OF Sec] unfolding no_leakage'_def by blast show ?A proof (intro ballI) fix l \ assume l: "l \ L" and \: "\ \ reachable_constraints (map (transaction_proj l) P)" show "no_leakage \" proof (rule ccontr) assume "\no_leakage \" then obtain \\<^sub>\ \' s where \': "prefix \' \" "\l ts. suffix [(l, receive\ts\)] \'" "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\" "welltyped_constraint_model \\<^sub>\ (\'@[(k, send\[s]\)])" unfolding no_leakage_def by blast have s: "\{} \\<^sub>c s" "fv s = {}" using \'(3) Sec by auto have \\<^sub>\: "constr_sem_stateful \\<^sub>\ (unlabel \'@[send\[s]\])" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)" using \'(4) unfolding welltyped_constraint_model_def constraint_model_def by auto show False using 2[OF 1(1) 0(1) s \ \'(1,2) \\<^sub>\(1)] l lbl \'(3) \\<^sub>\(2,3,4) singleton_lst_proj(4)[of k "send\[s]\"] unlabel_append[of _ "[(k, send\[s]\)]"] unfolding no_leakage'_def welltyped_constraint_model_def constraint_model_def by metis qed qed show ?B proof (intro ballI) fix \ assume \: "\ \ reachable_constraints (map transaction_star_proj P)" show "no_leakage \" proof (rule ccontr) assume "\no_leakage \" then obtain \\<^sub>\ \' s where \': "prefix \' \" "\l ts. suffix [(l, receive\ts\)] \'" "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\" "welltyped_constraint_model \\<^sub>\ (\'@[(k, send\[s]\)])" unfolding no_leakage_def by blast have s: "\{} \\<^sub>c s" "fv s = {}" using \'(3) Sec by auto have \\<^sub>\: "constr_sem_stateful \\<^sub>\ (unlabel \'@[send\[s]\])" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)" using \'(4) unfolding welltyped_constraint_model_def constraint_model_def by auto show False using 2[OF 1(2) 0(2) s \ \'(1,2) \\<^sub>\(1)] 3 \'(3) \\<^sub>\(2,3,4) singleton_lst_proj(4)[of k "send\[s]\"] unlabel_append[of _ "[(k, send\[s]\)]"] unfolding no_leakage'_def welltyped_constraint_model_def constraint_model_def by metis qed qed qed lemma composable_protocols_par_comp_prot_alt: fixes S f Sec lbls Ps defines "f \ \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}" and "Sec \ (f (set S)) - {m. {} \\<^sub>c m}" and "lbls \ \T. map (the_LabelN o fst) (filter (Not o has_LabelS) (transaction_strand T))" and "L \ set (remdups (concat (map lbls (concat Ps))))" and "no_leakage \ \\. \\\<^sub>\ \' s. prefix \' \ \ (\l ts. suffix [(l, receive\ts\)] \') \ s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\ \ welltyped_constraint_model \\<^sub>\ (\'@[\\, send\[s]\\])" assumes Ps_pc: "wellformed_composable_protocols Ps N" "wellformed_composable_protocols' Ps" "composable_protocols Ps Ms S" and component_secure: "\\ \ reachable_constraints (map (transaction_proj l) (concat Ps)). \\. welltyped_constraint_model \ (\@[\l, send\[attack\ln l\]\\])" and no_leakage: "\l \ L. \\ \ reachable_constraints (map (transaction_proj l) (concat Ps)). no_leakage \" shows "\\ \ reachable_constraints (concat Ps). \\. constraint_model \ (\@[\l, send\[attack\ln l\]\\])" proof fix \ assume \: "\ \ reachable_constraints (concat Ps)" let ?att = "[\l, send\[attack\ln l\]\\]" define Q where "Q \ \\\<^sub>\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)" define R where "R \ \\ \\<^sub>\. \\' l t. prefix \' \ \ suffix [(l, receive\t\)] \' \ strand_leaks\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' Sec \\<^sub>\" define M where "M \ \T\set (concat Ps). trms_transaction T \ pair' Pair ` setops_transaction T" have Sec: "\s \ Sec. \{} \\<^sub>c s" "ground Sec" unfolding Sec_def f_def by auto have par_comp': "\\ \ reachable_constraints (concat Ps). \\. constraint_model \ \ \ (\\\<^sub>\. welltyped_constraint_model \\<^sub>\ \ \ ((\n. welltyped_constraint_model \\<^sub>\ (proj n \)) \ R \ \\<^sub>\))" using \ composable_protocols_par_comp_constr[OF Ps_pc] unfolding Sec_def f_def R_def by fast have "\l. \\ \ reachable_constraints (map (transaction_proj l) (concat Ps)). \\\<^sub>\. Q \\<^sub>\ \ R \ \\<^sub>\" using reachable_constraints_no_leakage_alt[OF Sec no_leakage[unfolded no_leakage_def L_def lbls_def]] unfolding Q_def R_def by blast hence no_leakage': "\\ \ reachable_constraints (concat Ps). \\\<^sub>\. Q \\<^sub>\ \ R \ \\<^sub>\" using reachable_constraints_component_leaks_if_composed_leaks[OF Sec, of "concat Ps" "\\\<^sub>\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)"] unfolding Q_def R_def by blast have M: "has_all_wt_instances_of \ M (set N)" "finite (set N)" "tfr\<^sub>s\<^sub>e\<^sub>t (set N)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set N)" and P: "\T \ set (concat Ps). wellformed_transaction T" "\T \ set (concat Ps). admissible_transaction_terms T" "\T \ set (concat Ps). \x \ vars_transaction T. \TAtom AttackType \ \\<^sub>v x" "\T \ set (concat Ps). \s \ set (transaction_strand T). is_Send (snd s) \ length (the_msgs (snd s)) = 1 \ is_Fun_Attack (hd (the_msgs (snd s))) \ the_Attack_label (the_Fun (hd (the_msgs (snd s)))) = fst s" "\T \ set (concat Ps). list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (transaction_strand T))" using Ps_pc(1,2) tfr\<^sub>s\<^sub>e\<^sub>t_if_comp_tfr\<^sub>s\<^sub>e\<^sub>t tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_comp_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p unfolding wellformed_composable_protocols_def wellformed_composable_protocols'_def list_all_iff Let_def M_def wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s'_def wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code unlabel_def \\<^sub>v_TAtom''(1,2) by (force, force, fast, fast, fast, fast, fast, simp, simp) have P_fresh: "\T \ set (concat Ps). \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" (is "\T \ ?P. \x \ ?frsh T. ?Q x") proof (intro ballI) fix T x assume T: "T \ ?P" "x \ ?frsh T" hence "\\<^sub>v x = TAtom Value \ (is_Var (\\<^sub>v x) \ is_Atom (the_Var (\\<^sub>v x)))" using Ps_pc(1) unfolding wellformed_composable_protocols_def list_all_iff Let_def by fastforce thus "?Q x" by (metis prot_atom.is_Atom_def term.collapse(1)) qed have P': "\T \ set (concat Ps). wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (trms_transaction T)" using P(2) admissible_transaction_terms_def by fast have "\welltyped_constraint_model \ (\@?att)" for \ proof assume "welltyped_constraint_model \ (\@?att)" hence \: "welltyped_constraint_model \ \" "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ attack\ln l\" using strand_sem_append_stateful[of "{}" "{}" "unlabel \" "unlabel ?att"] unlabel_append[of \ ?att] unfolding welltyped_constraint_model_def constraint_model_def by auto obtain \\<^sub>\ where \\<^sub>\: "welltyped_constraint_model \\<^sub>\ \" "welltyped_constraint_model \\<^sub>\ (proj l \)" using \ \ no_leakage' par_comp' unfolding Q_def welltyped_constraint_model_def constraint_model_def by metis have "\l, receive\[attack\ln l\]\\ \ set \" using reachable_constraints_receive_attack_if_attack(3)[OF \ P(1-2) P_fresh P(3) \ P(4)] by auto hence "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj l \) \\<^sub>s\<^sub>e\<^sub>t \\<^sub>\ \ attack\ln l\" using in_proj_set[of l "receive\[attack\ln l\]\" \] in_ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t_iff[of "attack\ln l\" "proj l \"] intruder_deduct.Axiom[of "attack\ln l\" "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj l \) \\<^sub>s\<^sub>e\<^sub>t \\<^sub>\"] by fastforce hence "welltyped_constraint_model \\<^sub>\ (proj l \@?att)" using \\<^sub>\ strand_sem_append_stateful[of "{}" "{}" "unlabel (proj l \)" "unlabel ?att" \\<^sub>\] unfolding welltyped_constraint_model_def constraint_model_def by auto thus False using component_secure reachable_constraints_transaction_proj[OF \, of l] by simp qed thus "\\. constraint_model \ (\@?att)" using reachable_constraints_typing_result'[OF M_def M P(1) P' P(5) \] by blast qed lemma composable_protocols_par_comp_prot: fixes S f Sec lbls Ps defines "f \ \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}" and "Sec \ (f (set S)) - {m. {} \\<^sub>c m}" and "lbls \ \T. map (the_LabelN o fst) (filter (Not o has_LabelS) (transaction_strand T))" and "L \ set (remdups (concat (map lbls (concat Ps))))" and "no_leakage \ \\. \\\<^sub>\ s. (\l ts. suffix [(l, receive\ts\)] \) \ s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>\ \ welltyped_constraint_model \\<^sub>\ (\@[\\, send\[s]\\])" assumes Ps_pc: "wellformed_composable_protocols Ps N" "wellformed_composable_protocols' Ps" "composable_protocols Ps Ms S" "list_all ((list_all (Not \ has_LabelS)) \ tl \ transaction_send) (concat Ps)" and component_secure: "\\ \ reachable_constraints (map (transaction_proj l) (concat Ps)). \\. welltyped_constraint_model \ (\@[\l, send\[attack\ln l\]\\])" and no_leakage: "\l \ L. \\ \ reachable_constraints (map (transaction_proj l) (concat Ps)). no_leakage \" shows "\\ \ reachable_constraints (concat Ps). \\. constraint_model \ (\@[\l, send\[attack\ln l\]\\])" proof - have P': "list_all wellformed_transaction (concat Ps)" using Ps_pc(2) unfolding wellformed_composable_protocols'_def by meson have Sec: "\s \ Sec. \{} \\<^sub>c s" "ground Sec" unfolding Sec_def f_def by auto note 0 = composable_protocols_par_comp_prot_alt[ OF Ps_pc(1-3) component_secure, unfolded lbls_def[symmetric] L_def[symmetric]] note 1 = reachable_constraints_no_leakage_alt'[ OF P' Ps_pc(4) Sec no_leakage[unfolded no_leakage_def]] show ?thesis using 0 1 unfolding f_def Sec_def by argo qed lemma composable_protocols_par_comp_prot': assumes P_defs: "Pc = concat Ps" "set Ps_with_stars = (\n. map (transaction_proj n) Pc) ` set (remdups (concat (map (\T. map (the_LabelN \ fst) (filter (Not \ has_LabelS) (transaction_strand T))) Pc)))" and Ps_wellformed: "list_all (list_all (Not \ has_LabelS) \ tl \ transaction_send) Pc" "wellformed_composable_protocols Ps N" "wellformed_composable_protocols' Ps" "composable_protocols Ps Ms S" and Ps_no_leakage: "list_all (welltyped_leakage_free_protocol S) Ps_with_stars" and P_def: "P = map (transaction_proj n) Pc" and P_wt_secure: "\\ \ reachable_constraints P. \\. welltyped_constraint_model \ (\@[\n, send\[attack\ln n\]\\])" shows "\\ \ reachable_constraints Pc. \\. constraint_model \ (\@[\n, send\[attack\ln n\]\\])" by (rule composable_protocols_par_comp_prot[ OF Ps_wellformed(2,3,4,1)[unfolded P_defs(1)] P_wt_secure[unfolded P_def[unfolded P_defs(1)]] transaction_proj_ball_subst[ OF P_defs(2)[unfolded P_defs(1)] Ps_no_leakage(1)[ unfolded list_all_iff welltyped_leakage_free_protocol_def Let_def]], unfolded P_defs(1)[symmetric]]) end context begin lemma welltyped_constraint_model_leakage_model_swap: fixes I \ \::"('fun,'atom,'sets,'lbl) prot_subst" and s assumes A: "welltyped_constraint_model I (A@[\\, send\[s \ \]\\])" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "subst_domain \ = fv s" "ground (subst_range \)" obtains J where "welltyped_constraint_model J (A@[\\, send\[s \ \]\\])" and "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t J \ s \ \ \ J" proof note defs = welltyped_constraint_model_def constraint_model_def note \_s = subst_fv_dom_ground_if_ground_img[OF equalityD2[OF \(3)] \(4)] note \' = transaction_renaming_subst_is_renaming(2)[OF \] inj_on_subset[OF transaction_renaming_subst_is_injective[OF \] subset_UNIV[of "fv s"]] transaction_renaming_subst_var_obtain(2)[OF \, of _ s] transaction_renaming_subst_is_renaming(6)[OF \, of s] transaction_renaming_subst_vars_disj(8)[OF \] transaction_renaming_subst_wt[OF \] define \inv where "\inv \ subst_var_inv \ (fv s)" define \' where "\' \ rm_vars (UNIV - fv (s \ \)) (\inv \\<^sub>s \)" define J where "J \ \x. if x \ fv (s \ \) then \' x else I x" have \_invertible: "s = s \ \ \\<^sub>s \inv" using \'(1) inj_var_ran_subst_is_invertible'[of \ s] inj_on_subset[OF \'(2)] unfolding \inv_def by blast have \'_domain: "subst_domain \' = fv (s \ \)" proof - have "x \ subst_domain (\inv \\<^sub>s \)" when x: "x \ fv (s \ \)" for x proof - obtain y where y: "y \ fv s" "\ y = Var x" using \'(3)[OF x] by blast have "y \ subst_domain \" using y(1) \(3) by blast moreover have "(\inv \\<^sub>s \) x = \ y" using y \'(3)[OF x] \_invertible vars_term_subset_subst_eq[of "Var y" s "\ \\<^sub>s \inv" Var] unfolding \'_def \inv_def by (metis (no_types, lifting) fv_subst_subset eval_term.simps(1) subst_apply_term_empty subst_compose) ultimately show ?thesis using \(4) by fastforce qed thus ?thesis using rm_vars_dom[of "UNIV - fv (s \ \)" "\inv \\<^sub>s \"] unfolding \'_def by blast qed have \'_range: "fv t = {}" when t: "t \ (subst_range \')" for t proof - obtain x where "x \ fv (s \ \)" "\' x = t" using t \'_domain by auto thus ?thesis by (metis (no_types, lifting) \'_def subst_compose_def \(3,4) \_invertible fv_subst_subset subst_fv_dom_ground_if_ground_img subst_subst_compose Diff_iff) qed have J0: "x \ fv (s \ \) \ J x = \' x" "x \ fv (s \ \) \ J x = I x" for x unfolding J_def by (cases "x \ fv (s \ \)") (simp_all add: subst_compose) have J1: "subst_range J \ subst_range \' \ subst_range I" proof fix t assume "t \ subst_range J" then obtain x where x: "x \ subst_domain J" "J x = t" by auto hence "t = \' x \ x \ subst_domain \'" "t = I x \ x \ subst_domain I" by (metis subst_domI subst_dom_vars_in_subst)+ thus "t \ subst_range \' \ subst_range I" using x(2) J0[of x] by auto qed have "x \ fv (s \ \)" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@[\\, send\[s \ \]\\])" for x using x \_s \'(4) \'(5) by auto hence "I x = J x" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@[\\, send\[s \ \]\\])" for x using x unfolding J_def \'_def by auto hence "constr_sem_stateful J (unlabel (A@[\\, send\[s \ \]\\]))" using A strand_sem_model_swap[of "unlabel (A@[\\, send\[s \ \]\\])" I J "{}" "{}"] unfolding defs by blast moreover have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t J" using A subst_var_inv_wt[OF \'(6), of "fv s"] wt_subst_trm''[OF \(1)] subst_compose[of "subst_var_inv \ (fv s)" \] unfolding defs J_def \'_def \inv_def wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by presburger moreover have "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t J" proof - have "fv t = {}" when t: "t \ (subst_range J)" for t using t A J1 \'_range unfolding defs by auto moreover have "x \ subst_domain J" for x proof (cases "x \ fv (s \ \)") case True thus ?thesis using J0(1)[of x] \'_domain unfolding subst_domain_def by auto next case False have "subst_domain I = UNIV" using A unfolding defs by fast thus ?thesis using J0(2)[OF False] unfolding subst_domain_def by auto qed ultimately show ?thesis by auto qed moreover have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \')" using wf_trms_subst_compose[OF subst_var_inv_wf_trms[of \ "fv s"] \(2)] unfolding \'_def \inv_def by force hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range J)" using A J1 unfolding defs by fast ultimately show "welltyped_constraint_model J (A@[\\, send\[s \ \]\\])" unfolding defs by blast hence "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t J \ s \ \" using \_s strand_sem_append_stateful[of "{}" "{}" "unlabel A" "[send\[s \ \]\]" J] unfolding defs by (simp add: subst_ground_ident) moreover have "s \ \ \ J = s \ \" proof - have "J x = \' x" when x: "x \ fv (s \ \)" for x using x unfolding J_def by argo hence "s \ \ \ J = s \ \ \ \'" using subst_agreement[of "s \ \" J \'] by force thus ?thesis using \_invertible unfolding \'_def rm_vars_subst_eq'[symmetric] by (metis subst_subst_compose) qed hence "s \ \ \ J = s \ \" by auto ultimately show "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t J \ s \ \ \ J" by argo qed lemma welltyped_leakage_free_protocol_pointwise: "welltyped_leakage_free_protocol S P \ list_all (\s. welltyped_leakage_free_protocol [s] P) S" unfolding welltyped_leakage_free_protocol_def list_all_iff Let_def by fastforce lemma welltyped_leakage_free_no_deduct_constI: fixes c defines "s \ Fun c []::('fun,'atom,'sets,'lbl) prot_term" assumes s: "\\ \ reachable_constraints P. \\\<^sub>\. welltyped_constraint_model \\<^sub>\ (\@[\\, send\[s]\\])" shows "welltyped_leakage_free_protocol [s] P" using s unfolding welltyped_leakage_free_protocol_def s_def by auto lemma welltyped_leakage_free_pub_termI: assumes s: "{} \\<^sub>c s" shows "welltyped_leakage_free_protocol [s] P" proof - define Q where "Q \ \M t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}" define f where "f \ \M. {t \ \ | t \. Q M t \}" define Sec where "Sec \ f (set [s]) - {m. {} \\<^sub>c m}" have 0: "fv s = {}" using s pgwt_ground pgwt_is_empty_synth by blast have 1: "s \ \ = s" for \ by (rule subst_ground_ident[OF 0]) have 2: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t Var" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range Var)" using wt_subst_Var wf_trm_subst_range_Var by (blast,blast) have "f (set [s]) = {s}" proof show "f (set [s]) \ {s}" using 0 1 unfolding f_def Q_def by auto have "Q {s} s Var" using 0 2 unfolding Q_def by auto thus "{s} \ f (set [s])" using 1[of Var] unfolding f_def by force qed hence "Sec = {}" using s unfolding Sec_def by simp thus ?thesis unfolding welltyped_leakage_free_protocol_def Let_def Sec_def f_def Q_def by blast qed lemma welltyped_leakage_free_pub_constI: assumes c: "public\<^sub>f c" "arity\<^sub>f c = 0" shows "welltyped_leakage_free_protocol [\c\\<^sub>c] P" using c welltyped_leakage_free_pub_termI[OF intruder_synth.ComposeC[of "[]" "Fu c" "{}"]] by simp lemma welltyped_leakage_free_long_term_secretI: fixes n defines "Tatt \ \s'. Transaction (\(). []) [] [\n, receive\[s']\\] [] [] [\n, send\[attack\ln n\]\\]" assumes P_wt_secure: "\\ \ reachable_constraints P. \\. welltyped_constraint_model \ (\@[\n, send\[attack\ln n\]\\])" and s_long_term_secret: "\\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ inj_on \ (fv s) \ \ ` fv s \ range Var \ Tatt (s \ \) \ set P" shows "welltyped_leakage_free_protocol [s] P" proof (rule ccontr) define Q where "Q \ \M t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}" define f where "f \ \M. {t \ \ | t \. Q M t \}" define Sec where "Sec \ f (set [s]) - {m. {} \\<^sub>c m}" note defs = Sec_def f_def Q_def note defs' = welltyped_constraint_model_def constraint_model_def assume "\welltyped_leakage_free_protocol [s] P" then obtain A I s' where A: "A \ reachable_constraints P" "s' \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I" "welltyped_constraint_model I (A@[\\, send\[s']\\])" unfolding welltyped_leakage_free_protocol_def defs by fastforce obtain \ where \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "\ ` fv s \ range Var" "inj_on \ (fv s)" "Tatt (s \ \) \ set P" using s_long_term_secret by blast obtain \ where \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "subst_domain \ = fv (s \ \)" "ground (subst_range \)" "s' = s \ \ \ \" proof - obtain \ where *: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "fv s' = {}" "s' = s \ \" using A(2) unfolding defs by auto define \ where "\ \ subst_var_inv \ (fv s) \\<^sub>s \" define \' where "\' \ rm_vars (UNIV - fv (s \ \)) \" have **: "s' = s \ \ \ \" using *(4) inj_var_ran_subst_is_invertible[OF \(3,2)] unfolding \_def by simp have "s' = s \ \ \ \'" using ** rm_vars_subst_eq'[of "s \ \" \] unfolding \'_def by simp moreover have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using \(1) *(1) subst_var_inv_wt wt_subst_compose unfolding \_def by presburger hence "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \'" using wt_subst_rm_vars unfolding \'_def by blast moreover have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using wf_trms_subst_compose[OF subst_var_inv_wf_trms *(2)] unfolding \_def by blast hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \')" using wf_trms_subst_rm_vars'[of \] unfolding \'_def by blast moreover have "fv (s \ \) \ subst_domain \" using *(3) ** ground_term_subst_domain_fv_subset unfolding \_def by blast hence "subst_domain \' = fv (s \ \)" using rm_vars_dom[of "UNIV - fv (s \ \)" \] unfolding \'_def by blast moreover have "ground (subst_range \')" proof - { fix t assume "t \ subst_range \'" then obtain x where "x \ fv (s \ \)" "\' x = t" using \subst_domain \' = fv (s \ \)\ by auto hence "t \ s \ \ \ \'" by (meson subst_mono_fv) hence "fv t = {}" using \s' = s \ \ \ \'\ *(3) ground_subterm by blast } thus ?thesis by force qed ultimately show thesis using that[of \'] by fast qed have \: "transaction_decl_subst Var (Tatt t)" and \: "transaction_fresh_subst Var (Tatt t) (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" for t unfolding transaction_decl_subst_def transaction_fresh_subst_def Tatt_def by simp_all obtain \::"('fun,'atom,'sets,'lbl) prot_subst" where \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" unfolding transaction_renaming_subst_def by blast obtain J where J: "welltyped_constraint_model J (A@[\\, send\[s \ \ \ \]\\])" "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t J \ s \ \ \ \ \ J" using welltyped_constraint_model_leakage_model_swap[OF A(3)[unfolded \(5)] \ \(1-4)] by blast define T where "T = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand (Tatt (s \ \)) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" define B where "B \ A@T" have "transaction_receive (Tatt t) = [\n, receive\[t]\\]" "transaction_checks (Tatt t) = []" "transaction_updates (Tatt t) = []" "transaction_send (Tatt t) = [\n, send\[attack\ln n\]\\]" for t unfolding Tatt_def by simp_all hence T_def': "T = [\n, send\[s \ \ \ \]\\, \n, receive\[attack\ln n\]\\]" using subst_lsst_append[of "transaction_receive (Tatt (s \ \))" _ \] subst_lsst_singleton[of "ln n" "receive\[s \ \]\" \] subst_lsst_singleton[of "ln n" "send\[attack\ln n\]\" \] unfolding transaction_strand_def T_def by fastforce have B0: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t B \\<^sub>s\<^sub>e\<^sub>t J \ attack\ln n\" using in_ik\<^sub>s\<^sub>s\<^sub>t_iff[of "attack\ln n\" "unlabel T"] unfolding B_def T_def' by (force intro!: intruder_deduct.Axiom) have B1: "B \ reachable_constraints P" using reachable_constraints.step[OF A(1) \(4) \ \ \] unfolding B_def T_def by simp have "welltyped_constraint_model J B" using J strand_sem_append_stateful[of "{}" "{}" "unlabel A" _ J] unfolding defs' B_def T_def' by fastforce hence B2: "welltyped_constraint_model J (B@[\n, send\[attack\ln n\]\\])" using B0 strand_sem_append_stateful[of "{}" "{}" "unlabel B" "[send\[attack\ln n\]\]" J] unfolding defs' B_def by auto show False using P_wt_secure B1 B2 by blast qed lemma welltyped_leakage_free_value_constI: assumes P: "\T \ set P. wellformed_transaction T" "\T \ set P. admissible_transaction_terms T" "\T \ set P. transaction_decl T () = []" "\T \ set P. bvars_transaction T = {}" and P_fresh_declass: "\T \ set P. transaction_fresh T \ [] \ (transaction_send T \ [] \ (let (l,a) = hd (transaction_send T) in l = \ \ is_Send a \ Var ` set (transaction_fresh T) \ set (the_msgs a)))" shows "welltyped_leakage_free_protocol [\m: value\\<^sub>v] P" proof (rule ccontr) define Q where "Q \ \M t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}" define f where "f \ \M. {t \ \ | t \. Q M t \}" define Sec where "Sec \ f (set [\m: value\\<^sub>v]) - {m. {} \\<^sub>c m}" note defs = Sec_def f_def Q_def note defs' = welltyped_constraint_model_def constraint_model_def assume "\welltyped_leakage_free_protocol [\m: value\\<^sub>v] P" then obtain A I s where A: "A \ reachable_constraints P" "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I" "welltyped_constraint_model I (A@[\\, send\[s]\\])" unfolding welltyped_leakage_free_protocol_def defs by fastforce have "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I \ s \ I" using welltyped_constraint_model_deduct_split[OF A(3)] by simp moreover have "s \ I = s" using A(2) unfolding defs by fast ultimately have s_deduct: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I) \ s" by (metis ik\<^sub>s\<^sub>s\<^sub>t_subst unlabel_subst) note I0 = welltyped_constraint_model_prefix[OF A(3)] have I1: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" using A(3) unfolding defs' by blast obtain f ts \ where f: "s = Fun f ts" "s = \m: value\\<^sub>v \ \" "\{} \\<^sub>c s" "s \ declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "fv s = {}" using A(2) unfolding defs by (cases s) auto have s1: "\ s = TAtom Value" by (metis \.simps(1) \\<^sub>v_TAtom f(2) wt_subst_trm''[OF \(1)]) have s2: "wf\<^sub>t\<^sub>r\<^sub>m s" using f(2) \(2) by force have s3: "ts = []" using f(1) s1 s2 const_type_inv_wf by blast obtain sn where sn: "s = Fun (Val sn) []" using s1 f(3) \_Fu_simps(4) \_Set_simps(3) unfolding f(1) s3 by (cases f) auto have "s \\<^sub>s\<^sub>e\<^sub>t ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I" using private_fun_deduct_in_ik'[OF s_deduct[unfolded sn]] by (metis sn public.simps(3) ik\<^sub>s\<^sub>s\<^sub>t_subst unlabel_subst) hence s4: "s \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" using constraint_model_Val_const_in_constr_prefix[OF A(1) I0 P(1,2)] unfolding sn by presburger obtain B T \ \ \ where B: "prefix (B@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) A" "B \ reachable_constraints P" "T \ set P" "transaction_decl_subst \ T" "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" "s \ subst_range \" using constraint_model_Value_in_constr_prefix_fresh_action'[OF A(1) P(2-) s4[unfolded sn]] sn by blast obtain Tts Tsnds sx where T: "transaction_send T = \\, send\Tts\\#Tsnds" "Var ` set (transaction_fresh T) \ set Tts" and sx: "Var sx \ set Tts" "\ sx = s" using P_fresh_declass B(3,5,7) unfolding transaction_fresh_subst_def is_Send_def by (cases "transaction_send T") (fastforce,fastforce) have \_elim: "\ \\<^sub>s \ \\<^sub>s \ = \ \\<^sub>s \" using admissible_transaction_decl_subst_empty'[OF bspec[OF P(3) B(3)] B(4)] by simp have s5: "s \ set (Tts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \ \\<^sub>s I)" using sx unfolding \_elim sn by force have s6: "\\, receive\Tts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \ \\<^sub>s I\\ \ set (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I)" proof - have "\\, send\Tts\\ \ set (transaction_send T)" using T(1) by simp hence "\\, send\Tts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\ \ set (transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" for \ unfolding subst_apply_labeled_stateful_strand_def by force hence "\\, send\Tts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\ \ set (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" for \ using transaction_strand_subst_subsets(4)[of T \] by fast hence *: "\\, receive\Tts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\ \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" for \ using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_steps_iff(1)[of \ "Tts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \"] by blast have "\\, receive\Tts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \\\ \ set A" using B(1) *[of "\ \\<^sub>s \ \\<^sub>s \"] unfolding prefix_def by force thus ?thesis unfolding subst_apply_labeled_stateful_strand_def by force qed show False using s6 f(4) ideduct_mono[OF Axiom[OF s5], of "\{set ts|ts. \\,receive\ts\\ \ set (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I)}"] unfolding declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by blast qed lemma welltyped_leakage_free_priv_constI: fixes c defines "s \ Fun c []::('fun,'atom,'sets,'lbl) prot_term" assumes c: "\public c" "arity c = 0" "\ s = TAtom ca" "ca \ Value" and P: "\T \ set P. \x \ vars_transaction T. is_Var (\\<^sub>v x)" "\T \ set P. \x \ vars_transaction T \ set (transaction_fresh T). \ s \ \\<^sub>v x" "\T \ set P. \t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)). s \ set (snd (Ana t))" "\T \ set P. s \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = \a\\<^sub>\\<^sub>a)" "\T \ set P. wellformed_transaction T" shows "\\ \ reachable_constraints P. \\\<^sub>\. welltyped_constraint_model \\<^sub>\ (\@[\\, send\[s]\\])" (is "\\ \ ?R. ?P \") and "welltyped_leakage_free_protocol [s] P" proof - show "\\ \ ?R. ?P \" proof fix A assume A: "A \ reachable_constraints P" define Q where "Q \ \M t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}" define f where "f \ \M. {t \ \ | t \. Q M t \}" define Sec where "Sec \ f (set [s]) - {m. {} \\<^sub>c m}" define f' where "f' \ \(T,\,\::('fun,'atom,'sets,'lbl) prot_subst,\). dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" define g' where "g' \ concat \ map f'" let ?P_s_cases = "\M. s \ M \ (\m \ subterms\<^sub>s\<^sub>e\<^sub>t M. s \ set (snd (Ana m)))" let ?P_s_cases' = "\M \. s \ M \\<^sub>s\<^sub>e\<^sub>t \ \ (\m \ subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \. s \ set (snd (Ana m)))" note defs = Sec_def f_def Q_def note defs' = welltyped_constraint_model_def constraint_model_def show "?P A" proof (rule ccontr) assume "\?P A" then obtain I where I: "welltyped_constraint_model I (A@[\\, send\[s]\\])" by blast obtain Ts where Ts: "A = g' Ts" "\B. prefix B Ts \ g' B \ reachable_constraints P" "\B T \ \ \. prefix (B@[(T,\,\,\)]) Ts \ T \ set P \ transaction_decl_subst \ T \ transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (g' B)) \ transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (g' B))" using reachable_constraints_as_transaction_lists[OF A(1)] unfolding g'_def f'_def by blast have "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I \ s \ I" and I_s: "s \ I = s" using welltyped_constraint_model_deduct_split[OF I] unfolding s_def by simp_all hence s_deduct: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I) \ s" "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I \ s" by (metis ik\<^sub>s\<^sub>s\<^sub>t_subst unlabel_subst, metis) have I_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" and I_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)" and I_grounds: "ground (subst_range I)" and I_interp: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" using I unfolding defs' by (blast,blast,blast,blast) have Sec_unfold: "Sec = {s}" proof have "\{} \\<^sub>c s" using ideduct_synth_priv_const_in_ik[OF _ c(1)] unfolding s_def by blast thus "{s} \ Sec" unfolding defs s_def by fastforce qed (auto simp add: defs s_def) have s2: "wf\<^sub>t\<^sub>r\<^sub>m s" using c(1,2) unfolding s_def by fastforce have A_ik_fv: "\a. \\<^sub>v x = TAtom a \ a \ ca" when x: "x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" for x proof - obtain T where T: "T \ set P" "\\<^sub>v x \ \\<^sub>v ` fv_transaction T" using fv_ik\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t[OF x] reachable_constraints_var_types_in_transactions(1)[OF A P(5)] by fast then obtain y where y: "y \ vars_transaction T" "\\<^sub>v y = \\<^sub>v x" using vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_strand T)"] by fastforce then obtain a where a: "\\<^sub>v y = TAtom a" using P(1) T(1) by blast hence "\\<^sub>v x = TAtom a" "\ s \ \\<^sub>v x" "\ s = TAtom ca" using y P(2) T(1) c(3) by auto thus ?thesis by force qed have I_s_x: "\s \ I x" when x: "x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" for x proof - obtain a where a: "\\<^sub>v x = TAtom a" "a \ ca" using A_ik_fv[OF x] by blast hence a': "\ (I x) = TAtom a" using wt_subst_trm''[OF I_wt, of "Var x"] by simp obtain f ts where f: "I x = Fun f ts" by (meson empty_fv_exists_fun interpretation_grounds_all[OF I_interp]) hence ts: "ts = []" using I_wf const_type_inv_wf[OF a'[unfolded f]] by fastforce have "c \ f" using f[unfolded ts] a a' c(3)[unfolded s_def] by force thus ?thesis using f ts unfolding s_def by simp qed have A_ik_I_const: "\f. arity f = 0 \ I x = Fun f []" when x: "x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" for x using x A_ik_fv I_wt empty_fv_exists_fun[OF interpretation_grounds_all[OF I_interp, of x]] wf_trm_subst_rangeD[OF I_wf, of x] const_type_inv const_type_inv_wf by (metis (no_types, lifting) \.simps(1) wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) hence A_ik_subst: "subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I) = subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \\<^sub>s\<^sub>e\<^sub>t I" using subterms_subst''[of "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" I] by blast have sublmm1: "s \ set (snd (Ana m))" when m: "m \\<^sub>s\<^sub>e\<^sub>t M" "s \ set (snd (Ana (m \ \)))" and M: "\y. y \ fv\<^sub>s\<^sub>e\<^sub>t M \ \s \ \ y" for m M \ proof - have m_fun: "is_Fun m" using m M Ana_subterm' vars_iff_subtermeq_set unfolding s_def is_Var_def by (metis eval_term.simps(1)) obtain f K R ts i where f: "m = \f ts\\<^sub>t" "arity\<^sub>f f = length ts" "arity\<^sub>f f > 0" "Ana\<^sub>f f = (K, R)" and i: "i < length R" "s = ts ! (R ! i) \ \" and R_i: "\i < length R. map ((!) ts) R ! i = ts ! (R ! i) \ R ! i < length ts" proof - obtain f ts K R where f: "m \ \ = \f ts\\<^sub>t" "arity\<^sub>f f = length ts" "arity\<^sub>f f > 0" "Ana\<^sub>f f = (K, R)" "Ana (m \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) ts, map ((!) ts) R)" using m(2) Ana_nonempty_inv[of "m \ \"] by force obtain ts' where m': "m = \f ts'\\<^sub>t" "ts = ts' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" using f(1) m_fun by auto have R_i: "map ((!) ts) R ! i = ts ! (R ! i)" "R ! i < length ts" when i: "i < length R" for i using i Ana\<^sub>f_assm2_alt[OF f(4), of "R ! i"] f(2) by simp_all then obtain i where i: "s = ts ! (R ! i)" "i < length R" by (metis (no_types, lifting) m(2) f(5) in_set_conv_nth length_map snd_conv) have ts': "arity\<^sub>f f = length ts'" "length ts = length ts'" using m'(2) f(2) by simp_all have s': "s = ts' ! (R ! i) \ \" using R_i(2)[OF i(2)] i(1) unfolding ts'(2) m'(2) by simp show thesis using that f m' R_i ts' s' i by auto qed have "s = ts ! (R ! i)" proof (cases "ts ! (R ! i)") case (Var x) hence "Var x \ set ts" using R_i i nth_mem by fastforce hence "x \ fv\<^sub>s\<^sub>e\<^sub>t M" using m(1) f(1) fv_subterms_set by fastforce thus ?thesis using i M Var by fastforce qed (use i s_def in fastforce) thus "s \ set (snd (Ana m))" using f(1) Ana_Fu_intro[OF f(2-4)] i(1) by simp qed have "\s \ \ y" when m: "m \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "s \ set (snd (Ana (m \ \)))" and T: "T \ set P" and \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and \_ran: "\t. t \ subst_range \ \ (\c. t = Fun c [] \ arity c = 0) \ (\x. t = Var x)" and y: "y \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))" for m T \ y proof assume "s \ \ y" hence "\\<^sub>v y = \ s" using wt_subst_trm''[OF \_wt, of "Var y"] \_ran[of "\ y"] by fastforce moreover have "y \ vars_transaction T" using y trms\<^sub>s\<^sub>s\<^sub>t_fv_vars\<^sub>s\<^sub>s\<^sub>t_subset unfolding vars_transaction_unfold[of T] by fastforce ultimately show False using P(2) T by force qed hence sublmm2: "s \ set (snd (Ana m))" when m: "m \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "s \ set (snd (Ana (m \ \)))" and T: "T \ set P" and \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and \_ran: "\t. t \ subst_range \ \ (\c. t = Fun c [] \ arity c = 0) \ (\x. t = Var x)" for m T \ using sublmm1[OF m] m T \_wt \_ran by blast have "s \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ (\m \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \\<^sub>s\<^sub>e\<^sub>t I. s \ set (snd (Ana m)))" using private_const_deduct[OF c(1) s_deduct(2)[unfolded s_def]] I_s_x const_mem_subst_cases[of c] A_ik_subst unfolding s_def by blast hence "?P_s_cases (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" using sublmm1[of _ "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A"] I_s_x by blast then obtain T \ \ \ where T: "(T,\,\,\) \ set Ts" "?P_s_cases (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (f' (T,\,\,\)))" using ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t_concat[of "map f' Ts"] Ts(1)[unfolded g'_def] by fastforce obtain B where "prefix (B@[(T, \, \, \)]) Ts" by (metis T(1) prefix_snoc_in_iff) hence T_in_P: "T \ set P" and T_wf: "wellformed_transaction T" and \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (concat (map f' B)))" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (concat (map f' B)))" using P(6) Ts(3)[unfolded g'_def] unfolding comp_def by (metis,metis,metis,metis,metis) note \\\_wt = transaction_decl_fresh_renaming_substs_wt[OF \ \ \] note \\\_ran = transaction_decl_fresh_renaming_substs_range'(1)[OF \ \ \] have "subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \) = subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \" for M using \\\_ran subterms_subst''[of _ "\ \\<^sub>s \ \\<^sub>s \"] by (meson subst_imgI) hence s_cases: "?P_s_cases' (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) (\ \\<^sub>s \ \\<^sub>s \)" using T(2) dual_transaction_ik_is_transaction_send'[OF T_wf, of "\ \\<^sub>s \ \\<^sub>s \"] unfolding f'_def by auto from s_cases show False proof assume "s \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \" - then obtain t where t: "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "s = t \ \ \\<^sub>s \ \\<^sub>s \" by moura + then obtain t where t: "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "s = t \ \ \\<^sub>s \ \\<^sub>s \" by force have "s \ t" using P(4) T_in_P t(1) by blast then obtain x where x: "t = Var x" using t(2) unfolding s_def by (cases t) auto have "\\<^sub>v x = \ s" using x t(2) wt_subst_trm''[OF \\\_wt, of "Var x"] by simp moreover have "x \ vars_transaction T" using t(1) trms\<^sub>s\<^sub>s\<^sub>t_fv_vars\<^sub>s\<^sub>s\<^sub>t_subset unfolding x vars_transaction_unfold[of T] by fastforce ultimately show False using P(2) T_in_P by force qed (use sublmm2[OF _ _ T_in_P \\\_wt \\\_ran] P(3) T_in_P in blast) qed qed thus "welltyped_leakage_free_protocol [s] P" using welltyped_leakage_free_no_deduct_constI[of P c] unfolding s_def by blast qed lemma welltyped_leakage_free_priv_constI': assumes c: "\public\<^sub>f c" "arity\<^sub>f c = 0" "\\<^sub>f c = Some ca" and P: "\T \ set P. wellformed_transaction T" "\T \ set P. \x \ vars_transaction T \ set (transaction_fresh T). \ \c\\<^sub>c \ \\<^sub>v x" "\T \ set P. \x \ vars_transaction T. is_Var (\\<^sub>v x)" "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = \a\\<^sub>\\<^sub>a)" "\T \ set P. \t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)). \c\\<^sub>c \ set (snd (Ana t))" "\T \ set P. \c\\<^sub>c \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" shows "\\ \ reachable_constraints P. \\\<^sub>\. welltyped_constraint_model \\<^sub>\ (\@[\\, send\[\c\\<^sub>c]\\])" and "welltyped_leakage_free_protocol [\c\\<^sub>c] P" using c welltyped_leakage_free_priv_constI[OF _ _ _ _ P(3,2,5,6,4,1), of "Atom ca"] by (force, force) lemma welltyped_leakage_free_set_constI: assumes P: "\T \ set P. wellformed_transaction T" "\T \ set P. \f \ \(funs_term ` (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))). \is_Set f" "\T \ set P. \x \ vars_transaction T \ set (transaction_fresh T). \\<^sub>v x \ TAtom SetType" "\T \ set P. \x \ vars_transaction T. is_Var (\\<^sub>v x)" "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = \a\\<^sub>\\<^sub>a)" and c: "arity\<^sub>s c = 0" shows "\\ \ reachable_constraints P. \\\<^sub>\. welltyped_constraint_model \\<^sub>\ (\@[\\, send\[\c\\<^sub>s]\\])" and "welltyped_leakage_free_protocol [\c\\<^sub>s] P" proof - have c'': "\c\\<^sub>s \ subterms t" when T: "T \ set P" and t: "t \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" for T t using t bspec[OF P(2) T] subtermeq_imp_funs_term_subset[of t] funs_term_Fun_subterm'[of "Set c" "[]::('fun,'atom,'sets,'lbl) prot_term list"] by fastforce have P': "\T \ set P. \t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)). \c\\<^sub>s \ set (snd (Ana t))" "\T \ set P. \c\\<^sub>s \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" subgoal using Ana_subterm' c'' by fast subgoal using c'' by fast done have P'': "\T \ set P. \x \ vars_transaction T \ set (transaction_fresh T). \ \c\\<^sub>s \ \\<^sub>v x" using P(3) \_consts_simps(4)[OF c] by fastforce show "\\ \ reachable_constraints P. \\\<^sub>\. welltyped_constraint_model \\<^sub>\ (\@[\\, send\[\c\\<^sub>s]\\])" "welltyped_leakage_free_protocol [\c\\<^sub>s] P" using c welltyped_leakage_free_priv_constI[OF _ _ _ _ P(4) P'' P' P(5,1), of SetType] by (force, force) qed lemma welltyped_leakage_free_occurssec_constI: defines "s \ Fun OccursSec []" assumes P: "\T \ set P. wellformed_transaction T" "\T \ set P. \x \ vars_transaction T \ set (transaction_fresh T). \\<^sub>v x \ TAtom OccursSecType" "\T \ set P. \t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)). Fun OccursSec [] \ set (snd (Ana t))" "\T \ set P. Fun OccursSec [] \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "\T \ set P. \x \ vars_transaction T. is_Var (\\<^sub>v x)" "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = \a\\<^sub>\\<^sub>a)" shows "\\ \ reachable_constraints P. \\\<^sub>\. welltyped_constraint_model \\<^sub>\ (\@[\\, send\[s]\\])" and "welltyped_leakage_free_protocol [s] P" proof - have P': "\T \ set P. \x \ vars_transaction T \ set (transaction_fresh T). \ (Fun OccursSec []) \ \\<^sub>v x" using P(2) by auto show "\\ \ reachable_constraints P. \\\<^sub>\. welltyped_constraint_model \\<^sub>\ (\@[\\, send\[s]\\])" "welltyped_leakage_free_protocol [s] P" using welltyped_leakage_free_priv_constI[OF _ _ _ _ P(5) P' P(3,4,6,1), of OccursSecType] unfolding s_def by auto qed lemma welltyped_leakage_free_occurs_factI: assumes P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and P_occ_star: "\T \ set P. \r \ set (transaction_send T). OccursFact \ \(funs_term ` (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd r))) \ fst r = \" shows "welltyped_leakage_free_protocol [occurs x] P" proof - define Q where "Q \ \M t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}" define f where "f \ \M. {t \ \ | t \. Q M t \}" define Sec where "Sec \ f (set [occurs x]) - {m. {} \\<^sub>c m}" define f' where "f' \ \(T,\,\::('fun,'atom,'sets,'lbl) prot_subst,\). dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" define g' where "g' \ concat \ map f'" note defs = Sec_def f_def Q_def note defs' = welltyped_constraint_model_def constraint_model_def show ?thesis proof (rule ccontr) assume "\welltyped_leakage_free_protocol [occurs x] P" then obtain A I k where A: "A \ reachable_constraints P" "occurs k \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I" "welltyped_constraint_model I (A@[\\, send\[occurs k]\\])" unfolding welltyped_leakage_free_protocol_def defs by fastforce note A' = welltyped_constraint_model_prefix[OF A(3)] have occ_I: "occurs k \ I = occurs k" using A(2) unfolding defs by auto hence occ_in_ik: "occurs k \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "occurs k \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I" using welltyped_constraint_model_deduct_split(2)[OF A(3)] reachable_constraints_occurs_fact_deduct_in_ik[OF A(1) A' P P_occ, of k] by (argo, argo) then obtain l ts where ts: "(l,receive\ts\) \ set A" "occurs k \ set ts" - using in_ik\<^sub>s\<^sub>s\<^sub>t_iff[of "occurs k" "unlabel A"] unfolding unlabel_def by moura + using in_ik\<^sub>s\<^sub>s\<^sub>t_iff[of "occurs k" "unlabel A"] unfolding unlabel_def by force obtain T a B \ \ \ where B: "prefix (B@f' (T,\,\,\)) A" and T: "T \ set P" "transaction_decl_subst \ T" "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" and a: "a \ set (transaction_strand T)" "fst (l,receive\ts\) = fst a" "(l,receive\ts\) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \\<^sub>s \ \\<^sub>s \" using reachable_constraints_transaction_action_obtain[OF A(1) ts(1), of thesis] unfolding f'_def by simp obtain ts' where ts': "a = (l,send\ts'\)" "ts = ts' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \" using surj_pair[of a] a(2,3) by (cases "snd a") force+ obtain t where t: "t \ set ts'" "occurs k = t \ \ \\<^sub>s \ \\<^sub>s \" using ts(2) unfolding ts'(2) by force have occ_t: "OccursFact \ funs_term t" proof (cases t) case (Var y) thus ?thesis using t(2) transaction_decl_fresh_renaming_substs_range'(1)[OF T(2-), of "occurs k"] by fastforce qed (use t(2) in simp) have P_wf: "\T \ set P. wellformed_transaction T" using P admissible_transaction_is_wellformed_transaction(1) by blast have l: "l = \" using wellformed_transaction_strand_memberD(8)[OF bspec[OF P_wf T(1)] a(1)[unfolded ts'(1)]] t(1) T(1) P_occ_star occ_t unfolding ts'(1) by fastforce have "occurs k \ \{set ts |ts. \\, receive\ts\\ \ set (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I)}" using subst_lsst_memI[OF ts(1), of I] subst_set_map[OF ts(2), of I] unfolding occ_I l by auto thus False using A(2) unfolding declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by simp qed qed lemma welltyped_leakage_free_setop_pairI: assumes P: "\T \ set P. wellformed_transaction T" "\T \ set P. \x \ vars_transaction T. \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = \a\\<^sub>\\<^sub>a)" "\T \ set P. \f \ \(funs_term ` (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))). \is_Set f" "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" "\T \ set P. transaction_decl T () = []" "\T \ set P. admissible_transaction_terms T" and c: "arity\<^sub>s c = 0" shows "welltyped_leakage_free_protocol [pair (x, \c\\<^sub>s)] P" proof - define Q where "Q \ \M t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}" define f where "f \ \M. {t \ \ | t \. Q M t \}" define Sec where "Sec \ f (set [pair (x, \c\\<^sub>s)]) - {m. {} \\<^sub>c m}" define f' where "f' \ \(T,\,\::('fun,'atom,'sets,'lbl) prot_subst,\). dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" define g' where "g' \ concat \ map f'" note defs = Sec_def f_def Q_def note defs' = welltyped_constraint_model_def constraint_model_def have P': "\T \ set P. \x \ vars_transaction T \ set (transaction_fresh T). \\<^sub>v x \ TAtom SetType" "\T \ set P. \x \ vars_transaction T. is_Var (\\<^sub>v x)" "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = \a\\<^sub>\\<^sub>a)" subgoal using P(2,4) by fastforce subgoal using P(2) by fastforce subgoal using P(4) by fast done note 0 = welltyped_leakage_free_set_constI(1)[OF P(1,3) P' c] show ?thesis proof (rule ccontr) assume "\welltyped_leakage_free_protocol [pair (x, \c\\<^sub>s)] P" then obtain A I k where A: "A \ reachable_constraints P" "pair (k, \c\\<^sub>s) \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I" "welltyped_constraint_model I (A@[\\, send\[pair (k, \c\\<^sub>s)]\\])" unfolding welltyped_leakage_free_protocol_def defs pair_def by fastforce note A' = welltyped_constraint_model_prefix[OF A(3)] have "pair (k, \c\\<^sub>s) \ I = pair (k, \c\\<^sub>s)" using A(2) unfolding defs by auto hence "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I \ pair (k, \c\\<^sub>s)" using welltyped_constraint_model_deduct_split(2)[OF A(3)] by argo then obtain n where n: "intruder_deduct_num (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I) n (pair (k, \c\\<^sub>s))" using deduct_num_if_deduct by fast have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)" "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" using A(3) ik\<^sub>s\<^sub>s\<^sub>t_trms\<^sub>s\<^sub>s\<^sub>t_subset unfolding defs' by simp_all hence "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I \ SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" by blast hence "Pair \ \(funs_term ` (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I))" using reachable_constraints_no_Pair_fun'[OF A(1) P(4-6)] P by blast hence 1: "\pair (k, \c\\<^sub>s) \\<^sub>s\<^sub>e\<^sub>t ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I" using funs_term_Fun_subterm'[of Pair] unfolding pair_def by auto have 2: "pair (k, \c\\<^sub>s) \ set (snd (Ana m))" when m: "m \\<^sub>s\<^sub>e\<^sub>t ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I" for m using m 1 term.dual_order.trans Ana_subterm'[of "pair (k, \c\\<^sub>s)" m] by auto have "\ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I \ \c\\<^sub>s" using 0 A(1) A' welltyped_constraint_model_deduct_iff[of I A \ "\c\\<^sub>s"] by force moreover have "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I \ \c\\<^sub>s" using 1 2 deduct_inv[OF n] deduct_if_deduct_num[of "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I" _ "\c\\<^sub>s"] unfolding pair_def by auto ultimately show False by blast qed qed lemma welltyped_leakage_free_short_term_secretI: fixes c x y f n d l l' defines "s \ \f [\c\\<^sub>c, \x: value\\<^sub>v]\\<^sub>t" and "Tatt \ Transaction (\(). []) [] [\\, receive\[occurs \y: value\\<^sub>v]\\, (l, receive\[\f [\c\\<^sub>c, \y: value\\<^sub>v]\\<^sub>t]\)] [(l', \\y: value\\<^sub>v not in \d\\<^sub>s\)] [] [\n, send\[attack\ln n\]\\]" assumes P: "\T \ set P. admissible_transaction' T" "\T \ set P. admissible_transaction_occurs_checks T" "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = \a\\<^sub>\\<^sub>a)" and subterms_sec: "\\ \ reachable_constraints P. \\\<^sub>\. welltyped_constraint_model \\<^sub>\ (\@[\\, send\[\c\\<^sub>c]\\])" and P_sec: "\\ \ reachable_constraints P. \\\<^sub>\. welltyped_constraint_model \\<^sub>\ (\@[\n, send\[attack\ln n\]\\])" and P_Tatt: "Tatt \ set P" and P_d: "\T \ set P. \a \ set (transaction_updates T). is_Insert (snd a) \ the_set_term (snd a) = \d\\<^sub>s \ transaction_send T \ [] \ (let (l,b) = hd (transaction_send T) in l = \ \ is_Send b \ \f [\c\\<^sub>c, the_elem_term (snd a)]\\<^sub>t \ set (the_msgs b))" shows "welltyped_leakage_free_protocol [s] P" proof - define Q where "Q \ \M t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}" define Sec where "Sec \ {t \ \ | t \. Q (set [s]) t \} - {m. {} \\<^sub>c m}" define f' where "f' \ \(T,\,\::('fun,'atom,'sets,'lbl) prot_subst,\). dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" define g' where "g' \ concat \ map f'" note defs = Sec_def Q_def note defs' = welltyped_constraint_model_def constraint_model_def show ?thesis proof (rule ccontr) assume "\welltyped_leakage_free_protocol [s] P" then obtain A I k where A: "A \ reachable_constraints P" "\f [\c\\<^sub>c, k]\\<^sub>t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I" "welltyped_constraint_model I (A@[\\, send\[\f [\c\\<^sub>c, k]\\<^sub>t]\\])" unfolding welltyped_leakage_free_protocol_def defs s_def by fastforce have I: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)" using A(3) unfolding defs' by (blast,blast,blast) note A' = welltyped_constraint_model_prefix[OF A(3)] have "strand_sem_stateful {} {} (unlabel A) I" using A' unfolding defs' by simp hence A'': "strand_sem_stateful {} {} (unlabel A) (I(z := k))" when z: "z \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" for z using z strand_sem_model_swap[of "unlabel A" I "I(z := k)" "{}" "{}"] by auto obtain \ where \: "\ (the_Var \x: value\\<^sub>v) = k" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "fv (\ (the_Var \x: value\\<^sub>v)) = {}" using A(2) unfolding defs s_def by auto have k: "\ k = TAtom Value" "fv k = {}" "wf\<^sub>t\<^sub>r\<^sub>m k" subgoal using \(1) wt_subst_trm''[OF \(2), of "\x: value\\<^sub>v"] by simp subgoal using \(1,4) by blast subgoal using \(1,3) by force done then obtain fk where fk: "k = Fun fk []" using const_type_inv_wf by (cases k) auto have fk': "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I \ \f [\c\\<^sub>c, k]\\<^sub>t" using fk welltyped_constraint_model_deduct_split(2)[OF A(3)] by auto have "\welltyped_constraint_model I (A@[\\, send\[\c\\<^sub>c]\\])" using subterms_sec(1) A(1) by blast hence "\ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I \ \c\\<^sub>c" using A' strand_sem_append_stateful[of "{}" "{}" "unlabel A" "unlabel [\\, send\[\c\\<^sub>c]\\]" I] unfolding defs' by auto hence "\f [\c\\<^sub>c, k]\\<^sub>t \\<^sub>s\<^sub>e\<^sub>t ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I" using fk' deduct_inv'[OF fk'] by force moreover have "k \ \f [\c\\<^sub>c, k]\\<^sub>t" by simp ultimately have k_in_ik: "k \\<^sub>s\<^sub>e\<^sub>t ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I" using in_subterms_subset_Union by blast hence "k \\<^sub>s\<^sub>e\<^sub>t ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ (\x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A). k \ I x)" using const_subterms_subst_cases[of fk I "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A"] unfolding fk by fast hence "fk \ \(funs_term ` ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \ (\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. k \ I x)" unfolding fk by (meson UN_iff funs_term_Fun_subterm' fv_ik\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t ) hence "fk \ \(funs_term ` trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \ (\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. k \ I x)" using ik\<^sub>s\<^sub>s\<^sub>t_trms\<^sub>s\<^sub>s\<^sub>t_subset by blast moreover have "\\<^sub>v x = TAtom Value" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" for x using x reachable_constraints_var_types_in_transactions(1)[OF A(1) P(3)] P(1,2) admissible_transaction_Value_vars by force ultimately have "fk \ \(funs_term ` trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \ (\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. \\<^sub>v x = TAtom Value \ k \ I x)" by blast hence "fk \ \(funs_term ` trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \ (\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. \\<^sub>v x = TAtom Value \ I x = k)" using I(1,3) wf_trm_TComp_subterm wf_trm_subst_rangeD unfolding fk wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by (metis \.simps(1) term.distinct(1)) then obtain kn where kn: "fk = Val kn" using reachable_constraints_val_funs_private[OF A(1) P(1), of fk] constraint_model_Value_term_is_Val[OF A(1) A' P(1,2)] Fun_Value_type_inv[OF k(1)[unfolded fk]] unfolding fk is_PubConstValue_def by (cases fk) force+ obtain \::"('fun,'atom,'sets,'lbl) prot_subst" where \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" unfolding transaction_renaming_subst_def by blast obtain y' yn' where y': "\ (the_Var \y: value\\<^sub>v) = Var y'" "y \ yn'" "Var y' = \yn': value\\<^sub>v" using transaction_renaming_subst_is_renaming(1,2)[OF \] by force define B where "B \ A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand Tatt \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" define J where "J \ I(y' := k)" have "y' \ range_vars \" using y'(1) transaction_renaming_subst_is_renaming(3)[OF \] by (metis (no_types, lifting) in_mono subst_fv_imgI term.set_intros(3)) hence y'': "y' \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" using transaction_renaming_subst_vars_disj(6)[OF \] by blast have 0: "(k,\d\\<^sub>s) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I)" proof assume a: "(k,\d\\<^sub>s) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I)" obtain l t t' where t: "(l,insert\t,t'\) \ set A" "t \ I = k" "t' \ I = \d\\<^sub>s" using db\<^sub>s\<^sub>s\<^sub>t_in_cases[OF a[unfolded db\<^sub>s\<^sub>s\<^sub>t_def]] unfolding unlabel_def by auto obtain T b B \ \ \ where T: "prefix (B@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) A" "T \ set P" "transaction_decl_subst \ T" "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" "b \ set (transaction_strand T)" "(l, insert\t,t'\) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p b \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \\<^sub>s \ \\<^sub>s \" "fst (l, insert\t,t'\) = fst b" using reachable_constraints_transaction_action_obtain[OF A(1) t(1)] by metis define \ where "\ \ \ \\<^sub>s \ \\<^sub>s \" obtain b' where "b = (l,b')" using T(8) by (cases b) simp then obtain tb tb' where b': "b = (l,insert\tb,tb'\)" and tb: "t = tb \ \" and tb': "t' = tb' \ \" using T(7) unfolding \_def by (cases b') auto note T_adm = bspec[OF P(1) T(2)] note T_wf = admissible_transaction_is_wellformed_transaction(1,3)[OF T_adm] have b: "b \ set (transaction_updates T)" using transaction_strand_memberD[OF T(6)[unfolded b']] wellformed_transaction_cases[OF T_wf(1)] unfolding b' by blast have "\n. tb = \n: value\\<^sub>v" and *: "tb' = \d\\<^sub>s" using tb tb' T(6) t(3) transaction_inserts_are_Value_vars[OF T_wf, of tb tb'] unfolding b' unlabel_def by (force,force) have "is_Insert (snd b)" "the_set_term (snd b) = \d\\<^sub>s" "the_elem_term (snd b) = tb" unfolding b' * by simp_all hence "transaction_send T \ []" "let (l, a) = hd (transaction_send T) in l = \ \ is_Send a \ \f [\c\\<^sub>c, tb]\\<^sub>t \ set (the_msgs a)" using P_d T(2) b by (fast,fast) hence "\ts. \\,send\ts\\ \ set (transaction_send T) \ \f [\c\\<^sub>c, tb]\\<^sub>t \ set ts" unfolding is_Send_def by (cases "transaction_send T") auto then obtain ts where ts: "\\,send\ts\\ \ set (transaction_strand T)" "\f [\c\\<^sub>c, tb]\\<^sub>t \ set ts" unfolding transaction_strand_def by auto have "\\,receive\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\ \ set A" "\f [\c\\<^sub>c, t]\\<^sub>t \ set (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using subst_lsst_memI[OF ts(1), of \] subst_set_map[OF ts(2), of \] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_steps_iff(1)[of \ "ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] set_mono_prefix[OF T(1)[unfolded \_def[symmetric]]] unfolding tb by auto hence "\f [\c\\<^sub>c, k]\\<^sub>t \ \{set ts |ts. \\, receive\ts\\ \ set (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I)}" using t(2) subst_lsst_memI[of "\\,receive\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\" A I] by force thus False using A(2) unfolding declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by auto qed have "y' \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" using y'' fv_ik_subset_vars_sst'[of "unlabel A"] by blast hence 1: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t J" unfolding J_def by (metis (no_types, lifting) fv_subset image_cong in_mono repl_invariance) have "(k,\d\\<^sub>s) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I {}" using 0 db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t[of "unlabel A" I "[]"] unfolding db\<^sub>s\<^sub>s\<^sub>t_def by force hence "(k,\d\\<^sub>s) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) J {}" using y'' vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel A"] db\<^sub>s\<^sub>s\<^sub>t_subst_swap[of "unlabel A" I J "[]"] db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t[of "unlabel A" _ "[]"] unfolding db\<^sub>s\<^sub>s\<^sub>t_def J_def by (metis (no_types, lifting) Un_iff empty_set fun_upd_other) hence "((Var y' \ J, \d\\<^sub>s \ J) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) J {})" unfolding J_def fk by simp hence "strand_sem_stateful (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t J) (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) J {}) (unlabel [\n, \Var y' not in \d\\<^sub>s\\]) J" using stateful_strand_sem_NegChecks_no_bvars(1)[ of "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t J" "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) J {}" "Var y'" "\d\\<^sub>s" J] by simp hence 2: "strand_sem_stateful {} {} (unlabel (A@[\n, \Var y' not in \d\\<^sub>s\\])) J" using A'' y' y'' vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel A"] strand_sem_append_stateful[ of "{}" "{}" "unlabel A" "unlabel [\n, \Var y' not in \d\\<^sub>s\\]" J] unfolding J_def by simp have B: "B \ reachable_constraints P" using reachable_constraints.step[OF A(1) P_Tatt _ _ \, of Var Var] unfolding B_def Tatt_def transaction_decl_subst_def transaction_fresh_subst_def by simp have Tatt': "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand Tatt \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = [\\, send\[occurs (Var y')]\\, (l, send\[\f [\c\\<^sub>c, Var y']\\<^sub>t]\), (l', \Var y' not in \d\\<^sub>s\), \n, receive\[attack\ln n\]\\]" using y' unfolding Tatt_def transaction_strand_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def subst_apply_labeled_stateful_strand_def by auto have J: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t J" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t J" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range J)" unfolding J_def subgoal using wt_subst_subst_upd[OF I(1)] k(1) y'(3) by simp subgoal by (metis I(2) k(2) fun_upd_apply interpretation_grounds_all interpretation_substI) subgoal using I(3) k(3) by fastforce done have 3: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t J \ \f [\c\\<^sub>c, Var y']\\<^sub>t \ J" using 1 fk fk' unfolding J_def by auto have 4: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t J \ occurs (Var y') \ J" using reachable_constraints_occurs_fact_ik_case'[ OF A(1) P(1,2) constraint_model_Val_const_in_constr_prefix'[ OF A(1) A' P(1) k_in_ik[unfolded fk kn]]] intruder_deduct.Axiom[of "occurs k" "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t J"] unfolding J_def fk kn by fastforce have "strand_sem_stateful {} {} (unlabel A) J" using 2 strand_sem_append_stateful by force hence "strand_sem_stateful {} {} (unlabel (A@[\\, send\[occurs (Var y')]\\, \n, send\[\f [\c\\<^sub>c, Var y']\\<^sub>t]\\, \n, \Var y' not in \d\\<^sub>s\\])) J" using 2 3 4 strand_sem_append_stateful[of "{}" "{}" _ _ J] unfolding unlabel_def ik\<^sub>s\<^sub>s\<^sub>t_def by force hence "strand_sem_stateful {} {} (unlabel (B@[\n, send\[attack\ln n\]\\])) J" using strand_sem_receive_send_append[of "{}" "{}" _ J "attack\ln n\"] strand_sem_append_stateful[of "{}" "{}" _ _ J] unfolding B_def Tatt' by simp hence "welltyped_constraint_model J (B@[\n, send\[attack\ln n\]\\])" using B J unfolding defs' by blast thus False using B(1) P_sec by blast qed qed lemma welltyped_leakage_free_short_term_secretI': fixes c x y f n d l l' \ defines "s \ \f [\c\\<^sub>c, Var (TAtom \,x)]\\<^sub>t" and "Tatt \ Transaction (\(). []) [] [(l, receive\[\f [\c\\<^sub>c, Var (TAtom \,y)]\\<^sub>t]\)] [(l', \Var (TAtom \,y) not in \d\\<^sub>s\)] [] [\n, send\[attack\ln n\]\\]" assumes P: "\T \ set P. wellformed_transaction T" "\T \ set P. \x \ set (unlabel (transaction_updates T)). is_Update x \ is_Fun (the_set_term x)" and subterms_sec: "\\ \ reachable_constraints P. \\\<^sub>\. welltyped_constraint_model \\<^sub>\ (\@[\\, send\[\c\\<^sub>c]\\])" and P_sec: "\\ \ reachable_constraints P. \\\<^sub>\. welltyped_constraint_model \\<^sub>\ (\@[\n, send\[attack\ln n\]\\])" and P_Tatt: "Tatt \ set P" and P_d: "\T \ set P. \a \ set (transaction_updates T). is_Insert (snd a) \ the_set_term (snd a) = \d\\<^sub>s \ transaction_send T \ [] \ (let (l,b) = hd (transaction_send T) in l = \ \ is_Send b \ \f [\c\\<^sub>c, the_elem_term (snd a)]\\<^sub>t \ set (the_msgs b))" shows "welltyped_leakage_free_protocol [s] P" proof - define Q where "Q \ \M t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}" define Sec where "Sec \ {t \ \ | t \. Q (set [s]) t \} - {m. {} \\<^sub>c m}" define f' where "f' \ \(T,\,\::('fun,'atom,'sets,'lbl) prot_subst,\). dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" define g' where "g' \ concat \ map f'" note defs = Sec_def Q_def note defs' = welltyped_constraint_model_def constraint_model_def show ?thesis proof (rule ccontr) assume "\welltyped_leakage_free_protocol [s] P" then obtain A I k where A: "A \ reachable_constraints P" "\f [\c\\<^sub>c, k]\\<^sub>t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I" "welltyped_constraint_model I (A@[\\, send\[\f [\c\\<^sub>c, k]\\<^sub>t]\\])" unfolding welltyped_leakage_free_protocol_def defs s_def by fastforce have I: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)" using A(3) unfolding defs' by (blast,blast,blast) note A' = welltyped_constraint_model_prefix[OF A(3)] have "strand_sem_stateful {} {} (unlabel A) I" using A' unfolding defs' by simp hence A'': "strand_sem_stateful {} {} (unlabel A) (I(z := k))" when z: "z \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" for z using z strand_sem_model_swap[of "unlabel A" I "I(z := k)" "{}" "{}"] by auto obtain \ where \: "\ (TAtom \,x) = k" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "fv (\ (TAtom \,x)) = {}" using A(2) unfolding defs s_def by auto have k: "\ k = TAtom \" "fv k = {}" "wf\<^sub>t\<^sub>r\<^sub>m k" subgoal using \(1) wt_subst_trm''[OF \(2), of "Var (TAtom \,x)"] by simp subgoal using \(1,4) by blast subgoal using \(1,3) by force done then obtain fk where fk: "k = Fun fk []" using const_type_inv_wf by (cases k) auto have fk': "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I \ \f [\c\\<^sub>c, k]\\<^sub>t" using fk welltyped_constraint_model_deduct_split(2)[OF A(3)] by auto obtain \::"('fun,'atom,'sets,'lbl) prot_subst" where \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" unfolding transaction_renaming_subst_def by blast obtain y' yn' where y': "\ (TAtom \,y) = Var y'" "y \ yn'" "y' = (TAtom \,yn')" using transaction_renaming_subst_is_renaming(1,2)[OF \] by force define B where "B \ A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand Tatt \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" define J where "J \ I(y' := k)" have "y' \ range_vars \" using y'(1) transaction_renaming_subst_is_renaming(3)[OF \] by (metis (no_types, lifting) in_mono subst_fv_imgI term.set_intros(3)) hence y'': "y' \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" using transaction_renaming_subst_vars_disj(6)[OF \] by blast have 0: "(k,\d\\<^sub>s) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I)" proof assume a: "(k,\d\\<^sub>s) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I)" obtain l t t' where t: "(l,insert\t,t'\) \ set A" "t \ I = k" "t' \ I = \d\\<^sub>s" using db\<^sub>s\<^sub>s\<^sub>t_in_cases[OF a[unfolded db\<^sub>s\<^sub>s\<^sub>t_def]] unfolding unlabel_def by auto obtain T b B \ \ \ where T: "prefix (B@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) A" "T \ set P" "transaction_decl_subst \ T" "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" "b \ set (transaction_strand T)" "(l, insert\t,t'\) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p b \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \\<^sub>s \ \\<^sub>s \" "fst (l, insert\t,t'\) = fst b" using reachable_constraints_transaction_action_obtain[OF A(1) t(1)] by metis define \ where "\ \ \ \\<^sub>s \ \\<^sub>s \" obtain b' where "b = (l,b')" using T(8) by (cases b) simp then obtain tb tb' where b': "b = (l,insert\tb,tb'\)" and tb: "t = tb \ \" and tb': "t' = tb' \ \" using T(7) unfolding \_def by (cases b') auto note T_wf = bspec[OF P(1) T(2)] bspec[OF P(2) T(2)] have b: "b \ set (transaction_updates T)" using transaction_strand_memberD[OF T(6)[unfolded b']] wellformed_transaction_cases[OF T_wf(1)] unfolding b' by blast have "is_Fun tb'" using bspec[OF P(2) T(2)] wellformed_transaction_strand_unlabel_memberD(6)[ OF T_wf(1) unlabel_in[OF T(6)[unfolded b']]] by fastforce hence *: "tb' = \d\\<^sub>s" using t(3) unfolding b' tb' by force have "is_Insert (snd b)" "the_set_term (snd b) = \d\\<^sub>s" "the_elem_term (snd b) = tb" unfolding b' * by simp_all hence "transaction_send T \ []" "let (l, a) = hd (transaction_send T) in l = \ \ is_Send a \ \f [\c\\<^sub>c, tb]\\<^sub>t \ set (the_msgs a)" using P_d T(2) b by (fast,fast) hence "\ts. \\,send\ts\\ \ set (transaction_send T) \ \f [\c\\<^sub>c, tb]\\<^sub>t \ set ts" unfolding is_Send_def by (cases "transaction_send T") auto then obtain ts where ts: "\\,send\ts\\ \ set (transaction_strand T)" "\f [\c\\<^sub>c, tb]\\<^sub>t \ set ts" unfolding transaction_strand_def by auto have "\\,receive\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\ \ set A" "\f [\c\\<^sub>c, t]\\<^sub>t \ set (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using subst_lsst_memI[OF ts(1), of \] subst_set_map[OF ts(2), of \] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_steps_iff(1)[of \ "ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] set_mono_prefix[OF T(1)[unfolded \_def[symmetric]]] unfolding tb by auto hence "\f [\c\\<^sub>c, k]\\<^sub>t \ \{set ts |ts. \\, receive\ts\\ \ set (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I)}" using t(2) subst_lsst_memI[of "\\,receive\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\" A I] by force thus False using A(2) unfolding declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by auto qed have "y' \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" using y'' fv_ik_subset_vars_sst'[of "unlabel A"] by blast hence 1: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t J" unfolding J_def by (metis (no_types, lifting) fv_subset image_cong in_mono repl_invariance) have "(k,\d\\<^sub>s) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I {}" using 0 db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t[of "unlabel A" I "[]"] unfolding db\<^sub>s\<^sub>s\<^sub>t_def by force hence "(k,\d\\<^sub>s) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) J {}" using y'' vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel A"] db\<^sub>s\<^sub>s\<^sub>t_subst_swap[of "unlabel A" I J "[]"] db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t[of "unlabel A" _ "[]"] unfolding db\<^sub>s\<^sub>s\<^sub>t_def J_def by (metis (no_types, lifting) Un_iff empty_set fun_upd_other) hence "((Var y' \ J, \d\\<^sub>s \ J) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) J {})" unfolding J_def fk by simp hence "strand_sem_stateful (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t J) (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) J {}) (unlabel [\n, \Var y' not in \d\\<^sub>s\\]) J" using stateful_strand_sem_NegChecks_no_bvars(1)[ of "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t J" "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) J {}" "Var y'" "\d\\<^sub>s" J] by simp hence 2: "strand_sem_stateful {} {} (unlabel (A@[\n, \Var y' not in \d\\<^sub>s\\])) J" using A'' y' y'' vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel A"] strand_sem_append_stateful[ of "{}" "{}" "unlabel A" "unlabel [\n, \Var y' not in \d\\<^sub>s\\]" J] unfolding J_def by simp have B: "B \ reachable_constraints P" using reachable_constraints.step[OF A(1) P_Tatt _ _ \, of Var Var] unfolding B_def Tatt_def transaction_decl_subst_def transaction_fresh_subst_def by simp have Tatt': "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand Tatt \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = [(l, send\[\f [\c\\<^sub>c, Var y']\\<^sub>t]\), (l', \Var y' not in \d\\<^sub>s\), \n, receive\[attack\ln n\]\\]" using y' unfolding Tatt_def transaction_strand_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def subst_apply_labeled_stateful_strand_def by auto have J: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t J" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t J" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range J)" unfolding J_def subgoal using wt_subst_subst_upd[OF I(1)] k(1) y'(3) by simp subgoal by (metis I(2) k(2) fun_upd_apply interpretation_grounds_all interpretation_substI) subgoal using I(3) k(3) by fastforce done have 3: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t J \ \f [\c\\<^sub>c, Var y']\\<^sub>t \ J" using 1 fk fk' unfolding J_def by auto have "strand_sem_stateful {} {} (unlabel A) J" using 2 strand_sem_append_stateful by force hence "strand_sem_stateful {} {} (unlabel (A@[\n, send\[\f [\c\\<^sub>c, Var y']\\<^sub>t]\\, \n, \Var y' not in \d\\<^sub>s\\])) J" using 2 3 strand_sem_append_stateful[of "{}" "{}" _ _ J] unfolding unlabel_def ik\<^sub>s\<^sub>s\<^sub>t_def by force hence "strand_sem_stateful {} {} (unlabel (B@[\n, send\[attack\ln n\]\\])) J" using strand_sem_receive_send_append[of "{}" "{}" _ J "attack\ln n\"] strand_sem_append_stateful[of "{}" "{}" _ _ J] unfolding B_def Tatt' by simp hence "welltyped_constraint_model J (B@[\n, send\[attack\ln n\]\\])" using B J unfolding defs' by blast thus False using B(1) P_sec by blast qed qed definition welltyped_leakage_free_invkey_conditions' where "welltyped_leakage_free_invkey_conditions' invfun privfunsec declassifiedset S n P \ let f = \s. is_Var s \ fst (the_Var s) = TAtom Value; g = \s. is_Fun s \ args s = [] \ is_Set (the_Fun s) \ arity\<^sub>s (the_Set (the_Fun s)) = 0; h = \s. is_Fun s \ args s = [] \ is_Fu (the_Fun s) \ public\<^sub>f (the_Fu (the_Fun s)) \ arity\<^sub>f (the_Fu (the_Fun s)) = 0 in (\s\set S. f s \ (is_Fun s \ the_Fun s = Pair \ length (args s) = 2 \ g (args s ! 1)) \ g s \ h s \ s = \privfunsec\\<^sub>c \ s = Fun OccursSec [] \ (is_Fun s \ the_Fun s = OccursFact \ length (args s) = 2 \ args s ! 0 = Fun OccursSec []) \ (is_Fun s \ the_Fun s = Fu invfun \ length (args s) = 2 \ args s ! 0 = \privfunsec\\<^sub>c \ f (args s ! 1)) \ (is_Fun s \ is_Fu (the_Fun s) \ fv s = {} \ Transaction (\(). []) [] [\n, receive\[s]\\] [] [] [\n, send\[attack\ln n\]\\]\set P)) \ (\public\<^sub>f privfunsec \ arity\<^sub>f privfunsec = 0 \ \\<^sub>f privfunsec \ None) \ (\T\set P. transaction_fresh T \ [] \ transaction_send T \ [] \ (let (l, a) = hd (transaction_send T) in l = \ \ is_Send a \ Var ` set (transaction_fresh T) \ set (the_msgs a))) \ (\T\set P. \x\vars_transaction T. is_Var (\\<^sub>v x)) \ (\T\set P. \x\set (transaction_fresh T). \\<^sub>v x = Var Value \ (\a. \\<^sub>v x = \a\\<^sub>\\<^sub>a)) \ (\T\set P. \f\\(funs_term ` trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)). \is_Set f) \ (\T\set P. \r\set (transaction_send T). OccursFact \ \(funs_term ` trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd r)) \ has_LabelS r) \ (\T\set P. \t\subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)). \privfunsec\\<^sub>c \ set (snd (Ana t))) \ (\T\set P. \privfunsec\\<^sub>c \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \ (\T\set P. \a\set (transaction_updates T). is_Insert (snd a) \ the_set_term (snd a) = \declassifiedset\\<^sub>s \ transaction_send T \ [] \ (let (l, b) = hd (transaction_send T) in l = \ \ is_Send b \ \invfun [\privfunsec\\<^sub>c, the_elem_term (snd a)]\\<^sub>t \ set (the_msgs b)))" definition welltyped_leakage_free_invkey_conditions where "welltyped_leakage_free_invkey_conditions invfun privfunsec declassifiedset S n P \ let Tatt = \R. Transaction (\(). []) [] (R@[\n, receive\[\invfun [\privfunsec\\<^sub>c, \0: value\\<^sub>v]\\<^sub>t]\\]) [\\, \\0: value\\<^sub>v not in \declassifiedset\\<^sub>s\\] [] [\n, send\[attack\ln n\]\\] in welltyped_leakage_free_invkey_conditions' invfun privfunsec declassifiedset S n P \ has_initial_value_producing_transaction P \ (if Tatt [\\, receive\[occurs \0: value\\<^sub>v]\\] \ set P then \T\set P. admissible_transaction' T \ admissible_transaction_occurs_checks T else Tatt [] \ set P \ (\T\set P. wellformed_transaction T) \ (\T\set P. admissible_transaction_terms T) \ (\T\set P. bvars_transaction T = {}) \ (\T\set P. transaction_decl T () = []) \ (\T\set P. \x\set (transaction_fresh T). let \ = fst x in \ = TAtom Value \ \ \ \ \privfunsec\\<^sub>c) \ (\T\set P. \x\vars_transaction T. let \ = fst x in is_Var \ \ (the_Var \ = Value \ is_Atom (the_Var \)) \ \ \ \ \privfunsec\\<^sub>c) \ (\T\set P. \t\subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)). Fun OccursSec [] \ set (snd (Ana t))) \ (\T\set P. Fun OccursSec [] \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \ (\T\set P. \x\set (unlabel (transaction_updates T)). is_Update x \ is_Fun (the_set_term x)) \ (\s\set S. is_Fun s \ the_Fun s \ OccursFact))" lemma welltyped_leakage_free_invkeyI: assumes P_wt_secure: "\\ \ reachable_constraints P. \\. welltyped_constraint_model \ (\@[\n, send\[attack\ln n\]\\])" and a: "welltyped_leakage_free_invkey_conditions invfun privsec declassset S n P" shows "welltyped_leakage_free_protocol S P" proof - let ?Tatt' = "\R C. Transaction (\(). []) [] R C [] [\n, send\[attack\ln n\]\\] ::('fun,'atom,'sets,'lbl) prot_transaction" let ?Tatt = "\R. ?Tatt' (R@[\n, receive\[\invfun [\privsec\\<^sub>c, \0: value\\<^sub>v]\\<^sub>t]\\]) [\\, \\0: value\\<^sub>v not in \declassset\\<^sub>s\\]" define Tatt1 where "Tatt1 \ ?Tatt [\\, receive\[occurs \0: value\\<^sub>v]\\]" define Tatt2 where "Tatt2 \ ?Tatt []" define Tatt_lts where "Tatt_lts \ \s. ?Tatt' [\n, receive\[s]\\] []" note defs = welltyped_leakage_free_invkey_conditions_def Let_def note defs' = defs welltyped_leakage_free_invkey_conditions'_def note Tatts = Tatt1_def Tatt2_def Tatt_lts_def obtain at where 0: "\public\<^sub>f privsec" "arity\<^sub>f privsec = 0" "\\<^sub>f privsec = Some at" using a unfolding defs' by fast have *: "\T\set P. admissible_transaction' T" "\T\set P. admissible_transaction_occurs_checks T" when "Tatt1 \ set P" using a that unfolding defs Tatts by (meson,meson) have **: "Tatt1 \ set P \ Tatt2 \ set P" using a unfolding defs Tatts by argo have ***: "\T\set P. \x\set (transaction_fresh T). \\<^sub>v x = TAtom Value \ \\<^sub>v x \ \ \privsec\\<^sub>c" "\T\set P. \x\vars_transaction T. \a. \\<^sub>v x = TAtom a \ (a = Value \ (\b. a = Atom b)) \ TAtom a \ \ \privsec\\<^sub>c" when "Tatt1 \ set P" subgoal using a that \\<^sub>v_TAtom''(2) unfolding defs Tatts by metis subgoal using a that \\<^sub>v_TAtom''(1,2) unfolding defs Tatts[symmetric] is_Atom_def is_Var_def by fastforce done have ****: "s \ occurs x" when "Tatt1 \ set P" "s \ set S" for s x using a that ** unfolding defs Tatts the_Fun_def by fastforce have 1: "\T\set P. transaction_fresh T \ [] \ transaction_send T \ [] \ (let (l, a) = hd (transaction_send T) in l = \ \ is_Send a \ Var ` set (transaction_fresh T) \ set (the_msgs a))" "\T\set P. \x\vars_transaction T. is_Var (\\<^sub>v x)" "\T\set P. \x\set (transaction_fresh T). \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = \a\\<^sub>\\<^sub>a)" "\T\set P. \f\\(funs_term ` trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)). \is_Set f" "\T\set P. \r\set (transaction_send T). OccursFact \ \(funs_term ` trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd r)) \ has_LabelS r" "\T\set P. \t\subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)). \privsec\\<^sub>c \ set (snd (Ana t))" "\T\set P. \privsec\\<^sub>c \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "\T\set P. \a\set (transaction_updates T). is_Insert (snd a) \ the_set_term (snd a) = \declassset\\<^sub>s \ transaction_send T \ [] \ (let (l, b) = hd (transaction_send T) in l = \ \ is_Send b \ \invfun [\privsec\\<^sub>c, the_elem_term (snd a)]\\<^sub>t \ set (the_msgs b))" using a unfolding defs' by (meson,meson,meson,meson,meson,meson,meson,meson) have 2: "\T\set P. wellformed_transaction T" "\T\set P. \x\vars_transaction T \ set (transaction_fresh T). \ \privsec\\<^sub>c \ \\<^sub>v x" "\T\set P. admissible_transaction_terms T" "\T\set P. \x\set (transaction_fresh T). \\<^sub>v x = TAtom Value" "\T\set P. transaction_decl T () = []" "\T\set P. \x\vars_transaction T. \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = \a\\<^sub>\\<^sub>a)" "\T\set P. \x\vars_transaction T \ set (transaction_fresh T). \\<^sub>v x \ TAtom SetType" "\T\set P. \x\vars_transaction T \ set (transaction_fresh T). \\<^sub>v x \ TAtom OccursSecType" "\T\set P. \t\subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)). Fun OccursSec [] \ set (snd (Ana t))" "\T\set P. Fun OccursSec [] \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "\T\set P. bvars_transaction T = {}" "\T\set P. \x\set (unlabel (transaction_updates T)). is_Update x \ is_Fun (the_set_term x)" subgoal using a * unfolding defs by (metis admissible_transaction_is_wellformed_transaction(1)) subgoal apply (cases "Tatt1 \ set P") subgoal using a * admissible_transactionE(2,3) \_Fu_simps(4) unfolding defs Tatts by force subgoal using a \_Fu_simps(4) unfolding defs Tatts by fastforce done subgoal using a * admissible_transaction_is_wellformed_transaction(4) unfolding defs by metis subgoal using a * admissible_transactionE(2) unfolding defs Tatts[symmetric] by fastforce subgoal using a * admissible_transactionE(1) unfolding defs Tatts[symmetric] by metis subgoal using * *** admissible_transactionE(3) by fast subgoal using * *** admissible_transactionE(2,3) by (cases "Tatt1 \ set P") (force, fastforce) subgoal using * *** admissible_transactionE(2,3) by (cases "Tatt1 \ set P") (force, fastforce) subgoal using a * admissible_transaction_occurs_checksE6 unfolding defs by metis subgoal using a * admissible_transaction_occurs_checksE5 unfolding defs by metis subgoal using a * admissible_transaction_no_bvars(2) unfolding defs Tatts[symmetric] by fastforce subgoal using a * admissible_transaction_is_wellformed_transaction(3) unfolding defs Tatts[symmetric] admissible_transaction_updates_def by fastforce done have Tatt_lts_case: "\\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ inj_on \ (fv s) \ \ ` fv s \ range Var \ ?Tatt' [\n, receive\[s \ \]\\] [] \ set P" when s: "fv s = {}" "Tatt_lts s \ set P" for s proof - have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t Var" "inj_on Var (fv s)" "Var ` fv s \ range Var" "s \ Var = s" using s(1) by simp_all thus ?thesis using s(2) unfolding Tatts by metis qed have Tatt1_case: "?Tatt' [\\, receive\[occurs \0: value\\<^sub>v]\\, \n, receive\[\invfun [\privsec\\<^sub>c, \0: value\\<^sub>v]\\<^sub>t]\\] [\\, \\0: value\\<^sub>v not in \declassset\\<^sub>s\\] \ set P" when "Tatt1 \ set P" using that unfolding Tatts by auto have Tatt2_case: "?Tatt' [\n, receive\[\invfun [\privsec\\<^sub>c, \0: value\\<^sub>v]\\<^sub>t]\\] [\\, \\0: value\\<^sub>v not in \declassset\\<^sub>s\\] \ set P" when "Tatt2 \ set P" using that unfolding Tatts by auto note 3 = pair_def case_prod_conv note 4 = welltyped_leakage_free_priv_constI'[OF 0(1-3) 2(1,2) 1(2,3,6,7)] note 5 = welltyped_leakage_free_setop_pairI[OF 2(1,6) 1(4) 2(4,5,3), unfolded 3] welltyped_leakage_free_set_constI(2)[OF 2(1) 1(4) 2(7) 1(2,3), unfolded 3] welltyped_leakage_free_pub_constI 4(2) welltyped_leakage_free_occurssec_constI(2)[OF 2(1,8-10) 1(2,3)] welltyped_leakage_free_value_constI[OF 2(1,3,5,11) 1(1)] welltyped_leakage_free_short_term_secretI'[ OF 2(1,12) 4(1) P_wt_secure Tatt2_case 1(8)] welltyped_leakage_free_long_term_secretI[OF P_wt_secure Tatt_lts_case] welltyped_leakage_free_short_term_secretI[ OF * 1(3) 4(1) P_wt_secure Tatt1_case 1(8)] welltyped_leakage_free_occurs_factI[OF * 1(5)] ** **** have 6: "is_Fun s \ length (args s) = 2 \ (\f t u. s = Fun f [t, u])" for s::"('fun,'atom,'sets,'lbl) prot_term" by auto define pubconst_cond where "pubconst_cond \ \s::('fun,'atom,'sets,'lbl) prot_term. is_Fun s \ args s = [] \ is_Fu (the_Fun s) \ public\<^sub>f (the_Fu (the_Fun s)) \ arity\<^sub>f (the_Fu (the_Fun s)) = 0" define valuevar_cond where "valuevar_cond \ \s::('fun,'atom,'sets,'lbl) prot_term. is_Var s \ fst (the_Var s) = TAtom Value" define setconst_cond where "setconst_cond \ \s::('fun,'atom,'sets,'lbl) prot_term. is_Fun s \ args s = [] \ is_Set (the_Fun s) \ arity\<^sub>s (the_Set (the_Fun s)) = 0" define setop_pair_cond where "setop_pair_cond \ \s. is_Fun s \ the_Fun s = Pair \ length (args s) = 2 \ setconst_cond (args s ! 1)" define occursfact_cond where "occursfact_cond \ \s::('fun,'atom,'sets,'lbl) prot_term. is_Fun s \ the_Fun s = OccursFact \ length (args s) = 2 \ args s ! 0 = Fun OccursSec []" define invkey_cond where "invkey_cond \ \s. is_Fun s \ the_Fun s = Fu invfun \ length (args s) = 2 \ args s ! 0 = \privsec\\<^sub>c \ valuevar_cond (args s ! 1)" define ground_lts_cond where "ground_lts_cond \ \s. is_Fun s \ is_Fu (the_Fun s) \ fv s = {} \ Tatt_lts s \ set P" note cond_defs = pubconst_cond_def valuevar_cond_def setconst_cond_def setop_pair_cond_def occursfact_cond_def invkey_cond_def ground_lts_cond_def have "(\m. s = \m: value\\<^sub>v) \ valuevar_cond s" "(\x c. arity\<^sub>s c = 0 \ s = Fun Pair [x, \c\\<^sub>s]) \ setop_pair_cond s" "(\c. arity\<^sub>s c = 0 \ s = \c\\<^sub>s) \ setconst_cond s" "(\c. public\<^sub>f c \ arity\<^sub>f c = 0 \ s = \c\\<^sub>c) \ pubconst_cond s" "(\x. s = occurs x) \ occursfact_cond s" "(\x. s = \invfun [\privsec\\<^sub>c, \x: value\\<^sub>v]\\<^sub>t) \ invkey_cond s" "(\f ts. s = \f ts\\<^sub>t \ fv s = {} \ Tatt_lts s \ set P) \ ground_lts_cond s" for s::"('fun,'atom,'sets,'lbl) prot_term" unfolding is_Set_def the_Set_def is_Fu_def cond_defs by (fastforce,use 6[of s] in fastforce,fastforce,force,fastforce,fastforce,fastforce) moreover have "(\s \ set S. valuevar_cond s \ setop_pair_cond s \ setconst_cond s \ pubconst_cond s \ s = \privsec\\<^sub>c \ s = Fun OccursSec [] \ occursfact_cond s \ invkey_cond s \ ground_lts_cond s) \ (\public\<^sub>f privsec \ arity\<^sub>f privsec = 0 \ \\<^sub>f privsec \ None)" using a unfolding defs' cond_defs Tatts by meson ultimately have 7: "\s \ set S. (\x c. arity\<^sub>s c = 0 \ s = Fun Pair [x, \c\\<^sub>s]) \ (\c. arity\<^sub>s c = 0 \ s = \c\\<^sub>s) \ (\c. public\<^sub>f c \ arity\<^sub>f c = 0 \ s = \c\\<^sub>c) \ s = \privsec\\<^sub>c \ s = Fun OccursSec [] \ (\m. s = \m: value\\<^sub>v) \ (\x. s = occurs x) \ (\x. s = \invfun [\privsec\\<^sub>c, \x: value\\<^sub>v]\\<^sub>t) \ (\f ts. s = \f ts\\<^sub>t \ fv s = {} \ Tatt_lts s \ set P)" unfolding Let_def by fastforce show ?thesis by (rule iffD2[OF welltyped_leakage_free_protocol_pointwise]; unfold list_all_iff; intro ballI) (use bspec[OF 7] 5 in blast) qed end end locale composable_stateful_protocols = pm: stateful_protocol_model arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 for arity\<^sub>f::"'fun \ nat" and arity\<^sub>s::"'sets \ nat" and public\<^sub>f::"'fun \ bool" and Ana\<^sub>f::"'fun \ ((('fun,'atom::finite,'sets,nat) prot_fun, nat) term list \ nat list)" and \\<^sub>f::"'fun \ 'atom option" and label_witness1::"nat" and label_witness2::"nat" + fixes Pc::"('fun,'atom,'sets,nat) prot_transaction list" and Ps Ps_with_star_projs::"('fun,'atom,'sets,nat) prot_transaction list list" and Pc_SMP Sec_symbolic::"('fun,'atom,'sets,nat) prot_term list" and Ps_GSMPs::"(nat \ ('fun,'atom,'sets,nat) prot_term list) list" assumes Pc_def: "Pc = concat Ps" and Ps_with_star_projs_def: "let Pc' = Pc; L = [0..n. (map (transaction_proj n) Pc')) L \ set L = set (remdups (concat ( map (\T. map (the_LabelN \ fst) (filter (Not \ has_LabelS) (transaction_strand T))) Pc')))" and Pc_wellformed_composable: "list_all (list_all (Not \ has_LabelS) \ tl \ transaction_send) Pc" "pm.wellformed_composable_protocols Ps Pc_SMP" "pm.wellformed_composable_protocols' Ps" "pm.composable_protocols Ps Ps_GSMPs Sec_symbolic" begin theorem composed_protocol_preserves_component_goals: assumes components_leakage_free: "list_all (pm.welltyped_leakage_free_protocol Sec_symbolic) Ps_with_star_projs" and n_def: "n < length Ps_with_star_projs" and P_def: "P = Ps_with_star_projs ! n" and P_welltyped_secure: "\\ \ pm.reachable_constraints P. \\. pm.welltyped_constraint_model \ (\@[\n, send\[attack\ln n\]\\])" shows "\\ \ pm.reachable_constraints Pc. \\. pm.constraint_model \ (\@[\n, send\[attack\ln n\]\\])" proof - note 0 = Ps_with_star_projs_def[unfolded Let_def] have 1: "set Ps_with_star_projs = (\n. map (transaction_proj n) Pc) ` set (remdups (concat (map (\T. map (the_LabelN \ fst) (filter (Not \ has_LabelS) (transaction_strand T))) Pc)))" by (metis (mono_tags, lifting) 0 image_set) have 2: "Ps_with_star_projs ! n = map (transaction_proj n) Pc" using conjunct1[OF 0] n_def by fastforce show ?thesis by (rule pm.composable_protocols_par_comp_prot'[ OF Pc_def 1 Pc_wellformed_composable components_leakage_free 2 P_welltyped_secure[unfolded P_def]]) qed end end diff --git a/thys/Automated_Stateful_Protocol_Verification/Term_Implication.thy b/thys/Automated_Stateful_Protocol_Verification/Term_Implication.thy --- a/thys/Automated_Stateful_Protocol_Verification/Term_Implication.thy +++ b/thys/Automated_Stateful_Protocol_Verification/Term_Implication.thy @@ -1,2845 +1,2845 @@ (* Title: Term_Implication.thy Author: Andreas Viktor Hess, DTU Author: Sebastian A. Mödersheim, DTU Author: Achim D. Brucker, University of Exeter Author: Anders Schlichtkrull, DTU SPDX-License-Identifier: BSD-3-Clause *) section\Term Implication\ theory Term_Implication imports Stateful_Protocol_Model Term_Variants begin subsection \Single Term Implications\ definition timpl_apply_term ("\_ --\ _\\_\") where "\a --\ b\\t\ \ term_variants ((\_. [])(Abs a := [Abs b])) t" definition timpl_apply_terms ("\_ --\ _\\_\\<^sub>s\<^sub>e\<^sub>t") where "\a --\ b\\M\\<^sub>s\<^sub>e\<^sub>t \ \((set o timpl_apply_term a b) ` M)" lemma timpl_apply_Fun: assumes "\i. i < length T \ S ! i \ set \a --\ b\\T ! i\" and "length T = length S" shows "Fun f S \ set \a --\ b\\Fun f T\" using assms term_variants_Fun term_variants_pred_iff_in_term_variants by (metis timpl_apply_term_def) lemma timpl_apply_Abs: assumes "\i. i < length T \ S ! i \ set \a --\ b\\T ! i\" and "length T = length S" shows "Fun (Abs b) S \ set \a --\ b\\Fun (Abs a) T\" using assms(1) term_variants_P[OF assms(2), of "(\_. [])(Abs a := [Abs b])" "Abs b" "Abs a"] unfolding timpl_apply_term_def term_variants_pred_iff_in_term_variants[symmetric] by fastforce lemma timpl_apply_refl: "t \ set \a --\ b\\t\" unfolding timpl_apply_term_def by (metis term_variants_pred_refl term_variants_pred_iff_in_term_variants) lemma timpl_apply_const: "Fun (Abs b) [] \ set \a --\ b\\Fun (Abs a) []\" using term_variants_pred_iff_in_term_variants term_variants_pred_const unfolding timpl_apply_term_def by auto lemma timpl_apply_const': "c = a \ set \a --\ b\\Fun (Abs c) []\ = {Fun (Abs b) [], Fun (Abs c) []}" "c \ a \ set \a --\ b\\Fun (Abs c) []\ = {Fun (Abs c) []}" using term_variants_pred_const_cases[of "(\_. [])(Abs a := [Abs b])" "Abs c"] term_variants_pred_iff_in_term_variants[of "(\_. [])(Abs a := [Abs b])"] unfolding timpl_apply_term_def by auto lemma timpl_apply_term_subst: "s \ set \a --\ b\\t\ \ s \ \ \ set \a --\ b\\t \ \\" by (metis term_variants_pred_iff_in_term_variants term_variants_pred_subst timpl_apply_term_def) lemma timpl_apply_inv: assumes "Fun h S \ set \a --\ b\\Fun f T\" shows "length T = length S" and "\i. i < length T \ S ! i \ set \a --\ b\\T ! i\" and "f \ h \ f = Abs a \ h = Abs b" using assms term_variants_pred_iff_in_term_variants[of "(\_. [])(Abs a := [Abs b])"] unfolding timpl_apply_term_def by (metis (full_types) term_variants_pred_inv(1), metis (full_types) term_variants_pred_inv(2), fastforce dest: term_variants_pred_inv(3)) lemma timpl_apply_inv': assumes "s \ set \a --\ b\\Fun f T\" shows "\g S. s = Fun g S" proof - have *: "term_variants_pred ((\_. [])(Abs a := [Abs b])) (Fun f T) s" using assms term_variants_pred_iff_in_term_variants[of "(\_. [])(Abs a := [Abs b])"] unfolding timpl_apply_term_def by force show ?thesis using term_variants_pred.cases[OF *, of ?thesis] by fastforce qed lemma timpl_apply_term_Var_iff: "Var x \ set \a --\ b\\t\ \ t = Var x" using term_variants_pred_inv_Var term_variants_pred_iff_in_term_variants unfolding timpl_apply_term_def by metis subsection \Term Implication Closure\ inductive_set timpl_closure for t TI where FP: "t \ timpl_closure t TI" | TI: "\u \ timpl_closure t TI; (a,b) \ TI; term_variants_pred ((\_. [])(Abs a := [Abs b])) u s\ \ s \ timpl_closure t TI" definition "timpl_closure_set M TI \ (\t \ M. timpl_closure t TI)" inductive_set timpl_closure'_step for TI where "\(a,b) \ TI; term_variants_pred ((\_. [])(Abs a := [Abs b])) t s\ \ (t,s) \ timpl_closure'_step TI" definition "timpl_closure' TI \ (timpl_closure'_step TI)\<^sup>*" definition comp_timpl_closure where "comp_timpl_closure FP TI \ let f = \X. FP \ (\x \ X. \(a,b) \ TI. set \a --\ b\\x\) in while (\X. f X \ X) f {}" definition comp_timpl_closure_list where "comp_timpl_closure_list FP TI \ let f = \X. remdups (concat (map (\x. concat (map (\(a,b). \a --\ b\\x\) TI)) X)@X) in while (\X. set (f X) \ set X) f FP" lemma timpl_closure_setI: "t \ M \ t \ timpl_closure_set M TI" unfolding timpl_closure_set_def by (auto intro: timpl_closure.FP) lemma timpl_closure_set_empty_timpls: "timpl_closure t {} = {t}" (is "?A = ?B") proof (intro subset_antisym subsetI) fix s show "s \ ?A \ s \ ?B" by (induct s rule: timpl_closure.induct) auto qed (simp add: timpl_closure.FP) lemmas timpl_closure_set_is_timpl_closure_union = meta_eq_to_obj_eq[OF timpl_closure_set_def] lemma term_variants_pred_eq_case_Abs: fixes a b defines "P \ (\_. [])(Abs a := [Abs b])" assumes "term_variants_pred P t s" "\f \ funs_term s. \is_Abs f" shows "t = s" using assms(2,3) proof (induction t s rule: term_variants_pred.induct) case (term_variants_Fun T S f) have "\is_Abs h" when i: "i < length S" and h: "h \ funs_term (S ! i)" for i h using i h term_variants_Fun.prems by auto hence "T ! i = S ! i" when i: "i < length T" for i using i term_variants_Fun.hyps(1) term_variants_Fun.IH by auto hence "T = S" using term_variants_Fun.hyps(1) nth_equalityI[of T S] by fast thus ?case using term_variants_Fun.hyps(1) by blast qed (simp_all add: term_variants_pred_refl P_def) lemma timpl_closure'_step_inv: assumes "(t,s) \ timpl_closure'_step TI" obtains a b where "(a,b) \ TI" "term_variants_pred ((\_. [])(Abs a := [Abs b])) t s" using assms by (auto elim: timpl_closure'_step.cases) lemma timpl_closure_mono: assumes "TI \ TI'" shows "timpl_closure t TI \ timpl_closure t TI'" proof fix s show "s \ timpl_closure t TI \ s \ timpl_closure t TI'" apply (induct rule: timpl_closure.induct) using assms by (auto intro: timpl_closure.intros) qed lemma timpl_closure_set_mono: assumes "M \ M'" "TI \ TI'" shows "timpl_closure_set M TI \ timpl_closure_set M' TI'" using assms(1) timpl_closure_mono[OF assms(2)] unfolding timpl_closure_set_def by fast lemma timpl_closure_idem: "timpl_closure_set (timpl_closure t TI) TI = timpl_closure t TI" (is "?A = ?B") proof have "s \ timpl_closure t TI" when "s \ timpl_closure u TI" "u \ timpl_closure t TI" for s u using that by (induction rule: timpl_closure.induct) (auto intro: timpl_closure.intros) thus "?A \ ?B" unfolding timpl_closure_set_def by blast show "?B \ ?A" unfolding timpl_closure_set_def by (blast intro: timpl_closure.FP) qed lemma timpl_closure_set_idem: "timpl_closure_set (timpl_closure_set M TI) TI = timpl_closure_set M TI" using timpl_closure_idem[of _ TI]unfolding timpl_closure_set_def by auto lemma timpl_closure_set_mono_timpl_closure_set: assumes N: "N \ timpl_closure_set M TI" shows "timpl_closure_set N TI \ timpl_closure_set M TI" using timpl_closure_set_mono[OF N, of TI TI] timpl_closure_set_idem[of M TI] by simp lemma timpl_closure_is_timpl_closure': "s \ timpl_closure t TI \ (t,s) \ timpl_closure' TI" proof show "s \ timpl_closure t TI \ (t,s) \ timpl_closure' TI" unfolding timpl_closure'_def by (induct rule: timpl_closure.induct) (auto intro: rtrancl_into_rtrancl timpl_closure'_step.intros) show "(t,s) \ timpl_closure' TI \ s \ timpl_closure t TI" unfolding timpl_closure'_def by (induct rule: rtrancl_induct) (auto dest: timpl_closure'_step_inv intro: timpl_closure.FP timpl_closure.TI) qed lemma timpl_closure'_mono: assumes "TI \ TI'" shows "timpl_closure' TI \ timpl_closure' TI'" using timpl_closure_mono[OF assms] timpl_closure_is_timpl_closure'[of _ _ TI] timpl_closure_is_timpl_closure'[of _ _ TI'] by fast lemma timpl_closureton_is_timpl_closure: "timpl_closure_set {t} TI = timpl_closure t TI" by (simp add: timpl_closure_set_is_timpl_closure_union) lemma timpl_closure'_timpls_trancl_subset: "timpl_closure' (c\<^sup>+) \ timpl_closure' c" unfolding timpl_closure'_def proof fix s t::"(('a,'b,'c,'d) prot_fun,'e) term" show "(s,t) \ (timpl_closure'_step (c\<^sup>+))\<^sup>* \ (s,t) \ (timpl_closure'_step c)\<^sup>*" proof (induction rule: rtrancl_induct) case (step u t) obtain a b where ab: "(a,b) \ c\<^sup>+" "term_variants_pred ((\_. [])(Abs a := [Abs b])) u t" using step.hyps(2) timpl_closure'_step_inv by blast hence "(u,t) \ (timpl_closure'_step c)\<^sup>*" proof (induction arbitrary: t rule: trancl_induct) case (step d e) obtain s where s: "term_variants_pred ((\_. [])(Abs a := [Abs d])) u s" "term_variants_pred ((\_. [])(Abs d := [Abs e])) s t" using term_variants_pred_dense'[OF step.prems, of "Abs d"] by blast have "(u,s) \ (timpl_closure'_step c)\<^sup>*" "(s,t) \ timpl_closure'_step c" using step.hyps(2) s(2) step.IH[OF s(1)] by (auto intro: timpl_closure'_step.intros) thus ?case by simp qed (auto intro: timpl_closure'_step.intros) thus ?case using step.IH by simp qed simp qed lemma timpl_closure'_timpls_trancl_subset': "timpl_closure' {(a,b) \ c\<^sup>+. a \ b} \ timpl_closure' c" using timpl_closure'_timpls_trancl_subset timpl_closure'_mono[of "{(a,b) \ c\<^sup>+. a \ b}" "c\<^sup>+"] by fast lemma timpl_closure_set_timpls_trancl_subset: "timpl_closure_set M (c\<^sup>+) \ timpl_closure_set M c" using timpl_closure'_timpls_trancl_subset[of c] timpl_closure_is_timpl_closure'[of _ _ c] timpl_closure_is_timpl_closure'[of _ _ "c\<^sup>+"] timpl_closure_set_is_timpl_closure_union[of M c] timpl_closure_set_is_timpl_closure_union[of M "c\<^sup>+"] by fastforce lemma timpl_closure_set_timpls_trancl_subset': "timpl_closure_set M {(a,b) \ c\<^sup>+. a \ b} \ timpl_closure_set M c" using timpl_closure'_timpls_trancl_subset'[of c] timpl_closure_is_timpl_closure'[of _ _ c] timpl_closure_is_timpl_closure'[of _ _ "{(a,b) \ c\<^sup>+. a \ b}"] timpl_closure_set_is_timpl_closure_union[of M c] timpl_closure_set_is_timpl_closure_union[of M "{(a,b) \ c\<^sup>+. a \ b}"] by fastforce lemma timpl_closure'_timpls_trancl_supset': "timpl_closure' c \ timpl_closure' {(a,b) \ c\<^sup>+. a \ b}" unfolding timpl_closure'_def proof let ?cl = "{(a,b) \ c\<^sup>+. a \ b}" fix s t::"(('a,'b,'c,'d) prot_fun,'e) term" show "(s,t) \ (timpl_closure'_step c)\<^sup>* \ (s,t) \ (timpl_closure'_step ?cl)\<^sup>*" proof (induction rule: rtrancl_induct) case (step u t) obtain a b where ab: "(a,b) \ c" "term_variants_pred ((\_. [])(Abs a := [Abs b])) u t" using step.hyps(2) timpl_closure'_step_inv by blast hence "(a,b) \ c\<^sup>+" by simp hence "(u,t) \ (timpl_closure'_step ?cl)\<^sup>*" using ab(2) proof (induction arbitrary: t rule: trancl_induct) case (base d) show ?case proof (cases "a = d") case True thus ?thesis using base term_variants_pred_refl_inv[of _ u t] by force next case False thus ?thesis using base timpl_closure'_step.intros[of a d ?cl] by fast qed next case (step d e) obtain s where s: "term_variants_pred ((\_. [])(Abs a := [Abs d])) u s" "term_variants_pred ((\_. [])(Abs d := [Abs e])) s t" using term_variants_pred_dense'[OF step.prems, of "Abs d"] by blast show ?case proof (cases "d = e") case True thus ?thesis using step.prems step.IH[of t] by blast next case False hence "(u,s) \ (timpl_closure'_step ?cl)\<^sup>*" "(s,t) \ timpl_closure'_step ?cl" using step.hyps(2) s(2) step.IH[OF s(1)] by (auto intro: timpl_closure'_step.intros) thus ?thesis by simp qed qed thus ?case using step.IH by simp qed simp qed lemma timpl_closure'_timpls_trancl_supset: "timpl_closure' c \ timpl_closure' (c\<^sup>+)" using timpl_closure'_timpls_trancl_supset'[of c] timpl_closure'_mono[of "{(a,b) \ c\<^sup>+. a \ b}" "c\<^sup>+"] by fast lemma timpl_closure'_timpls_trancl_eq: "timpl_closure' (c\<^sup>+) = timpl_closure' c" using timpl_closure'_timpls_trancl_subset timpl_closure'_timpls_trancl_supset by blast lemma timpl_closure'_timpls_trancl_eq': "timpl_closure' {(a,b) \ c\<^sup>+. a \ b} = timpl_closure' c" using timpl_closure'_timpls_trancl_subset' timpl_closure'_timpls_trancl_supset' by blast lemma timpl_closure'_timpls_rtrancl_subset: "timpl_closure' (c\<^sup>*) \ timpl_closure' c" unfolding timpl_closure'_def proof fix s t::"(('a,'b,'c,'d) prot_fun,'e) term" show "(s,t) \ (timpl_closure'_step (c\<^sup>*))\<^sup>* \ (s,t) \ (timpl_closure'_step c)\<^sup>*" proof (induction rule: rtrancl_induct) case (step u t) obtain a b where ab: "(a,b) \ c\<^sup>*" "term_variants_pred ((\_. [])(Abs a := [Abs b])) u t" using step.hyps(2) timpl_closure'_step_inv by blast hence "(u,t) \ (timpl_closure'_step c)\<^sup>*" proof (induction arbitrary: t rule: rtrancl_induct) case base hence "u = t" using term_variants_pred_refl_inv by fastforce thus ?case by simp next case (step d e) obtain s where s: "term_variants_pred ((\_. [])(Abs a := [Abs d])) u s" "term_variants_pred ((\_. [])(Abs d := [Abs e])) s t" using term_variants_pred_dense'[OF step.prems, of "Abs d"] by blast have "(u,s) \ (timpl_closure'_step c)\<^sup>*" "(s,t) \ timpl_closure'_step c" using step.hyps(2) s(2) step.IH[OF s(1)] by (auto intro: timpl_closure'_step.intros) thus ?case by simp qed thus ?case using step.IH by simp qed simp qed lemma timpl_closure'_timpls_rtrancl_supset: "timpl_closure' c \ timpl_closure' (c\<^sup>*)" unfolding timpl_closure'_def proof fix s t::"(('a,'b,'c,'d) prot_fun,'e) term" show "(s,t) \ (timpl_closure'_step c)\<^sup>* \ (s,t) \ (timpl_closure'_step (c\<^sup>*))\<^sup>*" proof (induction rule: rtrancl_induct) case (step u t) obtain a b where ab: "(a,b) \ c" "term_variants_pred ((\_. [])(Abs a := [Abs b])) u t" using step.hyps(2) timpl_closure'_step_inv by blast hence "(a,b) \ c\<^sup>*" by simp hence "(u,t) \ (timpl_closure'_step (c\<^sup>*))\<^sup>*" using ab(2) proof (induction arbitrary: t rule: rtrancl_induct) case (base t) thus ?case using term_variants_pred_refl_inv[of _ u t] by fastforce next case (step d e) obtain s where s: "term_variants_pred ((\_. [])(Abs a := [Abs d])) u s" "term_variants_pred ((\_. [])(Abs d := [Abs e])) s t" using term_variants_pred_dense'[OF step.prems, of "Abs d"] by blast show ?case proof (cases "d = e") case True thus ?thesis using step.prems step.IH[of t] by blast next case False hence "(u,s) \ (timpl_closure'_step (c\<^sup>*))\<^sup>*" "(s,t) \ timpl_closure'_step (c\<^sup>*)" using step.hyps(2) s(2) step.IH[OF s(1)] by (auto intro: timpl_closure'_step.intros) thus ?thesis by simp qed qed thus ?case using step.IH by simp qed simp qed lemma timpl_closure'_timpls_rtrancl_eq: "timpl_closure' (c\<^sup>*) = timpl_closure' c" using timpl_closure'_timpls_rtrancl_subset timpl_closure'_timpls_rtrancl_supset by blast lemma timpl_closure_timpls_trancl_eq: "timpl_closure t (c\<^sup>+) = timpl_closure t c" using timpl_closure'_timpls_trancl_eq[of c] timpl_closure_is_timpl_closure'[of _ _ c] timpl_closure_is_timpl_closure'[of _ _ "c\<^sup>+"] by fastforce lemma timpl_closure_set_timpls_trancl_eq: "timpl_closure_set M (c\<^sup>+) = timpl_closure_set M c" using timpl_closure_timpls_trancl_eq timpl_closure_set_is_timpl_closure_union[of M c] timpl_closure_set_is_timpl_closure_union[of M "c\<^sup>+"] by fastforce lemma timpl_closure_set_timpls_trancl_eq': "timpl_closure_set M {(a,b) \ c\<^sup>+. a \ b} = timpl_closure_set M c" using timpl_closure'_timpls_trancl_eq'[of c] timpl_closure_is_timpl_closure'[of _ _ c] timpl_closure_is_timpl_closure'[of _ _ "{(a,b) \ c\<^sup>+. a \ b}"] timpl_closure_set_is_timpl_closure_union[of M c] timpl_closure_set_is_timpl_closure_union[of M "{(a,b) \ c\<^sup>+. a \ b}"] by fastforce lemma timpl_closure_Var_in_iff: "Var x \ timpl_closure t TI \ t = Var x" (is "?A \ ?B") proof have "s \ timpl_closure t TI \ s = Var x \ s = t" for s apply (induction rule: timpl_closure.induct) by (simp, metis term_variants_pred_inv_Var(2)) thus "?A \ ?B" by blast qed (blast intro: timpl_closure.FP) lemma timpl_closure_set_Var_in_iff: "Var x \ timpl_closure_set M TI \ Var x \ M" unfolding timpl_closure_set_def by (simp add: timpl_closure_Var_in_iff[of x _ TI]) lemma timpl_closure_Var_inv: assumes "t \ timpl_closure (Var x) TI" shows "t = Var x" using assms proof (induction rule: timpl_closure.induct) case (TI u a b s) thus ?case using term_variants_pred_inv_Var by fast qed simp lemma timpls_Un_mono: "mono (\X. FP \ (\x \ X. \(a,b) \ TI. set \a --\ b\\x\))" by (auto intro!: monoI) lemma timpl_closure_set_lfp: fixes M TI defines "f \ \X. M \ (\x \ X. \(a,b) \ TI. set \a --\ b\\x\)" shows "lfp f = timpl_closure_set M TI" proof note 0 = timpls_Un_mono[of M TI, unfolded f_def[symmetric]] let ?N = "timpl_closure_set M TI" show "lfp f \ ?N" proof (induction rule: lfp_induct) case 2 thus ?case proof fix t assume "t \ f (lfp f \ ?N)" hence "t \ M \ t \ (\x \ ?N. \(a,b) \ TI. set \a --\ b\\x\)" (is "?A \ ?B") unfolding f_def by blast thus "t \ ?N" proof assume ?B - then obtain s a b where s: "s \ ?N" "(a,b) \ TI" "t \ set \a --\ b\\s\" by moura - thus ?thesis + then obtain s a b where s: "s \ ?N" "(a,b) \ TI" "t \ set \a --\ b\\s\" by blast + thus ?thesis using term_variants_pred_iff_in_term_variants[of "(\_. [])(Abs a := [Abs b])" s] unfolding timpl_closure_set_def timpl_apply_term_def by (auto intro: timpl_closure.intros) qed (auto simp add: timpl_closure_set_def intro: timpl_closure.intros) qed qed (rule 0) have "t \ lfp f" when t: "t \ timpl_closure s TI" and s: "s \ M" for t s using t proof (induction t rule: timpl_closure.induct) case (TI u a b v) thus ?case using term_variants_pred_iff_in_term_variants[of "(\_. [])(Abs a := [Abs b])"] lfp_fixpoint[OF 0] unfolding timpl_apply_term_def f_def by fastforce qed (use s lfp_fixpoint[OF 0] f_def in blast) thus "?N \ lfp f" unfolding timpl_closure_set_def by blast qed lemma timpl_closure_set_supset: assumes "\t \ FP. t \ closure" and "\t \ closure. \(a,b) \ TI. \s \ set \a --\ b\\t\. s \ closure" shows "timpl_closure_set FP TI \ closure" proof - have "t \ closure" when t: "t \ timpl_closure s TI" and s: "s \ FP" for t s using t proof (induction rule: timpl_closure.induct) case FP thus ?case using s assms(1) by blast next case (TI u a b s') thus ?case using assms(2) term_variants_pred_iff_in_term_variants[of "(\_. [])(Abs a := [Abs b])"] unfolding timpl_apply_term_def by fastforce qed thus ?thesis unfolding timpl_closure_set_def by blast qed lemma timpl_closure_set_supset': assumes "\t \ FP. \(a,b) \ TI. \s \ set \a --\ b\\t\. s \ FP" shows "timpl_closure_set FP TI \ FP" using timpl_closure_set_supset[OF _ assms] by blast lemma timpl_closure'_param: assumes "(t,s) \ timpl_closure' c" and fg: "f = g \ (\a b. (a,b) \ c \ f = Abs a \ g = Abs b)" shows "(Fun f (S@t#T), Fun g (S@s#T)) \ timpl_closure' c" using assms(1) unfolding timpl_closure'_def proof (induction rule: rtrancl_induct) case base thus ?case proof (cases "f = g") case False then obtain a b where ab: "(a,b) \ c" "f = Abs a" "g = Abs b" - using fg by moura + using fg by blast show ?thesis using term_variants_pred_param[OF term_variants_pred_refl[of "(\_. [])(Abs a := [Abs b])" t]] timpl_closure'_step.intros[OF ab(1)] ab(2,3) by fastforce qed simp next case (step u s) obtain a b where ab: "(a,b) \ c" "term_variants_pred ((\_. [])(Abs a := [Abs b])) u s" using timpl_closure'_step_inv[OF step.hyps(2)] by blast have "(Fun g (S@u#T), Fun g (S@s#T)) \ timpl_closure'_step c" using ab(1) term_variants_pred_param[OF ab(2), of g g S T] by (auto simp add: timpl_closure'_def intro: timpl_closure'_step.intros) thus ?case using rtrancl_into_rtrancl[OF step.IH] fg by blast qed lemma timpl_closure'_param': assumes "(t,s) \ timpl_closure' c" shows "(Fun f (S@t#T), Fun f (S@s#T)) \ timpl_closure' c" using timpl_closure'_param[OF assms] by simp lemma timpl_closure_FunI: assumes IH: "\i. i < length T \ (T ! i, S ! i) \ timpl_closure' c" and len: "length T = length S" and fg: "f = g \ (\a b. (a,b) \ c\<^sup>+ \ f = Abs a \ g = Abs b)" shows "(Fun f T, Fun g S) \ timpl_closure' c" proof - have aux: "(Fun f T, Fun g (take n S@drop n T)) \ timpl_closure' c" when "n \ length T" for n using that proof (induction n) case 0 have "(T ! n, T ! n) \ timpl_closure' c" when n: "n < length T" for n using n unfolding timpl_closure'_def by simp hence "(Fun f T, Fun g T) \ timpl_closure' c" proof (cases "f = g") case False then obtain a b where ab: "(a, b) \ c\<^sup>+" "f = Abs a" "g = Abs b" - using fg by moura + using fg by blast show ?thesis using timpl_closure'_step.intros[OF ab(1), of "Fun f T" "Fun g T"] ab(2,3) term_variants_P[OF _ term_variants_pred_refl[of "(\_. [])(Abs a := [Abs b])"], of T g f] timpl_closure'_timpls_trancl_eq unfolding timpl_closure'_def by (metis fun_upd_same list.set_intros(1) r_into_rtrancl) qed (simp add: timpl_closure'_def) thus ?case by simp next case (Suc n) hence IH': "(Fun f T, Fun g (take n S@drop n T)) \ timpl_closure' c" and n: "n < length T" "n < length S" by (simp_all add: len) obtain T1 T2 where T: "T = T1@T ! n#T2" "length T1 = n" using length_prefix_ex'[OF n(1)] by auto obtain S1 S2 where S: "S = S1@S ! n#S2" "length S1 = n" using length_prefix_ex'[OF n(2)] by auto have "take n S@drop n T = S1@T ! n#T2" "take (Suc n) S@drop (Suc n) T = S1@S ! n#T2" using n T S append_eq_conv_conj by (metis, metis (no_types, opaque_lifting) Cons_nth_drop_Suc append.assoc append_Cons append_Nil take_Suc_conv_app_nth) moreover have "(T ! n, S ! n) \ timpl_closure' c" using IH Suc.prems by simp ultimately show ?case using timpl_closure'_param IH'(1) by (metis (no_types, lifting) timpl_closure'_def rtrancl_trans) qed show ?thesis using aux[of "length T"] len by simp qed lemma timpl_closure_FunI': assumes IH: "\i. i < length T \ (T ! i, S ! i) \ timpl_closure' c" and len: "length T = length S" shows "(Fun f T, Fun f S) \ timpl_closure' c" using timpl_closure_FunI[OF IH len] by simp lemma timpl_closure_FunI2: fixes f g::"('a, 'b, 'c, 'd) prot_fun" assumes IH: "\i. i < length T \ \u. (T!i, u) \ timpl_closure' c \ (S!i, u) \ timpl_closure' c" and len: "length T = length S" and fg: "f = g \ (\a b d. (a, d) \ c\<^sup>+ \ (b, d) \ c\<^sup>+ \ f = Abs a \ g = Abs b)" shows "\h U. (Fun f T, Fun h U) \ timpl_closure' c \ (Fun g S, Fun h U) \ timpl_closure' c" proof - let ?P = "\i u. (T ! i, u) \ timpl_closure' c \ (S ! i, u) \ timpl_closure' c" define U where "U \ map (\i. SOME u. ?P i u) [0.. timpl_closure' c \ (S ! i, U ! i) \ timpl_closure' c" when i: "i < length U" for i using i someI_ex[of "?P i"] IH[of i] U1 len unfolding U_def by simp show ?thesis proof (cases "f = g") case False then obtain a b d where abd: "(a, d) \ c\<^sup>+" "(b, d) \ c\<^sup>+" "f = Abs a" "g = Abs b" - using fg by moura + using fg by blast define h::"('a, 'b, 'c, 'd) prot_fun" where "h = Abs d" have "f = h \ (\a b. (a, b) \ c\<^sup>+ \ f = Abs a \ h = Abs b)" "g = h \ (\a b. (a, b) \ c\<^sup>+ \ g = Abs a \ h = Abs b)" using abd unfolding h_def by blast+ thus ?thesis by (metis timpl_closure_FunI len U1 U2) qed (metis timpl_closure_FunI' len U1 U2) qed lemma timpl_closure_FunI3: fixes f g::"('a, 'b, 'c, 'd) prot_fun" assumes IH: "\i. i < length T \ \u. (T!i, u) \ timpl_closure' c \ (S!i, u) \ timpl_closure' c" and len: "length T = length S" and fg: "f = g \ (\a b d. (a, d) \ c \ (b, d) \ c \ f = Abs a \ g = Abs b)" shows "\h U. (Fun f T, Fun h U) \ timpl_closure' c \ (Fun g S, Fun h U) \ timpl_closure' c" using timpl_closure_FunI2[OF IH len] fg unfolding timpl_closure'_timpls_trancl_eq by blast lemma timpl_closure_fv_eq: assumes "s \ timpl_closure t T" shows "fv s = fv t" using assms by (induct rule: timpl_closure.induct) (metis, metis term_variants_pred_fv_eq) lemma (in stateful_protocol_model) timpl_closure_subst: assumes t: "wf\<^sub>t\<^sub>r\<^sub>m t" "\x \ fv t. \a. \\<^sub>v x = TAtom (Atom a)" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "timpl_closure (t \ \) T = timpl_closure t T \\<^sub>s\<^sub>e\<^sub>t \" proof have "s \ timpl_closure t T \\<^sub>s\<^sub>e\<^sub>t \" when "s \ timpl_closure (t \ \) T" for s using that proof (induction s rule: timpl_closure.induct) case FP thus ?case using timpl_closure.FP[of t T] by simp next case (TI u a b s) - then obtain u' where u': "u' \ timpl_closure t T" "u = u' \ \" by moura + then obtain u' where u': "u' \ timpl_closure t T" "u = u' \ \" by blast have u'_fv: "\x \ fv u'. \a. \\<^sub>v x = TAtom (Atom a)" using timpl_closure_fv_eq[OF u'(1)] t(2) by simp hence u_fv: "\x \ fv u. \a. \\<^sub>v x = TAtom (Atom a)" using u'(2) wt_subst_trm''[OF \(1)] wt_subst_const_fv_type_eq[OF _ \(1,2), of u'] by fastforce have "\x \ fv u' \ fv s. (\y. \ x = Var y) \ (\f. \ x = Fun f [] \ Abs a \ f)" proof (intro ballI) fix x assume x: "x \ fv u' \ fv s" then obtain c where c: "\\<^sub>v x = TAtom (Atom c)" using u'_fv u_fv term_variants_pred_fv_eq[OF TI.hyps(3)] by blast show "(\y. \ x = Var y) \ (\f. \ x = Fun f [] \ Abs a \ f)" proof (cases "\ x") case (Fun f T) hence **: "\ (Fun f T) = TAtom (Atom c)" and "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" using c wt_subst_trm''[OF \(1), of "Var x"] \(2) by fastforce+ hence "\ x = Fun f []" using Fun const_type_inv_wf by metis thus ?thesis using ** by force qed metis qed hence *: "\x \ fv u' \ fv s. (\y. \ x = Var y) \ (\f. \ x = Fun f [] \ ((\_. [])(Abs a := [Abs b])) f = [])" by fastforce obtain s' where s': "term_variants_pred ((\_. [])(Abs a := [Abs b])) u' s'" "s = s' \ \" using term_variants_pred_subst'[OF _ *] u'(2) TI.hyps(3) by blast show ?case using timpl_closure.TI[OF u'(1) TI.hyps(2) s'(1)] s'(2) by blast qed thus "timpl_closure (t \ \) T \ timpl_closure t T \\<^sub>s\<^sub>e\<^sub>t \" by fast have "s \ timpl_closure (t \ \) T" when s: "s \ timpl_closure t T \\<^sub>s\<^sub>e\<^sub>t \" for s proof - - obtain s' where s': "s' \ timpl_closure t T" "s = s' \ \" using s by moura + obtain s' where s': "s' \ timpl_closure t T" "s = s' \ \" using s by blast have "s' \ \ \ timpl_closure (t \ \) T" using s'(1) proof (induction s' rule: timpl_closure.induct) case FP thus ?case using timpl_closure.FP[of "t \ \" T] by simp next case (TI u' a b s') show ?case using timpl_closure.TI[OF TI.IH TI.hyps(2)] term_variants_pred_subst[OF TI.hyps(3)] by blast qed thus ?thesis using s'(2) by metis qed thus "timpl_closure t T \\<^sub>s\<^sub>e\<^sub>t \ \ timpl_closure (t \ \) T" by fast qed lemma (in stateful_protocol_model) timpl_closure_subst_subset: assumes t: "t \ M" and M: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" "\x \ fv\<^sub>s\<^sub>e\<^sub>t M. \a. \\<^sub>v x = TAtom (Atom a)" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "ground (subst_range \)" "subst_domain \ \ fv\<^sub>s\<^sub>e\<^sub>t M" and M_supset: "timpl_closure t T \ M" shows "timpl_closure (t \ \) T \ M \\<^sub>s\<^sub>e\<^sub>t \" proof - have t': "wf\<^sub>t\<^sub>r\<^sub>m t" "\x \ fv t. \a. \\<^sub>v x = TAtom (Atom a)" using t M by auto show ?thesis using timpl_closure_subst[OF t' \(1,2), of T] M_supset by blast qed lemma (in stateful_protocol_model) timpl_closure_set_subst_subset: assumes M: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" "\x \ fv\<^sub>s\<^sub>e\<^sub>t M. \a. \\<^sub>v x = TAtom (Atom a)" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "ground (subst_range \)" "subst_domain \ \ fv\<^sub>s\<^sub>e\<^sub>t M" and M_supset: "timpl_closure_set M T \ M" shows "timpl_closure_set (M \\<^sub>s\<^sub>e\<^sub>t \) T \ M \\<^sub>s\<^sub>e\<^sub>t \" using timpl_closure_subst_subset[OF _ M \, of _ T] M_supset timpl_closure_set_is_timpl_closure_union[of "M \\<^sub>s\<^sub>e\<^sub>t \" T] timpl_closure_set_is_timpl_closure_union[of M T] by auto lemma timpl_closure_set_Union: "timpl_closure_set (\Ms) T = (\M \ Ms. timpl_closure_set M T)" using timpl_closure_set_is_timpl_closure_union[of "\Ms" T] timpl_closure_set_is_timpl_closure_union[of _ T] by force lemma timpl_closure_set_Union_subst_set: assumes "s \ timpl_closure_set (\{M \\<^sub>s\<^sub>e\<^sub>t \ | \. P \}) T" shows "\\. P \ \ s \ timpl_closure_set (M \\<^sub>s\<^sub>e\<^sub>t \) T" using assms timpl_closure_set_is_timpl_closure_union[of "(\{M \\<^sub>s\<^sub>e\<^sub>t \ | \. P \})" T] timpl_closure_set_is_timpl_closure_union[of _ T] by blast lemma timpl_closure_set_Union_subst_singleton: assumes "s \ timpl_closure_set {t \ \ | \. P \} T" shows "\\. P \ \ s \ timpl_closure_set {t \ \} T" using assms timpl_closure_set_is_timpl_closure_union[of "{t \ \ |\. P \}" T] timpl_closureton_is_timpl_closure[of _ T] by fast lemma timpl_closure'_inv: assumes "(s, t) \ timpl_closure' TI" shows "(\x. s = Var x \ t = Var x) \ (\f g S T. s = Fun f S \ t = Fun g T \ length S = length T)" using assms unfolding timpl_closure'_def proof (induction rule: rtrancl_induct) case base thus ?case by (cases s) auto next case (step t u) obtain a b where ab: "(a, b) \ TI" "term_variants_pred ((\_. [])(Abs a := [Abs b])) t u" using timpl_closure'_step_inv[OF step.hyps(2)] by blast show ?case using step.IH proof assume "\x. s = Var x \ t = Var x" thus ?case using step.hyps(2) term_variants_pred_inv_Var ab by fastforce next assume "\f g S T. s = Fun f S \ t = Fun g T \ length S = length T" - then obtain f g S T where st: "s = Fun f S" "t = Fun g T" "length S = length T" by moura + then obtain f g S T where st: "s = Fun f S" "t = Fun g T" "length S = length T" by blast thus ?case using ab step.hyps(2) term_variants_pred_inv'[of "(\_. [])(Abs a := [Abs b])" g T u] by auto qed qed lemma timpl_closure'_inv': assumes "(s, t) \ timpl_closure' TI" shows "(\x. s = Var x \ t = Var x) \ (\f g S T. s = Fun f S \ t = Fun g T \ length S = length T \ (\i < length T. (S ! i, T ! i) \ timpl_closure' TI) \ (f \ g \ is_Abs f \ is_Abs g \ (the_Abs f, the_Abs g) \ TI\<^sup>+))" (is "?A s t \ ?B s t (timpl_closure' TI)") using assms unfolding timpl_closure'_def proof (induction rule: rtrancl_induct) case base thus ?case by (cases s) auto next case (step t u) obtain a b where ab: "(a, b) \ TI" "term_variants_pred ((\_. [])(Abs a := [Abs b])) t u" using timpl_closure'_step_inv[OF step.hyps(2)] by blast show ?case using step.IH proof assume "?A s t" thus ?case using step.hyps(2) term_variants_pred_inv_Var ab by fastforce next assume "?B s t ((timpl_closure'_step TI)\<^sup>*)" then obtain f g S T where st: "s = Fun f S" "t = Fun g T" "length S = length T" "\i. i < length T \ (S ! i, T ! i) \ (timpl_closure'_step TI)\<^sup>*" "f \ g \ is_Abs f \ is_Abs g \ (the_Abs f, the_Abs g) \ TI\<^sup>+" - by moura + by blast obtain h U where u: "u = Fun h U" "length T = length U" "\i. i < length T \ term_variants_pred ((\_. [])(Abs a := [Abs b])) (T ! i) (U ! i)" "g \ h \ is_Abs g \ is_Abs h \ (the_Abs g, the_Abs h) \ TI\<^sup>+" using ab(2) st(2) r_into_trancl[OF ab(1)] term_variants_pred_inv'(1,2,3,4)[of "(\_. [])(Abs a := [Abs b])" g T u] term_variants_pred_inv'(5)[of "(\_. [])(Abs a := [Abs b])" g T u "Abs a" "Abs b"] unfolding is_Abs_def the_Abs_def by force have "(S ! i, U ! i) \ (timpl_closure'_step TI)\<^sup>*" when i: "i < length U" for i using u(2) i rtrancl.rtrancl_into_rtrancl[OF st(4)[of i] timpl_closure'_step.intros[OF ab(1) u(3)[of i]]] by argo moreover have "length S = length U" using st u by argo moreover have "is_Abs f \ is_Abs h \ (the_Abs f, the_Abs h) \ TI\<^sup>+" when fh: "f \ h" using fh st u by fastforce ultimately show ?case using st(1) u(1) by blast qed qed lemma timpl_closure'_inv'': assumes "(Fun f S, Fun g T) \ timpl_closure' TI" shows "length S = length T" and "\i. i < length T \ (S ! i, T ! i) \ timpl_closure' TI" and "f \ g \ is_Abs f \ is_Abs g \ (the_Abs f, the_Abs g) \ TI\<^sup>+" using assms timpl_closure'_inv' by auto lemma timpl_closure_Fun_inv: assumes "s \ timpl_closure (Fun f T) TI" shows "\g S. s = Fun g S" using assms timpl_closure_is_timpl_closure' timpl_closure'_inv by fastforce lemma timpl_closure_Fun_inv': assumes "Fun g S \ timpl_closure (Fun f T) TI" shows "length S = length T" and "\i. i < length S \ S ! i \ timpl_closure (T ! i) TI" and "f \ g \ is_Abs f \ is_Abs g \ (the_Abs f, the_Abs g) \ TI\<^sup>+" using assms timpl_closure_is_timpl_closure' by (metis timpl_closure'_inv''(1), metis timpl_closure'_inv''(2), metis timpl_closure'_inv''(3)) lemma timpl_closure_Fun_not_Var[simp]: "Fun f T \ timpl_closure (Var x) TI" using timpl_closure_Var_inv by fast lemma timpl_closure_Var_not_Fun[simp]: "Var x \ timpl_closure (Fun f T) TI" using timpl_closure_Fun_inv by fast lemma (in stateful_protocol_model) timpl_closure_wf_trms: assumes m: "wf\<^sub>t\<^sub>r\<^sub>m m" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (timpl_closure m TI)" proof fix t assume "t \ timpl_closure m TI" thus "wf\<^sub>t\<^sub>r\<^sub>m t" proof (induction t rule: timpl_closure.induct) case TI thus ?case using term_variants_pred_wf_trms by force qed (rule m) qed lemma (in stateful_protocol_model) timpl_closure_set_wf_trms: assumes M: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (timpl_closure_set M TI)" proof fix t assume "t \ timpl_closure_set M TI" then obtain m where "t \ timpl_closure m TI" "m \ M" "wf\<^sub>t\<^sub>r\<^sub>m m" using M timpl_closure_set_is_timpl_closure_union by blast thus "wf\<^sub>t\<^sub>r\<^sub>m t" using timpl_closure_wf_trms by blast qed lemma timpl_closure_Fu_inv: assumes "t \ timpl_closure (Fun (Fu f) T) TI" shows "\S. length S = length T \ t = Fun (Fu f) S" using assms proof (induction t rule: timpl_closure.induct) case (TI u a b s) then obtain U where U: "length U = length T" "u = Fun (Fu f) U" - by moura + by blast hence *: "term_variants_pred ((\_. [])(Abs a := [Abs b])) (Fun (Fu f) U) s" using TI.hyps(3) by meson show ?case using term_variants_pred_inv'(1,2,4)[OF *] U by force qed simp lemma timpl_closure_Fu_inv': assumes "Fun (Fu f) T \ timpl_closure t TI" shows "\S. length S = length T \ t = Fun (Fu f) S" using assms proof (induction "Fun (Fu f) T" arbitrary: T rule: timpl_closure.induct) case (TI u a b) obtain g U where U: "u = Fun g U" "length U = length T" "Fu f \ g \ Abs a = g \ Fu f = Abs b" using term_variants_pred_inv''[OF TI.hyps(4)] by fastforce have g: "g = Fu f" using U(3) by blast show ?case using TI.hyps(2)[OF U(1)[unfolded g]] U(2) by auto qed simp lemma timpl_closure_no_Abs_eq: assumes "t \ timpl_closure s TI" and "\f \ funs_term t. \is_Abs f" shows "t = s" using assms proof (induction t rule: timpl_closure.induct) case (TI t a b s) thus ?case using term_variants_pred_eq_case_Abs[of a b t s] unfolding timpl_apply_term_def term_variants_pred_iff_in_term_variants[symmetric] by metis qed simp lemma timpl_closure_set_no_Abs_in_set: assumes "t \ timpl_closure_set FP TI" and "\f \ funs_term t. \is_Abs f" shows "t \ FP" using assms timpl_closure_no_Abs_eq unfolding timpl_closure_set_def by blast lemma timpl_closure_funs_term_subset: "\(funs_term ` (timpl_closure t TI)) \ funs_term t \ Abs ` snd ` TI" (is "?A \ ?B \ ?C") proof fix f assume "f \ ?A" - then obtain s where "s \ timpl_closure t TI" "f \ funs_term s" by moura + then obtain s where "s \ timpl_closure t TI" "f \ funs_term s" by blast thus "f \ ?B \ ?C" proof (induction s rule: timpl_closure.induct) case (TI u a b s) have "Abs b \ Abs ` snd ` TI" using TI.hyps(2) by force thus ?case using term_variants_pred_funs_term[OF TI.hyps(3) TI.prems] TI.IH by force qed blast qed lemma timpl_closure_set_funs_term_subset: "\(funs_term ` (timpl_closure_set FP TI)) \ \(funs_term ` FP) \ Abs ` snd ` TI" using timpl_closure_funs_term_subset[of _ TI] timpl_closure_set_is_timpl_closure_union[of FP TI] by auto lemma funs_term_OCC_TI_subset: defines "absc \ \a. Fun (Abs a) []" assumes OCC1: "\t \ FP. \f \ funs_term t. is_Abs f \ f \ Abs ` OCC" and OCC2: "snd ` TI \ OCC" shows "\t \ timpl_closure_set FP TI. \f \ funs_term t. is_Abs f \ f \ Abs ` OCC" (is ?A) and "\t \ absc ` OCC. \(a,b) \ TI. \s \ set \a --\ b\\t\. s \ absc ` OCC" (is ?B) proof - let ?F = "\(funs_term ` FP)" let ?G = "Abs ` snd ` TI" show ?A proof (intro ballI impI) fix t f assume t: "t \ timpl_closure_set FP TI" and f: "f \ funs_term t" "is_Abs f" hence "f \ ?F \ f \ ?G" using timpl_closure_set_funs_term_subset[of FP TI] by auto thus "f \ Abs ` OCC" proof assume "f \ ?F" thus ?thesis using OCC1 f(2) by fast next assume "f \ ?G" thus ?thesis using OCC2 by auto qed qed { fix s t a b assume t: "t \ absc ` OCC" and ab: "(a, b) \ TI" and s: "s \ set \a --\ b\\t\" - obtain c where c: "t = absc c" "c \ OCC" using t by moura + obtain c where c: "t = absc c" "c \ OCC" using t by blast hence "s = absc b \ s = absc c" using ab s timpl_apply_const'[of c a b] unfolding absc_def by auto moreover have "b \ OCC" using ab OCC2 by auto ultimately have "s \ absc ` OCC" using c(2) by blast } thus ?B by blast qed lemma (in stateful_protocol_model) intruder_synth_timpl_closure_set: fixes M::"('fun,'atom,'sets,'lbl) prot_terms" and t::"('fun,'atom,'sets,'lbl) prot_term" assumes "M \\<^sub>c t" and "s \ timpl_closure t TI" shows "timpl_closure_set M TI \\<^sub>c s" using assms proof (induction t arbitrary: s rule: intruder_synth_induct) case (AxiomC t) hence "s \ timpl_closure_set M TI" using timpl_closure_set_is_timpl_closure_union[of M TI] by blast thus ?case by simp next case (ComposeC T f) obtain g S where s: "s = Fun g S" - using timpl_closure_Fun_inv[OF ComposeC.prems] by moura + using timpl_closure_Fun_inv[OF ComposeC.prems] by blast hence s': "f = g" "length S = length T" "\i. i < length S \ S ! i \ timpl_closure (T ! i) TI" using timpl_closure_Fun_inv'[of g S f T TI] ComposeC.prems ComposeC.hyps(2) unfolding is_Abs_def by fastforce+ have "timpl_closure_set M TI \\<^sub>c u" when u: "u \ set S" for u using ComposeC.IH u s'(2,3) in_set_conv_nth[of _ T] in_set_conv_nth[of u S] by auto thus ?case using s s'(1,2) ComposeC.hyps(1,2) intruder_synth.ComposeC[of S g "timpl_closure_set M TI"] by argo qed lemma (in stateful_protocol_model) intruder_synth_timpl_closure': fixes M::"('fun,'atom,'sets,'lbl) prot_terms" and t::"('fun,'atom,'sets,'lbl) prot_term" assumes "timpl_closure_set M TI \\<^sub>c t" and "s \ timpl_closure t TI" shows "timpl_closure_set M TI \\<^sub>c s" by (metis intruder_synth_timpl_closure_set[OF assms] timpl_closure_set_idem) lemma timpl_closure_set_absc_subset_in: defines "absc \ \a. Fun (Abs a) []" assumes A: "timpl_closure_set (absc ` A) TI \ absc ` A" and a: "a \ A" "(a,b) \ TI\<^sup>+" shows "b \ A" proof - have "timpl_closure (absc a) (TI\<^sup>+) \ absc ` A" using a(1) A timpl_closure_timpls_trancl_eq unfolding timpl_closure_set_def by fast thus ?thesis using timpl_closure.TI[OF timpl_closure.FP[of "absc a"] a(2), of "absc b"] term_variants_P[of "[]" "[]" "(\_. [])(Abs a := [Abs b])" "Abs b" "Abs a"] unfolding absc_def by auto qed lemma timpl_closure_Abs_ex: assumes t: "s \ timpl_closure t TI" and a: "Abs a \ funs_term t" shows "\b ts. (a,b) \ TI\<^sup>* \ Fun (Abs b) ts \ s" using t proof (induction rule: timpl_closure.induct) case (TI u b c s) obtain d ts where d: "(a,d) \ TI\<^sup>*" "Fun (Abs d) ts \ u" using TI.IH by blast note 0 = TI.hyps(2) d(1) term_variants_pred_inv'(5)[OF term_variants_pred_const] show ?case using TI.hyps(3) d(2) proof (induction rule: term_variants_pred.induct) case (term_variants_P T S g f) note hyps = term_variants_P.hyps note prems = term_variants_P.prems note IH = term_variants_P.IH show ?case proof (cases "Fun (Abs d) ts = Fun f T") case False hence "\t \ set T. Fun (Abs d) ts \ t" using prems(1) by force then obtain i where i: "i < length T" "Fun (Abs d) ts \ T ! i" by (metis in_set_conv_nth) show ?thesis by (metis IH[OF i] i(1) hyps(1) nth_mem subtermeqI'' term.order.trans) qed (metis hyps(3) 0 prot_fun.sel(4) r_into_rtrancl rtrancl_trans term.eq_refl term.sel(2)) next case (term_variants_Fun T S f) note hyps = term_variants_Fun.hyps note prems = term_variants_Fun.prems note IH = term_variants_Fun.IH show ?case proof (cases "Fun (Abs d) ts = Fun f T") case False hence "\t \ set T. Fun (Abs d) ts \ t" using prems(1) by force then obtain i where i: "i < length T" "Fun (Abs d) ts \ T ! i" by (metis in_set_conv_nth) show ?thesis by (metis IH[OF i] i(1) hyps(1) nth_mem subtermeqI'' term.order.trans) qed (metis 0(2) term.eq_refl term.sel(2)) qed simp qed (meson a funs_term_Fun_subterm rtrancl_eq_or_trancl) lemma timpl_closure_trans: assumes "s \ timpl_closure t TI" and "u \ timpl_closure s TI" shows "u \ timpl_closure t TI" using assms unfolding timpl_closure_is_timpl_closure' timpl_closure'_def by simp lemma (in stateful_protocol_model) term_variants_pred_Ana\<^sub>f_keys: assumes "length ss = length ts" "\x \ fv k. x < length ss" "\i. i < length ss \ term_variants_pred P (ss ! i) (ts ! i)" shows "term_variants_pred P (k \ (!) ss) (k \ (!) ts)" using assms by (meson term_variants_pred_subst'') lemma (in stateful_protocol_model) term_variants_pred_Ana_keys: fixes a b and s t::"('fun,'atom,'sets,'lbl) prot_term" defines "P \ ((\_. [])(Abs a := [Abs b]))" assumes ab: "term_variants_pred P s t" and s: "Ana s = (Ks, Rs)" and t: "Ana t = (Kt, Rt)" shows "length Kt = length Ks" (is ?A) and "\i < length Ks. term_variants_pred P (Ks ! i) (Kt ! i)" (is ?B) proof - have ?A ?B when "s = Var x" for x using that ab s t iffD1[OF term_variants_pred_inv_Var(1) ab[unfolded that]] by auto moreover have ?A ?B when s': "s = Fun f ss" for f ss proof - obtain g ts where t': "t = Fun g ts" "length ss = length ts" "\i. i < length ss \ term_variants_pred P (ss ! i) (ts ! i)" "f \ g \ f = Abs a \ g = Abs b" using term_variants_pred_inv'[OF ab[unfolded s']] unfolding P_def by fastforce have ?A ?B when "f \ g" using that s s' t t'(1,2,4) by auto moreover have A: ?A when fg: "f = g" using s t Ana_Fu_keys_length_eq[OF t'(2)] Ana_nonempty_inv[of s] Ana_nonempty_inv[of t] unfolding fg s' t'(1) by (metis (no_types) fst_conv term.inject(2)) moreover have ?B when fg: "f = g" proof (cases "Ana s = ([],[])") case True thus ?thesis using s t t'(2,3) A[OF fg] unfolding fg s' t'(1) by auto next case False then obtain h Kh Rh where h: "f = Fu h" "g = Fu h" "arity\<^sub>f h = length ss" "arity\<^sub>f h > 0" "Ana\<^sub>f h = (Kh, Rh)" "Ana s = (Kh \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) ss, map ((!) ss) Rh)" "Ana t = (Kh \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) ts, map ((!) ts) Rh)" using A[OF fg] s t t'(2,3) Ana_nonempty_inv[of s] Ana_nonempty_inv[of t] unfolding fg s' t'(1) by fastforce show ?thesis proof (intro allI impI) fix i assume i: "i < length Ks" have Ks: "Ks = Kh \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) ss" and Kt: "Kt = Kh \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) ts" using h(6,7) s t by auto have 0: "Kh ! i \ set Kh" using Ks i by simp have 1: "\x \ fv (Kh ! i). x < length ss" using 0 Ana\<^sub>f_assm2_alt[OF h(5)] unfolding h(1-3) by fastforce have "term_variants_pred P (Kh ! i \ (!) ss) (Kh ! i \ (!) ts)" using term_variants_pred_Ana\<^sub>f_keys[OF t'(2) 1 t'(3)] unfolding P_def by fast thus "term_variants_pred P (Ks ! i) (Kt ! i)" using i unfolding Ks Kt by simp qed qed ultimately show ?A ?B by fast+ qed ultimately show ?A ?B by (cases s; simp_all)+ qed lemma (in stateful_protocol_model) timpl_closure_Ana_keys: fixes s t::"('fun,'atom,'sets,'lbl) prot_term" assumes "t \ timpl_closure s TI" and "Ana s = (Ks, Rs)" and "Ana t = (Kt, Rt)" shows "length Kt = length Ks" (is ?A) and "\i < length Ks. Kt ! i \ timpl_closure (Ks ! i) TI" (is ?B) using assms proof (induction arbitrary: Ks Rs Kt Rt rule: timpl_closure.induct) case FP { case 1 thus ?case by simp } { case 2 thus ?case using FP timpl_closure.FP by force } next case (TI u a b t) obtain Ku Ru where u: "Ana u = (Ku, Ru)" by (metis surj_pair) note 0 = term_variants_pred_Ana_keys[OF TI.hyps(3) u] { case 1 thus ?case using 0(1) TI.IH(1) u by fastforce } { case 2 thus ?case by (metis 0(2)[OF 2(2)] TI.IH[OF 2(1) u] timpl_closure.TI[OF _ TI.hyps(2)]) } qed lemma (in stateful_protocol_model) timpl_closure_Ana_keys_length_eq: fixes s t::"('fun,'atom,'sets,'lbl) prot_term" assumes "t \ timpl_closure s TI" and "Ana s = (Ks, Rs)" and "Ana t = (Kt, Rt)" shows "length Kt = length Ks" by (rule timpl_closure_Ana_keys(1,2)[OF assms]) lemma (in stateful_protocol_model) timpl_closure_Ana_keys_subset: fixes s t::"('fun,'atom,'sets,'lbl) prot_term" assumes "t \ timpl_closure s TI" and "Ana s = (Ks, Rs)" and "Ana t = (Kt, Rt)" shows "set Kt \ timpl_closure_set (set Ks) TI" proof - note 0 = timpl_closure_Ana_keys[OF assms] have "\i < length Ks. Kt ! i \ timpl_closure_set (set Ks) TI" using in_set_conv_nth 0(2) unfolding timpl_closure_set_def by auto thus "set Kt \ timpl_closure_set (set Ks) TI" using 0(1) by (metis subsetI in_set_conv_nth) qed subsection \Composition-only Intruder Deduction Modulo Term Implication Closure of the Intruder Knowledge\ context stateful_protocol_model begin fun in_trancl where "in_trancl TI a b = ( if (a,b) \ set TI then True else list_ex (\(c,d). c = a \ in_trancl (removeAll (c,d) TI) d b) TI)" definition in_rtrancl where "in_rtrancl TI a b \ a = b \ in_trancl TI a b" declare in_trancl.simps[simp del] fun timpls_transformable_to where "timpls_transformable_to TI (Var x) (Var y) = (x = y)" | "timpls_transformable_to TI (Fun f T) (Fun g S) = ( (f = g \ (is_Abs f \ is_Abs g \ (the_Abs f, the_Abs g) \ set TI)) \ list_all2 (timpls_transformable_to TI) T S)" | "timpls_transformable_to _ _ _ = False" fun timpls_transformable_to' where "timpls_transformable_to' TI (Var x) (Var y) = (x = y)" | "timpls_transformable_to' TI (Fun f T) (Fun g S) = ( (f = g \ (is_Abs f \ is_Abs g \ in_trancl TI (the_Abs f) (the_Abs g))) \ list_all2 (timpls_transformable_to' TI) T S)" | "timpls_transformable_to' _ _ _ = False" fun equal_mod_timpls where "equal_mod_timpls TI (Var x) (Var y) = (x = y)" | "equal_mod_timpls TI (Fun f T) (Fun g S) = ( (f = g \ (is_Abs f \ is_Abs g \ ((the_Abs f, the_Abs g) \ set TI \ (the_Abs g, the_Abs f) \ set TI \ (\ti \ set TI. (the_Abs f, snd ti) \ set TI \ (the_Abs g, snd ti) \ set TI)))) \ list_all2 (equal_mod_timpls TI) T S)" | "equal_mod_timpls _ _ _ = False" fun intruder_synth_mod_timpls where "intruder_synth_mod_timpls M TI (Var x) = List.member M (Var x)" | "intruder_synth_mod_timpls M TI (Fun f T) = ( (list_ex (\t. timpls_transformable_to TI t (Fun f T)) M) \ (public f \ length T = arity f \ list_all (intruder_synth_mod_timpls M TI) T))" fun intruder_synth_mod_timpls' where "intruder_synth_mod_timpls' M TI (Var x) = List.member M (Var x)" | "intruder_synth_mod_timpls' M TI (Fun f T) = ( (list_ex (\t. timpls_transformable_to' TI t (Fun f T)) M) \ (public f \ length T = arity f \ list_all (intruder_synth_mod_timpls' M TI) T))" fun intruder_synth_mod_eq_timpls where "intruder_synth_mod_eq_timpls M TI (Var x) = (Var x \ M)" | "intruder_synth_mod_eq_timpls M TI (Fun f T) = ( (\t \ M. equal_mod_timpls TI t (Fun f T)) \ (public f \ length T = arity f \ list_all (intruder_synth_mod_eq_timpls M TI) T))" definition analyzed_closed_mod_timpls where "analyzed_closed_mod_timpls M TI \ let ti = intruder_synth_mod_timpls M TI; cl = \ts. comp_timpl_closure ts (set TI); f = list_all ti; g = \t. if f (fst (Ana t)) then f (snd (Ana t)) else if list_all (\t. \f \ funs_term t. \is_Abs f) (fst (Ana t)) then True else if \s \ cl (set (fst (Ana t))). \ti s then True else \s \ cl {t}. case Ana s of (K,R) \ f K \ f R in list_all g M" definition analyzed_closed_mod_timpls' where "analyzed_closed_mod_timpls' M TI \ let f = list_all (intruder_synth_mod_timpls' M TI); g = \t. if f (fst (Ana t)) then f (snd (Ana t)) else if list_all (\t. \f \ funs_term t. \is_Abs f) (fst (Ana t)) then True else \s \ comp_timpl_closure {t} (set TI). case Ana s of (K,R) \ f K \ f R in list_all g M" lemma term_variants_pred_Abs_Ana_keys: fixes a b defines "P \ ((\_. [])(Abs a := [Abs b]))" assumes st: "term_variants_pred P s t" shows "length (fst (Ana s)) = length (fst (Ana t))" (is "?P s t") and "\i < length (fst (Ana s)). term_variants_pred P (fst (Ana s) ! i) (fst (Ana t) ! i)" (is "?Q s t") proof - show "?P s t" using st proof (induction s t rule: term_variants_pred.induct) case (term_variants_Fun T S f) show ?case proof (cases f) case (Fu g) thus ?thesis using term_variants_Fun Ana_Fu_keys_length_eq by blast qed simp_all qed (simp_all add: P_def) show "?Q s t" using st proof (induction s t rule: term_variants_pred.induct) case (term_variants_Fun T S f) note hyps = term_variants_Fun.hyps let ?K = "\U. fst (Ana (Fun f U))" show ?case proof (cases f) case (Fu g) show ?thesis proof (cases "arity\<^sub>f g = length T \ arity\<^sub>f g > 0") case True hence *: "?K T = fst (Ana\<^sub>f g) \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T" "?K S = fst (Ana\<^sub>f g) \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) S" using Fu Ana_Fu_intro fst_conv prod.collapse by (metis (mono_tags, lifting), metis (mono_tags, lifting) hyps(1)) have K: "j < length T" when j: "j \ fv\<^sub>s\<^sub>e\<^sub>t (set (fst (Ana\<^sub>f g)))" for j using True Ana\<^sub>f_assm2_alt[of g "fst (Ana\<^sub>f g)" _ j ] by (metis UnI1 prod.collapse that) show ?thesis proof (intro allI impI) fix i assume i: "i < length (?K T)" let ?u = "fst (Ana\<^sub>f g) ! i" have **: "?K T ! i = ?u \ (!) T" "?K S ! i = ?u \ (!) S" using * i by simp_all have ***: "x < length T" when "x \ fv (fst (Ana\<^sub>f g) ! i)" for x using that K Ana\<^sub>f_assm2_alt[of g "fst (Ana\<^sub>f g)" _ x] i hyps(1) unfolding * by force show "term_variants_pred P (?K T ! i) (?K S ! i)" using i hyps K *** term_variants_pred_subst''[of ?u P "(!) T" "(!) S"] unfolding * by auto qed qed (auto simp add: Fu) qed simp_all qed (simp_all add: P_def) qed lemma term_variants_pred_Abs_eq_case: assumes t: "term_variants_pred ((\_. [])(Abs a := [Abs b])) s t" (is "?R s t") and s: "\f \ funs_term s. \is_Abs f" (is "?P s") shows "s = t" using s term_variants_pred_eq_case[OF t] by fastforce lemma term_variants_Ana_keys_no_Abs_eq_case: fixes s t::"(('fun,'atom,'sets,'lbl) prot_fun,'v) term" assumes t: "term_variants_pred ((\_. [])(Abs a := [Abs b])) s t" (is "?R s t") and s: "\t \ set (fst (Ana s)). \f \ funs_term t. \is_Abs f" (is "?P s") shows "fst (Ana t) = fst (Ana s)" (is "?Q t s") using s term_variants_pred_Abs_Ana_keys[OF t] term_variants_pred_Abs_eq_case[of a b] by (metis nth_equalityI nth_mem) lemma timpl_closure_Ana_keys_no_Abs_eq_case: assumes t: "t \ timpl_closure s TI" and s: "\t \ set (fst (Ana s)). \f \ funs_term t. \is_Abs f" (is "?P s") shows "fst (Ana t) = fst (Ana s)" using t proof (induction t rule: timpl_closure.induct) case (TI u a b t) thus ?case using s term_variants_Ana_keys_no_Abs_eq_case by fastforce qed simp lemma in_trancl_closure_iff_in_trancl_fun: "(a,b) \ (set TI)\<^sup>+ \ in_trancl TI a b" (is "?A TI a b \ ?B TI a b") proof show "?A TI a b \ ?B TI a b" proof (induction rule: trancl_induct) case (step c d) show ?case using step.IH step.hyps(2) proof (induction TI a c rule: in_trancl.induct) case (1 TI a b) have "(a,b) \ set TI \ (\e. (a,e) \ set TI \ in_trancl (removeAll (a,e) TI) e b)" using "1.prems"(1) in_trancl.simps[of TI a b] unfolding list_ex_iff by (cases "(a,b) \ set TI") auto show ?case proof (cases "(a,b) \ set TI") case True thus ?thesis by (metis (mono_tags, lifting) in_trancl.simps "1.prems"(2) list_ex_iff case_prodI member_remove prod.inject remove_code(1)) next case F: False then obtain e where e: "(a,e) \ set TI" "in_trancl (removeAll (a,e) TI) e b" using "1.prems"(1) in_trancl.simps[of TI a b] unfolding list_ex_iff by (cases "(a,b) \ set TI") auto show ?thesis proof (cases "(b, d) \ set (removeAll (a, e) TI)") case True thus ?thesis using in_trancl.simps[of TI a d] e(1) "1.prems"(2) "1.IH"[OF F e(1) _ e(2)] unfolding list_ex_iff by auto next case False thus ?thesis using in_trancl.simps[of TI a d] "1.prems"(2) by simp qed qed qed qed (metis in_trancl.simps) show "?B TI a b \ ?A TI a b" proof (induction TI a b rule: in_trancl.induct) case (1 TI a b) let ?P = "\TI a b c d. in_trancl (List.removeAll (c,d) TI) d b" have *: "\(c,d) \ set TI. c = a \ ?P TI a b c d" when "(a,b) \ set TI" using that "1.prems" list_ex_iff[of _ TI] in_trancl.simps[of TI a b] by auto show ?case proof (cases "(a,b) \ set TI") case False hence "\(c,d) \ set TI. c = a \ ?P TI a b c d" using * by blast then obtain d where d: "(a,d) \ set TI" "?P TI a b a d" by blast have "(d,b) \ (set (removeAll (a,d) TI))\<^sup>+" using "1.IH"[OF False d(1)] d(2) by blast moreover have "set (removeAll (a,d) TI) \ set TI" by simp ultimately have "(d,b) \ (set TI)\<^sup>+" using trancl_mono by blast thus ?thesis using d(1) by fastforce qed simp qed qed lemma in_rtrancl_closure_iff_in_rtrancl_fun: "(a,b) \ (set TI)\<^sup>* \ in_rtrancl TI a b" by (metis rtrancl_eq_or_trancl in_trancl_closure_iff_in_trancl_fun in_rtrancl_def) lemma in_trancl_mono: assumes "set TI \ set TI'" and "in_trancl TI a b" shows "in_trancl TI' a b" by (metis assms in_trancl_closure_iff_in_trancl_fun trancl_mono) lemma equal_mod_timpls_refl: "equal_mod_timpls TI t t" proof (induction t) case (Fun f T) thus ?case using list_all2_conv_all_nth[of "equal_mod_timpls TI" T T] by force qed simp lemma equal_mod_timpls_inv_Var: "equal_mod_timpls TI (Var x) t \ t = Var x" (is "?A \ ?C") "equal_mod_timpls TI t (Var x) \ t = Var x" (is "?B \ ?C") proof - show "?A \ ?C" by (cases t) auto show "?B \ ?C" by (cases t) auto qed lemma equal_mod_timpls_inv: assumes "equal_mod_timpls TI (Fun f T) (Fun g S)" shows "length T = length S" and "\i. i < length T \ equal_mod_timpls TI (T ! i) (S ! i)" and "f \ g \ (is_Abs f \ is_Abs g \ ( (the_Abs f, the_Abs g) \ set TI \ (the_Abs g, the_Abs f) \ set TI \ (\ti \ set TI. (the_Abs f, snd ti) \ set TI \ (the_Abs g, snd ti) \ set TI)))" using assms list_all2_conv_all_nth[of "equal_mod_timpls TI" T S] by (auto elim: equal_mod_timpls.cases) lemma equal_mod_timpls_inv': assumes "equal_mod_timpls TI (Fun f T) t" shows "is_Fun t" and "length T = length (args t)" and "\i. i < length T \ equal_mod_timpls TI (T ! i) (args t ! i)" and "f \ the_Fun t \ (is_Abs f \ is_Abs (the_Fun t) \ ( (the_Abs f, the_Abs (the_Fun t)) \ set TI \ (the_Abs (the_Fun t), the_Abs f) \ set TI \ (\ti \ set TI. (the_Abs f, snd ti) \ set TI \ (the_Abs (the_Fun t), snd ti) \ set TI)))" and "\is_Abs f \ f = the_Fun t" using assms list_all2_conv_all_nth[of "equal_mod_timpls TI" T] by (cases t; auto)+ lemma equal_mod_timpls_if_term_variants: fixes s t::"(('a, 'b, 'c, 'd) prot_fun, 'e) term" and a b::"'c set" defines "P \ (\_. [])(Abs a := [Abs b])" assumes st: "term_variants_pred P s t" and ab: "(a,b) \ set TI" shows "equal_mod_timpls TI s t" using st P_def proof (induction rule: term_variants_pred.induct) case (term_variants_P T S f) thus ?case using ab list_all2_conv_all_nth[of "equal_mod_timpls TI" T S] in_trancl_closure_iff_in_trancl_fun[of _ _ TI] by auto next case (term_variants_Fun T S f) thus ?case using ab list_all2_conv_all_nth[of "equal_mod_timpls TI" T S] in_trancl_closure_iff_in_trancl_fun[of _ _ TI] by auto qed simp lemma equal_mod_timpls_mono: assumes "set TI \ set TI'" and "equal_mod_timpls TI s t" shows "equal_mod_timpls TI' s t" using assms proof (induction TI s t rule: equal_mod_timpls.induct) case (2 TI f T g S) have *: "f = g \ (is_Abs f \ is_Abs g \ ((the_Abs f, the_Abs g) \ set TI \ (the_Abs g, the_Abs f) \ set TI \ (\ti \ set TI. (the_Abs f, snd ti) \ set TI \ (the_Abs g, snd ti) \ set TI)))" "list_all2 (equal_mod_timpls TI) T S" using "2.prems" by simp_all show ?case using "2.IH"[OF _ _ "2.prems"(1)] list.rel_mono_strong[OF *(2), of "equal_mod_timpls TI'"] *(1) in_trancl_mono[OF "2.prems"(1)] set_rev_mp[OF _ "2.prems"(1)] equal_mod_timpls.simps(2)[of TI' f T g S] by metis qed auto lemma equal_mod_timpls_refl_minus_eq: "equal_mod_timpls TI s t \ equal_mod_timpls (filter (\(a,b). a \ b) TI) s t" (is "?A \ ?B") proof show ?A when ?B using that equal_mod_timpls_mono[of "filter (\(a,b). a \ b) TI" TI] by auto show ?B when ?A using that proof (induction TI s t rule: equal_mod_timpls.induct) case (2 TI f T g S) define TI' where "TI' \ filter (\(a,b). a \ b) TI" let ?P = "\X Y. f = g \ (is_Abs f \ is_Abs g \ ((the_Abs f, the_Abs g) \ set X \ (the_Abs g, the_Abs f) \ set X \ (\ti \ set Y. (the_Abs f, snd ti) \ set X \ (the_Abs g, snd ti) \ set X)))" have *: "?P TI TI" "list_all2 (equal_mod_timpls TI) T S" using "2.prems" by simp_all have "?P TI' TI" using *(1) unfolding TI'_def is_Abs_def by auto hence "?P TI' TI'" by (metis (no_types, lifting) snd_conv) moreover have "list_all2 (equal_mod_timpls TI') T S" using *(2) "2.IH" list.rel_mono_strong unfolding TI'_def by blast ultimately show ?case unfolding TI'_def by force qed auto qed lemma timpls_transformable_to_refl: "timpls_transformable_to TI t t" (is ?A) "timpls_transformable_to' TI t t" (is ?B) by (induct t) (auto simp add: list_all2_conv_all_nth) lemma timpls_transformable_to_inv_Var: "timpls_transformable_to TI (Var x) t \ t = Var x" (is "?A \ ?C") "timpls_transformable_to TI t (Var x) \ t = Var x" (is "?B \ ?C") "timpls_transformable_to' TI (Var x) t \ t = Var x" (is "?A' \ ?C") "timpls_transformable_to' TI t (Var x) \ t = Var x" (is "?B' \ ?C") by (cases t; auto)+ lemma timpls_transformable_to_inv: assumes "timpls_transformable_to TI (Fun f T) (Fun g S)" shows "length T = length S" and "\i. i < length T \ timpls_transformable_to TI (T ! i) (S ! i)" and "f \ g \ (is_Abs f \ is_Abs g \ (the_Abs f, the_Abs g) \ set TI)" using assms list_all2_conv_all_nth[of "timpls_transformable_to TI" T S] by auto lemma timpls_transformable_to'_inv: assumes "timpls_transformable_to' TI (Fun f T) (Fun g S)" shows "length T = length S" and "\i. i < length T \ timpls_transformable_to' TI (T ! i) (S ! i)" and "f \ g \ (is_Abs f \ is_Abs g \ in_trancl TI (the_Abs f) (the_Abs g))" using assms list_all2_conv_all_nth[of "timpls_transformable_to' TI" T S] by auto lemma timpls_transformable_to_inv': assumes "timpls_transformable_to TI (Fun f T) t" shows "is_Fun t" and "length T = length (args t)" and "\i. i < length T \ timpls_transformable_to TI (T ! i) (args t ! i)" and "f \ the_Fun t \ ( is_Abs f \ is_Abs (the_Fun t) \ (the_Abs f, the_Abs (the_Fun t)) \ set TI)" and "\is_Abs f \ f = the_Fun t" using assms list_all2_conv_all_nth[of "timpls_transformable_to TI" T] by (cases t; auto)+ lemma timpls_transformable_to'_inv': assumes "timpls_transformable_to' TI (Fun f T) t" shows "is_Fun t" and "length T = length (args t)" and "\i. i < length T \ timpls_transformable_to' TI (T ! i) (args t ! i)" and "f \ the_Fun t \ ( is_Abs f \ is_Abs (the_Fun t) \ in_trancl TI (the_Abs f) (the_Abs (the_Fun t)))" and "\is_Abs f \ f = the_Fun t" using assms list_all2_conv_all_nth[of "timpls_transformable_to' TI" T] by (cases t; auto)+ lemma timpls_transformable_to_size_eq: fixes s t::"(('a, 'b, 'c, 'd) prot_fun, 'e) term" shows "timpls_transformable_to TI s t \ size s = size t" (is "?A \ ?C") and "timpls_transformable_to' TI s t \ size s = size t" (is "?B \ ?C") proof - have *: "size_list size T = size_list size S" when "length T = length S" "\i. i < length T \ size (T ! i) = size (S ! i)" for S T::"(('a, 'b, 'c, 'd) prot_fun, 'e) term list" using that proof (induction T arbitrary: S) case (Cons x T') then obtain y S' where y: "S = y#S'" by (cases S) auto hence "size_list size T' = size_list size S'" "size x = size y" using Cons.prems Cons.IH[of S'] by force+ thus ?case using y by simp qed simp show ?C when ?A using that proof (induction rule: timpls_transformable_to.induct) case (2 TI f T g S) hence "length T = length S" "\i. i < length T \ size (T ! i) = size (S ! i)" using timpls_transformable_to_inv(1,2)[of TI f T g S] by auto thus ?case using *[of S T] by simp qed simp_all show ?C when ?B using that proof (induction rule: timpls_transformable_to.induct) case (2 TI f T g S) hence "length T = length S" "\i. i < length T \ size (T ! i) = size (S ! i)" using timpls_transformable_to'_inv(1,2)[of TI f T g S] by auto thus ?case using *[of S T] by simp qed simp_all qed lemma timpls_transformable_to_if_term_variants: fixes s t::"(('a, 'b, 'c, 'd) prot_fun, 'e) term" and a b::"'c set" defines "P \ (\_. [])(Abs a := [Abs b])" assumes st: "term_variants_pred P s t" and ab: "(a,b) \ set TI" shows "timpls_transformable_to TI s t" using st P_def proof (induction rule: term_variants_pred.induct) case (term_variants_P T S f) thus ?case using ab list_all2_conv_all_nth[of "timpls_transformable_to TI" T S] by auto next case (term_variants_Fun T S f) thus ?case using ab list_all2_conv_all_nth[of "timpls_transformable_to TI" T S] by auto qed simp lemma timpls_transformable_to'_if_term_variants: fixes s t::"(('a, 'b, 'c, 'd) prot_fun, 'e) term" and a b::"'c set" defines "P \ (\_. [])(Abs a := [Abs b])" assumes st: "term_variants_pred P s t" and ab: "(a,b) \ (set TI)\<^sup>+" shows "timpls_transformable_to' TI s t" using st P_def proof (induction rule: term_variants_pred.induct) case (term_variants_P T S f) thus ?case using ab list_all2_conv_all_nth[of "timpls_transformable_to' TI" T S] in_trancl_closure_iff_in_trancl_fun[of _ _ TI] by auto next case (term_variants_Fun T S f) thus ?case using ab list_all2_conv_all_nth[of "timpls_transformable_to' TI" T S] in_trancl_closure_iff_in_trancl_fun[of _ _ TI] by auto qed simp lemma timpls_transformable_to_trans: assumes TI_trancl: "\(a,b) \ (set TI)\<^sup>+. a \ b \ (a,b) \ set TI" and st: "timpls_transformable_to TI s t" and tu: "timpls_transformable_to TI t u" shows "timpls_transformable_to TI s u" using st tu proof (induction s arbitrary: t u) case (Var x) thus ?case using tu timpls_transformable_to_inv_Var(1) by fast next case (Fun f T) obtain g S where t: "t = Fun g S" "length T = length S" "\i. i < length T \ timpls_transformable_to TI (T ! i) (S ! i)" "f \ g \ is_Abs f \ is_Abs g \ (the_Abs f, the_Abs g) \ set TI" - using timpls_transformable_to_inv'[OF Fun.prems(1)] TI_trancl by moura + using timpls_transformable_to_inv'[OF Fun.prems(1)] TI_trancl by auto obtain h U where u: "u = Fun h U" "length S = length U" "\i. i < length S \ timpls_transformable_to TI (S ! i) (U ! i)" "g \ h \ is_Abs g \ is_Abs h \ (the_Abs g, the_Abs h) \ set TI" - using timpls_transformable_to_inv'[OF Fun.prems(2)[unfolded t(1)]] TI_trancl by moura + using timpls_transformable_to_inv'[OF Fun.prems(2)[unfolded t(1)]] TI_trancl by auto have "list_all2 (timpls_transformable_to TI) T U" using t(1,2,3) u(1,2,3) Fun.IH list_all2_conv_all_nth[of "timpls_transformable_to TI" T S] list_all2_conv_all_nth[of "timpls_transformable_to TI" S U] list_all2_conv_all_nth[of "timpls_transformable_to TI" T U] by force moreover have "(the_Abs f, the_Abs h) \ set TI" when "(the_Abs f, the_Abs g) \ set TI" "(the_Abs g, the_Abs h) \ set TI" "f \ h" "is_Abs f" "is_Abs h" using that(3,4,5) TI_trancl trancl_into_trancl[OF r_into_trancl[OF that(1)] that(2)] unfolding is_Abs_def the_Abs_def by force hence "is_Abs f \ is_Abs h \ (the_Abs f, the_Abs h) \ set TI" when "f \ h" using that TI_trancl t(4) u(4) by fast ultimately show ?case using t(1) u(1) by force qed lemma timpls_transformable_to'_trans: assumes st: "timpls_transformable_to' TI s t" and tu: "timpls_transformable_to' TI t u" shows "timpls_transformable_to' TI s u" using st tu proof (induction s arbitrary: t u) case (Var x) thus ?case using tu timpls_transformable_to_inv_Var(3) by fast next case (Fun f T) note 0 = in_trancl_closure_iff_in_trancl_fun[of _ _ TI] obtain g S where t: "t = Fun g S" "length T = length S" "\i. i < length T \ timpls_transformable_to' TI (T ! i) (S ! i)" "f \ g \ is_Abs f \ is_Abs g \ (the_Abs f, the_Abs g) \ (set TI)\<^sup>+" - using timpls_transformable_to'_inv'[OF Fun.prems(1)] 0 by moura + using timpls_transformable_to'_inv'[OF Fun.prems(1)] 0 by auto obtain h U where u: "u = Fun h U" "length S = length U" "\i. i < length S \ timpls_transformable_to' TI (S ! i) (U ! i)" "g \ h \ is_Abs g \ is_Abs h \ (the_Abs g, the_Abs h) \ (set TI)\<^sup>+" - using timpls_transformable_to'_inv'[OF Fun.prems(2)[unfolded t(1)]] 0 by moura + using timpls_transformable_to'_inv'[OF Fun.prems(2)[unfolded t(1)]] 0 by auto have "list_all2 (timpls_transformable_to' TI) T U" using t(1,2,3) u(1,2,3) Fun.IH list_all2_conv_all_nth[of "timpls_transformable_to' TI" T S] list_all2_conv_all_nth[of "timpls_transformable_to' TI" S U] list_all2_conv_all_nth[of "timpls_transformable_to' TI" T U] by force moreover have "(the_Abs f, the_Abs h) \ (set TI)\<^sup>+" when "(the_Abs f, the_Abs g) \ (set TI)\<^sup>+" "(the_Abs g, the_Abs h) \ (set TI)\<^sup>+" using that by simp hence "is_Abs f \ is_Abs h \ (the_Abs f, the_Abs h) \ (set TI)\<^sup>+" when "f \ h" by (metis that t(4) u(4)) ultimately show ?case using t(1) u(1) 0 by force qed lemma timpls_transformable_to_mono: assumes "set TI \ set TI'" and "timpls_transformable_to TI s t" shows "timpls_transformable_to TI' s t" using assms proof (induction TI s t rule: timpls_transformable_to.induct) case (2 TI f T g S) have *: "f = g \ (is_Abs f \ is_Abs g \ (the_Abs f, the_Abs g) \ set TI)" "list_all2 (timpls_transformable_to TI) T S" using "2.prems" by simp_all show ?case using "2.IH" "2.prems"(1) list.rel_mono_strong[OF *(2)] *(1) in_trancl_mono[of TI TI'] by (metis (no_types, lifting) timpls_transformable_to.simps(2) set_rev_mp) qed auto lemma timpls_transformable_to'_mono: assumes "set TI \ set TI'" and "timpls_transformable_to' TI s t" shows "timpls_transformable_to' TI' s t" using assms proof (induction TI s t rule: timpls_transformable_to'.induct) case (2 TI f T g S) have *: "f = g \ (is_Abs f \ is_Abs g \ in_trancl TI (the_Abs f) (the_Abs g))" "list_all2 (timpls_transformable_to' TI) T S" using "2.prems" by simp_all show ?case using "2.IH" "2.prems"(1) list.rel_mono_strong[OF *(2)] *(1) in_trancl_mono[of TI TI'] by (metis (no_types, lifting) timpls_transformable_to'.simps(2)) qed auto lemma timpls_transformable_to_refl_minus_eq: "timpls_transformable_to TI s t \ timpls_transformable_to (filter (\(a,b). a \ b) TI) s t" (is "?A \ ?B") proof let ?TI' = "\TI. filter (\(a,b). a \ b) TI" show ?A when ?B using that timpls_transformable_to_mono[of "?TI' TI" TI] by auto show ?B when ?A using that proof (induction TI s t rule: timpls_transformable_to.induct) case (2 TI f T g S) have *: "f = g \ (is_Abs f \ is_Abs g \ (the_Abs f, the_Abs g) \ set TI)" "list_all2 (timpls_transformable_to TI) T S" using "2.prems" by simp_all have "f = g \ (is_Abs f \ is_Abs g \ (the_Abs f, the_Abs g) \ set (?TI' TI))" using *(1) unfolding is_Abs_def by auto moreover have "list_all2 (timpls_transformable_to (?TI' TI)) T S" using *(2) "2.IH" list.rel_mono_strong by blast ultimately show ?case by force qed auto qed lemma timpls_transformable_to_iff_in_timpl_closure: assumes "set TI' = {(a,b) \ (set TI)\<^sup>+. a \ b}" shows "timpls_transformable_to TI' s t \ t \ timpl_closure s (set TI)" (is "?A s t \ ?B s t") proof show "?A s t \ ?B s t" using assms proof (induction s t rule: timpls_transformable_to.induct) case (2 TI f T g S) note prems = "2.prems" note IH = "2.IH" have 1: "length T = length S" "\i timpl_closure' (set TI')" when i: "i < length S" for i proof - have "timpls_transformable_to TI' (T ! i) (S ! i)" using i 1 by presburger hence "S ! i \ timpl_closure (T ! i) (set TI)" using IH[of "T ! i" "S ! i"] i 1(1) prems(2) by force thus ?thesis using 2[of "S ! i" "T ! i" "set TI"] 4 by blast qed have 5: "f = g \ (\a b. (a, b) \ (set TI')\<^sup>+ \ f = Abs a \ g = Abs b)" using prems(1) the_Abs_def[of f] the_Abs_def[of g] is_Abs_def[of f] is_Abs_def[of g] by fastforce show ?case using 2 4 timpl_closure_FunI[OF IH' 1(1) 5] 1(1) by auto qed (simp_all add: timpl_closure.FP) show "?B s t \ ?A s t" proof (induction t rule: timpl_closure.induct) case (TI u a b v) show ?case proof (cases "a = b") case True thus ?thesis using TI.hyps(3) TI.IH term_variants_pred_refl_inv by fastforce next case False hence 1: "timpls_transformable_to TI' u v" using TI.hyps(2) assms timpls_transformable_to_if_term_variants[OF TI.hyps(3), of TI'] by blast have 2: "(c,d) \ set TI'" when cd: "(c,d) \ (set TI')\<^sup>+" "c \ d" for c d proof - let ?cl = "\X. {(a,b) \ X\<^sup>+. a \ b}" have "?cl (set TI') = ?cl (?cl (set TI))" using assms by presburger hence "set TI' = ?cl (set TI')" using assms trancl_minus_refl_idem[of "set TI"] by argo thus ?thesis using cd by blast qed show ?thesis using timpls_transformable_to_trans[OF _ TI.IH 1] 2 by blast qed qed (use timpls_transformable_to_refl in fast) qed lemma timpls_transformable_to'_iff_in_timpl_closure: "timpls_transformable_to' TI s t \ t \ timpl_closure s (set TI)" (is "?A s t \ ?B s t") proof show "?A s t \ ?B s t" proof (induction s t rule: timpls_transformable_to'.induct) case (2 TI f T g S) note prems = "2.prems" note IH = "2.IH" have 1: "length T = length S" "\i timpl_closure' (set TI)" when i: "i < length S" for i proof - have "timpls_transformable_to' TI (T ! i) (S ! i)" using i 1 by presburger hence "S ! i \ timpl_closure (T ! i) (set TI)" using IH[of "T ! i" "S ! i"] i 1(1) by force thus ?thesis using 2[of "S ! i" "T ! i" "set TI"] by blast qed have 4: "f = g \ (\a b. (a, b) \ (set TI)\<^sup>+ \ f = Abs a \ g = Abs b)" using prems the_Abs_def[of f] the_Abs_def[of g] is_Abs_def[of f] is_Abs_def[of g] in_trancl_closure_iff_in_trancl_fun[of _ _ TI] by auto show ?case using 2 timpl_closure_FunI[OF IH' 1(1) 4] 1(1) by auto qed (simp_all add: timpl_closure.FP) show "?B s t \ ?A s t" proof (induction t rule: timpl_closure.induct) case (TI u a b v) thus ?case using timpls_transformable_to'_trans timpls_transformable_to'_if_term_variants by blast qed (use timpls_transformable_to_refl(2) in fast) qed lemma equal_mod_timpls_iff_ex_in_timpl_closure: assumes "set TI' = {(a,b) \ TI\<^sup>+. a \ b}" shows "equal_mod_timpls TI' s t \ (\u. u \ timpl_closure s TI \ u \ timpl_closure t TI)" (is "?A s t \ ?B s t") proof show "?A s t \ ?B s t" using assms proof (induction s t rule: equal_mod_timpls.induct) case (2 TI' f T g S) note prems = "2.prems" note IH = "2.IH" have 1: "length T = length S" "\iu. (T ! i, u) \ timpl_closure' TI \ (S ! i, u) \ timpl_closure' TI" when i: "i < length S" for i proof - have "equal_mod_timpls TI' (T ! i) (S ! i)" using i 1 by presburger hence "\u. u \ timpl_closure (T ! i) TI \ u \ timpl_closure (S ! i) TI" using IH[of "T ! i" "S ! i"] i 1(1) prems by force thus ?thesis using 4 unfolding 2 by blast qed let ?P = "\G. f = g \ (\a b. (a, b) \ G \ f = Abs a \ g = Abs b) \ (\a b. (a, b) \ G \ f = Abs b \ g = Abs a) \ (\a b c. (a, c) \ G \ (b, c) \ G \ f = Abs a \ g = Abs b)" have "?P (set TI')" using prems the_Abs_def[of f] the_Abs_def[of g] is_Abs_def[of f] is_Abs_def[of g] by fastforce hence "?P (TI\<^sup>+)" unfolding prems by blast hence "?P (rtrancl TI)" by (metis (no_types, lifting) trancl_into_rtrancl) hence 5: "f = g \ (\a b c. (a, c) \ TI\<^sup>* \ (b, c) \ TI\<^sup>* \ f = Abs a \ g = Abs b)" by blast show ?case using timpl_closure_FunI3[OF _ 1(1) 5] IH 1(1) unfolding timpl_closure'_timpls_rtrancl_eq 2 by auto qed (use timpl_closure.FP in auto) show "?A s t" when B: "?B s t" proof - obtain u where u: "u \ timpl_closure s TI" "u \ timpl_closure t TI" - using B by moura + using B by blast thus ?thesis using assms proof (induction u arbitrary: s t rule: term.induct) case (Var x s t) thus ?case using timpl_closure_Var_in_iff[of x s TI] timpl_closure_Var_in_iff[of x t TI] equal_mod_timpls.simps(1)[of TI' x x] by blast next case (Fun f U s t) obtain g S where s: "s = Fun g S" "length U = length S" "\i. i < length U \ U ! i \ timpl_closure (S ! i) TI" "g \ f \ is_Abs g \ is_Abs f \ (the_Abs g, the_Abs f) \ TI\<^sup>+" using Fun.prems(1) timpl_closure_Fun_inv'[of f U _ _ TI] by (cases s) auto obtain h T where t: "t = Fun h T" "length U = length T" "\i. i < length U \ U ! i \ timpl_closure (T ! i) TI" "h \ f \ is_Abs h \ is_Abs f \ (the_Abs h, the_Abs f) \ TI\<^sup>+" using Fun.prems(2) timpl_closure_Fun_inv'[of f U _ _ TI] by (cases t) auto have g: "(the_Abs g, the_Abs f) \ set TI'" "is_Abs f" "is_Abs g" when neq_f: "g \ f" proof - obtain ga fa where a: "g = Abs ga" "f = Abs fa" using s(4)[OF neq_f] unfolding is_Abs_def by presburger hence "the_Abs g \ the_Abs f" using neq_f by simp thus "(the_Abs g, the_Abs f) \ set TI'" "is_Abs f" "is_Abs g" using s(4)[OF neq_f] Fun.prems by blast+ qed have h: "(the_Abs h, the_Abs f) \ set TI'" "is_Abs f" "is_Abs h" when neq_f: "h \ f" proof - obtain ha fa where a: "h = Abs ha" "f = Abs fa" using t(4)[OF neq_f] unfolding is_Abs_def by presburger hence "the_Abs h \ the_Abs f" using neq_f by simp thus "(the_Abs h, the_Abs f) \ set TI'" "is_Abs f" "is_Abs h" using t(4)[OF neq_f] Fun.prems by blast+ qed have "equal_mod_timpls TI' (S ! i) (T ! i)" when i: "i < length U" for i using i Fun.IH s(1,2,3) t(1,2,3) nth_mem[OF i] Fun.prems by meson hence "list_all2 (equal_mod_timpls TI') S T" using list_all2_conv_all_nth[of "equal_mod_timpls TI'" S T] s(2) t(2) by presburger thus ?case using s(1) t(1) g h by fastforce qed qed qed (* lemma equal_mod_timpls_iff_ex_in_timpl_closure': "equal_mod_timpls (TI\<^sup>+) s t \ (\u. u \ timpl_closure s TI \ u \ timpl_closure t TI)" using equal_mod_timpls_iff_ex_in_timpl_closure equal_mod_timpls_refl_minus_eq by blast *) context begin private inductive timpls_transformable_to_pred where Var: "timpls_transformable_to_pred A (Var x) (Var x)" | Fun: "\\is_Abs f; length T = length S; \i. i < length T \ timpls_transformable_to_pred A (T ! i) (S ! i)\ \ timpls_transformable_to_pred A (Fun f T) (Fun f S)" | Abs: "b \ A \ timpls_transformable_to_pred A (Fun (Abs a) []) (Fun (Abs b) [])" private lemma timpls_transformable_to_pred_inv_Var: assumes "timpls_transformable_to_pred A (Var x) t" shows "t = Var x" using assms by (auto elim: timpls_transformable_to_pred.cases) private lemma timpls_transformable_to_pred_inv: assumes "timpls_transformable_to_pred A (Fun f T) t" shows "is_Fun t" and "length T = length (args t)" and "\i. i < length T \ timpls_transformable_to_pred A (T ! i) (args t ! i)" and "\is_Abs f \ f = the_Fun t" and "is_Abs f \ (is_Abs (the_Fun t) \ the_Abs (the_Fun t) \ A)" using assms by (auto elim!: timpls_transformable_to_pred.cases[of A]) private lemma timpls_transformable_to_pred_finite_aux1: assumes f: "\is_Abs f" shows "{s. timpls_transformable_to_pred A (Fun f T) s} \ (\S. Fun f S) ` {S. length T = length S \ (\s \ set S. \t \ set T. timpls_transformable_to_pred A t s)}" (is "?B \ ?C") proof fix s assume s: "s \ ?B" hence *: "timpls_transformable_to_pred A (Fun f T) s" by blast obtain S where S: "s = Fun f S" "length T = length S" "\i. i < length T \ timpls_transformable_to_pred A (T ! i) (S ! i)" using f timpls_transformable_to_pred_inv[OF *] unfolding the_Abs_def is_Abs_def by auto have "\s\set S. \t\set T. timpls_transformable_to_pred A t s" using S(2,3) in_set_conv_nth by metis thus "s \ ?C" using S(1,2) by blast qed private lemma timpls_transformable_to_pred_finite_aux2: "{s. timpls_transformable_to_pred A (Fun (Abs a) []) s} \ (\b. Fun (Abs b) []) ` A" (is "?B \ ?C") proof fix s assume s: "s \ ?B" hence *: "timpls_transformable_to_pred A (Fun (Abs a) []) s" by blast obtain b where b: "s = Fun (Abs b) []" "b \ A" using timpls_transformable_to_pred_inv[OF *] unfolding the_Abs_def is_Abs_def by auto thus "s \ ?C" by blast qed private lemma timpls_transformable_to_pred_finite: fixes t::"(('fun,'atom,'sets,'lbl) prot_fun, 'a) term" assumes A: "finite A" and t: "wf\<^sub>t\<^sub>r\<^sub>m t" shows "finite {s. timpls_transformable_to_pred A t s}" using t proof (induction t) case (Var x) have "{s::(('fun,'atom,'sets,'lbl) prot_fun, 'a) term. timpls_transformable_to_pred A (Var x) s} = {Var x}" by (auto intro: timpls_transformable_to_pred.Var elim: timpls_transformable_to_pred_inv_Var) thus ?case by simp next case (Fun f T) have IH: "finite {s. timpls_transformable_to_pred A t s}" when t: "t \ set T" for t using Fun.IH[OF t] wf_trm_param[OF Fun.prems t] by blast show ?case proof (cases "is_Abs f") case True then obtain a where a: "f = Abs a" unfolding is_Abs_def by presburger hence "T = []" using wf_trm_arity[OF Fun.prems] by simp_all hence "{a. timpls_transformable_to_pred A (Fun f T) a} \ (\b. Fun (Abs b) []) ` A" using timpls_transformable_to_pred_finite_aux2[of A a] a by auto thus ?thesis using A finite_subset by fast next case False thus ?thesis using IH finite_lists_length_eq' timpls_transformable_to_pred_finite_aux1[of f A T] finite_subset by blast qed qed private lemma timpls_transformable_to_pred_if_timpls_transformable_to: assumes s: "timpls_transformable_to TI t s" and t: "wf\<^sub>t\<^sub>r\<^sub>m t" "\f \ funs_term t. is_Abs f \ the_Abs f \ A" shows "timpls_transformable_to_pred (A \ fst ` (set TI)\<^sup>+ \ snd ` (set TI)\<^sup>+) t s" using s t proof (induction rule: timpls_transformable_to.induct) case (2 TI f T g S) let ?A = "A \ fst ` (set TI)\<^sup>+ \ snd ` (set TI)\<^sup>+" note prems = "2.prems" note IH = "2.IH" note 0 = timpls_transformable_to_inv[OF prems(1)] have 1: "T = []" "S = []" when f: "f = Abs a" for a using f wf_trm_arity[OF prems(2)] 0(1) by simp_all have "\f \ funs_term t. is_Abs f \ the_Abs f \ A" when t: "t \ set T" for t using t prems(3) funs_term_subterms_eq(1)[of "Fun f T"] by blast hence 2: "timpls_transformable_to_pred ?A (T ! i) (S ! i)" when i: "i < length T" for i using i IH 0(1,2) wf_trm_param[OF prems(2)] by (metis (no_types) in_set_conv_nth) have 3: "the_Abs f \ ?A" when f: "is_Abs f" using prems(3) f by force show ?case proof (cases "f = g") case True note fg = True show ?thesis proof (cases "is_Abs f") case True - then obtain a where a: "f = Abs a" unfolding is_Abs_def by moura + then obtain a where a: "f = Abs a" unfolding is_Abs_def by blast thus ?thesis using fg 1[OF a] timpls_transformable_to_pred.Abs[of a ?A a] 3 by simp qed (use fg timpls_transformable_to_pred.Fun[OF _ 0(1) 2, of f] in blast) next case False then obtain a b where ab: "f = Abs a" "g = Abs b" "(a, b) \ (set TI)\<^sup>+" using 0(3) in_trancl_closure_iff_in_trancl_fun[of _ _ TI] unfolding is_Abs_def the_Abs_def by fastforce hence "a \ ?A" "b \ ?A" by force+ thus ?thesis using timpls_transformable_to_pred.Abs ab(1,2) 1[OF ab(1)] by metis qed qed (simp_all add: timpls_transformable_to_pred.Var) private lemma timpls_transformable_to_pred_if_timpls_transformable_to': assumes s: "timpls_transformable_to' TI t s" and t: "wf\<^sub>t\<^sub>r\<^sub>m t" "\f \ funs_term t. is_Abs f \ the_Abs f \ A" shows "timpls_transformable_to_pred (A \ fst ` (set TI)\<^sup>+ \ snd ` (set TI)\<^sup>+) t s" using s t proof (induction rule: timpls_transformable_to.induct) case (2 TI f T g S) let ?A = "A \ fst ` (set TI)\<^sup>+ \ snd ` (set TI)\<^sup>+" note prems = "2.prems" note IH = "2.IH" note 0 = timpls_transformable_to'_inv[OF prems(1)] have 1: "T = []" "S = []" when f: "f = Abs a" for a using f wf_trm_arity[OF prems(2)] 0(1) by simp_all have "\f \ funs_term t. is_Abs f \ the_Abs f \ A" when t: "t \ set T" for t using t prems(3) funs_term_subterms_eq(1)[of "Fun f T"] by blast hence 2: "timpls_transformable_to_pred ?A (T ! i) (S ! i)" when i: "i < length T" for i using i IH 0(1,2) wf_trm_param[OF prems(2)] by (metis (no_types) in_set_conv_nth) have 3: "the_Abs f \ ?A" when f: "is_Abs f" using prems(3) f by force show ?case proof (cases "f = g") case True note fg = True show ?thesis proof (cases "is_Abs f") case True - then obtain a where a: "f = Abs a" unfolding is_Abs_def by moura + then obtain a where a: "f = Abs a" unfolding is_Abs_def by blast thus ?thesis using fg 1[OF a] timpls_transformable_to_pred.Abs[of a ?A a] 3 by simp qed (use fg timpls_transformable_to_pred.Fun[OF _ 0(1) 2, of f] in blast) next case False then obtain a b where ab: "f = Abs a" "g = Abs b" "(a, b) \ (set TI)\<^sup>+" using 0(3) in_trancl_closure_iff_in_trancl_fun[of _ _ TI] unfolding is_Abs_def the_Abs_def by fastforce hence "a \ ?A" "b \ ?A" by force+ thus ?thesis using timpls_transformable_to_pred.Abs ab(1,2) 1[OF ab(1)] by metis qed qed (simp_all add: timpls_transformable_to_pred.Var) private lemma timpls_transformable_to_pred_if_equal_mod_timpls: assumes s: "equal_mod_timpls TI t s" and t: "wf\<^sub>t\<^sub>r\<^sub>m t" "\f \ funs_term t. is_Abs f \ the_Abs f \ A" shows "timpls_transformable_to_pred (A \ fst ` (set TI)\<^sup>+ \ snd ` (set TI)\<^sup>+) t s" using s t proof (induction rule: equal_mod_timpls.induct) case (2 TI f T g S) let ?A = "A \ fst ` (set TI)\<^sup>+ \ snd ` (set TI)\<^sup>+" note prems = "2.prems" note IH = "2.IH" note 0 = equal_mod_timpls_inv[OF prems(1)] have 1: "T = []" "S = []" when f: "f = Abs a" for a using f wf_trm_arity[OF prems(2)] 0(1) by simp_all have "\f \ funs_term t. is_Abs f \ the_Abs f \ A" when t: "t \ set T" for t using t prems(3) funs_term_subterms_eq(1)[of "Fun f T"] by blast hence 2: "timpls_transformable_to_pred ?A (T ! i) (S ! i)" when i: "i < length T" for i using i IH 0(1,2) wf_trm_param[OF prems(2)] by (metis (no_types) in_set_conv_nth) have 3: "the_Abs f \ ?A" when f: "is_Abs f" using prems(3) f by force show ?case proof (cases "f = g") case True note fg = True show ?thesis proof (cases "is_Abs f") case True - then obtain a where a: "f = Abs a" unfolding is_Abs_def by moura + then obtain a where a: "f = Abs a" unfolding is_Abs_def by blast thus ?thesis using fg 1[OF a] timpls_transformable_to_pred.Abs[of a ?A a] 3 by simp qed (use fg timpls_transformable_to_pred.Fun[OF _ 0(1) 2, of f] in blast) next case False then obtain a b where ab: "f = Abs a" "g = Abs b" "(a, b) \ (set TI)\<^sup>+ \ (b, a) \ (set TI)\<^sup>+ \ (\ti \ set TI. (a, snd ti) \ (set TI)\<^sup>+ \ (b, snd ti) \ (set TI)\<^sup>+)" using 0(3) in_trancl_closure_iff_in_trancl_fun[of _ _ TI] unfolding is_Abs_def the_Abs_def by fastforce hence "a \ ?A" "b \ ?A" by force+ thus ?thesis using timpls_transformable_to_pred.Abs ab(1,2) 1[OF ab(1)] by metis qed qed (simp_all add: timpls_transformable_to_pred.Var) lemma timpls_transformable_to_finite: assumes t: "wf\<^sub>t\<^sub>r\<^sub>m t" shows "finite {s. timpls_transformable_to TI t s}" (is ?P) and "finite {s. timpls_transformable_to' TI t s}" (is ?Q) proof - let ?A = "the_Abs ` {f \ funs_term t. is_Abs f} \ fst ` (set TI)\<^sup>+ \ snd ` (set TI)\<^sup>+" have 0: "finite ?A" by auto have 1: "{s. timpls_transformable_to TI t s} \ {s. timpls_transformable_to_pred ?A t s}" using timpls_transformable_to_pred_if_timpls_transformable_to[OF _ t] by auto have 2: "{s. timpls_transformable_to' TI t s} \ {s. timpls_transformable_to_pred ?A t s}" using timpls_transformable_to_pred_if_timpls_transformable_to'[OF _ t] by auto show ?P using timpls_transformable_to_pred_finite[OF 0 t] finite_subset[OF 1] by blast show ?Q using timpls_transformable_to_pred_finite[OF 0 t] finite_subset[OF 2] by blast qed lemma equal_mod_timpls_finite: assumes t: "wf\<^sub>t\<^sub>r\<^sub>m t" shows "finite {s. equal_mod_timpls TI t s}" proof - let ?A = "the_Abs ` {f \ funs_term t. is_Abs f} \ fst ` (set TI)\<^sup>+ \ snd ` (set TI)\<^sup>+" have 0: "finite ?A" by auto have 1: "{s. equal_mod_timpls TI t s} \ {s. timpls_transformable_to_pred ?A t s}" using timpls_transformable_to_pred_if_equal_mod_timpls[OF _ t] by auto show ?thesis using timpls_transformable_to_pred_finite[OF 0 t] finite_subset[OF 1] by blast qed end lemma intruder_synth_mod_timpls_is_synth_timpl_closure_set: fixes t::"(('fun,'atom,'sets,'lbl) prot_fun, 'a) term" and TI TI' assumes "set TI' = {(a,b) \ (set TI)\<^sup>+. a \ b}" shows "intruder_synth_mod_timpls M TI' t \ timpl_closure_set (set M) (set TI) \\<^sub>c t" (is "?C t \ ?D t") proof - have *: "(\m \ M. timpls_transformable_to TI' m t) \ t \ timpl_closure_set M (set TI)" when "set TI' = {(a,b) \ (set TI)\<^sup>+. a \ b}" for M TI TI' and t::"(('fun,'atom,'sets,'lbl) prot_fun, 'a) term" using timpls_transformable_to_iff_in_timpl_closure[OF that] timpl_closure_set_is_timpl_closure_union[of M "set TI"] timpl_closure_set_timpls_trancl_eq[of M "set TI"] timpl_closure_set_timpls_trancl_eq'[of M "set TI"] by auto show "?C t \ ?D t" proof show "?C t \ ?D t" using assms proof (induction t arbitrary: M TI TI' rule: intruder_synth_mod_timpls.induct) case (1 M TI' x) hence "Var x \ timpl_closure_set (set M) (set TI)" using timpl_closure.FP member_def unfolding timpl_closure_set_def by force thus ?case by simp next case (2 M TI f T) show ?case proof (cases "\m \ set M. timpls_transformable_to TI' m (Fun f T)") case True thus ?thesis using "2.prems" *[of TI' TI "set M" "Fun f T"] intruder_synth.AxiomC[of "Fun f T" "timpl_closure_set (set M) (set TI)"] by blast next case False hence "\(list_ex (\t. timpls_transformable_to TI' t (Fun f T)) M)" unfolding list_ex_iff by blast hence "public f" "length T = arity f" "list_all (intruder_synth_mod_timpls M TI') T" using "2.prems"(1) by force+ thus ?thesis using "2.IH"[OF _ _ "2.prems"(2)] unfolding list_all_iff by force qed qed show "?D t \ ?C t" proof (induction t rule: intruder_synth_induct) case (AxiomC t) thus ?case using timpl_closure_set_Var_in_iff[of _ "set M" "set TI"] *[OF assms, of "set M" t] by (cases t rule: term.exhaust) (force simp add: member_def list_ex_iff)+ next case (ComposeC T f) thus ?case using list_all_iff[of "intruder_synth_mod_timpls M TI'" T] intruder_synth_mod_timpls.simps(2)[of M TI' f T] by blast qed qed qed lemma intruder_synth_mod_timpls'_is_synth_timpl_closure_set: fixes t::"(('fun,'atom,'sets,'lbl) prot_fun, 'a) term" and TI shows "intruder_synth_mod_timpls' M TI t \ timpl_closure_set (set M) (set TI) \\<^sub>c t" (is "?A t \ ?B t") proof - have *: "(\m \ M. timpls_transformable_to' TI m t) \ t \ timpl_closure_set M (set TI)" for M TI and t::"(('fun,'atom,'sets,'lbl) prot_fun, 'a) term" using timpls_transformable_to'_iff_in_timpl_closure[of TI _ t] timpl_closure_set_is_timpl_closure_union[of M "set TI"] by blast+ show "?A t \ ?B t" proof show "?A t \ ?B t" proof (induction t arbitrary: M TI rule: intruder_synth_mod_timpls'.induct) case (1 M TI x) hence "Var x \ timpl_closure_set (set M) (set TI)" using timpl_closure.FP List.member_def[of M] unfolding timpl_closure_set_def by auto thus ?case by simp next case (2 M TI f T) show ?case proof (cases "\m \ set M. timpls_transformable_to' TI m (Fun f T)") case True thus ?thesis using "2.prems" *[of "set M" TI "Fun f T"] intruder_synth.AxiomC[of "Fun f T" "timpl_closure_set (set M) (set TI)"] by blast next case False hence "public f" "length T = arity f" "list_all (intruder_synth_mod_timpls' M TI) T" using "2.prems" list_ex_iff[of _ M] by force+ thus ?thesis using "2.IH"[of _ M TI] list_all_iff[of "intruder_synth_mod_timpls' M TI" T] by force qed qed show "?B t \ ?A t" proof (induction t rule: intruder_synth_induct) case (AxiomC t) thus ?case using AxiomC timpl_closure_set_Var_in_iff[of _ "set M" "set TI"] *[of "set M" TI t] list_ex_iff[of _ M] List.member_def[of M] by (cases t rule: term.exhaust) force+ next case (ComposeC T f) thus ?case using list_all_iff[of "intruder_synth_mod_timpls' M TI" T] intruder_synth_mod_timpls'.simps(2)[of M TI f T] by blast qed qed qed lemma intruder_synth_mod_eq_timpls_is_synth_timpl_closure_set: fixes t::"(('fun,'atom,'sets,'lbl) prot_fun, 'a) term" and TI defines "cl \ \TI. {(a,b) \ TI\<^sup>+. a \ b}" shows (* "set TI' = (set TI)\<^sup>+ \ intruder_synth_mod_eq_timpls M TI' t \ (\s \ timpl_closure t (set TI). timpl_closure_set M (set TI) \\<^sub>c s)" (is "?P TI TI' \ ?A t \ ?B t") and *) "set TI' = {(a,b) \ (set TI)\<^sup>+. a \ b} \ intruder_synth_mod_eq_timpls M TI' t \ (\s \ timpl_closure t (set TI). timpl_closure_set M (set TI) \\<^sub>c s)" (is "?Q TI TI' \ ?C t \ ?D t") proof - (* have *: "(\m \ M. equal_mod_timpls TI' m t) \ (\s \ timpl_closure t (set TI). s \ timpl_closure_set M (set TI))" when P: "?P TI TI'" for M TI TI' and t::"(('fun,'atom,'sets,'lbl) prot_fun, 'a) term" using equal_mod_timpls_iff_ex_in_timpl_closure'[OF P] timpl_closure_set_is_timpl_closure_union[of M "set TI"] timpl_closure_set_timpls_trancl_eq[of M "set TI"] by blast *) have **: "(\m \ M. equal_mod_timpls TI' m t) \ (\s \ timpl_closure t (set TI). s \ timpl_closure_set M (set TI))" when Q: "?Q TI TI'" for M TI TI' and t::"(('fun,'atom,'sets,'lbl) prot_fun, 'a) term" using equal_mod_timpls_iff_ex_in_timpl_closure[OF Q] timpl_closure_set_is_timpl_closure_union[of M "set TI"] timpl_closure_set_timpls_trancl_eq'[of M "set TI"] by fastforce (* show "?A t \ ?B t" when P: "?P TI TI'" proof show "?A t \ ?B t" proof (induction t arbitrary: M TI rule: intruder_synth_mod_eq_timpls.induct) case (1 M TI x) hence "Var x \ timpl_closure_set M TI" "Var x \ timpl_closure (Var x) TI" using timpl_closure.FP unfolding timpl_closure_set_def by auto thus ?case by force next case (2 M TI f T) show ?case proof (cases "\m \ M. equal_mod_timpls (TI\<^sup>+) m (Fun f T)") case True thus ?thesis using "2.prems" *[of M TI "Fun f T"] intruder_synth.AxiomC[of _ "timpl_closure_set M TI"] by blast next case False hence f: "public f" "length T = arity f" "list_all (intruder_synth_mod_eq_timpls M (TI\<^sup>+)) T" using "2.prems" by force+ let ?sy = "intruder_synth (timpl_closure_set M TI)" have IH: "\u \ timpl_closure (T ! i) TI. ?sy u" when i: "i < length T" for i using "2.IH"[of _ M TI] f(3) nth_mem[OF i] unfolding list_all_iff by blast define S where "S \ map (\u. SOME v. v \ timpl_closure u TI \ ?sy v) T" have S1: "length T = length S" unfolding S_def by simp have S2: "S ! i \ timpl_closure (T ! i) TI" "timpl_closure_set M TI \\<^sub>c S ! i" when i: "i < length S" for i using i IH someI_ex[of "\v. v \ timpl_closure (T ! i) TI \ ?sy v"] unfolding S_def by auto have "Fun f S \ timpl_closure (Fun f T) TI" using timpl_closure_FunI[of T S TI f f] S1 S2(1) unfolding timpl_closure_is_timpl_closure' by presburger thus ?thesis by (metis intruder_synth.ComposeC[of S f] f(1,2) S1 S2(2) in_set_conv_nth[of _ S]) qed qed show "?A t" when B: "?B t" proof - obtain s where "timpl_closure_set M TI \\<^sub>c s" "s \ timpl_closure t TI" - using B by moura + using B by blast thus ?thesis proof (induction s arbitrary: t rule: intruder_synth_induct) case (AxiomC s t) note 1 = timpl_closure_set_Var_in_iff[of _ M TI] timpl_closure_Var_inv[of s _ TI] note 2 = *[of M TI] show ?case proof (cases t) case Var thus ?thesis using 1 AxiomC by auto next case Fun thus ?thesis using 2 AxiomC by auto qed next case (ComposeC T f t) obtain g S where gS: "t = Fun g S" "length S = length T" "\i < length T. T ! i \ timpl_closure (S ! i) TI" "g \ f \ is_Abs g \ is_Abs f \ (the_Abs g, the_Abs f) \ TI\<^sup>+" using ComposeC.prems(1) timpl_closure'_inv'[of t "Fun f T" TI] timpl_closure_is_timpl_closure'[of _ _ TI] by fastforce have IH: "intruder_synth_mod_eq_timpls M (TI\<^sup>+) u" when u: "u \ set S" for u by (metis u gS(2,3) ComposeC.IH in_set_conv_nth) note 0 = list_all_iff[of "intruder_synth_mod_eq_timpls M (TI\<^sup>+)" S] intruder_synth_mod_eq_timpls.simps(2)[of M "TI\<^sup>+" g S] have "f = g" using ComposeC.hyps gS(4) unfolding is_Abs_def by fastforce thus ?case by (metis ComposeC.hyps(1,2) gS(1,2) IH 0) qed qed qed *) show "?C t \ ?D t" when Q: "?Q TI TI'" proof show "?C t \ ?D t" using Q proof (induction t arbitrary: M TI rule: intruder_synth_mod_eq_timpls.induct) case (1 M TI' x M TI) hence "Var x \ timpl_closure_set M (set TI)" "Var x \ timpl_closure (Var x) (set TI)" using timpl_closure.FP unfolding timpl_closure_set_def by auto thus ?case by force next case (2 M TI' f T M TI) show ?case proof (cases "\m \ M. equal_mod_timpls TI' m (Fun f T)") case True thus ?thesis using **[OF "2.prems"(2), of M "Fun f T"] intruder_synth.AxiomC[of _ "timpl_closure_set M (set TI)"] by blast next case False hence f: "public f" "length T = arity f" "list_all (intruder_synth_mod_eq_timpls M TI') T" using "2.prems" by force+ let ?sy = "intruder_synth (timpl_closure_set M (set TI))" have IH: "\u \ timpl_closure (T ! i) (set TI). ?sy u" when i: "i < length T" for i using "2.IH"[of _ M TI] f(3) nth_mem[OF i] "2.prems"(2) unfolding list_all_iff by blast define S where "S \ map (\u. SOME v. v \ timpl_closure u (set TI) \ ?sy v) T" have S1: "length T = length S" unfolding S_def by simp have S2: "S ! i \ timpl_closure (T ! i) (set TI)" "timpl_closure_set M (set TI) \\<^sub>c S ! i" when i: "i < length S" for i using i IH someI_ex[of "\v. v \ timpl_closure (T ! i) (set TI) \ ?sy v"] unfolding S_def by auto have "Fun f S \ timpl_closure (Fun f T) (set TI)" using timpl_closure_FunI[of T S "set TI" f f] S1 S2(1) unfolding timpl_closure_is_timpl_closure' by presburger thus ?thesis by (metis intruder_synth.ComposeC[of S f] f(1,2) S1 S2(2) in_set_conv_nth[of _ S]) qed qed show "?C t" when D: "?D t" proof - obtain s where "timpl_closure_set M (set TI) \\<^sub>c s" "s \ timpl_closure t (set TI)" - using D by moura + using D by blast thus ?thesis proof (induction s arbitrary: t rule: intruder_synth_induct) case (AxiomC s t) note 1 = timpl_closure_set_Var_in_iff[of _ M "set TI"] timpl_closure_Var_inv[of s _ "set TI"] note 2 = **[OF Q, of M] show ?case proof (cases t) case Var thus ?thesis using 1 AxiomC by auto next case Fun thus ?thesis using 2 AxiomC by auto qed next case (ComposeC T f t) obtain g S where gS: "t = Fun g S" "length S = length T" "\i < length T. T ! i \ timpl_closure (S ! i) (set TI)" "g \ f \ is_Abs g \ is_Abs f \ (the_Abs g, the_Abs f) \ (set TI)\<^sup>+" using ComposeC.prems(1) timpl_closure'_inv'[of t "Fun f T" "set TI"] timpl_closure_is_timpl_closure'[of _ _ "set TI"] by fastforce have IH: "intruder_synth_mod_eq_timpls M TI' u" when u: "u \ set S" for u by (metis u gS(2,3) ComposeC.IH in_set_conv_nth) note 0 = list_all_iff[of "intruder_synth_mod_eq_timpls M TI'" S] intruder_synth_mod_eq_timpls.simps(2)[of M TI' g S] have "f = g" using ComposeC.hyps gS(4) unfolding is_Abs_def by fastforce thus ?case by (metis ComposeC.hyps(1,2) gS(1,2) IH 0) qed qed qed qed lemma timpl_closure_finite: assumes t: "wf\<^sub>t\<^sub>r\<^sub>m t" shows "finite (timpl_closure t (set TI))" using timpls_transformable_to'_iff_in_timpl_closure[of TI t] timpls_transformable_to_finite[OF t, of TI] by auto lemma timpl_closure_set_finite: fixes TI::"('sets set \ 'sets set) list" assumes M_finite: "finite M" and M_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" shows "finite (timpl_closure_set M (set TI))" using timpl_closure_set_is_timpl_closure_union[of M "set TI"] timpl_closure_finite[of _ TI] M_finite M_wf finite by auto lemma comp_timpl_closure_is_timpl_closure_set: fixes M and TI::"('sets set \ 'sets set) list" assumes M_finite: "finite M" and M_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" shows "comp_timpl_closure M (set TI) = timpl_closure_set M (set TI)" using lfp_while''[OF timpls_Un_mono[of M "set TI"]] timpl_closure_set_finite[OF M_finite M_wf] timpl_closure_set_lfp[of M "set TI"] unfolding comp_timpl_closure_def Let_def by presburger context begin private lemma analyzed_closed_mod_timpls_is_analyzed_closed_timpl_closure_set_aux1: fixes M::"('fun,'atom,'sets,'lbl) prot_terms" assumes f: "arity\<^sub>f f = length T" "arity\<^sub>f f > 0" "Ana\<^sub>f f = (K, R)" and i: "i < length R" and M: "timpl_closure_set M TI \\<^sub>c T ! (R ! i)" and m: "Fun (Fu f) T \ M" and t: "Fun (Fu f) S \ timpl_closure (Fun (Fu f) T) TI" shows "timpl_closure_set M TI \\<^sub>c S ! (R ! i)" proof - have "R ! i < length T" using i Ana\<^sub>f_assm2_alt[OF f(3)] f(1) by simp thus ?thesis using timpl_closure_Fun_inv'(1,2)[OF t] intruder_synth_timpl_closure'[OF M] by presburger qed private lemma analyzed_closed_mod_timpls_is_analyzed_closed_timpl_closure_set_aux2: fixes M::"('fun,'atom,'sets,'lbl) prot_terms" assumes M: "\s \ set (snd (Ana m)). timpl_closure_set M TI \\<^sub>c s" and m: "m \ M" and t: "t \ timpl_closure m TI" and s: "s \ set (snd (Ana t))" shows "timpl_closure_set M TI \\<^sub>c s" proof - obtain f S K N where fS: "t = Fun (Fu f) S" "arity\<^sub>f f = length S" "0 < arity\<^sub>f f" and Ana_f: "Ana\<^sub>f f = (K, N)" and Ana_t: "Ana t = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) S, map ((!) S) N)" using Ana_nonempty_inv[of t] s by fastforce then obtain T where T: "m = Fun (Fu f) T" "length T = length S" using t timpl_closure_Fu_inv'[of f S m TI] - by moura + by blast hence Ana_m: "Ana m = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T, map ((!) T) N)" using fS(2,3) Ana_f by auto obtain i where i: "i < length N" "s = S ! (N ! i)" using s[unfolded fS(1)] Ana_t[unfolded fS(1)] T(2) in_set_conv_nth[of s "map (\i. S ! i) N"] by auto hence "timpl_closure_set M TI \\<^sub>c T ! (N ! i)" using M[unfolded T(1)] Ana_m[unfolded T(1)] T(2) by simp thus ?thesis using analyzed_closed_mod_timpls_is_analyzed_closed_timpl_closure_set_aux1[ OF fS(2)[unfolded T(2)[symmetric]] fS(3) Ana_f i(1) _ m[unfolded T(1)] t[unfolded fS(1) T(1)]] i(2) by argo qed lemma analyzed_closed_mod_timpls_is_analyzed_timpl_closure_set: fixes M::"('fun,'atom,'sets,'lbl) prot_term list" assumes TI': "set TI' = {(a,b) \ (set TI)\<^sup>+. a \ b}" and M_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set M)" shows "analyzed_closed_mod_timpls M TI' \ analyzed (timpl_closure_set (set M) (set TI))" (is "?A \ ?B") proof let ?C = "\t \ timpl_closure_set (set M) (set TI). analyzed_in t (timpl_closure_set (set M) (set TI))" let ?P = "\T. \t \ set T. timpl_closure_set (set M) (set TI) \\<^sub>c t" let ?Q = "\t. \s \ comp_timpl_closure {t} (set TI'). case Ana s of (K, R) \ ?P K \ ?P R" let ?W = "\t. \t \ set (fst (Ana t)). \f \ funs_term t. \is_Abs f" let ?V = "\t. \s \ comp_timpl_closure (set (fst (Ana t))) (set TI'). \timpl_closure_set (set M) (set TI) \\<^sub>c s" note defs = analyzed_closed_mod_timpls_def analyzed_in_code note 0 = intruder_synth_mod_timpls_is_synth_timpl_closure_set[OF TI', of M] note 1 = timpl_closure_set_is_timpl_closure_union[of _ "set TI"] have 2: "comp_timpl_closure N (set TI') = timpl_closure_set N (set TI)" when "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s N" "finite N" for N::"('fun,'atom,'sets,'lbl) prot_terms" using that timpl_closure_set_timpls_trancl_eq'[of N "set TI"] comp_timpl_closure_is_timpl_closure_set[of N TI'] unfolding TI'[symmetric] by presburger hence 3: "comp_timpl_closure {t} (set TI') \ timpl_closure_set (set M) (set TI)" when t: "t \ set M" "wf\<^sub>t\<^sub>r\<^sub>m t" for t using t timpl_closure_set_mono[of "{t}" "set M"] by simp have ?A when C: ?C unfolding analyzed_closed_mod_timpls_def intruder_synth_mod_timpls_is_synth_timpl_closure_set[OF TI'] list_all_iff Let_def proof (intro ballI) fix t assume t: "t \ set M" show "if ?P (fst (Ana t)) then ?P (snd (Ana t)) else if ?W t then True else if ?V t then True else ?Q t" (is ?R) proof (cases "?P (fst (Ana t))") case True hence "?P (snd (Ana t))" using C timpl_closure_setI[OF t, of "set TI"] prod.exhaust_sel unfolding analyzed_in_def by blast thus ?thesis using True by simp next case False have "?Q t" using 3[OF t] C M_wf t unfolding analyzed_in_def by auto thus ?thesis using False by argo qed qed thus ?A when B: ?B using B analyzed_is_all_analyzed_in by metis have ?C when A: ?A unfolding analyzed_in_def Let_def proof (intro ballI allI impI; elim conjE) fix t K T s assume t: "t \ timpl_closure_set (set M) (set TI)" and s: "s \ set T" and Ana_t: "Ana t = (K, T)" and K: "\k \ set K. timpl_closure_set (set M) (set TI) \\<^sub>c k" obtain m where m: "m \ set M" "t \ timpl_closure m (set TI)" - using timpl_closure_set_is_timpl_closure_union t by moura + using timpl_closure_set_is_timpl_closure_union t by blast show "timpl_closure_set (set M) (set TI) \\<^sub>c s" proof (cases "\k \ set (fst (Ana m)). timpl_closure_set (set M) (set TI) \\<^sub>c k") case True hence *: "\r \ set (snd (Ana m)). timpl_closure_set (set M) (set TI) \\<^sub>c r" using m(1) A unfolding analyzed_closed_mod_timpls_def intruder_synth_mod_timpls_is_synth_timpl_closure_set[OF TI'] list_all_iff Let_def by simp show ?thesis using K s Ana_t A analyzed_closed_mod_timpls_is_analyzed_closed_timpl_closure_set_aux2[OF * m] by simp next case False note F = this have *: "comp_timpl_closure {m} (set TI') = timpl_closure m (set TI)" using 2[of "{m}"] timpl_closureton_is_timpl_closure M_wf m(1) by blast have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set (fst (Ana m)))" using Ana_keys_wf'[of m "fst (Ana m)"] M_wf m(1) surj_pair[of "Ana m"] by fastforce hence **: "comp_timpl_closure (set (fst (Ana m))) (set TI') = timpl_closure_set (set (fst (Ana m))) (set TI)" using 2[of "set (fst (Ana m))"] by blast have ***: "set K \ timpl_closure_set (set (fst (Ana m))) (set TI)" "length K = length (fst (Ana m))" using timpl_closure_Ana_keys_subset[OF m(2) _ Ana_t] timpl_closure_Ana_keys_length_eq[OF m(2) _ Ana_t] surj_pair[of "Ana m"] by fastforce+ show ?thesis proof (cases "?W m") case True hence "fst (Ana t) = fst (Ana m)" using m timpl_closure_Ana_keys_no_Abs_eq_case by fast thus ?thesis using F K Ana_t by simp next case False note F' = this show ?thesis proof (cases "?V m") case True hence "\k \ set K. \timpl_closure_set (set M) (set TI) \\<^sub>c k" using F K Ana_t m s *** unfolding ** by blast thus ?thesis using K F' *** by simp next case False hence "?Q m" using m(1) A F F' unfolding analyzed_closed_mod_timpls_def intruder_synth_mod_timpls_is_synth_timpl_closure_set[OF TI'] list_all_iff Let_def by auto thus ?thesis using * m(2) K s Ana_t unfolding Let_def by auto qed qed qed qed thus ?B when A: ?A using A analyzed_is_all_analyzed_in by metis qed lemma analyzed_closed_mod_timpls'_is_analyzed_timpl_closure_set: fixes M::"('fun,'atom,'sets,'lbl) prot_term list" assumes M_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set M)" shows "analyzed_closed_mod_timpls' M TI \ analyzed (timpl_closure_set (set M) (set TI))" (is "?A \ ?B") proof let ?C = "\t \ timpl_closure_set (set M) (set TI). analyzed_in t (timpl_closure_set (set M) (set TI))" let ?P = "\T. \t \ set T. timpl_closure_set (set M) (set TI) \\<^sub>c t" let ?Q = "\t. \s \ comp_timpl_closure {t} (set TI). case Ana s of (K, R) \ ?P K \ ?P R" let ?W = "\t. \t \ set (fst (Ana t)). \f \ funs_term t. \is_Abs f" note defs = analyzed_closed_mod_timpls'_def analyzed_in_code note 0 = intruder_synth_mod_timpls'_is_synth_timpl_closure_set[of M TI] note 1 = timpl_closure_set_is_timpl_closure_union[of _ "set TI"] have 2: "comp_timpl_closure {t} (set TI) = timpl_closure_set {t} (set TI)" when t: "t \ set M" "wf\<^sub>t\<^sub>r\<^sub>m t" for t using t timpl_closure_set_timpls_trancl_eq[of "{t}" "set TI"] comp_timpl_closure_is_timpl_closure_set[of "{t}"] by blast hence 3: "comp_timpl_closure {t} (set TI) \ timpl_closure_set (set M) (set TI)" when t: "t \ set M" "wf\<^sub>t\<^sub>r\<^sub>m t" for t using t timpl_closure_set_mono[of "{t}" "set M"] by fast have ?A when C: ?C unfolding analyzed_closed_mod_timpls'_def intruder_synth_mod_timpls'_is_synth_timpl_closure_set list_all_iff Let_def proof (intro ballI) fix t assume t: "t \ set M" show "if ?P (fst (Ana t)) then ?P (snd (Ana t)) else if ?W t then True else ?Q t" (is ?R) proof (cases "?P (fst (Ana t))") case True hence "?P (snd (Ana t))" using C timpl_closure_setI[OF t, of "set TI"] prod.exhaust_sel unfolding analyzed_in_def by blast thus ?thesis using True by simp next case False have "?Q t" using 3[OF t] C M_wf t unfolding analyzed_in_def by auto thus ?thesis using False by argo qed qed thus ?A when B: ?B using B analyzed_is_all_analyzed_in by metis have ?C when A: ?A unfolding analyzed_in_def Let_def proof (intro ballI allI impI; elim conjE) fix t K T s assume t: "t \ timpl_closure_set (set M) (set TI)" and s: "s \ set T" and Ana_t: "Ana t = (K, T)" and K: "\k \ set K. timpl_closure_set (set M) (set TI) \\<^sub>c k" obtain m where m: "m \ set M" "t \ timpl_closure m (set TI)" - using timpl_closure_set_is_timpl_closure_union t by moura + using timpl_closure_set_is_timpl_closure_union t by blast show "timpl_closure_set (set M) (set TI) \\<^sub>c s" proof (cases "\k \ set (fst (Ana m)). timpl_closure_set (set M) (set TI) \\<^sub>c k") case True hence *: "\r \ set (snd (Ana m)). timpl_closure_set (set M) (set TI) \\<^sub>c r" using m(1) A unfolding analyzed_closed_mod_timpls'_def intruder_synth_mod_timpls'_is_synth_timpl_closure_set list_all_iff by simp show ?thesis using K s Ana_t A analyzed_closed_mod_timpls_is_analyzed_closed_timpl_closure_set_aux2[OF * m] by simp next case False note F = this have *: "comp_timpl_closure {m} (set TI) = timpl_closure m (set TI)" using 2[OF m(1)] timpl_closureton_is_timpl_closure M_wf m(1) by blast show ?thesis proof (cases "?W m") case True hence "fst (Ana t) = fst (Ana m)" using m timpl_closure_Ana_keys_no_Abs_eq_case by fast thus ?thesis using F K Ana_t by simp next case False hence "?Q m" using m(1) A F unfolding analyzed_closed_mod_timpls'_def intruder_synth_mod_timpls'_is_synth_timpl_closure_set list_all_iff Let_def by auto thus ?thesis using * m(2) K s Ana_t unfolding Let_def by auto qed qed qed thus ?B when A: ?A using A analyzed_is_all_analyzed_in by metis qed end end end diff --git a/thys/Automated_Stateful_Protocol_Verification/Term_Variants.thy b/thys/Automated_Stateful_Protocol_Verification/Term_Variants.thy --- a/thys/Automated_Stateful_Protocol_Verification/Term_Variants.thy +++ b/thys/Automated_Stateful_Protocol_Verification/Term_Variants.thy @@ -1,424 +1,424 @@ (* Title: Term_Variants.thy Author: Andreas Viktor Hess, DTU Author: Sebastian A. Mödersheim, DTU Author: Achim D. Brucker, University of Exeter Author: Anders Schlichtkrull, DTU SPDX-License-Identifier: BSD-3-Clause *) section\Term Variants\ theory Term_Variants imports Stateful_Protocol_Composition_and_Typing.Intruder_Deduction begin fun term_variants where "term_variants P (Var x) = [Var x]" | "term_variants P (Fun f T) = ( let S = product_lists (map (term_variants P) T) in map (Fun f) S@concat (map (\g. map (Fun g) S) (P f)))" inductive term_variants_pred for P where term_variants_Var: "term_variants_pred P (Var x) (Var x)" | term_variants_P: "\length T = length S; \i. i < length T \ term_variants_pred P (T ! i) (S ! i); g \ set (P f)\ \ term_variants_pred P (Fun f T) (Fun g S)" | term_variants_Fun: "\length T = length S; \i. i < length T \ term_variants_pred P (T ! i) (S ! i)\ \ term_variants_pred P (Fun f T) (Fun f S)" lemma term_variants_pred_inv: assumes "term_variants_pred P (Fun f T) (Fun h S)" shows "length T = length S" and "\i. i < length T \ term_variants_pred P (T ! i) (S ! i)" and "f \ h \ h \ set (P f)" using assms by (auto elim: term_variants_pred.cases) lemma term_variants_pred_inv': assumes "term_variants_pred P (Fun f T) t" shows "is_Fun t" and "length T = length (args t)" and "\i. i < length T \ term_variants_pred P (T ! i) (args t ! i)" and "f \ the_Fun t \ the_Fun t \ set (P f)" and "P \ (\_. [])(g := [h]) \ f \ the_Fun t \ f = g \ the_Fun t = h" using assms by (auto elim: term_variants_pred.cases) lemma term_variants_pred_inv'': assumes "term_variants_pred P t (Fun f T)" shows "is_Fun t" and "length T = length (args t)" and "\i. i < length T \ term_variants_pred P (args t ! i) (T ! i)" and "f \ the_Fun t \ f \ set (P (the_Fun t))" and "P \ (\_. [])(g := [h]) \ f \ the_Fun t \ f = h \ the_Fun t = g" using assms by (auto elim: term_variants_pred.cases) lemma term_variants_pred_inv_Var: "term_variants_pred P (Var x) t \ t = Var x" "term_variants_pred P t (Var x) \ t = Var x" by (auto intro: term_variants_Var elim: term_variants_pred.cases) lemma term_variants_pred_inv_const: "term_variants_pred P (Fun c []) t \ ((\g \ set (P c). t = Fun g []) \ (t = Fun c []))" by (auto intro: term_variants_P term_variants_Fun elim: term_variants_pred.cases) lemma term_variants_pred_refl: "term_variants_pred P t t" by (induct t) (auto intro: term_variants_pred.intros) lemma term_variants_pred_refl_inv: assumes st: "term_variants_pred P s t" and P: "\f. \g \ set (P f). f = g" shows "s = t" using st P proof (induction s t rule: term_variants_pred.induct) case (term_variants_Var x) thus ?case by blast next case (term_variants_P T S g f) hence "T ! i = S ! i" when i: "i < length T" for i using i by blast hence "T = S" using term_variants_P.hyps(1) by (simp add: nth_equalityI) thus ?case using term_variants_P.prems term_variants_P.hyps(3) by fast next case (term_variants_Fun T S f) hence "T ! i = S ! i" when i: "i < length T" for i using i by blast hence "T = S" using term_variants_Fun.hyps(1) by (simp add: nth_equalityI) thus ?case by fast qed lemma term_variants_pred_const: assumes "b \ set (P a)" shows "term_variants_pred P (Fun a []) (Fun b [])" using term_variants_P[of "[]" "[]"] assms by simp lemma term_variants_pred_const_cases: "P a \ [] \ term_variants_pred P (Fun a []) t \ (t = Fun a [] \ (\b \ set (P a). t = Fun b []))" "P a = [] \ term_variants_pred P (Fun a []) t \ t = Fun a []" using term_variants_pred_inv_const[of P] by auto lemma term_variants_pred_param: assumes "term_variants_pred P t s" and fg: "f = g \ g \ set (P f)" shows "term_variants_pred P (Fun f (S@t#T)) (Fun g (S@s#T))" proof - have 1: "length (S@t#T) = length (S@s#T)" by simp have "term_variants_pred P (T ! i) (T ! i)" "term_variants_pred P (S ! i) (S ! i)" for i by (metis term_variants_pred_refl)+ hence 2: "term_variants_pred P ((S@t#T) ! i) ((S@s#T) ! i)" for i by (simp add: assms nth_Cons' nth_append) show ?thesis by (metis term_variants_Fun[OF 1 2] term_variants_P[OF 1 2] fg) qed lemma term_variants_pred_Cons: assumes t: "term_variants_pred P t s" and T: "term_variants_pred P (Fun f T) (Fun f S)" and fg: "f = g \ g \ set (P f)" shows "term_variants_pred P (Fun f (t#T)) (Fun g (s#S))" proof - have 1: "length (t#T) = length (s#S)" and "\i. i < length T \ term_variants_pred P (T ! i) (S ! i)" using term_variants_pred_inv[OF T] by simp_all hence 2: "\i. i < length (t#T) \ term_variants_pred P ((t#T) ! i) ((s#S) ! i)" by (metis t One_nat_def diff_less length_Cons less_Suc_eq less_imp_diff_less nth_Cons' zero_less_Suc) show ?thesis using 1 2 fg by (auto intro: term_variants_pred.intros) qed lemma term_variants_pred_dense: fixes P Q::"'a set" and fs gs::"'a list" defines "P_fs x \ if x \ P then fs else []" and "P_gs x \ if x \ P then gs else []" and "Q_fs x \ if x \ Q then fs else []" assumes ut: "term_variants_pred P_fs u t" and g: "g \ Q" "g \ set gs" shows "\s. term_variants_pred P_gs u s \ term_variants_pred Q_fs s t" proof - define F where "F \ \(P::'a set) (fs::'a list) x. if x \ P then fs else []" show ?thesis using ut g P_fs_def unfolding P_gs_def Q_fs_def proof (induction u t arbitrary: g gs rule: term_variants_pred.induct) case (term_variants_Var h x) thus ?case by (auto intro: term_variants_pred.term_variants_Var) next case (term_variants_P T S h' h g gs) note hyps = term_variants_P.hyps(1,2,4,5,6,7) note IH = term_variants_P.hyps(3) have "\s. term_variants_pred (F P gs) (T ! i) s \ term_variants_pred (F Q fs) s (S ! i)" when i: "i < length T" for i using IH[OF i hyps(4,5,6)] unfolding F_def by presburger then obtain U where U: "length T = length U" "\i. i < length T \ term_variants_pred (F P gs) (T ! i) (U ! i)" "length U = length S" "\i. i < length U \ term_variants_pred (F Q fs) (U ! i) (S ! i)" using hyps(1) Skolem_list_nth[of _ "\i s. term_variants_pred (F P gs) (T ! i) s \ term_variants_pred (F Q fs) s (S ! i)"] - by moura + by (metis (no_types)) show ?case using term_variants_pred.term_variants_P[OF U(1,2), of g h] term_variants_pred.term_variants_P[OF U(3,4), of h' g] hyps(3)[unfolded hyps(6)] hyps(4,5) unfolding F_def by force next case (term_variants_Fun T S h' g gs) note hyps = term_variants_Fun.hyps(1,2,4,5,6) note IH = term_variants_Fun.hyps(3) have "\s. term_variants_pred (F P gs) (T ! i) s \ term_variants_pred (F Q fs) s (S ! i)" when i: "i < length T" for i using IH[OF i hyps(3,4,5)] unfolding F_def by presburger then obtain U where U: "length T = length U" "\i. i < length T \ term_variants_pred (F P gs) (T ! i) (U ! i)" "length U = length S" "\i. i < length U \ term_variants_pred (F Q fs) (U ! i) (S ! i)" using hyps(1) Skolem_list_nth[of _ "\i s. term_variants_pred (F P gs) (T ! i) s \ term_variants_pred (F Q fs) s (S ! i)"] - by moura + by (metis (no_types)) thus ?case using term_variants_pred.term_variants_Fun[OF U(1,2)] term_variants_pred.term_variants_Fun[OF U(3,4)] unfolding F_def by meson qed qed lemma term_variants_pred_dense': assumes ut: "term_variants_pred ((\_. [])(a := [b])) u t" shows "\s. term_variants_pred ((\_. [])(a := [c])) u s \ term_variants_pred ((\_. [])(c := [b])) s t" using ut term_variants_pred_dense[of "{a}" "[b]" u t c "{c}" "[c]"] unfolding fun_upd_def by simp lemma term_variants_pred_eq_case: fixes t s::"('a,'b) term" assumes "term_variants_pred P t s" "\f \ funs_term t. P f = []" shows "t = s" using assms proof (induction t s rule: term_variants_pred.induct) case (term_variants_Fun T S f) thus ?case using subtermeq_imp_funs_term_subset[OF Fun_param_in_subterms[OF nth_mem], of _ T f] nth_equalityI[of T S] by blast qed (simp_all add: term_variants_pred_refl) lemma term_variants_pred_subst: assumes "term_variants_pred P t s" shows "term_variants_pred P (t \ \) (s \ \)" using assms proof (induction t s rule: term_variants_pred.induct) case (term_variants_P T S f g) have 1: "length (map (\t. t \ \) T) = length (map (\t. t \ \) S)" using term_variants_P.hyps by simp have 2: "term_variants_pred P ((map (\t. t \ \) T) ! i) ((map (\t. t \ \) S) ! i)" when "i < length (map (\t. t \ \) T)" for i using term_variants_P that by fastforce show ?case using term_variants_pred.term_variants_P[OF 1 2 term_variants_P.hyps(3)] by fastforce next case (term_variants_Fun T S f) have 1: "length (map (\t. t \ \) T) = length (map (\t. t \ \) S)" using term_variants_Fun.hyps by simp have 2: "term_variants_pred P ((map (\t. t \ \) T) ! i) ((map (\t. t \ \) S) ! i)" when "i < length (map (\t. t \ \) T)" for i using term_variants_Fun that by fastforce show ?case using term_variants_pred.term_variants_Fun[OF 1 2] by fastforce qed (simp add: term_variants_pred_refl) lemma term_variants_pred_subst': fixes t s::"('a,'b) term" and \::"('a,'b) subst" assumes "term_variants_pred P (t \ \) s" and "\x \ fv t \ fv s. (\y. \ x = Var y) \ (\f. \ x = Fun f [] \ P f = [])" shows "\u. term_variants_pred P t u \ s = u \ \" using assms proof (induction "t \ \" s arbitrary: t rule: term_variants_pred.induct) case (term_variants_Var x g) thus ?case using term_variants_pred_refl by fast next case (term_variants_P T S g f) show ?case proof (cases t) case (Var x) thus ?thesis using term_variants_P.hyps(4,5) term_variants_P.prems by fastforce next case (Fun h U) hence 1: "h = f" "T = map (\s. s \ \) U" "length U = length T" using term_variants_P.hyps(5) by simp_all hence 2: "T ! i = U ! i \ \" when "i < length T" for i using that by simp have "\x \ fv (U ! i) \ fv (S ! i). (\y. \ x = Var y) \ (\f. \ x = Fun f [] \ P f = [])" when "i < length U" for i using that Fun term_variants_P.prems term_variants_P.hyps(1) 1(3) by force hence IH: "\i < length U. \u. term_variants_pred P (U ! i) u \ S ! i = u \ \" by (metis 1(3) term_variants_P.hyps(3)[OF _ 2]) have "\V. length U = length V \ S = map (\v. v \ \) V \ (\i < length U. term_variants_pred P (U ! i) (V ! i))" using term_variants_P.hyps(1) 1(3) subst_term_list_obtain[OF IH] by metis then obtain V where V: "length U = length V" "S = map (\v. v \ \) V" "\i. i < length U \ term_variants_pred P (U ! i) (V ! i)" - by moura + by blast have "term_variants_pred P (Fun f U) (Fun g V)" by (metis term_variants_pred.term_variants_P[OF V(1,3) term_variants_P.hyps(4)]) moreover have "Fun g S = Fun g V \ \" using V(2) by simp ultimately show ?thesis using term_variants_P.hyps(1,4) Fun 1 by blast qed next case (term_variants_Fun T S f t) show ?case proof (cases t) case (Var x) hence "T = []" "P f = []" using term_variants_Fun.hyps(4) term_variants_Fun.prems by fastforce+ thus ?thesis using term_variants_pred_refl Var term_variants_Fun.hyps(1,4) by fastforce next case (Fun h U) hence 1: "h = f" "T = map (\s. s \ \) U" "length U = length T" using term_variants_Fun.hyps(4) by simp_all hence 2: "T ! i = U ! i \ \" when "i < length T" for i using that by simp have "\x \ fv (U ! i) \ fv (S ! i). (\y. \ x = Var y) \ (\f. \ x = Fun f [] \ P f = [])" when "i < length U" for i using that Fun term_variants_Fun.prems term_variants_Fun.hyps(1) 1(3) by force hence IH: "\i < length U. \u. term_variants_pred P (U ! i) u \ S ! i = u \ \" by (metis 1(3) term_variants_Fun.hyps(3)[OF _ 2 ]) have "\V. length U = length V \ S = map (\v. v \ \) V \ (\i < length U. term_variants_pred P (U ! i) (V ! i))" using term_variants_Fun.hyps(1) 1(3) subst_term_list_obtain[OF IH] by metis then obtain V where V: "length U = length V" "S = map (\v. v \ \) V" "\i. i < length U \ term_variants_pred P (U ! i) (V ! i)" - by moura + by blast have "term_variants_pred P (Fun f U) (Fun f V)" by (metis term_variants_pred.term_variants_Fun[OF V(1,3)]) moreover have "Fun f S = Fun f V \ \" using V(2) by simp ultimately show ?thesis using term_variants_Fun.hyps(1) Fun 1 by blast qed qed lemma term_variants_pred_subst'': assumes "\x \ fv t. term_variants_pred P (\ x) (\ x)" shows "term_variants_pred P (t \ \) (t \ \)" using assms proof (induction t) case (Fun f ts) thus ?case using term_variants_Fun[of "map (\t. t \ \) ts" "map (\t. t \ \) ts" P f] by force qed simp lemma term_variants_pred_iff_in_term_variants: fixes t::"('a,'b) term" shows "term_variants_pred P t s \ s \ set (term_variants P t)" (is "?A t s \ ?B t s") proof define U where "U \ \P (T::('a,'b) term list). product_lists (map (term_variants P) T)" have a: "g \ set (P f) \ set (map (Fun g) (U P T)) \ set (term_variants P (Fun f T))" "set (map (Fun f) (U P T)) \ set (term_variants P (Fun f T))" for f P g and T::"('a,'b) term list" using term_variants.simps(2)[of P f T] unfolding U_def Let_def by auto have b: "\S \ set (U P T). s = Fun f S \ (\g \ set (P f). s = Fun g S)" when "s \ set (term_variants P (Fun f T))" for P T f s using that by (cases "P f") (auto simp add: U_def Let_def) have c: "length T = length S" when "S \ set (U P T)" for S P T using that unfolding U_def by (simp add: in_set_product_lists_length) show "?A t s \ ?B t s" proof (induction t s rule: term_variants_pred.induct) case (term_variants_P T S g f) note hyps = term_variants_P.hyps note IH = term_variants_P.IH have "S \ set (U P T)" using IH hyps(1) product_lists_in_set_nth'[of _ S] unfolding U_def by simp thus ?case using a(1)[of _ P, OF hyps(3)] by auto next case (term_variants_Fun T S f) note hyps = term_variants_Fun.hyps note IH = term_variants_Fun.IH have "S \ set (U P T)" using IH hyps(1) product_lists_in_set_nth'[of _ S] unfolding U_def by simp thus ?case using a(2)[of f P T] by (cases "P f") auto qed (simp add: term_variants_Var) show "?B t s \ ?A t s" proof (induction P t arbitrary: s rule: term_variants.induct) case (2 P f T) obtain S where S: "s = Fun f S \ (\g \ set (P f). s = Fun g S)" "S \ set (U P T)" "length T = length S" - using c b[OF "2.prems"] by moura + using c b[OF "2.prems"] by blast have "\i < length T. term_variants_pred P (T ! i) (S ! i)" using "2.IH" S product_lists_in_set_nth by (fastforce simp add: U_def) thus ?case using S by (auto intro: term_variants_pred.intros) qed (simp add: term_variants_Var) qed lemma term_variants_pred_finite: "finite {s. term_variants_pred P t s}" using term_variants_pred_iff_in_term_variants[of P t] by simp lemma term_variants_pred_fv_eq: assumes "term_variants_pred P s t" shows "fv s = fv t" using assms by (induct rule: term_variants_pred.induct) (metis, metis fv_eq_FunI, metis fv_eq_FunI) lemma (in intruder_model) term_variants_pred_wf_trms: assumes "term_variants_pred P s t" and "\f g. g \ set (P f) \ arity f = arity g" and "wf\<^sub>t\<^sub>r\<^sub>m s" shows "wf\<^sub>t\<^sub>r\<^sub>m t" using assms apply (induction rule: term_variants_pred.induct, simp) by (metis (no_types) wf_trmI wf_trm_arity in_set_conv_nth wf_trm_param_idx)+ lemma term_variants_pred_funs_term: assumes "term_variants_pred P s t" and "f \ funs_term t" shows "f \ funs_term s \ (\g \ funs_term s. f \ set (P g))" using assms proof (induction rule: term_variants_pred.induct) case (term_variants_P T S g h) thus ?case proof (cases "f = g") case False then obtain s where "s \ set S" "f \ funs_term s" using funs_term_subterms_eq(1)[of "Fun g S"] term_variants_P.prems by auto thus ?thesis using term_variants_P.IH term_variants_P.hyps(1) in_set_conv_nth[of s S] by force qed simp next case (term_variants_Fun T S h) thus ?case proof (cases "f = h") case False then obtain s where "s \ set S" "f \ funs_term s" using funs_term_subterms_eq(1)[of "Fun h S"] term_variants_Fun.prems by auto thus ?thesis using term_variants_Fun.IH term_variants_Fun.hyps(1) in_set_conv_nth[of s S] by force qed simp qed fast end diff --git a/thys/Automated_Stateful_Protocol_Verification/Transactions.thy b/thys/Automated_Stateful_Protocol_Verification/Transactions.thy --- a/thys/Automated_Stateful_Protocol_Verification/Transactions.thy +++ b/thys/Automated_Stateful_Protocol_Verification/Transactions.thy @@ -1,898 +1,898 @@ (* Title: Transactions.thy Author: Andreas Viktor Hess, DTU Author: Sebastian A. Mödersheim, DTU Author: Achim D. Brucker, University of Exeter Author: Anders Schlichtkrull, DTU SPDX-License-Identifier: BSD-3-Clause *) section\Protocol Transactions\ theory Transactions imports Stateful_Protocol_Composition_and_Typing.Typed_Model Stateful_Protocol_Composition_and_Typing.Labeled_Stateful_Strands begin subsection \Definitions\ datatype 'b prot_atom = is_Atom: Atom 'b | Value | SetType | AttackType | Bottom | OccursSecType | AbsValue datatype ('a,'b,'c,'d) prot_fun = Fu (the_Fu: 'a) | Set (the_Set: 'c) | Val (the_Val: "nat") | Abs (the_Abs: "'c set") | Attack (the_Attack_label: "'d strand_label") | Pair | PubConst (the_PubConst_type: "'b prot_atom") nat | OccursFact | OccursSec definition "is_Fun_Set t \ is_Fun t \ args t = [] \ is_Set (the_Fun t)" definition "is_Fun_Attack t \ is_Fun t \ args t = [] \ is_Attack (the_Fun t)" definition "is_PubConstValue f \ is_PubConst f \ the_PubConst_type f = Value" abbreviation occurs where "occurs t \ Fun OccursFact [Fun OccursSec [], t]" type_synonym ('a,'b,'c,'d) prot_term_type = "(('a,'b,'c,'d) prot_fun,'b prot_atom) term_type" type_synonym ('a,'b,'c,'d) prot_var = "('a,'b,'c,'d) prot_term_type \ nat" type_synonym ('a,'b,'c,'d) prot_term = "(('a,'b,'c,'d) prot_fun,('a,'b,'c,'d) prot_var) term" type_synonym ('a,'b,'c,'d) prot_terms = "('a,'b,'c,'d) prot_term set" type_synonym ('a,'b,'c,'d) prot_subst = "(('a,'b,'c,'d) prot_fun, ('a,'b,'c,'d) prot_var) subst" type_synonym ('a,'b,'c,'d) prot_strand_step = "(('a,'b,'c,'d) prot_fun, ('a,'b,'c,'d) prot_var, 'd) labeled_stateful_strand_step" type_synonym ('a,'b,'c,'d) prot_strand = "('a,'b,'c,'d) prot_strand_step list" type_synonym ('a,'b,'c,'d) prot_constr = "('a,'b,'c,'d) prot_strand_step list" (* TODO: rename transaction_decl *) datatype ('a,'b,'c,'d) prot_transaction = Transaction (transaction_decl: "unit \ (('a,'b,'c,'d) prot_var \ 'a set) list") (transaction_fresh: "('a,'b,'c,'d) prot_var list") (transaction_receive: "('a,'b,'c,'d) prot_strand") (transaction_checks: "('a,'b,'c,'d) prot_strand") (transaction_updates: "('a,'b,'c,'d) prot_strand") (transaction_send: "('a,'b,'c,'d) prot_strand") definition transaction_strand where "transaction_strand T \ transaction_receive T@transaction_checks T@ transaction_updates T@transaction_send T" fun transaction_proj where "transaction_proj l (Transaction A B C D E F) = ( let f = proj l in Transaction A B (f C) (f D) (f E) (f F))" fun transaction_star_proj where "transaction_star_proj (Transaction A B C D E F) = ( let f = filter has_LabelS in Transaction A B (f C) (f D) (f E) (f F))" abbreviation fv_transaction where "fv_transaction T \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)" abbreviation bvars_transaction where "bvars_transaction T \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)" abbreviation vars_transaction where "vars_transaction T \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)" abbreviation trms_transaction where "trms_transaction T \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)" abbreviation setops_transaction where "setops_transaction T \ setops\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T))" definition wellformed_transaction where "wellformed_transaction T \ list_all is_Receive (unlabel (transaction_receive T)) \ list_all is_Check_or_Assignment (unlabel (transaction_checks T)) \ list_all is_Update (unlabel (transaction_updates T)) \ list_all is_Send (unlabel (transaction_send T)) \ distinct (map fst (transaction_decl T ())) \ distinct (transaction_fresh T) \ set (transaction_fresh T) \ fst ` set (transaction_decl T ()) = {} \ set (transaction_fresh T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) = {} \ set (transaction_fresh T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) = {} \ set (transaction_fresh T) \ bvars_transaction T = {} \ fv_transaction T \ bvars_transaction T = {} \ wf'\<^sub>s\<^sub>s\<^sub>t (fst ` set (transaction_decl T ()) \ set (transaction_fresh T)) (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)))" type_synonym ('a,'b,'c,'d) prot = "('a,'b,'c,'d) prot_transaction list" abbreviation Var_Value_term ("\_: value\\<^sub>v") where "\n: value\\<^sub>v \ Var (Var Value, n)::('a,'b,'c,'d) prot_term" abbreviation Var_SetType_term ("\_: SetType\\<^sub>v") where "\n: SetType\\<^sub>v \ Var (Var SetType, n)::('a,'b,'c,'d) prot_term" abbreviation Var_AttackType_term ("\_: AttackType\\<^sub>v") where "\n: AttackType\\<^sub>v \ Var (Var AttackType, n)::('a,'b,'c,'d) prot_term" abbreviation Var_Atom_term ("\_: _\\<^sub>v") where "\n: a\\<^sub>v \ Var (Var (Atom a), n)::('a,'b,'c,'d) prot_term" abbreviation Var_Comp_Fu_term ("\_: _\_\\\<^sub>v") where "\n: f\T\\\<^sub>v \ Var (Fun (Fu f) T, n)::('a,'b,'c,'d) prot_term" abbreviation TAtom_Atom_term ("\_\\<^sub>\\<^sub>a") where "\a\\<^sub>\\<^sub>a \ Var (Atom a)::('a,'b,'c,'d) prot_term_type" abbreviation TComp_Fu_term ("\_ _\\<^sub>\") where "\f T\\<^sub>\ \ Fun (Fu f) T::('a,'b,'c,'d) prot_term_type" abbreviation Fun_Fu_term ("\_ _\\<^sub>t") where "\f T\\<^sub>t \ Fun (Fu f) T::('a,'b,'c,'d) prot_term" abbreviation Fun_Fu_const_term ("\_\\<^sub>c") where "\c\\<^sub>c \ Fun (Fu c) []::('a,'b,'c,'d) prot_term" abbreviation Fun_Set_const_term ("\_\\<^sub>s") where "\f\\<^sub>s \ Fun (Set f) []::('a,'b,'c,'d) prot_term" abbreviation Fun_Set_composed_term ("\_\_\\\<^sub>s") where "\f\T\\\<^sub>s \ Fun (Set f) T::('a,'b,'c,'d) prot_term" abbreviation Fun_Abs_const_term ("\_\\<^sub>a\<^sub>b\<^sub>s") where "\a\\<^sub>a\<^sub>b\<^sub>s \ Fun (Abs a) []::('a,'b,'c,'d) prot_term" abbreviation Fun_Attack_const_term ("attack\_\") where "attack\n\ \ Fun (Attack n) []::('a,'b,'c,'d) prot_term" abbreviation prot_transaction1 ("transaction\<^sub>1 _ _ new _ _ _") where "transaction\<^sub>1 (S1::('a,'b,'c,'d) prot_strand) S2 new (B::('a,'b,'c,'d) prot_term list) S3 S4 \ Transaction (\(). []) (map the_Var B) S1 S2 S3 S4" abbreviation prot_transaction2 ("transaction\<^sub>2 _ _ _ _") where "transaction\<^sub>2 (S1::('a,'b,'c,'d) prot_strand) S2 S3 S4 \ Transaction (\(). []) [] S1 S2 S3 S4" subsection \Lemmata\ lemma prot_atom_UNIV: "(UNIV::'b prot_atom set) = range Atom \ {Value, SetType, AttackType, Bottom, OccursSecType, AbsValue}" proof - have "a \ range Atom \ a = Value \ a = SetType \ a = AttackType \ a = Bottom \ a = OccursSecType \ a = AbsValue" for a::"'b prot_atom" by (cases a) auto thus ?thesis by auto qed instance prot_atom::(finite) finite by intro_classes (simp add: prot_atom_UNIV) instantiation prot_atom::(enum) enum begin definition "enum_prot_atom == map Atom enum_class.enum@[Value, SetType, AttackType, Bottom, OccursSecType, AbsValue]" definition "enum_all_prot_atom P == list_all P (map Atom enum_class.enum@[Value, SetType, AttackType, Bottom, OccursSecType, AbsValue])" definition "enum_ex_prot_atom P == list_ex P (map Atom enum_class.enum@[Value, SetType, AttackType, Bottom, OccursSecType, AbsValue])" instance proof intro_classes have *: "set (map Atom (enum_class.enum::'a list)) = range Atom" "distinct (enum_class.enum::'a list)" using UNIV_enum enum_distinct by auto show "(UNIV::'a prot_atom set) = set enum_class.enum" using *(1) by (simp add: prot_atom_UNIV enum_prot_atom_def) have "set (map Atom enum_class.enum) \ set [Value, SetType, AttackType, Bottom, OccursSecType, AbsValue] = {}" by auto moreover have "inj_on Atom (set (enum_class.enum::'a list))" unfolding inj_on_def by auto hence "distinct (map Atom (enum_class.enum::'a list))" by (metis *(2) distinct_map) ultimately show "distinct (enum_class.enum::'a prot_atom list)" by (simp add: enum_prot_atom_def) have "Ball UNIV P \ Ball (range Atom) P \ Ball {Value, SetType, AttackType, Bottom, OccursSecType, AbsValue} P" for P::"'a prot_atom \ bool" by (metis prot_atom_UNIV UNIV_I UnE) thus "enum_class.enum_all P = Ball (UNIV::'a prot_atom set) P" for P using *(1) Ball_set[of "map Atom enum_class.enum" P] by (auto simp add: enum_all_prot_atom_def) have "Bex UNIV P \ Bex (range Atom) P \ Bex {Value, SetType, AttackType, Bottom, OccursSecType, AbsValue} P" for P::"'a prot_atom \ bool" by (metis prot_atom_UNIV UNIV_I UnE) thus "enum_class.enum_ex P = Bex (UNIV::'a prot_atom set) P" for P using *(1) Bex_set[of "map Atom enum_class.enum" P] by (auto simp add: enum_ex_prot_atom_def) qed end lemma wellformed_transaction_cases: assumes "wellformed_transaction T" shows "(l,x) \ set (transaction_receive T) \ \t. x = receive\t\" (is "?A \ ?A'") "(l,x) \ set (transaction_checks T) \ (\ac t s. x = \ac: t \ s\) \ (\ac t s. x = \ac: t \ s\) \ (\X F G. x = \X\\\: F \\: G\)" (is "?B \ ?B'") "(l,x) \ set (transaction_updates T) \ (\t s. x = insert\t,s\) \ (\t s. x = delete\t,s\)" (is "?C \ ?C'") "(l,x) \ set (transaction_send T) \ \t. x = send\t\" (is "?D \ ?D'") proof - have a: "list_all is_Receive (unlabel (transaction_receive T))" "list_all is_Check_or_Assignment (unlabel (transaction_checks T))" "list_all is_Update (unlabel (transaction_updates T))" "list_all is_Send (unlabel (transaction_send T))" using assms unfolding wellformed_transaction_def by metis+ note b = Ball_set unlabel_in note c = stateful_strand_step.collapse show "?A \ ?A'" by (metis (mono_tags, lifting) a(1) b c(2)) show "?B \ ?B'" by (metis (no_types, lifting) a(2) b c(3,6,7)) show "?C \ ?C'" by (metis (mono_tags, lifting) a(3) b c(4,5)) show "?D \ ?D'" by (metis (mono_tags, lifting) a(4) b c(1)) qed lemma wellformed_transaction_unlabel_cases: assumes "wellformed_transaction T" shows "x \ set (unlabel (transaction_receive T)) \ \t. x = receive\t\" (is "?A \ ?A'") "x \ set (unlabel (transaction_checks T)) \ (\ac t s. x = \ac: t \ s\) \ (\ac t s. x = \ac: t \ s\) \ (\X F G. x = \X\\\: F \\: G\)" (is "?B \ ?B'") "x \ set (unlabel (transaction_updates T)) \ (\t s. x = insert\t,s\) \ (\t s. x = delete\t,s\)" (is "?C \ ?C'") "x \ set (unlabel (transaction_send T)) \ \t. x = send\t\" (is "?D \ ?D'") proof - have a: "list_all is_Receive (unlabel (transaction_receive T))" "list_all is_Check_or_Assignment (unlabel (transaction_checks T))" "list_all is_Update (unlabel (transaction_updates T))" "list_all is_Send (unlabel (transaction_send T))" using assms unfolding wellformed_transaction_def by metis+ note b = Ball_set note c = stateful_strand_step.collapse show "?A \ ?A'" by (metis (mono_tags, lifting) a(1) b c(2)) show "?B \ ?B'" by (metis (no_types, lifting) a(2) b c(3,6,7)) show "?C \ ?C'" by (metis (mono_tags, lifting) a(3) b c(4,5)) show "?D \ ?D'" by (metis (mono_tags, lifting) a(4) b c(1)) qed lemma transaction_strand_subsets[simp]: "set (transaction_receive T) \ set (transaction_strand T)" "set (transaction_checks T) \ set (transaction_strand T)" "set (transaction_updates T) \ set (transaction_strand T)" "set (transaction_send T) \ set (transaction_strand T)" "set (unlabel (transaction_receive T)) \ set (unlabel (transaction_strand T))" "set (unlabel (transaction_checks T)) \ set (unlabel (transaction_strand T))" "set (unlabel (transaction_updates T)) \ set (unlabel (transaction_strand T))" "set (unlabel (transaction_send T)) \ set (unlabel (transaction_strand T))" unfolding transaction_strand_def unlabel_def by force+ lemma transaction_strand_subst_subsets[simp]: "set (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ set (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "set (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ set (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "set (transaction_updates T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ set (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "set (transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ set (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "set (unlabel (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ set (unlabel (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" "set (unlabel (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ set (unlabel (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" "set (unlabel (transaction_updates T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ set (unlabel (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" "set (unlabel (transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ set (unlabel (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" unfolding transaction_strand_def unlabel_def subst_apply_labeled_stateful_strand_def by force+ lemma transaction_strand_dual_unfold: defines "f \ \S. dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S" shows "f (transaction_strand T) = f (transaction_receive T)@f (transaction_checks T)@ f (transaction_updates T)@f (transaction_send T)" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_append unfolding f_def transaction_strand_def by auto lemma transaction_strand_unlabel_dual_unfold: defines "f \ \S. unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S)" shows "f (transaction_strand T) = f (transaction_receive T)@f (transaction_checks T)@ f (transaction_updates T)@f (transaction_send T)" using unlabel_append dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_append unfolding f_def transaction_strand_def by auto lemma transaction_dual_subst_unfold: "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by (simp add: transaction_strand_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_append subst_lsst_append) lemma transaction_dual_subst_unlabel_unfold: "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) = unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))@ unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))@ unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))@ unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" by (simp add: transaction_dual_subst_unfold unlabel_append) lemma trms_transaction_unfold: "trms_transaction T = trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T) \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" by (metis trms\<^sub>s\<^sub>s\<^sub>t_append unlabel_append append_assoc transaction_strand_def) lemma trms_transaction_subst_unfold: "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by (metis trms\<^sub>s\<^sub>s\<^sub>t_append unlabel_append append_assoc transaction_strand_def subst_lsst_append) lemma vars_transaction_unfold: "vars_transaction T = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" by (metis vars\<^sub>s\<^sub>s\<^sub>t_append unlabel_append append_assoc transaction_strand_def) lemma vars_transaction_subst_unfold: "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by (metis vars\<^sub>s\<^sub>s\<^sub>t_append unlabel_append append_assoc transaction_strand_def subst_lsst_append) lemma fv_transaction_unfold: "fv_transaction T = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" by (metis fv\<^sub>s\<^sub>s\<^sub>t_append unlabel_append append_assoc transaction_strand_def) lemma fv_transaction_subst_unfold: "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by (metis fv\<^sub>s\<^sub>s\<^sub>t_append unlabel_append append_assoc transaction_strand_def subst_lsst_append) lemma bvars_transaction_unfold: "bvars_transaction T = bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T) \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" by (metis bvars\<^sub>s\<^sub>s\<^sub>t_append unlabel_append append_assoc transaction_strand_def) lemma bvars_transaction_subst_unfold: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by (metis bvars\<^sub>s\<^sub>s\<^sub>t_append unlabel_append append_assoc transaction_strand_def subst_lsst_append) lemma bvars_wellformed_transaction_unfold: assumes "wellformed_transaction T" shows "bvars_transaction T = bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" (is ?A) and "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) = {}" (is ?B) and "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T) = {}" (is ?C) and "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) = {}" (is ?D) proof - have 0: "list_all is_Receive (unlabel (transaction_receive T))" "list_all is_Update (unlabel (transaction_updates T))" "list_all is_Send (unlabel (transaction_send T))" using assms unfolding wellformed_transaction_def by metis+ have "filter is_NegChecks (unlabel (transaction_receive T)) = []" "filter is_NegChecks (unlabel (transaction_updates T)) = []" "filter is_NegChecks (unlabel (transaction_send T)) = []" using list_all_filter_nil[OF 0(1), of is_NegChecks] list_all_filter_nil[OF 0(2), of is_NegChecks] list_all_filter_nil[OF 0(3), of is_NegChecks] stateful_strand_step.distinct_disc(11,21,29,35,39,41) by blast+ thus ?A ?B ?C ?D using bvars_transaction_unfold[of T] bvars\<^sub>s\<^sub>s\<^sub>t_NegChecks[of "unlabel (transaction_receive T)"] bvars\<^sub>s\<^sub>s\<^sub>t_NegChecks[of "unlabel (transaction_updates T)"] bvars\<^sub>s\<^sub>s\<^sub>t_NegChecks[of "unlabel (transaction_send T)"] by (metis bvars\<^sub>s\<^sub>s\<^sub>t_def UnionE emptyE list.set(1) list.simps(8) subsetI subset_Un_eq sup_commute)+ qed lemma transaction_strand_memberD[dest]: assumes "x \ set (transaction_strand T)" shows "x \ set (transaction_receive T) \ x \ set (transaction_checks T) \ x \ set (transaction_updates T) \ x \ set (transaction_send T)" using assms by (simp add: transaction_strand_def) lemma transaction_strand_unlabel_memberD[dest]: assumes "x \ set (unlabel (transaction_strand T))" shows "x \ set (unlabel (transaction_receive T)) \ x \ set (unlabel (transaction_checks T)) \ x \ set (unlabel (transaction_updates T)) \ x \ set (unlabel (transaction_send T))" using assms by (simp add: unlabel_def transaction_strand_def) lemma wellformed_transaction_strand_memberD[dest]: assumes "wellformed_transaction T" and "(l,x) \ set (transaction_strand T)" shows "x = receive\ts\ \ (l,x) \ set (transaction_receive T)" (is "?A \ ?A'") "x = select\t,s\ \ (l,x) \ set (transaction_checks T)" (is "?B \ ?B'") "x = \t == s\ \ (l,x) \ set (transaction_checks T)" (is "?C \ ?C'") "x = \t in s\ \ (l,x) \ set (transaction_checks T)" (is "?D \ ?D'") "x = \X\\\: F \\: G\ \ (l,x) \ set (transaction_checks T)" (is "?E \ ?E'") "x = insert\t,s\ \ (l,x) \ set (transaction_updates T)" (is "?F \ ?F'") "x = delete\t,s\ \ (l,x) \ set (transaction_updates T)" (is "?G \ ?G'") "x = send\ts\ \ (l,x) \ set (transaction_send T)" (is "?H \ ?H'") "x = \ac: t \ s\ \ (l,x) \ set (transaction_checks T)" (is "?I \ ?I'") "x = \ac: t \ s\ \ (l,x) \ set (transaction_checks T)" (is "?J \ ?J'") proof - have "(l,x) \ set (transaction_receive T) \ (l,x) \ set (transaction_checks T) \ (l,x) \ set (transaction_updates T) \ (l,x) \ set (transaction_send T)" using assms(2) by auto thus "?A \ ?A'" "?B \ ?B'" "?C \ ?C'" "?D \ ?D'" "?E \ ?E'" "?F \ ?F'" "?G \ ?G'" "?H \ ?H'" "?I \ ?I'" "?J \ ?J'" using wellformed_transaction_cases[OF assms(1)] by fast+ qed lemma wellformed_transaction_strand_unlabel_memberD[dest]: assumes "wellformed_transaction T" and "x \ set (unlabel (transaction_strand T))" shows "x = receive\ts\ \ x \ set (unlabel (transaction_receive T))" (is "?A \ ?A'") "x = select\t,s\ \ x \ set (unlabel (transaction_checks T))" (is "?B \ ?B'") "x = \t == s\ \ x \ set (unlabel (transaction_checks T))" (is "?C \ ?C'") "x = \t in s\ \ x \ set (unlabel (transaction_checks T))" (is "?D \ ?D'") "x = \X\\\: F \\: G\ \ x \ set (unlabel (transaction_checks T))" (is "?E \ ?E'") "x = insert\t,s\ \ x \ set (unlabel (transaction_updates T))" (is "?F \ ?F'") "x = delete\t,s\ \ x \ set (unlabel (transaction_updates T))" (is "?G \ ?G'") "x = send\ts\ \ x \ set (unlabel (transaction_send T))" (is "?H \ ?H'") "x = \ac: t \ s\ \ x \ set (unlabel (transaction_checks T))" (is "?I \ ?I'") "x = \ac: t \ s\ \ x \ set (unlabel (transaction_checks T))" (is "?J \ ?J'") proof - have "x \ set (unlabel (transaction_receive T)) \ x \ set (unlabel (transaction_checks T)) \ x \ set (unlabel (transaction_updates T)) \ x \ set (unlabel (transaction_send T))" using assms(2) by auto thus "?A \ ?A'" "?B \ ?B'" "?C \ ?C'" "?D \ ?D'" "?E \ ?E'" "?F \ ?F'" "?G \ ?G'" "?H \ ?H'" "?I \ ?I'" "?J \ ?J'" using wellformed_transaction_unlabel_cases[OF assms(1)] by fast+ qed lemma wellformed_transaction_send_receive_trm_cases: assumes T: "wellformed_transaction T" shows "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ \ts. t \ set ts \ receive\ts\ \ set (unlabel (transaction_receive T))" and "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \ \ts. t \ set ts \ send\ts\ \ set (unlabel (transaction_send T))" using wellformed_transaction_unlabel_cases(1,4)[OF T] trms\<^sub>s\<^sub>s\<^sub>t_in[of t "unlabel (transaction_receive T)"] trms\<^sub>s\<^sub>s\<^sub>t_in[of t "unlabel (transaction_send T)"] by fastforce+ lemma wellformed_transaction_send_receive_subst_trm_cases: assumes T: "wellformed_transaction T" shows "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \\<^sub>s\<^sub>e\<^sub>t \ \ \ts. t \ set ts \ receive\ts\ \ set (unlabel (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" and "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \ \ts. t \ set ts \ send\ts\ \ set (unlabel (transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" proof - assume "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \\<^sub>s\<^sub>e\<^sub>t \" then obtain s where s: "s \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" "t = s \ \" by blast hence "\ts. s \ set ts \ receive\ts\ \ set (unlabel (transaction_receive T))" using wellformed_transaction_send_receive_trm_cases(1)[OF T] by simp thus "\ts. t \ set ts \ receive\ts\ \ set (unlabel (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using s(2) unlabel_subst[of _ \] subst_set_map[of s _ \] stateful_strand_step_subst_inI(2)[of _ "unlabel (transaction_receive T)" \] by metis next assume "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \" then obtain s where s: "s \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "t = s \ \" by blast hence "\ts. s \ set ts \ send\ts\ \ set (unlabel (transaction_send T))" using wellformed_transaction_send_receive_trm_cases(2)[OF T] by simp thus "\ts. t \ set ts \ send\ts\ \ set (unlabel (transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using s(2) unlabel_subst[of _ \] subst_set_map[of s _ \] stateful_strand_step_subst_inI(1)[of _ "unlabel (transaction_send T)" \] by metis qed lemma wellformed_transaction_send_receive_fv_subset: assumes T: "wellformed_transaction T" shows "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ fv t \ fv_transaction T" (is "?A \ ?A'") and "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \ fv t \ fv_transaction T" (is "?B \ ?B'") proof - let ?P = "\ts. t \ set ts \ receive\ts\ \ set (unlabel (transaction_strand T))" let ?Q = "\ts. t \ set ts \ send\ts\ \ set (unlabel (transaction_strand T))" have *: "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ ?P" "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \ ?Q" using wellformed_transaction_send_receive_trm_cases[OF T, of t] unfolding transaction_strand_def by force+ show "?A \ ?A'" using *(1) by (induct "transaction_strand T") (simp, force) show "?B \ ?B'" using *(2) by (induct "transaction_strand T") (simp, force) qed lemma dual_wellformed_transaction_ident_cases[dest]: "list_all is_Assignment (unlabel S) \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S = S" "list_all is_Check (unlabel S) \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S = S" "list_all is_Update (unlabel S) \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S = S" proof (induction S) case (Cons s S) - obtain l x where s: "s = (l,x)" by moura + obtain l x where s: "s = (l,x)" by force { case 1 thus ?case using Cons s unfolding unlabel_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by (cases x) auto } { case 2 thus ?case using Cons s unfolding unlabel_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by (cases x) auto } { case 3 thus ?case using Cons s unfolding unlabel_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by (cases x) auto } qed simp_all lemma wellformed_transaction_wf\<^sub>s\<^sub>s\<^sub>t: fixes T::"('a, 'b, 'c, 'd) prot_transaction" assumes T: "wellformed_transaction T" shows "wf'\<^sub>s\<^sub>s\<^sub>t (fst ` set (transaction_decl T ()) \ set (transaction_fresh T)) (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)))" and "fv_transaction T \ bvars_transaction T = {}" using T unfolding wellformed_transaction_def by simp_all lemma dual_wellformed_transaction_ident_cases'[dest]: assumes "wellformed_transaction T" shows "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) = transaction_checks T" (is ?A) "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T) = transaction_updates T" (is ?B) proof - have "list_all is_Check_or_Assignment (unlabel (transaction_checks T))" "list_all is_Update (unlabel (transaction_updates T))" using assms is_Check_or_Assignment_iff unfolding wellformed_transaction_def by auto thus ?A ?B using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_list_all_same(9)[of "transaction_checks T"] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_list_all_same(8)[of "transaction_updates T"] by (blast, blast) qed lemma dual_transaction_strand: assumes "wellformed_transaction T" shows "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)@transaction_checks T@ transaction_updates T@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" using dual_wellformed_transaction_ident_cases'[OF assms] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_append unfolding transaction_strand_def by metis lemma dual_unlabel_transaction_strand: assumes "wellformed_transaction T" shows "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)) = (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)))@(unlabel (transaction_checks T))@ (unlabel (transaction_updates T))@(unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)))" using dual_transaction_strand[OF assms] by (simp add: unlabel_def) lemma dual_transaction_strand_subst: assumes "wellformed_transaction T" shows "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)@transaction_checks T@ transaction_updates T@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" proof - have "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst by metis thus ?thesis using dual_transaction_strand[OF assms] by argo qed lemma dual_transaction_ik_is_transaction_send: assumes "wellformed_transaction T" shows "ik\<^sub>s\<^sub>s\<^sub>t (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T))) = trms\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_send T))" (is "?A = ?B") proof - { fix t assume "t \ ?A" then obtain ts where ts: "t \ set ts" "receive\ts\ \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)))" by (auto simp add: ik\<^sub>s\<^sub>s\<^sub>t_def) hence *: "send\ts\ \ set (unlabel (transaction_strand T))" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(1) by metis have "t \ ?B" using ts(1) wellformed_transaction_strand_unlabel_memberD(8)[OF assms *, of ts] by force } moreover { fix t assume "t \ ?B" then obtain ts where ts: "t \ set ts" "send\ts\ \ set (unlabel (transaction_send T))" using wellformed_transaction_unlabel_cases(4)[OF assms] by fastforce hence "receive\ts\ \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)))" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(1) by metis hence "receive\ts\ \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)))" using dual_unlabel_transaction_strand[OF assms] by simp hence "t \ ?A" using ts(1) by (auto simp add: ik\<^sub>s\<^sub>s\<^sub>t_def) } ultimately show "?A = ?B" by auto qed lemma dual_transaction_ik_is_transaction_send': fixes \::"('a,'b,'c,'d) prot_subst" assumes "wellformed_transaction T" shows "ik\<^sub>s\<^sub>s\<^sub>t (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) = trms\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \" (is "?A = ?B") using dual_transaction_ik_is_transaction_send[OF assms] subst_lsst_unlabel[of "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)" \] ik\<^sub>s\<^sub>s\<^sub>t_subst[of "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T))" \] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst[of "transaction_strand T" \] by auto lemma db\<^sub>s\<^sub>s\<^sub>t_transaction_prefix_eq: assumes T: "wellformed_transaction T" and S: "prefix S (transaction_receive T@transaction_checks T)" shows "db\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = db\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" proof - let ?T1 = "transaction_receive T" let ?T2 = "transaction_checks T" have *: "prefix (unlabel S) (unlabel (?T1@?T2))" using S prefix_unlabel by blast have "list_all is_Receive (unlabel ?T1)" "list_all is_Check_or_Assignment (unlabel ?T2)" using T by (simp_all add: wellformed_transaction_def) hence "\b \ set (unlabel ?T1). \is_Insert b \ \is_Delete b" "\b \ set (unlabel ?T2). \is_Insert b \ \is_Delete b" by (metis (mono_tags, lifting) Ball_set stateful_strand_step.distinct_disc(16,18), metis (mono_tags, lifting) Ball_set stateful_strand_step.distinct_disc(24,26,33,35,37,39)) hence "\b \ set (unlabel (?T1@?T2)). \is_Insert b \ \is_Delete b" by (auto simp add: unlabel_def) hence "\b \ set (unlabel S). \is_Insert b \ \is_Delete b" using * unfolding prefix_def by fastforce hence "\b \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S) \\<^sub>s\<^sub>s\<^sub>t \). \is_Insert b \ \is_Delete b" proof (induction S) case (Cons a S) then obtain l b where "a = (l,b)" by (metis surj_pair) thus ?case using Cons unfolding dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def unlabel_def subst_apply_stateful_strand_def by (cases b) auto qed simp hence **: "\b \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))). \is_Insert b \ \is_Delete b" by (metis dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst_unlabel) show ?thesis using db\<^sub>s\<^sub>s\<^sub>t_no_upd_append[OF **] unlabel_append unfolding db\<^sub>s\<^sub>s\<^sub>t_def by metis qed lemma db\<^sub>l\<^sub>s\<^sub>s\<^sub>t_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_set_ex: assumes "d \ set (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ D)" "\t u. insert\t,u\ \ set (unlabel A) \ (\s. u = Fun (Set s) [])" "\t u. delete\t,u\ \ set (unlabel A) \ (\s. u = Fun (Set s) [])" "\d \ set D. \s. snd d = Fun (Set s) []" shows "\s. snd d = Fun (Set s) []" using assms proof (induction A arbitrary: D) case (Cons a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) have 1: "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = receive\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\#unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" when "b = send\ts\" for ts by (simp add: a that subst_lsst_unlabel_cons) have 2: "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = send\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\#unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" when "b = receive\ts\" for ts by (simp add: a that subst_lsst_unlabel_cons) have 3: "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)#unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" when "\ts. b = send\ts\ \ b = receive\ts\" using a that dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_Cons subst_lsst_unlabel_cons[of l b] by (cases b) auto show ?case using 1 2 3 a Cons by (cases b) fastforce+ qed simp lemma is_Fun_SetE[elim]: assumes t: "is_Fun_Set t" obtains s where "t = Fun (Set s) []" proof (cases t) case (Fun f T) - then obtain s where "f = Set s" using t unfolding is_Fun_Set_def by (cases f) moura+ + then obtain s where "f = Set s" using t unfolding is_Fun_Set_def by (cases f) force+ moreover have "T = []" using Fun t unfolding is_Fun_Set_def by (cases T) auto ultimately show ?thesis using Fun that by fast qed (use t is_Fun_Set_def in fast) lemma Fun_Set_InSet_iff: "(u = \a: Var x \ Fun (Set s) []\) \ (is_InSet u \ is_Var (the_elem_term u) \ is_Fun_Set (the_set_term u) \ the_Set (the_Fun (the_set_term u)) = s \ the_Var (the_elem_term u) = x \ the_check u = a)" (is "?A \ ?B") proof show "?A \ ?B" unfolding is_Fun_Set_def by auto assume B: ?B thus ?A proof (cases u) case (InSet b t t') hence "b = a" "t = Var x" "t' = Fun (Set s) []" using B by (simp, fastforce, fastforce) thus ?thesis using InSet by fast qed auto qed lemma Fun_Set_NotInSet_iff: "(u = \Var x not in Fun (Set s) []\) \ (is_NegChecks u \ bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p u = [] \ the_eqs u = [] \ length (the_ins u) = 1 \ is_Var (fst (hd (the_ins u))) \ is_Fun_Set (snd (hd (the_ins u)))) \ the_Set (the_Fun (snd (hd (the_ins u)))) = s \ the_Var (fst (hd (the_ins u))) = x" (is "?A \ ?B") proof show "?A \ ?B" unfolding is_Fun_Set_def by auto assume B: ?B show ?A proof (cases u) case (NegChecks X F F') hence "X = []" "F = []" using B by auto moreover have "fst (hd (the_ins u)) = Var x" "snd (hd (the_ins u)) = Fun (Set s) []" using B is_Fun_SetE[of "snd (hd (the_ins u))"] by (force, fastforce) hence "F' = [(Var x, Fun (Set s) [])]" using NegChecks B by (cases "the_ins u") auto ultimately show ?thesis using NegChecks by fast qed (use B in auto) qed lemma is_Fun_Set_exi: "is_Fun_Set x \ (\s. x = Fun (Set s) [])" by (metis prot_fun.collapse(2) term.collapse(2) prot_fun.disc(11) term.disc(2) term.sel(2,4) is_Fun_Set_def un_Fun1_def) lemma is_Fun_Set_subst: assumes "is_Fun_Set S'" shows "is_Fun_Set (S' \ \)" using assms by (fastforce simp add: is_Fun_Set_def) lemma is_Update_in_transaction_updates: assumes tu: "is_Update t" assumes t: "t \ set (unlabel (transaction_strand TT))" assumes vt: "wellformed_transaction TT" shows "t \ set (unlabel (transaction_updates TT))" using t tu vt unfolding transaction_strand_def wellformed_transaction_def list_all_iff by (auto simp add: unlabel_append) lemma transaction_proj_member: assumes "T \ set P" shows "transaction_proj n T \ set (map (transaction_proj n) P)" using assms by simp lemma transaction_strand_proj: "transaction_strand (transaction_proj n T) = proj n (transaction_strand T)" proof - obtain A B C D E F where "T = Transaction A B C D E F" by (cases T) simp thus ?thesis using transaction_proj.simps[of n A B C D E F] unfolding transaction_strand_def proj_def Let_def by auto qed lemma transaction_proj_decl_eq: "transaction_decl (transaction_proj n T) = transaction_decl T" proof - obtain A B C D E F where "T = Transaction A B C D E F" by (cases T) simp thus ?thesis using transaction_proj.simps[of n A B C D E F] unfolding proj_def Let_def by auto qed lemma transaction_proj_fresh_eq: "transaction_fresh (transaction_proj n T) = transaction_fresh T" proof - obtain A B C D E F where "T = Transaction A B C D E F" by (cases T) simp thus ?thesis using transaction_proj.simps[of n A B C D E F] unfolding proj_def Let_def by auto qed lemma transaction_proj_trms_subset: "trms_transaction (transaction_proj n T) \ trms_transaction T" proof - obtain A B C D E F where "T = Transaction A B C D E F" by (cases T) simp thus ?thesis using transaction_proj.simps[of n A B C D E F] trms\<^sub>s\<^sub>s\<^sub>t_proj_subset(1)[of n] unfolding transaction_fresh_def Let_def transaction_strand_def by auto qed lemma transaction_proj_vars_subset: "vars_transaction (transaction_proj n T) \ vars_transaction T" proof - obtain A B C D E F where "T = Transaction A B C D E F" by (cases T) simp thus ?thesis using transaction_proj.simps[of n A B C D E F] sst_vars_proj_subset(3)[of n "transaction_strand T"] unfolding transaction_fresh_def Let_def transaction_strand_def by simp qed lemma transaction_proj_labels: fixes T::"('a,'b,'c,'d) prot_transaction" shows "list_all (\a. has_LabelN l a \ has_LabelS a) (transaction_strand (transaction_proj l T))" proof - define h where "h \ \a::('a,'b,'c,'d) prot_strand_step. has_LabelN l a \ has_LabelS a" let ?f = "filter h" let ?g = "list_all h" obtain T1 T2 T3 T4 T5 T6 where T: "T = Transaction T1 T2 T3 T4 T5 T6" by (cases T) simp note 0 = transaction_proj.simps[unfolded Let_def, of l T1 T2 T3 T4 T5 T6] show ?thesis using T 0 unfolding list_all_iff proj_def by auto qed lemma transaction_proj_ident_iff: fixes T::"('a,'b,'c,'d) prot_transaction" shows "list_all (\a. has_LabelN l a \ has_LabelS a) (transaction_strand T) \ transaction_proj l T = T" (is "?A \ ?B") proof obtain T1 T2 T3 T4 T5 T6 where T: "T = Transaction T1 T2 T3 T4 T5 T6" by (cases T) simp hence "transaction_strand T = T3@T4@T5@T6" unfolding transaction_strand_def by simp thus "?A \ ?B" using T transaction_proj.simps[unfolded Let_def, of l T1 T2 T3 T4 T5 T6] unfolding list_all_iff proj_def by auto show "?B \ ?A" using transaction_proj_labels[of l T] by presburger qed lemma transaction_proj_idem: fixes T::"('a,'b,'c,'d) prot_transaction" shows "transaction_proj l (transaction_proj l T) = transaction_proj l T" by (meson transaction_proj_ident_iff transaction_proj_labels) lemma transaction_proj_ball_subst: assumes "set Ps = (\n. map (transaction_proj n) P) ` set L" "\p \ set Ps. Q p" shows "\l \ set L. Q (map (transaction_proj l) P)" using assms by auto lemma transaction_star_proj_has_star_labels: "list_all has_LabelS (transaction_strand (transaction_star_proj T))" proof - let ?f = "filter has_LabelS" obtain T1 T2 T3 T4 T5 T6 where T: "T = Transaction T1 T2 T3 T4 T5 T6" by (cases T) simp hence T': "transaction_strand (transaction_star_proj T) = ?f T3@?f T4@?f T5@?f T6" using transaction_star_proj.simps[unfolded Let_def, of T1 T2 T3 T4 T5 T6] unfolding transaction_strand_def by auto show ?thesis using Ball_set unfolding T' by fastforce qed lemma transaction_star_proj_ident_iff: "list_all has_LabelS (transaction_strand T) \ transaction_star_proj T = T" (is "?A \ ?B") proof obtain T1 T2 T3 T4 T5 T6 where T: "T = Transaction T1 T2 T3 T4 T5 T6" by (cases T) simp hence T': "transaction_strand T = T3@T4@T5@T6" unfolding transaction_strand_def by simp show "?A \ ?B" using T T' unfolding list_all_iff by auto show "?B \ ?A" using transaction_star_proj_has_star_labels[of T] by auto qed lemma transaction_star_proj_negates_transaction_proj: "transaction_star_proj (transaction_proj l T) = transaction_star_proj T" (is "?A l T") "k \ l \ transaction_proj k (transaction_proj l T) = transaction_star_proj T" (is "?B \ ?B'") proof - show "?A l T" for l T proof - obtain T1 T2 T3 T4 T5 T6 where T: "T = Transaction T1 T2 T3 T4 T5 T6" by (cases T) simp thus ?thesis by (metis dbproj_def transaction_proj.simps transaction_star_proj.simps proj_dbproj(2)) qed thus "?B \ ?B'" by (metis (no_types) has_LabelS_proj_iff_not_has_LabelN proj_elims_label transaction_star_proj_ident_iff transaction_strand_proj) qed lemma transaction_updates_send_ex_iff: fixes T::"('a,'b,'c,'d) prot_transaction" assumes "list_all is_Receive (unlabel (transaction_receive T))" "list_all is_Check_or_Assignment (unlabel (transaction_checks T))" "list_all is_Update (unlabel (transaction_updates T))" "list_all is_Send (unlabel (transaction_send T))" shows "transaction_updates T \ [] \ transaction_send T \ [] \ list_ex (\a. is_Send (snd a) \ is_Update (snd a)) (transaction_strand T)" proof - define f where "f \ \a::('a,'b,'c,'d) prot_strand_step. is_Send (snd a) \ is_Update (snd a)" have 0: "list_all (\a. \(f a)) (transaction_receive T)" "list_all (\a. \(f a)) (transaction_checks T)" "list_all (\a. (f a)) (transaction_updates T)" "list_all (\a. (f a)) (transaction_send T)" using assms unfolding list_all_iff unlabel_def f_def by auto have 1: "list_ex f (transaction_strand T) \ list_ex f (transaction_updates T) \ list_ex f (transaction_send T)" using 0 unfolding list_all_iff list_ex_iff transaction_strand_def by auto have 2: "A \ [] \ list_ex f A" when "list_all f A" for A using that by (induct A) auto thus ?thesis using 1 2[OF 0(3)] 2[OF 0(4)] unfolding list_all_iff list_ex_iff f_def[symmetric] by meson qed end diff --git a/thys/Banach_Steinhaus/Banach_Steinhaus_Missing.thy b/thys/Banach_Steinhaus/Banach_Steinhaus_Missing.thy --- a/thys/Banach_Steinhaus/Banach_Steinhaus_Missing.thy +++ b/thys/Banach_Steinhaus/Banach_Steinhaus_Missing.thy @@ -1,898 +1,898 @@ (* File: Banach_Steinhaus_Missing.thy Author: Dominique Unruh, University of Tartu Author: Jose Manuel Rodriguez Caballero, University of Tartu *) section \Missing results for the proof of Banach-Steinhaus theorem\ theory Banach_Steinhaus_Missing imports "HOL-Analysis.Bounded_Linear_Function" "HOL-Analysis.Line_Segment" begin subsection \Results missing for the proof of Banach-Steinhaus theorem\ text \ The results proved here are preliminaries for the proof of Banach-Steinhaus theorem using Sokal's approach, but they do not explicitly appear in Sokal's paper \<^cite>\sokal2011really\. \ text\Notation for the norm\ bundle notation_norm begin notation norm ("\_\") end bundle no_notation_norm begin no_notation norm ("\_\") end unbundle notation_norm text\Notation for apply bilinear function\ bundle notation_blinfun_apply begin notation blinfun_apply (infixr "*\<^sub>v" 70) end bundle no_notation_blinfun_apply begin no_notation blinfun_apply (infixr "*\<^sub>v" 70) end unbundle notation_blinfun_apply lemma bdd_above_plus: fixes f::\'a \ real\ assumes \bdd_above (f ` S)\ and \bdd_above (g ` S)\ shows \bdd_above ((\ x. f x + g x) ` S)\ text \ Explanation: If the images of two real-valued functions \<^term>\f\,\<^term>\g\ are bounded above on a set \<^term>\S\, then the image of their sum is bounded on \<^term>\S\. \ proof- obtain M where \\ x. x\S \ f x \ M\ using \bdd_above (f ` S)\ unfolding bdd_above_def by blast obtain N where \\ x. x\S \ g x \ N\ using \bdd_above (g ` S)\ unfolding bdd_above_def by blast have \\ x. x\S \ f x + g x \ M + N\ using \\x. x \ S \ f x \ M\ \\x. x \ S \ g x \ N\ by fastforce thus ?thesis unfolding bdd_above_def by blast qed text\The maximum of two functions\ definition pointwise_max:: "('a \ 'b::ord) \ ('a \ 'b) \ ('a \ 'b)" where \pointwise_max f g = (\x. max (f x) (g x))\ lemma max_Sup_absorb_left: fixes f g::\'a \ real\ assumes \X \ {}\ and \bdd_above (f ` X)\ and \bdd_above (g ` X)\ and \Sup (f ` X) \ Sup (g ` X)\ shows \Sup ((pointwise_max f g) ` X) = Sup (f ` X)\ text \Explanation: For real-valued functions \<^term>\f\ and \<^term>\g\, if the supremum of \<^term>\f\ is greater-equal the supremum of \<^term>\g\, then the supremum of \<^term>\max f g\ equals the supremum of \<^term>\f\. (Under some technical conditions.)\ proof- have y_Sup: \y \ ((\ x. max (f x) (g x)) ` X) \ y \ Sup (f ` X)\ for y proof- assume \y \ ((\ x. max (f x) (g x)) ` X)\ then obtain x where \y = max (f x) (g x)\ and \x \ X\ by blast have \f x \ Sup (f ` X)\ by (simp add: \x \ X\ \bdd_above (f ` X)\ cSUP_upper) moreover have \g x \ Sup (g ` X)\ by (simp add: \x \ X\ \bdd_above (g ` X)\ cSUP_upper) ultimately have \max (f x) (g x) \ Sup (f ` X)\ using \Sup (f ` X) \ Sup (g ` X)\ by auto thus ?thesis by (simp add: \y = max (f x) (g x)\) qed have y_f_X: \y \ f ` X \ y \ Sup ((\ x. max (f x) (g x)) ` X)\ for y proof- assume \y \ f ` X\ then obtain x where \x \ X\ and \y = f x\ by blast have \bdd_above ((\ \. max (f \) (g \)) ` X)\ by (metis (no_types) \bdd_above (f ` X)\ \bdd_above (g ` X)\ bdd_above_image_sup sup_max) moreover have \e > 0 \ \ k \ (\ \. max (f \) (g \)) ` X. y \ k + e\ for e::real using \Sup (f ` X) \ Sup (g ` X)\ by (smt (verit, best) \x \ X\ \y = f x\ imageI) ultimately show ?thesis using \x \ X\ \y = f x\ cSUP_upper by fastforce qed have \Sup ((\ x. max (f x) (g x)) ` X) \ Sup (f ` X)\ using y_Sup by (simp add: \X \ {}\ cSup_least) moreover have \Sup ((\ x. max (f x) (g x)) ` X) \ Sup (f ` X)\ using y_f_X by (metis (mono_tags) cSup_least calculation empty_is_image) ultimately show ?thesis unfolding pointwise_max_def by simp qed lemma max_Sup_absorb_right: fixes f g::\'a \ real\ assumes \X \ {}\ and \bdd_above (f ` X)\ and \bdd_above (g ` X)\ and \Sup (f ` X) \ Sup (g ` X)\ shows \Sup ((pointwise_max f g) ` X) = Sup (g ` X)\ text \ Explanation: For real-valued functions \<^term>\f\ and \<^term>\g\ and a nonempty set \<^term>\X\, such that the \<^term>\f\ and \<^term>\g\ are bounded above on \<^term>\X\, if the supremum of \<^term>\f\ on \<^term>\X\ is lower-equal the supremum of \<^term>\g\ on \<^term>\X\, then the supremum of \<^term>\pointwise_max f g\ on \<^term>\X\ equals the supremum of \<^term>\g\. This is the right analog of @{text max_Sup_absorb_left}. \ proof- have \Sup ((pointwise_max g f) ` X) = Sup (g ` X)\ using assms by (simp add: max_Sup_absorb_left) moreover have \pointwise_max g f = pointwise_max f g\ unfolding pointwise_max_def by auto ultimately show ?thesis by simp qed lemma max_Sup: fixes f g::\'a \ real\ assumes \X \ {}\ and \bdd_above (f ` X)\ and \bdd_above (g ` X)\ shows \Sup ((pointwise_max f g) ` X) = max (Sup (f ` X)) (Sup (g ` X))\ text \ Explanation: Let \<^term>\X\ be a nonempty set. Two supremum over \<^term>\X\ of the maximum of two real-value functions is equal to the maximum of their suprema over \<^term>\X\, provided that the functions are bounded above on \<^term>\X\. \ proof(cases \Sup (f ` X) \ Sup (g ` X)\) case True thus ?thesis by (simp add: assms(1) assms(2) assms(3) max_Sup_absorb_left) next case False have f1: "\ 0 \ Sup (f ` X) + - 1 * Sup (g ` X)" using False by linarith hence "Sup (Banach_Steinhaus_Missing.pointwise_max f g ` X) = Sup (g ` X)" by (simp add: assms(1) assms(2) assms(3) max_Sup_absorb_right) thus ?thesis using f1 by linarith qed lemma identity_telescopic: fixes x :: \_ \ 'a::real_normed_vector\ assumes \x \ l\ shows \(\ N. sum (\ k. x (Suc k) - x k) {n..N}) \ l - x n\ text\ Expression of a limit as a telescopic series. Explanation: If \<^term>\x\ converges to \<^term>\l\ then the sum \<^term>\sum (\ k. x (Suc k) - x k) {n..N}\ converges to \<^term>\l - x n\ as \<^term>\N\ goes to infinity. \ proof- have \(\ p. x (p + Suc n)) \ l\ using \x \ l\ by (rule LIMSEQ_ignore_initial_segment) hence \(\ p. x (Suc n + p)) \ l\ by (simp add: add.commute) hence \(\ p. x (Suc (n + p))) \ l\ by simp hence \(\ t. (- (x n)) + (\ p. x (Suc (n + p))) t ) \ (- (x n)) + l\ using tendsto_add_const_iff by metis hence f1: \(\ p. x (Suc (n + p)) - x n)\ l - x n\ by simp have \sum (\ k. x (Suc k) - x k) {n..n+p} = x (Suc (n+p)) - x n\ for p by (simp add: sum_Suc_diff) moreover have \(\ N. sum (\ k. x (Suc k) - x k) {n..N}) (n + t) = (\ p. sum (\ k. x (Suc k) - x k) {n..n+p}) t\ for t by blast ultimately have \(\ p. (\ N. sum (\ k. x (Suc k) - x k) {n..N}) (n + p)) \ l - x n\ using f1 by simp hence \(\ p. (\ N. sum (\ k. x (Suc k) - x k) {n..N}) (p + n)) \ l - x n\ by (simp add: add.commute) hence \(\ p. (\ N. sum (\ k. x (Suc k) - x k) {n..N}) p) \ l - x n\ using Topological_Spaces.LIMSEQ_offset[where f = "(\ N. sum (\ k. x (Suc k) - x k) {n..N})" and a = "l - x n" and k = n] by blast hence \(\ M. (\ N. sum (\ k. x (Suc k) - x k) {n..N}) M) \ l - x n\ by simp thus ?thesis by blast qed lemma bound_Cauchy_to_lim: assumes \y \ x\ and \\n. \y (Suc n) - y n\ \ c^n\ and \y 0 = 0\ and \c < 1\ shows \\x - y (Suc n)\ \ (c / (1 - c)) * c ^ n\ text\ Inequality about a sequence of approximations assuming that the sequence of differences is bounded by a geometric progression. Explanation: Let \<^term>\y\ be a sequence converging to \<^term>\x\. If \<^term>\y\ satisfies the inequality \\y (Suc n) - y n\ \ c ^ n\ for some \<^term>\c < 1\ and assuming \<^term>\y 0 = 0\ then the inequality \\x - y (Suc n)\ \ (c / (1 - c)) * c ^ n\ holds. \ proof- have \c \ 0\ using \\ n. \y (Suc n) - y n\ \ c^n\ by (metis dual_order.trans norm_ge_zero power_one_right) have norm_1: \norm (\k = Suc n..N. y (Suc k) - y k) \ (c ^ Suc n)/(1 - c)\ for N proof(cases \N < Suc n\) case True hence \\sum (\k. y (Suc k) - y k) {Suc n .. N}\ = 0\ by auto thus ?thesis using \c \ 0\ \c < 1\ by auto next case False hence \N \ Suc n\ by auto have \c^(Suc N) \ 0\ using \c \ 0\ by auto have \1 - c > 0\ by (simp add: \c < 1\) hence \(1 - c)/(1 - c) = 1\ by auto have \\sum (\k. y (Suc k) - y k) {Suc n .. N}\ \ (sum (\k. \y (Suc k) - y k\) {Suc n .. N})\ by (simp add: sum_norm_le) hence \\sum (\k. y (Suc k) - y k) {Suc n .. N}\ \ (sum (power c) {Suc n .. N})\ by (simp add: assms(2) sum_norm_le) hence \(1 - c) * \sum (\k. y (Suc k) - y k) {Suc n .. N}\ \ (1 - c) * (sum (power c) {Suc n .. N})\ using \0 < 1 - c\ mult_le_cancel_left_pos by blast also have \\ = c^(Suc n) - c^(Suc N)\ using Set_Interval.sum_gp_multiplied \Suc n \ N\ by blast also have \\ \ c^(Suc n)\ using \c^(Suc N) \ 0\ by auto finally have \(1 - c) * \\k = Suc n..N. y (Suc k) - y k\ \ c ^ Suc n\ by blast hence \((1 - c) * \\k = Suc n..N. y (Suc k) - y k\)/(1 - c) \ (c ^ Suc n)/(1 - c)\ using \0 < 1 - c\ divide_le_cancel by fastforce thus \\\k = Suc n..N. y (Suc k) - y k\ \ (c ^ Suc n)/(1 - c)\ using \0 < 1 - c\ by auto qed have \(\ N. (sum (\k. y (Suc k) - y k) {Suc n .. N})) \ x - y (Suc n)\ by (metis (no_types) \y \ x\ identity_telescopic) hence \(\ N. \sum (\k. y (Suc k) - y k) {Suc n .. N}\) \ \x - y (Suc n)\\ using tendsto_norm by blast hence \\x - y (Suc n)\ \ (c ^ Suc n)/(1 - c)\ using norm_1 Lim_bounded by blast hence \\x - y (Suc n)\ \ (c ^ Suc n)/(1 - c)\ by auto moreover have \(c ^ Suc n)/(1 - c) = (c / (1 - c)) * (c ^ n)\ by (simp add: divide_inverse_commute) ultimately show \\x - y (Suc n)\ \ (c / (1 - c)) * (c ^ n)\ by linarith qed lemma onorm_open_ball: includes notation_norm shows \\f\ = Sup { \f *\<^sub>v x\ | x. \x\ < 1 }\ text \ Explanation: Let \<^term>\f\ be a bounded linear operator. The operator norm of \<^term>\f\ is the supremum of \<^term>\norm (f x)\ for \<^term>\x\ such that \<^term>\norm x < 1\. \ proof(cases \(UNIV::'a set) = 0\) case True hence \x = 0\ for x::'a by auto hence \f *\<^sub>v x = 0\ for x by (metis (full_types) blinfun.zero_right) hence \\f\ = 0\ by (simp add: blinfun_eqI zero_blinfun.rep_eq) have \{ \f *\<^sub>v x\ | x. \x\ < 1} = {0}\ by (smt (verit, ccfv_SIG) Collect_cong \\x. f *\<^sub>v x = 0\ norm_zero singleton_conv) hence \Sup { \f *\<^sub>v x\ | x. \x\ < 1} = 0\ by simp thus ?thesis using \\f\ = 0\ by auto next case False hence \(UNIV::'a set) \ 0\ by simp have nonnegative: \\f *\<^sub>v x\ \ 0\ for x by simp have \\ x::'a. x \ 0\ using \UNIV \ 0\ by auto then obtain x::'a where \x \ 0\ by blast hence \\x\ \ 0\ by auto define y where \y = x /\<^sub>R \x\\ have \norm y = \ x /\<^sub>R \x\ \\ unfolding y_def by auto also have \\ = \x\ /\<^sub>R \x\\ by auto also have \\ = 1\ using \\x\ \ 0\ by auto finally have \\y\ = 1\ by blast hence norm_1_non_empty: \{ \f *\<^sub>v x\ | x. \x\ = 1} \ {}\ by blast have norm_1_bounded: \bdd_above { \f *\<^sub>v x\ | x. \x\ = 1}\ unfolding bdd_above_def apply auto by (metis norm_blinfun) have norm_less_1_non_empty: \{\f *\<^sub>v x\ | x. \x\ < 1} \ {}\ by (metis (mono_tags, lifting) Collect_empty_eq_bot bot_empty_eq empty_iff norm_zero zero_less_one) have norm_less_1_bounded: \bdd_above {\f *\<^sub>v x\ | x. \x\ < 1}\ proof- have \\r. \a r\ < 1 \ \f *\<^sub>v (a r)\ \ r\ for a :: "real \ 'a" proof- obtain r :: "('a \\<^sub>L 'b) \ real" where "\f x. 0 \ r f \ (bounded_linear f \ \f *\<^sub>v x\ \ \x\ * r f)" - using bounded_linear.nonneg_bounded by moura + by (metis mult.commute norm_blinfun norm_ge_zero) have \\ \f\ < 0\ by simp hence "(\r. \f\ * \a r\ \ r) \ (\r. \a r\ < 1 \ \f *\<^sub>v a r\ \ r)" by (meson less_eq_real_def mult_le_cancel_left2) thus ?thesis using dual_order.trans norm_blinfun by blast qed hence \\ M. \ x. \x\ < 1 \ \f *\<^sub>v x\ \ M\ by metis thus ?thesis by auto qed have Sup_non_neg: \Sup {\f *\<^sub>v x\ |x. \x\ = 1} \ 0\ by (metis (mono_tags, lifting) \\y\ = 1\ cSup_upper2 mem_Collect_eq norm_1_bounded norm_ge_zero) have \{0::real} \ {}\ by simp have \bdd_above {0::real}\ by simp show \\f\ = Sup {\f *\<^sub>v x\ | x. \x\ < 1}\ proof(cases \\x. f *\<^sub>v x = 0\) case True have \\f *\<^sub>v x\ = 0\ for x by (simp add: True) hence \{\f *\<^sub>v x\ | x. \x\ < 1 } \ {0}\ by blast moreover have \{\f *\<^sub>v x\ | x. \x\ < 1 } \ {0}\ using calculation norm_less_1_non_empty by fastforce ultimately have \{\f *\<^sub>v x\ | x. \x\ < 1 } = {0}\ by blast hence Sup1: \Sup {\f *\<^sub>v x\ | x. \x\ < 1 } = 0\ by simp have \\f\ = 0\ by (simp add: True blinfun_eqI) moreover have \Sup {\f *\<^sub>v x\ | x. \x\ < 1} = 0\ using Sup1 by blast ultimately show ?thesis by simp next case False have norm_f_eq_leq: \y \ {\f *\<^sub>v x\ | x. \x\ = 1} \ y \ Sup {\f *\<^sub>v x\ | x. \x\ < 1}\ for y proof- assume \y \ {\f *\<^sub>v x\ | x. \x\ = 1}\ hence \\ x. y = \f *\<^sub>v x\ \ \x\ = 1\ by blast then obtain x where \y = \f *\<^sub>v x\\ and \\x\ = 1\ by auto define y' where \y' n = (1 - (inverse (real (Suc n)))) *\<^sub>R y\ for n have \y' n \ {\f *\<^sub>v x\ | x. \x\ < 1}\ for n proof- have \y' n = (1 - (inverse (real (Suc n)))) *\<^sub>R \f *\<^sub>v x\\ using y'_def \y = \f *\<^sub>v x\\ by blast also have \... = \(1 - (inverse (real (Suc n))))\ *\<^sub>R \f *\<^sub>v x\\ by (metis (mono_tags, opaque_lifting) \y = \f *\<^sub>v x\\ abs_1 abs_le_self_iff abs_of_nat abs_of_nonneg add_diff_cancel_left' add_eq_if cancel_comm_monoid_add_class.diff_cancel diff_ge_0_iff_ge eq_iff_diff_eq_0 inverse_1 inverse_le_iff_le nat.distinct(1) of_nat_0 of_nat_Suc of_nat_le_0_iff zero_less_abs_iff zero_neq_one) also have \... = \f *\<^sub>v ((1 - (inverse (real (Suc n)))) *\<^sub>R x)\\ by (simp add: blinfun.scaleR_right) finally have y'_1: \y' n = \f *\<^sub>v ( (1 - (inverse (real (Suc n)))) *\<^sub>R x)\\ by blast have \\(1 - (inverse (Suc n))) *\<^sub>R x\ = (1 - (inverse (real (Suc n)))) * \x\\ by (simp add: linordered_field_class.inverse_le_1_iff) hence \\(1 - (inverse (Suc n))) *\<^sub>R x\ < 1\ by (simp add: \\x\ = 1\) thus ?thesis using y'_1 by blast qed have \(\n. (1 - (inverse (real (Suc n)))) ) \ 1\ using Limits.LIMSEQ_inverse_real_of_nat_add_minus by simp hence \(\n. (1 - (inverse (real (Suc n)))) *\<^sub>R y) \ 1 *\<^sub>R y\ using Limits.tendsto_scaleR by blast hence \(\n. (1 - (inverse (real (Suc n)))) *\<^sub>R y) \ y\ by simp hence \(\n. y' n) \ y\ using y'_def by simp hence \y' \ y\ by simp have \y' n \ Sup {\f *\<^sub>v x\ | x. \x\ < 1}\ for n using cSup_upper \\n. y' n \ {\f *\<^sub>v x\ |x. \x\ < 1}\ norm_less_1_bounded by blast hence \y \ Sup {\f *\<^sub>v x\ | x. \x\ < 1}\ using \y' \ y\ Topological_Spaces.Sup_lim by (meson LIMSEQ_le_const2) thus ?thesis by blast qed hence \Sup {\f *\<^sub>v x\ | x. \x\ = 1} \ Sup {\f *\<^sub>v x\ | x. \x\ < 1}\ by (metis (lifting) cSup_least norm_1_non_empty) have \y \ {\f *\<^sub>v x\ | x. \x\ < 1} \ y \ Sup {\f *\<^sub>v x\ | x. \x\ = 1}\ for y proof(cases \y = 0\) case True thus ?thesis by (simp add: Sup_non_neg) next case False hence \y \ 0\ by blast assume \y \ {\f *\<^sub>v x\ | x. \x\ < 1}\ hence \\ x. y = \f *\<^sub>v x\ \ \x\ < 1\ by blast then obtain x where \y = \f *\<^sub>v x\\ and \\x\ < 1\ by blast have \(1/\x\) * y = (1/\x\) * \f x\\ by (simp add: \y = \f *\<^sub>v x\\) also have \... = \1/\x\\ * \f *\<^sub>v x\\ by simp also have \... = \(1/\x\) *\<^sub>R (f *\<^sub>v x)\\ by simp also have \... = \f *\<^sub>v ((1/\x\) *\<^sub>R x)\\ by (simp add: blinfun.scaleR_right) finally have \(1/\x\) * y = \f *\<^sub>v ((1/\x\) *\<^sub>R x)\\ by blast have \x \ 0\ using \y \ 0\ \y = \f *\<^sub>v x\\ blinfun.zero_right by auto have \\ (1/\x\) *\<^sub>R x \ = \ (1/\x\) \ * \x\\ by simp also have \... = (1/\x\) * \x\\ by simp finally have \\(1/\x\) *\<^sub>R x\ = 1\ using \x \ 0\ by simp hence \(1/\x\) * y \ { \f *\<^sub>v x\ | x. \x\ = 1}\ using \1 / \x\ * y = \f *\<^sub>v (1 / \x\) *\<^sub>R x\\ by blast hence \(1/\x\) * y \ Sup { \f *\<^sub>v x\ | x. \x\ = 1}\ by (simp add: cSup_upper norm_1_bounded) moreover have \y \ (1/\x\) * y\ by (metis \\x\ < 1\ \y = \f *\<^sub>v x\\ mult_le_cancel_right1 norm_not_less_zero order.strict_implies_order \x \ 0\ less_divide_eq_1_pos zero_less_norm_iff) ultimately show ?thesis by linarith qed hence \Sup { \f *\<^sub>v x\ | x. \x\ < 1} \ Sup { \f *\<^sub>v x\ | x. \x\ = 1}\ by (smt (verit, del_insts) less_cSupD norm_less_1_non_empty) hence \Sup { \f *\<^sub>v x\ | x. \x\ = 1} = Sup { \f *\<^sub>v x\ | x. \x\ < 1}\ using \Sup {\f *\<^sub>v x\ |x. norm x = 1} \ Sup { \f *\<^sub>v x\ |x. \x\ < 1}\ by linarith have f1: \(SUP x. \f *\<^sub>v x\ / \x\) = Sup { \f *\<^sub>v x\ / \x\ | x. True}\ by (simp add: full_SetCompr_eq) have \y \ { \f *\<^sub>v x\ / \x\ |x. True} \ y \ { \f *\<^sub>v x\ |x. \x\ = 1} \ {0}\ for y proof- assume \y \ { \f *\<^sub>v x\ / \x\ |x. True}\ show ?thesis proof(cases \y = 0\) case True thus ?thesis by simp next case False have \\ x. y = \f *\<^sub>v x\ / \x\\ using \y \ { \f *\<^sub>v x\ / \x\ |x. True}\ by auto then obtain x where \y = \f *\<^sub>v x\ / \x\\ by blast hence \y = \(1/\x\)\ * \ f *\<^sub>v x \\ by simp hence \y = \(1/\x\) *\<^sub>R (f *\<^sub>v x)\\ by simp hence \y = \f ((1/\x\) *\<^sub>R x)\\ by (simp add: blinfun.scaleR_right) moreover have \\ (1/\x\) *\<^sub>R x \ = 1\ using False \y = \f *\<^sub>v x\ / \x\\ by auto ultimately have \y \ {\f *\<^sub>v x\ |x. \x\ = 1}\ by blast thus ?thesis by blast qed qed moreover have \y \ {\f x\ |x. \x\ = 1} \ {0} \ y \ {\f *\<^sub>v x\ / \x\ |x. True}\ for y proof(cases \y = 0\) case True thus ?thesis by auto next case False hence \y \ {0}\ by simp moreover assume \y \ {\f *\<^sub>v x\ |x. \x\ = 1} \ {0}\ ultimately have \y \ {\f *\<^sub>v x\ |x. \x\ = 1}\ by simp then obtain x where \\x\ = 1\ and \y = \f *\<^sub>v x\\ by auto have \y = \f *\<^sub>v x\ / \x\\ using \\x\ = 1\ \y = \f *\<^sub>v x\\ by simp thus ?thesis by auto qed ultimately have \{\f *\<^sub>v x\ / \x\ |x. True} = {\f *\<^sub>v x\ |x. \x\ = 1} \ {0}\ by blast hence \Sup {\f *\<^sub>v x\ / \x\ |x. True} = Sup ({\f *\<^sub>v x\ |x. \x\ = 1} \ {0})\ by simp have "\r s. \ (r::real) \ s \ sup r s = s" by (metis (lifting) sup.absorb_iff1 sup_commute) hence \Sup ({\f *\<^sub>v x\ |x. \x\ = 1} \ {(0::real)}) = max (Sup {\f *\<^sub>v x\ |x. \x\ = 1}) (Sup {0::real})\ using \0 \ Sup {\f *\<^sub>v x\ |x. \x\ = 1}\ \bdd_above {0}\ \{0} \ {}\ cSup_singleton cSup_union_distrib max.absorb_iff1 sup_commute norm_1_bounded norm_1_non_empty by (metis (no_types, lifting) ) moreover have \Sup {(0::real)} = (0::real)\ by simp ultimately have \Sup ({\f *\<^sub>v x\ |x. \x\ = 1} \ {0}) = Sup {\f *\<^sub>v x\ |x. \x\ = 1}\ using Sup_non_neg by linarith moreover have \Sup ( {\f *\<^sub>v x\ |x. \x\ = 1} \ {0}) = max (Sup {\f *\<^sub>v x\ |x. \x\ = 1}) (Sup {0}) \ using Sup_non_neg \Sup ({\f *\<^sub>v x\ |x. \x\ = 1} \ {0}) = max (Sup {\f *\<^sub>v x\ |x. \x\ = 1}) (Sup {0})\ by auto ultimately have f2: \Sup {\f *\<^sub>v x\ / \x\ | x. True} = Sup {\f *\<^sub>v x\ | x. \x\ = 1}\ using \Sup {\f *\<^sub>v x\ / \x\ |x. True} = Sup ({\f *\<^sub>v x\ |x. \x\ = 1} \ {0})\ by linarith have \(SUP x. \f *\<^sub>v x\ / \x\) = Sup {\f *\<^sub>v x\ | x. \x\ = 1}\ using f1 f2 by linarith hence \(SUP x. \f *\<^sub>v x\ / \x\) = Sup {\f *\<^sub>v x\ | x. \x\ < 1 }\ by (simp add: \Sup {\f *\<^sub>v x\ |x. \x\ = 1} = Sup {\f *\<^sub>v x\ |x. \x\ < 1}\) thus ?thesis apply transfer by (simp add: onorm_def) qed qed lemma onorm_r: includes notation_norm assumes \r > 0\ shows \\f\ = Sup ((\x. \f *\<^sub>v x\) ` (ball 0 r)) / r\ text \ Explanation: The norm of \<^term>\f\ is \<^term>\1/r\ of the supremum of the norm of \<^term>\f *\<^sub>v x\ for \<^term>\x\ in the ball of radius \<^term>\r\ centered at the origin. \ proof- have \\f\ = Sup {\f *\<^sub>v x\ |x. \x\ < 1}\ using onorm_open_ball by blast moreover have *: \{\f *\<^sub>v x\ |x. \x\ < 1} = (\x. \f *\<^sub>v x\) ` (ball 0 1)\ unfolding ball_def by auto ultimately have onorm_f: \\f\ = Sup ((\x. \f *\<^sub>v x\) ` (ball 0 1))\ by simp have s2: \x \ (\t. r *\<^sub>R \f *\<^sub>v t\) ` ball 0 1 \ x \ r * Sup ((\t. \f *\<^sub>v t\) ` ball 0 1)\ for x proof- assume \x \ (\t. r *\<^sub>R \f *\<^sub>v t\) ` ball 0 1\ hence \\ t. x = r *\<^sub>R \f *\<^sub>v t\ \ \t\ < 1\ by auto then obtain t where t: \x = r *\<^sub>R \f *\<^sub>v t\\ \\t\ < 1\ by blast define y where \y = x /\<^sub>R r\ have \x = r * (inverse r * x)\ using \x = r *\<^sub>R norm (f t)\ by auto hence \x - (r * (inverse r * x)) \ 0\ by linarith hence \x \ r * (x /\<^sub>R r)\ by auto have \y \ (\k. \f *\<^sub>v k\) ` ball 0 1\ unfolding y_def using assms t * by fastforce moreover have \x \ r * y\ using \x \ r * (x /\<^sub>R r)\ y_def by blast ultimately have y_norm_f: \y \ (\t. \f *\<^sub>v t\) ` ball 0 1 \ x \ r * y\ by blast have \(\t. \f *\<^sub>v t\) ` ball 0 1 \ {}\ by simp moreover have \bdd_above ((\t. \f *\<^sub>v t\) ` ball 0 1)\ by (simp add: bounded_linear_image blinfun.bounded_linear_right bounded_imp_bdd_above bounded_norm_comp) moreover have \\ y. y \ (\t. \f *\<^sub>v t\) ` ball 0 1 \ x \ r * y\ using y_norm_f by blast ultimately show ?thesis by (meson assms cSup_upper dual_order.trans mult_le_cancel_left_pos) qed have s3: \(\x. x \ (\t. r * \f *\<^sub>v t\) ` ball 0 1 \ x \ y) \ r * Sup ((\t. \f *\<^sub>v t\) ` ball 0 1) \ y\ for y proof- assume \\x. x \ (\t. r * \f *\<^sub>v t\) ` ball 0 1 \ x \ y\ have x_leq: \x \ (\t. \f *\<^sub>v t\) ` ball 0 1 \ x \ y / r\ for x proof- assume \x \ (\t. \f *\<^sub>v t\) ` ball 0 1\ then obtain t where \t \ ball (0::'a) 1\ and \x = \f *\<^sub>v t\\ by auto define x' where \x' = r *\<^sub>R x\ have \x' = r * \f *\<^sub>v t\\ by (simp add: \x = \f *\<^sub>v t\\ x'_def) hence \x' \ (\t. r * \f *\<^sub>v t\) ` ball 0 1\ using \t \ ball (0::'a) 1\ by auto hence \x' \ y\ using \\x. x \ (\t. r * \f *\<^sub>v t\) ` ball 0 1 \ x \ y\ by blast thus \x \ y / r\ unfolding x'_def using \r > 0\ by (simp add: mult.commute pos_le_divide_eq) qed have \(\t. \f *\<^sub>v t\) ` ball 0 1 \ {}\ by simp moreover have \bdd_above ((\t. \f *\<^sub>v t\) ` ball 0 1)\ by (simp add: bounded_linear_image blinfun.bounded_linear_right bounded_imp_bdd_above bounded_norm_comp) ultimately have \Sup ((\t. \f *\<^sub>v t\) ` ball 0 1) \ y/r\ using x_leq by (simp add: \bdd_above ((\t. \f *\<^sub>v t\) ` ball 0 1)\ cSup_least) thus ?thesis using \r > 0\ by (simp add: mult.commute pos_le_divide_eq) qed have norm_scaleR: \norm \ ((*\<^sub>R) r) = ((*\<^sub>R) \r\) \ (norm::'a \ real)\ by auto have f_x1: \f (r *\<^sub>R x) = r *\<^sub>R f x\ for x by (simp add: blinfun.scaleR_right) have \ball (0::'a) r = ((*\<^sub>R) r) ` (ball 0 1)\ by (smt assms ball_scale nonzero_mult_div_cancel_left right_inverse_eq scale_zero_right) hence \Sup ((\t. \f *\<^sub>v t\) ` (ball 0 r)) = Sup ((\t. \f *\<^sub>v t\) ` (((*\<^sub>R) r) ` (ball 0 1)))\ by simp also have \\ = Sup (((\t. \f *\<^sub>v t\) \ ((*\<^sub>R) r)) ` (ball 0 1))\ using Sup.SUP_image by auto also have \\ = Sup ((\t. \f *\<^sub>v (r *\<^sub>R t)\) ` (ball 0 1))\ using f_x1 by (simp add: comp_assoc) also have \\ = Sup ((\t. \r\ *\<^sub>R \f *\<^sub>v t\) ` (ball 0 1))\ using norm_scaleR f_x1 by auto also have \\ = Sup ((\t. r *\<^sub>R \f *\<^sub>v t\) ` (ball 0 1))\ using \r > 0\ by auto also have \\ = r * Sup ((\t. \f *\<^sub>v t\) ` (ball 0 1))\ apply (rule cSup_eq_non_empty) apply simp using s2 apply auto using s3 by auto also have \\ = r * \f\\ using onorm_f by auto finally have \Sup ((\t. \f *\<^sub>v t\) ` ball 0 r) = r * \f\\ by blast thus \\f\ = Sup ((\x. \f *\<^sub>v x\) ` (ball 0 r)) / r\ using \r > 0\ by simp qed text\Pointwise convergence\ definition pointwise_convergent_to :: \( nat \ ('a \ 'b::topological_space) ) \ ('a \ 'b) \ bool\ (\((_)/ \pointwise\ (_))\ [60, 60] 60) where \pointwise_convergent_to x l = (\ t::'a. (\ n. (x n) t) \ l t)\ lemma linear_limit_linear: fixes f :: \_ \ ('a::real_vector \ 'b::real_normed_vector)\ assumes \\n. linear (f n)\ and \f \pointwise\ F\ shows \linear F\ text\ Explanation: If a family of linear operators converges pointwise, then the limit is also a linear operator. \ proof show "F (x + y) = F x + F y" for x y proof- have "\a. F a = lim (\n. f n a)" using \f \pointwise\ F\ unfolding pointwise_convergent_to_def by (metis (full_types) limI) moreover have "\f b c g. (lim (\n. g n + f n) = (b::'b) + c \ \ f \ c) \ \ g \ b" by (metis (no_types) limI tendsto_add) moreover have "\a. (\n. f n a) \ F a" using assms(2) pointwise_convergent_to_def by force ultimately have lim_sum: \lim (\ n. (f n) x + (f n) y) = lim (\ n. (f n) x) + lim (\ n. (f n) y)\ by metis have \(f n) (x + y) = (f n) x + (f n) y\ for n using \\ n. linear (f n)\ unfolding linear_def using Real_Vector_Spaces.linear_iff assms(1) by auto hence \lim (\ n. (f n) (x + y)) = lim (\ n. (f n) x + (f n) y)\ by simp hence \lim (\ n. (f n) (x + y)) = lim (\ n. (f n) x) + lim (\ n. (f n) y)\ using lim_sum by simp moreover have \(\ n. (f n) (x + y)) \ F (x + y)\ using \f \pointwise\ F\ unfolding pointwise_convergent_to_def by blast moreover have \(\ n. (f n) x) \ F x\ using \f \pointwise\ F\ unfolding pointwise_convergent_to_def by blast moreover have \(\ n. (f n) y) \ F y\ using \f \pointwise\ F\ unfolding pointwise_convergent_to_def by blast ultimately show ?thesis by (metis limI) qed show "F (r *\<^sub>R x) = r *\<^sub>R F x" for r and x proof- have \(f n) (r *\<^sub>R x) = r *\<^sub>R (f n) x\ for n using \\ n. linear (f n)\ by (simp add: Real_Vector_Spaces.linear_def real_vector.linear_scale) hence \lim (\ n. (f n) (r *\<^sub>R x)) = lim (\ n. r *\<^sub>R (f n) x)\ by simp have \convergent (\ n. (f n) x)\ by (metis assms(2) convergentI pointwise_convergent_to_def) moreover have \isCont (\ t::'b. r *\<^sub>R t) tt\ for tt by (simp add: bounded_linear_scaleR_right) ultimately have \lim (\ n. r *\<^sub>R ((f n) x)) = r *\<^sub>R lim (\ n. (f n) x)\ using \f \pointwise\ F\ unfolding pointwise_convergent_to_def by (metis (mono_tags) isCont_tendsto_compose limI) hence \lim (\ n. (f n) (r *\<^sub>R x)) = r *\<^sub>R lim (\ n. (f n) x)\ using \lim (\ n. (f n) (r *\<^sub>R x)) = lim (\ n. r *\<^sub>R (f n) x)\ by simp moreover have \(\ n. (f n) x) \ F x\ using \f \pointwise\ F\ unfolding pointwise_convergent_to_def by blast moreover have \(\ n. (f n) (r *\<^sub>R x)) \ F (r *\<^sub>R x)\ using \f \pointwise\ F\ unfolding pointwise_convergent_to_def by blast ultimately show ?thesis by (metis limI) qed qed lemma non_Cauchy_unbounded: fixes a ::\_ \ real\ assumes \\n. a n \ 0\ and \e > 0\ and \\M. \m. \n. m \ M \ n \ M \ m > n \ sum a {Suc n..m} \ e\ shows \(\n. (sum a {0..n})) \ \\ text\ Explanation: If the sequence of partial sums of nonnegative terms is not Cauchy, then it converges to infinite. \ proof- define S::"ereal set" where \S = range (\n. sum a {0..n})\ have \\s\S. k*e \ s\ for k::nat proof(induction k) case 0 from \\M. \m. \n. m \ M \ n \ M \ m > n \ sum a {Suc n..m} \ e\ obtain m n where \m \ 0\ and \n \ 0\ and \m > n\ and \sum a {Suc n..m} \ e\ by blast have \n < Suc n\ by simp hence \{0..n} \ {Suc n..m} = {0..m}\ using Set_Interval.ivl_disj_un(7) \n < m\ by auto moreover have \finite {0..n}\ by simp moreover have \finite {Suc n..m}\ by simp moreover have \{0..n} \ {Suc n..m} = {}\ by simp ultimately have \sum a {0..n} + sum a {Suc n..m} = sum a {0..m}\ by (metis sum.union_disjoint) moreover have \sum a {Suc n..m} > 0\ using \e > 0\ \sum a {Suc n..m} \ e\ by linarith moreover have \sum a {0..n} \ 0\ by (simp add: assms(1) sum_nonneg) ultimately have \sum a {0..m} > 0\ by linarith moreover have \sum a {0..m} \ S\ unfolding S_def by blast ultimately have \\s\S. 0 \ s\ using ereal_less_eq(5) by fastforce thus ?case by (simp add: zero_ereal_def) next case (Suc k) assume \\s\S. k*e \ s\ then obtain s where \s\S\ and \ereal (k * e) \ s\ by blast have \\N. s = sum a {0..N}\ using \s\S\ unfolding S_def by blast then obtain N where \s = sum a {0..N}\ by blast from \\M. \m. \n. m \ M \ n \ M \ m > n \ sum a {Suc n..m} \ e\ obtain m n where \m \ Suc N\ and \n \ Suc N\ and \m > n\ and \sum a {Suc n..m} \ e\ by blast have \finite {Suc N..n}\ by simp moreover have \finite {Suc n..m}\ by simp moreover have \{Suc N..n} \ {Suc n..m} = {Suc N..m}\ using Set_Interval.ivl_disj_un by (metis \Suc N \ n\ \n < m\ atLeastSucAtMost_greaterThanAtMost order_less_imp_le) moreover have \{} = {Suc N..n} \ {Suc n..m}\ by simp ultimately have \sum a {Suc N..m} = sum a {Suc N..n} + sum a {Suc n..m}\ by (metis sum.union_disjoint) moreover have \sum a {Suc N..n} \ 0\ using \\n. a n \ 0\ by (simp add: sum_nonneg) ultimately have \sum a {Suc N..m} \ e\ using \e \ sum a {Suc n..m}\ by linarith have \finite {0..N}\ by simp have \finite {Suc N..m}\ by simp moreover have \{0..N} \ {Suc N..m} = {0..m}\ using Set_Interval.ivl_disj_un(7) \Suc N \ m\ by auto moreover have \{0..N} \ {Suc N..m} = {}\ by simp ultimately have \sum a {0..N} + sum a {Suc N..m} = sum a {0..m}\ by (metis \finite {0..N}\ sum.union_disjoint) hence \e + k * e \ sum a {0..m}\ using \ereal (real k * e) \ s\ \s = ereal (sum a {0..N})\ \e \ sum a {Suc N..m}\ by auto moreover have \e + k * e = (Suc k) * e\ by (simp add: semiring_normalization_rules(3)) ultimately have \(Suc k) * e \ sum a {0..m}\ by linarith hence \ereal ((Suc k) * e) \ sum a {0..m}\ by auto moreover have \sum a {0..m}\S\ unfolding S_def by blast ultimately show ?case by blast qed hence \\s\S. (real n) \ s\ for n by (meson assms(2) ereal_le_le ex_less_of_nat_mult less_le_not_le) hence \Sup S = \\ using Sup_le_iff Sup_subset_mono dual_order.strict_trans1 leD less_PInf_Ex_of_nat subsetI by metis hence Sup: \Sup ((range (\ n. (sum a {0..n})))::ereal set) = \\ using S_def by blast have \incseq (\n. (sum a {.. using \\n. a n \ 0\ using Extended_Real.incseq_sumI by auto hence \incseq (\n. (sum a {..< Suc n}))\ by (meson incseq_Suc_iff) hence \incseq (\n. (sum a {0..n})::ereal)\ using incseq_ereal by (simp add: atLeast0AtMost lessThan_Suc_atMost) hence \(\n. sum a {0..n}) \ Sup (range (\n. (sum a {0..n})::ereal))\ using LIMSEQ_SUP by auto thus ?thesis using Sup PInfty_neq_ereal by auto qed lemma sum_Cauchy_positive: fixes a ::\_ \ real\ assumes \\n. a n \ 0\ and \\K. \n. (sum a {0..n}) \ K\ shows \Cauchy (\n. sum a {0..n})\ text\ Explanation: If a series of nonnegative reals is bounded, then the series is Cauchy. \ proof (unfold Cauchy_altdef2, rule, rule) fix e::real assume \e>0\ have \\M. \m\M. \n\M. m > n \ sum a {Suc n..m} < e\ proof(rule classical) assume \\(\M. \m\M. \n\M. m > n \ sum a {Suc n..m} < e)\ hence \\M. \m. \n. m \ M \ n \ M \ m > n \ \(sum a {Suc n..m} < e)\ by blast hence \\M. \m. \n. m \ M \ n \ M \ m > n \ sum a {Suc n..m} \ e\ by fastforce hence \(\n. (sum a {0..n}) ) \ \\ using non_Cauchy_unbounded \0 < e\ assms(1) by blast from \\K. \n. sum a {0..n} \ K\ obtain K where \\n. sum a {0..n} \ K\ by blast from \(\n. sum a {0..n}) \ \\ have \\B. \N. \n\N. (\ n. (sum a {0..n}) ) n \ B\ using Lim_PInfty by simp hence \\n. (sum a {0..n}) \ K+1\ using ereal_less_eq(3) by blast thus ?thesis using \\n. (sum a {0..n}) \ K\ by (smt (verit, best)) qed have \sum a {Suc n..m} = sum a {0..m} - sum a {0..n}\ if "m > n" for m n by (metis add_diff_cancel_left' atLeast0AtMost less_imp_add_positive sum_up_index_split that) hence \\M. \m\M. \n\M. m > n \ sum a {0..m} - sum a {0..n} < e\ using \\M. \m\M. \n\M. m > n \ sum a {Suc n..m} < e\ by presburger then obtain M where \\m\M. \n\M. m > n \ sum a {0..m} - sum a {0..n} < e\ by blast moreover have \m > n \ sum a {0..m} \ sum a {0..n}\ for m n using \\ n. a n \ 0\ by (simp add: sum_mono2) ultimately have \\M. \m\M. \n\M. m > n \ \sum a {0..m} - sum a {0..n}\ < e\ by auto hence \\M. \m\M. \n\M. m \ n \ \sum a {0..m} - sum a {0..n}\ < e\ by (metis \0 < e\ abs_zero cancel_comm_monoid_add_class.diff_cancel diff_is_0_eq' less_irrefl_nat linorder_neqE_nat zero_less_diff) hence \\M. \m\M. \n\M. \sum a {0..m} - sum a {0..n}\ < e\ by (metis abs_minus_commute nat_le_linear) hence \\M. \m\M. \n\M. dist (sum a {0..m}) (sum a {0..n}) < e\ by (simp add: dist_real_def) hence \\M. \m\M. \n\M. dist (sum a {0..m}) (sum a {0..n}) < e\ by blast thus \\N. \n\N. dist (sum a {0..n}) (sum a {0..N}) < e\ by auto qed lemma convergent_series_Cauchy: fixes a::\nat \ real\ and \::\nat \ 'a::metric_space\ assumes \\M. \n. sum a {0..n} \ M\ and \\n. dist (\ (Suc n)) (\ n) \ a n\ shows \Cauchy \\ text\ Explanation: Let \<^term>\a\ be a real-valued sequence and let \<^term>\\\ be sequence in a metric space. If the partial sums of \<^term>\a\ are uniformly bounded and the distance between consecutive terms of \<^term>\\\ are bounded by the sequence \<^term>\a\, then \<^term>\\\ is Cauchy.\ proof (unfold Cauchy_altdef2, rule, rule) fix e::real assume \e > 0\ have \\k. a k \ 0\ using \\n. dist (\ (Suc n)) (\ n) \ a n\ dual_order.trans zero_le_dist by blast hence \Cauchy (\k. sum a {0..k})\ using \\M. \n. sum a {0..n} \ M\ sum_Cauchy_positive by blast hence \\M. \m\M. \n\M. dist (sum a {0..m}) (sum a {0..n}) < e\ unfolding Cauchy_def using \e > 0\ by blast hence \\M. \m\M. \n\M. m > n \ dist (sum a {0..m}) (sum a {0..n}) < e\ by blast have \dist (sum a {0..m}) (sum a {0..n}) = sum a {Suc n..m}\ if \n for m n proof - have \n < Suc n\ by simp have \finite {0..n}\ by simp moreover have \finite {Suc n..m}\ by simp moreover have \{0..n} \ {Suc n..m} = {0..m}\ using \n < Suc n\ \n < m\ by auto moreover have \{0..n} \ {Suc n..m} = {}\ by simp ultimately have sum_plus: \(sum a {0..n}) + sum a {Suc n..m} = (sum a {0..m})\ by (metis sum.union_disjoint) have \dist (sum a {0..m}) (sum a {0..n}) = \(sum a {0..m}) - (sum a {0..n})\\ using dist_real_def by blast moreover have \(sum a {0..m}) - (sum a {0..n}) = sum a {Suc n..m}\ using sum_plus by linarith ultimately show ?thesis by (simp add: \\k. 0 \ a k\ sum_nonneg) qed hence sum_a: \\M. \m\M. \n\M. m > n \ sum a {Suc n..m} < e\ by (metis \\M. \m\M. \n\M. dist (sum a {0..m}) (sum a {0..n}) < e\) obtain M where \\m\M. \n\M. m > n \ sum a {Suc n..m} < e\ using sum_a \e > 0\ by blast hence \\m. \n. Suc m \ Suc M \ Suc n \ Suc M \ Suc m > Suc n \ sum a {Suc n..Suc m - 1} < e\ by simp hence \\m\1. \n\1. m \ Suc M \ n \ Suc M \ m > n \ sum a {n..m - 1} < e\ by (metis Suc_le_D) hence sum_a2: \\M. \m\M. \n\M. m > n \ sum a {n..m-1} < e\ by (meson add_leE) have \dist (\ (n+p+1)) (\ n) \ sum a {n..n+p}\ for p n :: nat proof(induction p) case 0 thus ?case by (simp add: assms(2)) next case (Suc p) thus ?case by (smt(verit, ccfv_SIG) Suc_eq_plus1 add_Suc_right add_less_same_cancel1 assms(2) dist_self dist_triangle2 gr_implies_not0 sum.cl_ivl_Suc) qed hence \m > n \ dist (\ m) (\ n) \ sum a {n..m-1}\ for m n :: nat by (metis Suc_eq_plus1 Suc_le_D diff_Suc_1 gr0_implies_Suc less_eq_Suc_le less_imp_Suc_add zero_less_Suc) hence \\M. \m\M. \n\M. m > n \ dist (\ m) (\ n) < e\ using sum_a2 \e > 0\ by smt thus "\N. \n\N. dist (\ n) (\ N) < e" using \0 < e\ by fastforce qed unbundle notation_blinfun_apply unbundle no_notation_norm end diff --git a/thys/BenOr_Kozen_Reif/Matrix_Equation_Construction.thy b/thys/BenOr_Kozen_Reif/Matrix_Equation_Construction.thy --- a/thys/BenOr_Kozen_Reif/Matrix_Equation_Construction.thy +++ b/thys/BenOr_Kozen_Reif/Matrix_Equation_Construction.thy @@ -1,884 +1,865 @@ theory Matrix_Equation_Construction imports BKR_Algorithm begin section "Results with Sturm's Theorem" lemma relprime: fixes q::"real poly" assumes "coprime p q" assumes "p \ 0" assumes "q \ 0" shows "changes_R_smods p (pderiv p) = card {x. poly p x = 0 \ poly q x > 0} + card {x. poly p x = 0 \ poly q x < 0}" proof - have 1: "{x. poly p x = 0 \ poly q x = 0} = {}" using assms(1) coprime_poly_0 by auto have 2: "changes_R_smods p (pderiv p) = int (card {x . poly p x = 0})" using sturm_R by auto have 3: "{x. poly p x = 0 \ poly q x > 0} \ {x. poly p x = 0 \ poly q x < 0} = {}" by auto have "{x . poly p x = 0} = {x. poly p x = 0 \ poly q x > 0} \{x. poly p x = 0 \ poly q x < 0} \ {x. poly p x = 0 \ poly q x = 0}" by force then have "{x . poly p x = 0} = {x. poly p x = 0 \ poly q x > 0} \{x. poly p x = 0 \ poly q x < 0}" using 1 by auto then have "(card {x . poly p x = 0}) = (card ({x. poly p x = 0 \ poly q x > 0} \{x. poly p x = 0 \ poly q x < 0}))" by presburger then have 4: "(card {x . poly p x = 0}) = card {x. poly p x = 0 \ poly q x > 0} + card {x. poly p x = 0 \ poly q x < 0}" using 3 by (simp add: card_Un_disjoint assms(2) poly_roots_finite) show ?thesis by (simp add: "2" "4") qed (* This is the same proof as card_eq_sum *) lemma card_eq_const_sum: fixes k:: real assumes "finite A" shows "k*card A = sum (\x. k) A" proof - have "plus \ (\_. Suc 0) = (\_. Suc)" by (simp add: fun_eq_iff) then have "Finite_Set.fold (plus \ (\_. Suc 0)) = Finite_Set.fold (\_. Suc)" by (rule arg_cong) then have "Finite_Set.fold (plus \ (\_. Suc 0)) 0 A = Finite_Set.fold (\_. Suc) 0 A" by (blast intro: fun_cong) then show ?thesis by (simp add: card.eq_fold sum.eq_fold) qed lemma restate_tarski: fixes q::"real poly" assumes "coprime p q" assumes "p \ 0" assumes "q \ 0" shows "changes_R_smods p ((pderiv p) * q) = card {x. poly p x = 0 \ poly q x > 0} - int(card {x. poly p x = 0 \ poly q x < 0})" proof - have 3: "taq {x. poly p x=0} q \ \y\{x. poly p x=0}. Sturm_Tarski.sign (poly q y)" by (simp add: taq_def) have 4: "{x. poly p x=0} = {x. poly p x = 0 \ poly q x > 0} \ {x. poly p x = 0 \ poly q x < 0} \ {x. poly p x = 0 \ poly q x = 0}" by force then have 5: "{x. poly p x=0} = {x. poly p x = 0 \ poly q x > 0} \ {x. poly p x = 0 \ poly q x < 0}" using assms(1) coprime_poly_0 by auto then have 6: "\y\{x. poly p x=0}. Sturm_Tarski.sign (poly q y) \ \y\{x. poly p x = 0 \ poly q x > 0} \ {x. poly p x = 0 \ poly q x < 0}. Sturm_Tarski.sign (poly q y)" by presburger then have 12: "taq {x. poly p x=0} q \ \y\{x. poly p x = 0 \ poly q x > 0} \ {x. poly p x = 0 \ poly q x < 0}. Sturm_Tarski.sign (poly q y)" by (simp add: "3") have 7: "{x. poly p x = 0 \ poly q x > 0} \ {x. poly p x = 0 \ poly q x < 0} = {}" by auto then have 8: "\y\{x. poly p x = 0 \ poly q x > 0} \ {x. poly p x = 0 \ poly q x < 0}. Sturm_Tarski.sign (poly q y) \ (\y\{x. poly p x = 0 \ poly q x > 0}.Sturm_Tarski.sign (poly q y)) + (\y\{x. poly p x = 0 \ poly q x < 0}.Sturm_Tarski.sign(poly q y))" by (simp add: assms(2) poly_roots_finite sum.union_disjoint) then have 13: "taq {x. poly p x=0} q \ (\y\{x. poly p x = 0 \ poly q x > 0}.Sturm_Tarski.sign (poly q y)) + (\y\{x. poly p x = 0 \ poly q x < 0}.Sturm_Tarski.sign(poly q y))" by (simp add: "12") then have 9: "taq {x. poly p x = 0} q \ (\y\{x. poly p x = 0 \ poly q x > 0}.1) + (\y\{x. poly p x = 0 \ poly q x < 0}.(-1))" by simp have 10: "(\y\{x. poly p x = 0 \ poly q x > 0}.1) = card {x. poly p x = 0 \ poly q x > 0}" using card_eq_sum by auto have 11: " (\y\{x. poly p x = 0 \ poly q x < 0}.(-1)) = -1*card {x. poly p x = 0 \ poly q x < 0}" using card_eq_const_sum by simp have 14: "taq {x. poly p x = 0} q \ card {x. poly p x = 0 \ poly q x > 0} + -1*card {x. poly p x = 0 \ poly q x < 0}" using 9 10 11 by simp have 1: "changes_R_smods p (pderiv p * q) = taq {x. poly p x=0} q" using sturm_tarski_R by simp then have 15: "changes_R_smods p (pderiv p * q) = card {x. poly p x = 0 \ poly q x > 0} + (-1*card {x. poly p x = 0 \ poly q x < 0})" using 14 by linarith have 16: "(-1*card {x. poly p x = 0 \ poly q x < 0}) = - card {x. poly p x = 0 \ poly q x < 0}" by auto then show ?thesis using 15 by linarith qed lemma restate_tarski2: fixes q::"real poly" assumes "p \ 0" shows "changes_R_smods p ((pderiv p) * q) = int(card {x. poly p x = 0 \ poly q x > 0}) - int(card {x. poly p x = 0 \ poly q x < 0})" unfolding sturm_tarski_R[symmetric] taq_def proof - let ?all = "{x. poly p x=0}" let ?lt = "{x. poly p x=0 \ poly q x < 0}" let ?gt = "{x. poly p x=0 \ poly q x > 0}" let ?eq = "{x. poly p x=0 \ poly q x = 0}" have eq: "?all = ?lt \ ?gt \ ?eq" by force from poly_roots_finite[OF assms] have fin: "finite ?all" . show "(\x | poly p x = 0. Sturm_Tarski.sign (poly q x)) = int (card ?gt) - int (card ?lt)" unfolding eq apply (subst sum_Un) apply (auto simp add:fin) apply (subst sum_Un) by (auto simp add:fin) qed lemma coprime_set_prod: fixes I:: "real poly set" shows "finite I \ ((\ q \ I. (coprime p q)) \ (coprime p (\ I)))" proof (induct rule: finite_induct) case empty then show ?case by simp next case (insert x F) then show ?case using coprime_mult_right_iff by simp qed lemma finite_nonzero_set_prod: fixes I:: "real poly set" shows nonzero_hyp: "finite I \ ((\ q \ I. q \ 0) \ \ I \ 0)" proof (induct rule: finite_induct) case empty then show ?case by simp next case (insert x F) have h: "\ (insert x F) = x * (\ F)" by (simp add: insert.hyps(1) insert.hyps(2)) have h_xin: "x \ insert x F" by simp have hq: "(\ q \ (insert x F). q \ 0) \ x \ 0" using h_xin by blast show ?case using h hq using insert.hyps(3) by auto qed section "Setting up the construction: Definitions" definition characterize_root_list_p:: "real poly \ real list" where "characterize_root_list_p p \ sorted_list_of_set({x. poly p x = 0}::real set)" (************** Renegar's N(I); towards defining the RHS of the matrix equation **************) lemma construct_NofI_prop: fixes p:: "real poly" fixes I:: "real poly list" assumes nonzero: "p\0" shows "construct_NofI p I = rat_of_int (int (card {x. poly p x = 0 \ poly (prod_list I) x > 0}) - int (card {x. poly p x = 0 \ poly (prod_list I) x < 0}))" unfolding construct_NofI_def using assms restate_tarski2 nonzero rsquarefree_def by (simp add: rsquarefree_def) definition construct_s_vector:: "real poly \ real poly list list \ rat vec" where "construct_s_vector p Is = vec_of_list (map (\ I.(construct_NofI p I)) Is)" (* Consistent Sturm_Tarski.sign asSturm_Tarski.signments *) definition squash::"'a::linordered_field \ rat" where "squash x = (if x > 0 then 1 else if x < 0 then -1 else 0)" definition signs_at::"real poly list \ real \ rat list" where "signs_at qs x \ map (squash \ (\q. poly q x)) qs" definition characterize_consistent_signs_at_roots:: "real poly \ real poly list \ rat list list" where "characterize_consistent_signs_at_roots p qs = (remdups (map (signs_at qs) (characterize_root_list_p p)))" (* An alternate version designed to be used when every polynomial in qs is relatively prime to p*) definition consistent_sign_vec_copr::"real poly list \ real \ rat list" where "consistent_sign_vec_copr qs x \ map (\ q. if (poly q x > 0) then (1::rat) else (-1::rat)) qs" definition characterize_consistent_signs_at_roots_copr:: "real poly \ real poly list \ rat list list" where "characterize_consistent_signs_at_roots_copr p qss = (remdups (map (consistent_sign_vec_copr qss) (characterize_root_list_p p)))" lemma csa_list_copr_rel: fixes p:: "real poly" fixes qs:: "real poly list" assumes nonzero: "p\0" assumes pairwise_rel_prime: "\q. ((List.member qs q) \ (coprime p q))" shows "characterize_consistent_signs_at_roots p qs = characterize_consistent_signs_at_roots_copr p qs" proof - have "\q \ set(qs). \ x \ set (characterize_root_list_p p). poly q x \ 0" using pairwise_rel_prime using coprime_poly_0 in_set_member nonzero poly_roots_finite characterize_root_list_p_def by fastforce then have h: "\q \ set(qs). \ x \ set (characterize_root_list_p p). squash (poly q x) = (if (poly q x > 0) then (1::rat) else (-1::rat))" by (simp add: squash_def) have "map (\r. map (\p. if 0 < poly p r then 1 else - 1) qs) (characterize_root_list_p p) = map (\r. map (squash \ (\p. poly p r)) qs) (characterize_root_list_p p)" by (simp add: h) thus ?thesis unfolding characterize_consistent_signs_at_roots_def characterize_consistent_signs_at_roots_copr_def signs_at_def consistent_sign_vec_copr_def by presburger qed (************** Towards defining Renegar's polynomial function and the LHS of the matrix equation **************) definition list_constr:: "nat list \ nat \ bool" where "list_constr L n \ list_all (\x. x < n) L" definition all_list_constr:: "nat list list \ nat \ bool" where "all_list_constr L n \ (\x. List.member L x \ list_constr x n)" (* The first input is the subset; the second input is the consistent Sturm_Tarski.sign asSturm_Tarski.signment. We want to map over the first list and pull out all of the elements in the second list with corresponding positions, and then multiply those together. *) definition z:: "nat list \ rat list \ rat" where "z index_list sign_asg \ (prod_list (map (nth sign_asg) index_list))" definition mtx_row:: "rat list list \ nat list \ rat list" where "mtx_row sign_list index_list \ (map ( (z index_list)) sign_list)" definition matrix_A:: "rat list list \ nat list list \ rat mat" where "matrix_A sign_list subset_list = (mat_of_rows_list (length sign_list) (map (\i .(mtx_row sign_list i)) subset_list))" definition alt_matrix_A:: "rat list list \ nat list list \ rat mat" where "alt_matrix_A signs subsets = (mat (length subsets) (length signs) (\(i, j). z (subsets ! i) (signs ! j)))" lemma alt_matrix_char: "alt_matrix_A signs subsets = matrix_A signs subsets" proof - have h0: "(\i j. i < length subsets \ j < length signs \ map (\index_list. map (z index_list) signs) subsets ! i ! j = z (subsets ! i) (signs ! j))" proof - fix i fix j assume i_lt: "i < length subsets" assume j_lt: "j < length signs" show "((map (\index_list. map (z index_list) signs) subsets) ! i) ! j = z (subsets ! i) (signs ! j)" proof - have h0: "(map (\index_list. map (z index_list) signs) subsets) ! i = map (z (subsets ! i)) signs" using nth_map i_lt by blast then show ?thesis using nth_map j_lt by simp qed qed have h: " mat (length subsets) (length signs) (\(i, j). z (subsets ! i) (signs ! j)) = mat (length subsets) (length signs) (\(i, y). map (\index_list. map (z index_list) signs) subsets ! i ! y)" using h0 eq_matI[where A = "mat (length subsets) (length signs) (\(i, j). z (subsets ! i) (signs ! j))", where B = "mat (length subsets) (length signs) (\(i, y). map (\index_list. map (z index_list) signs) subsets ! i ! y)"] by auto show ?thesis unfolding alt_matrix_A_def matrix_A_def mat_of_rows_list_def apply (auto) unfolding mtx_row_def using h by blast qed lemma subsets_are_rows: "\i < (length subsets). row (alt_matrix_A signs subsets) i = vec (length signs) (\j. z (subsets ! i) (signs ! j))" unfolding row_def unfolding alt_matrix_A_def by auto lemma signs_are_cols: "\i < (length signs). col (alt_matrix_A signs subsets) i = vec (length subsets) (\j. z (subsets ! j) (signs ! i))" unfolding col_def unfolding alt_matrix_A_def by auto (* ith entry of LHS vector is the number of (distinct) real zeros of p where the Sturm_Tarski.sign vector of the qs is the ith entry of Sturm_Tarski.signs.*) definition construct_lhs_vector:: "real poly \ real poly list \ rat list list \ rat vec" where "construct_lhs_vector p qs signs \ vec_of_list (map (\w. rat_of_int (int (length (filter (\v. v = w) (map (consistent_sign_vec_copr qs) (characterize_root_list_p p)))))) signs)" (* Putting all of the pieces of the construction together *) definition satisfy_equation:: "real poly \ real poly list \ nat list list \ rat list list \ bool" where "satisfy_equation p qs subset_list sign_list = (mult_mat_vec (matrix_A sign_list subset_list) (construct_lhs_vector p qs sign_list) = (construct_rhs_vector p qs subset_list))" section "Setting up the construction: Proofs" (* Some matrix lemmas *) lemma row_mat_of_rows_list: assumes "list_all (\r. length r = nc) rs" assumes "i < length rs" shows "row (mat_of_rows_list nc rs) i = vec_of_list (nth rs i)" by (smt assms(1) assms(2) dim_col_mat(1) dim_vec_of_list eq_vecI index_row(2) index_vec list_all_length mat_of_rows_list_def row_mat split_conv vec_of_list_index) lemma mult_mat_vec_of_list: assumes "length ls = nc" assumes "list_all (\r. length r = nc) rs" shows "mat_of_rows_list nc rs *\<^sub>v vec_of_list ls = vec_of_list (map (\r. vec_of_list r \ vec_of_list ls) rs)" unfolding mult_mat_vec_def using row_mat_of_rows_list assms apply auto by (smt dim_row_mat(1) dim_vec dim_vec_of_list eq_vecI index_map_vec(1) index_map_vec(2) index_vec list_all_length mat_of_rows_list_def row_mat_of_rows_list vec_of_list_index) lemma mtx_row_length: "list_all (\r. length r = length signs) (map (mtx_row signs) ls)" apply (induction ls) by (auto simp add: mtx_row_def) thm construct_lhs_vector_def thm poly_roots_finite (* Recharacterize the LHS vector *) lemma construct_lhs_vector_clean: assumes "p \ 0" assumes "i < length signs" shows "(construct_lhs_vector p qs signs) $ i = card {x. poly p x = 0 \ ((consistent_sign_vec_copr qs x) = (nth signs i))}" proof - from poly_roots_finite[OF assms(1)] have "finite {x. poly p x = 0}" . then have eq: "(Collect ((\v. v = signs ! i) \ consistent_sign_vec_copr qs) \ set (sorted_list_of_set {x. poly p x = 0})) = {x. poly p x = 0 \ consistent_sign_vec_copr qs x = signs ! i}" by auto show ?thesis unfolding construct_lhs_vector_def vec_of_list_index characterize_root_list_p_def apply auto apply (subst nth_map[OF assms(2)]) apply auto apply (subst distinct_length_filter) using eq by auto qed lemma construct_lhs_vector_cleaner: assumes "p \ 0" shows "(construct_lhs_vector p qs signs) = vec_of_list (map (\s. rat_of_int (card {x. poly p x = 0 \ ((consistent_sign_vec_copr qs x) = s)})) signs)" apply (rule eq_vecI) apply (auto simp add: construct_lhs_vector_clean[OF assms] ) apply (simp add: vec_of_list_index) unfolding construct_lhs_vector_def using assms construct_lhs_vector_clean construct_lhs_vector_def apply auto[1] by simp (* Show that because our consistent Sturm_Tarski.sign vectors consist of 1 and -1's, z returns 1 or -1 when applied to a consistent Sturm_Tarski.sign vector *) lemma z_signs: assumes "list_all (\i. i < length signs) I" assumes "list_all (\s. s = 1 \ s = -1) signs" shows "(z I signs = 1) \ (z I signs = -1)" using assms proof (induction I) case Nil then show ?case by (auto simp add:z_def) next case (Cons a I) moreover have "signs ! a = 1 \ signs ! a = -1" by (metis (mono_tags, lifting) add_Suc_right calculation(2) calculation(3) gr0_conv_Suc list.size(4) list_all_length nth_Cons_0) ultimately show ?case by (auto simp add:z_def) qed lemma z_lemma: fixes I:: "nat list" fixes sign:: "rat list" assumes consistent: "sign \ set (characterize_consistent_signs_at_roots_copr p qs)" assumes welldefined: "list_constr I (length qs)" shows "(z I sign = 1) \ (z I sign = -1)" proof (rule z_signs) have "length sign = length qs" using consistent by (auto simp add: characterize_consistent_signs_at_roots_copr_def consistent_sign_vec_copr_def) thus "list_all (\i. i < length sign) I" using welldefined by (auto simp add: list_constr_def characterize_consistent_signs_at_roots_copr_def consistent_sign_vec_copr_def) show "list_all (\s. s = 1 \ s = - 1) sign" using consistent apply (auto simp add: list.pred_map characterize_consistent_signs_at_roots_copr_def consistent_sign_vec_copr_def) using Ball_set by force qed (* Show that all consistent sign vectors on roots of polynomials are in characterize_consistent_signs_at_roots_copr *) lemma in_set: fixes p:: "real poly" assumes nonzero: "p\0" fixes qs:: "real poly list" fixes I:: "nat list" fixes sign:: "rat list" fixes x:: "real" assumes root_p: "x \ {x. poly p x = 0}" assumes sign_fix: "sign = consistent_sign_vec_copr qs x" assumes welldefined: "list_constr I (length qs)" shows "sign \ set (characterize_consistent_signs_at_roots_copr p qs)" proof - have h1: "consistent_sign_vec_copr qs x \ set (remdups (map (consistent_sign_vec_copr qs) (sorted_list_of_set {x. poly p x = 0})))" using root_p apply auto apply (subst set_sorted_list_of_set) using nonzero poly_roots_finite rsquarefree_def apply blast by auto thus ?thesis unfolding characterize_consistent_signs_at_roots_copr_def characterize_root_list_p_def using sign_fix by blast qed (* Since all of the polynomials in qs are relatively prime to p, products of subsets of these polynomials are also relatively prime to p *) lemma nonzero_product: fixes p:: "real poly" assumes nonzero: "p\0" fixes qs:: "real poly list" assumes pairwise_rel_prime_1: "\q. ((List.member qs q) \ (coprime p q))" fixes I:: "nat list" fixes x:: "real" assumes root_p: "x \ {x. poly p x = 0}" assumes welldefined: "list_constr I (length qs)" shows "(poly (prod_list (retrieve_polys qs I)) x > 0) \ (poly (prod_list (retrieve_polys qs I)) x < 0)" proof - have "\x. x \ set (retrieve_polys qs I) \ coprime p x" unfolding retrieve_polys_def by (smt in_set_conv_nth in_set_member length_map list_all_length list_constr_def nth_map pairwise_rel_prime_1 welldefined) then have coprimeh: "coprime p (prod_list (retrieve_polys qs I))" using prod_list_coprime_right by auto thus ?thesis using root_p using coprime_poly_0 linorder_neqE_linordered_idom by blast qed (* The next few lemmas relate z to the signs of the product of subsets of polynomials of qs *) lemma horiz_vector_helper_pos_ind: fixes p:: "real poly" assumes nonzero: "p\0" fixes qs:: "real poly list" assumes pairwise_rel_prime_1: "\q. ((List.member qs q) \ (coprime p q))" fixes I:: "nat list" fixes sign:: "rat list" fixes x:: "real" assumes root_p: "x \ {x. poly p x = 0}" assumes sign_fix: "sign = consistent_sign_vec_copr qs x" shows "(list_constr I (length qs)) \ (poly (prod_list (retrieve_polys qs I)) x > 0) \ (z I sign = 1)" proof (induct I) case Nil then show ?case by (simp add: retrieve_polys_def z_def) next case (Cons a I) have welldef: "list_constr (a#I) (length qs) \ (list_constr I (length qs))" unfolding list_constr_def list_all_def by auto have set_hyp: "list_constr I (length qs) \ sign \ set (characterize_consistent_signs_at_roots_copr p qs)" using in_set using nonzero root_p sign_fix by blast have z_hyp: "list_constr I (length qs) \ ((z I sign = 1) \ (z I sign = -1))" using set_hyp z_lemma[where sign="sign", where I = "I", where p="p", where qs="qs"] by blast have sign_hyp: "sign = map (\ q. if (poly q x > 0) then 1 else -1) qs" using sign_fix unfolding consistent_sign_vec_copr_def by blast have ind_hyp_1: "list_constr (a#I) (length qs) \ ((0 < poly (prod_list (retrieve_polys qs I)) x) = (z I sign = 1))" using welldef Cons.hyps by auto have ind_hyp_2: "list_constr (a#I) (length qs) \ ((0 > poly (prod_list (retrieve_polys qs I)) x) = (z I sign = -1))" using welldef z_hyp Cons.hyps nonzero_product using pairwise_rel_prime_1 nonzero root_p by auto have h1: "prod_list (retrieve_polys qs (a # I)) = (nth qs a)*(prod_list (retrieve_polys qs I))" by (simp add: retrieve_polys_def) have h2: "(z (a # I) sign) = (nth sign a)*(z I sign)" by (metis (mono_tags, opaque_lifting) list.simps(9) prod_list.Cons z_def) have h3help: "list_constr (a#I) (length qs) \ a < length qs" unfolding list_constr_def by simp then have h3: "list_constr (a#I) (length qs) \ ((nth sign a) = (if (poly (nth qs a) x > 0) then 1 else -1))" using nth_map sign_hyp by auto have h2: "(0 < poly ((nth qs a)*(prod_list (retrieve_polys qs I))) x) \ ((0 < poly (nth qs a) x \ (0 < poly (prod_list (retrieve_polys qs I)) x)) \ (0 > poly (nth qs a) x \ (0 > poly (prod_list (retrieve_polys qs I)) x)))" by (simp add: zero_less_mult_iff) have final_hyp_a: "list_constr (a#I) (length qs) \ (((0 < poly (nth qs a) x \ (0 < poly (prod_list (retrieve_polys qs I)) x)) \ (0 > poly (nth qs a) x \ (0 > poly (prod_list (retrieve_polys qs I)) x))) = ((nth sign a)*(z I sign) = 1))" proof - have extra_hyp_a: "list_constr (a#I) (length qs) \ (0 < poly (nth qs a) x = ((nth sign a) = 1))" using h3 by simp have extra_hyp_b: "list_constr (a#I) (length qs) \ (0 > poly (nth qs a) x = ((nth sign a) = -1))" using h3 apply (auto) using coprime_poly_0 h3help in_set_member nth_mem pairwise_rel_prime_1 root_p by fastforce have ind_hyp_1: "list_constr (a#I) (length qs) \ (((0 < poly (nth qs a) x \ (z I sign = 1)) \ (0 > poly (nth qs a) x \ (z I sign = -1))) = ((nth sign a)*(z I sign) = 1))" using extra_hyp_a extra_hyp_b using zmult_eq_1_iff by (simp add: h3) then show ?thesis using ind_hyp_1 ind_hyp_2 by (simp add: Cons.hyps welldef) qed then show ?case using h1 z_def by (simp add: zero_less_mult_iff) qed lemma horiz_vector_helper_pos: fixes p:: "real poly" assumes nonzero: "p\0" fixes qs:: "real poly list" assumes pairwise_rel_prime_1: "\q. ((List.member qs q) \ (coprime p q))" fixes I:: "nat list" fixes sign:: "rat list" fixes x:: "real" assumes root_p: "x \ {x. poly p x = 0}" assumes sign_fix: "sign = consistent_sign_vec_copr qs x" assumes welldefined: "list_constr I (length qs)" shows "(poly (prod_list (retrieve_polys qs I)) x > 0) \ (z I sign = 1)" using horiz_vector_helper_pos_ind using pairwise_rel_prime_1 nonzero root_p sign_fix welldefined by blast lemma horiz_vector_helper_neg: fixes p:: "real poly" assumes nonzero: "p\0" fixes qs:: "real poly list" assumes pairwise_rel_prime_1: "\q. ((List.member qs q) \ (coprime p q))" fixes I:: "nat list" fixes sign:: "rat list" fixes x:: "real" assumes root_p: "x \ {x. poly p x = 0}" assumes sign_fix: "sign = consistent_sign_vec_copr qs x" assumes welldefined: "list_constr I (length qs)" shows "(poly (prod_list (retrieve_polys qs I)) x < 0) \ (z I sign = -1)" proof - have set_hyp: "list_constr I (length qs) \ sign \ set (characterize_consistent_signs_at_roots_copr p qs)" using in_set using nonzero root_p sign_fix by blast have z_hyp: "list_constr I (length qs) \ ((z I sign = 1) \ (z I sign = -1))" using set_hyp z_lemma[where sign="sign", where I = "I", where p="p", where qs="qs"] by blast have poly_hyp: "(poly (prod_list (retrieve_polys qs I)) x > 0) \ (poly (prod_list (retrieve_polys qs I)) x < 0)" using nonzero_product using pairwise_rel_prime_1 nonzero root_p using welldefined by blast have pos_hyp: "(poly (prod_list (retrieve_polys qs I)) x > 0) \ (z I sign = 1)" using horiz_vector_helper_pos using pairwise_rel_prime_1 nonzero root_p sign_fix welldefined by blast show ?thesis using z_hyp poly_hyp pos_hyp apply (auto) using welldefined by blast qed (* Recharacterize the dot product *) lemma vec_of_list_dot_rewrite: assumes "length xs = length ys" shows "vec_of_list xs \ vec_of_list ys = sum_list (map2 (*) xs ys)" using assms proof (induction xs arbitrary:ys) case Nil then show ?case by auto next case (Cons a xs) then show ?case apply auto by (smt (verit, best) Suc_length_conv list.simps(9) old.prod.case scalar_prod_vCons sum_list.Cons vec_of_list_Cons zip_Cons_Cons) qed lemma lhs_dot_rewrite: fixes p:: "real poly" fixes qs:: "real poly list" fixes I:: "nat list" fixes signs:: "rat list list" assumes nonzero: "p\0" shows "(vec_of_list (mtx_row signs I) \ (construct_lhs_vector p qs signs)) = sum_list (map (\s. (z I s) * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) signs)" proof - have "p \ 0" using nonzero by auto from construct_lhs_vector_cleaner[OF this] have rhseq: "construct_lhs_vector p qs signs = vec_of_list (map (\s. rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) signs)" by auto have "(vec_of_list (mtx_row signs I) \ (construct_lhs_vector p qs signs)) = sum_list (map2 (*) (mtx_row signs I) (map (\s. rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) signs))" unfolding rhseq apply (intro vec_of_list_dot_rewrite) by (auto simp add: mtx_row_def) thus ?thesis unfolding mtx_row_def using map2_map_map by (auto simp add: map2_map_map) qed lemma sum_list_distinct_filter: fixes f:: "'a \ int" assumes "distinct xs" "distinct ys" assumes "set ys \ set xs" assumes "\x. x \ set xs - set ys \ f x = 0" shows "sum_list (map f xs) = sum_list (map f ys)" by (metis List.finite_set assms(1) assms(2) assms(3) assms(4) sum.mono_neutral_cong_left sum_list_distinct_conv_sum_set) (* If we have a superset of the signs, we can drop to just the consistent ones *) lemma construct_lhs_vector_drop_consistent: fixes p:: "real poly" fixes qs:: "real poly list" fixes I:: "nat list" fixes signs:: "rat list list" assumes nonzero: "p\0" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" assumes welldefined: "list_constr I (length qs)" shows "(vec_of_list (mtx_row signs I) \ (construct_lhs_vector p qs signs)) = (vec_of_list (mtx_row (characterize_consistent_signs_at_roots_copr p qs) I) \ (construct_lhs_vector p qs (characterize_consistent_signs_at_roots_copr p qs)))" proof - have h0: "\ sgn. sgn \ set signs \ sgn \ consistent_sign_vec_copr qs ` set (characterize_root_list_p p) \ 0 < rat_of_nat (card {xa. poly p xa = 0 \ consistent_sign_vec_copr qs xa = sgn}) \ z I sgn = 0" proof - have "\ sgn. sgn \ set signs \ sgn \ consistent_sign_vec_copr qs ` set (characterize_root_list_p p) \ 0 < rat_of_int (card {xa. poly p xa = 0 \ consistent_sign_vec_copr qs xa = sgn}) \ {xa. poly p xa = 0 \ consistent_sign_vec_copr qs xa = sgn} \ {}" by fastforce then show ?thesis - proof - - { fix iis :: "rat list" - have ff1: "0 \ p" - using nonzero rsquarefree_def by blast - obtain rr :: "(real \ bool) \ real" where - ff2: "\p. p (rr p) \ Collect p = {}" - by moura - { assume "\is. is = iis \ {r. poly p r = 0 \ consistent_sign_vec_copr qs r = is} \ {}" - then have "\is. consistent_sign_vec_copr qs (rr (\r. poly p r = 0 \ consistent_sign_vec_copr qs r = is)) = iis \ {r. poly p r = 0 \ consistent_sign_vec_copr qs r = is} \ {}" - using ff2 - by (metis (mono_tags, lifting)) - then have "\r. poly p r = 0 \ consistent_sign_vec_copr qs r = iis" - using ff2 by smt - then have "iis \ consistent_sign_vec_copr qs ` set (sorted_list_of_set {r. poly p r = 0})" - using ff1 poly_roots_finite by fastforce } - then have "iis \ set signs \ iis \ consistent_sign_vec_copr qs ` set (characterize_root_list_p p) \ \ 0 < rat_of_int (int (card {r. poly p r = 0 \ consistent_sign_vec_copr qs r = iis}))" - by (metis (no_types) \\sgn. sgn \ set signs \ sgn \ consistent_sign_vec_copr qs ` set (characterize_root_list_p p) \ 0 < rat_of_int (int (card {xa. poly p xa = 0 \ consistent_sign_vec_copr qs xa = sgn})) \ {xa. poly p xa = 0 \ consistent_sign_vec_copr qs xa = sgn} \ {}\ characterize_root_list_p_def) } - then show ?thesis - by fastforce - qed + using characterize_consistent_signs_at_roots_copr_def in_set nonzero welldefined by auto qed then have "\ sgn. sgn \ set signs \ sgn \ consistent_sign_vec_copr qs ` set (characterize_root_list_p p) \ ((0 = rat_of_nat (card {xa. poly p xa = 0 \ consistent_sign_vec_copr qs xa = sgn}) \ z I sgn = 0))" by auto then have hyp: "\ s. s \ set signs \ s \ consistent_sign_vec_copr qs ` set (characterize_root_list_p p) \ (z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}) = 0)" by auto then have "(\s\ set(signs). z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) = (\s\(set (signs) \ (consistent_sign_vec_copr qs ` set (characterize_root_list_p p))). z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}))" proof - have "set(signs) =(set (signs) \ (consistent_sign_vec_copr qs ` set (characterize_root_list_p p))) \ (set(signs)-(consistent_sign_vec_copr qs ` set (characterize_root_list_p p)))" by blast then have sum_rewrite: "(\s\ set(signs). z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) = (\s\ (set (signs) \ (consistent_sign_vec_copr qs ` set (characterize_root_list_p p))) \ (set(signs)-(consistent_sign_vec_copr qs ` set (characterize_root_list_p p))). z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}))" by auto then have sum_split: "(\s\ (set (signs) \ (consistent_sign_vec_copr qs ` set (characterize_root_list_p p))) \ (set(signs)-(consistent_sign_vec_copr qs ` set (characterize_root_list_p p))). z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) = (\s\ (set (signs) \ (consistent_sign_vec_copr qs ` set (characterize_root_list_p p))). z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) + (\s\ (set(signs)-(consistent_sign_vec_copr qs ` set (characterize_root_list_p p))). z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}))" by (metis (no_types, lifting) List.finite_set sum.Int_Diff) have sum_zero: "(\s\ (set(signs)-(consistent_sign_vec_copr qs ` set (characterize_root_list_p p))). z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) = 0" using hyp by (simp add: hyp) show ?thesis using sum_rewrite sum_split sum_zero by linarith qed then have set_eq: "set (remdups (map (consistent_sign_vec_copr qs) (characterize_root_list_p p))) = set (signs) \ (consistent_sign_vec_copr qs ` set (characterize_root_list_p p))" using all_info by (simp add: characterize_consistent_signs_at_roots_copr_def subset_antisym) have hyp1: "(\s\signs. z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) = (\s\set (signs). z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}))" using distinct_signs sum_list_distinct_conv_sum_set by blast have hyp2: "(\s\remdups (map (consistent_sign_vec_copr qs) (characterize_root_list_p p)). z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) = (\s\ set (remdups (map (consistent_sign_vec_copr qs) (characterize_root_list_p p))). z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}))" using sum_list_distinct_conv_sum_set by blast have set_sum_eq: "(\s\(set (signs) \ (consistent_sign_vec_copr qs ` set (characterize_root_list_p p))). z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) = (\s\ set (remdups (map (consistent_sign_vec_copr qs) (characterize_root_list_p p))). z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}))" using set_eq by auto then have "(\s\signs. z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) = (\s\remdups (map (consistent_sign_vec_copr qs) (characterize_root_list_p p)). z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}))" using set_sum_eq hyp1 hyp2 using \(\s\set signs. z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) = (\s\set signs \ consistent_sign_vec_copr qs ` set (characterize_root_list_p p). z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}))\ by linarith then have "consistent_sign_vec_copr qs ` set (characterize_root_list_p p) \ set signs \ (\p qss. characterize_consistent_signs_at_roots_copr p qss = remdups (map (consistent_sign_vec_copr qss) (characterize_root_list_p p))) \ (\s\signs. z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) = (\s\remdups (map (consistent_sign_vec_copr qs) (characterize_root_list_p p)). z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}))" by linarith then show ?thesis unfolding lhs_dot_rewrite[OF nonzero] apply (auto intro!: sum_list_distinct_filter simp add: distinct_signs characterize_consistent_signs_at_roots_copr_def) using all_info characterize_consistent_signs_at_roots_copr_def by auto[1] qed (* Both matrix_equation_helper_step and matrix_equation_main_step relate the matrix construction to the Tarski queries, i.e. relate the product of a row of the matrix and the LHS vector to a Tarski query on the RHS *) lemma matrix_equation_helper_step: fixes p:: "real poly" fixes qs:: "real poly list" fixes I:: "nat list" fixes signs:: "rat list list" assumes nonzero: "p\0" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" assumes welldefined: "list_constr I (length qs)" assumes pairwise_rel_prime_1: "\q. ((List.member qs q) \ (coprime p q))" shows "(vec_of_list (mtx_row signs I) \ (construct_lhs_vector p qs signs)) = rat_of_int (card {x. poly p x = 0 \ poly (prod_list (retrieve_polys qs I)) x > 0}) - rat_of_int (card {x. poly p x = 0 \ poly (prod_list (retrieve_polys qs I)) x < 0})" proof - have "finite (set (map (consistent_sign_vec_copr qs) (characterize_root_list_p p)))" by auto let ?gt = "(set (map (consistent_sign_vec_copr qs) (characterize_root_list_p p)) \ {s. z I s = 1})" let ?lt = " (set (map (consistent_sign_vec_copr qs) (characterize_root_list_p p)) \ {s. z I s = -1})" have eq: "set (map (consistent_sign_vec_copr qs) (characterize_root_list_p p)) = ?gt \ ?lt" apply auto by (metis characterize_root_list_p_def horiz_vector_helper_neg horiz_vector_helper_pos_ind nonzero nonzero_product pairwise_rel_prime_1 poly_roots_finite sorted_list_of_set(1) welldefined) (* First, drop the signs that are irrelevant *) from construct_lhs_vector_drop_consistent[OF assms(1-4)] have "vec_of_list (mtx_row signs I) \ construct_lhs_vector p qs signs = vec_of_list (mtx_row (characterize_consistent_signs_at_roots_copr p qs) I) \ construct_lhs_vector p qs (characterize_consistent_signs_at_roots_copr p qs)" . (* Now we split the sum *) from lhs_dot_rewrite[OF assms(1)] moreover have "... = (\s\characterize_consistent_signs_at_roots_copr p qs. z I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}))" . moreover have "... = (\s\set (map (consistent_sign_vec_copr qs) (characterize_root_list_p p)). z I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}))" unfolding characterize_consistent_signs_at_roots_copr_def sum_code[symmetric] by (auto) ultimately have "... = (\s\?gt. z I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) + (\s\?lt. z I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}))" apply (subst eq) apply (rule sum.union_disjoint) by auto (* Now recharacterize lt, gt*) have setroots: "set (characterize_root_list_p p) = {x. poly p x = 0}" unfolding characterize_root_list_p_def using poly_roots_finite nonzero rsquarefree_def set_sorted_list_of_set by blast have *: "\s. {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s} = {x \{x. poly p x = 0}. consistent_sign_vec_copr qs x = s}" by auto have lem_e1: "\x. x \ {x. poly p x = 0} \ card {s \ consistent_sign_vec_copr qs ` {x. poly p x = 0} \ {s. z I s = 1}. consistent_sign_vec_copr qs x = s} = (if 0 < poly (prod_list (retrieve_polys qs I)) x then 1 else 0)" proof - fix x assume rt: "x \ {x. poly p x = 0}" then have 1: "{s \ consistent_sign_vec_copr qs ` {x. poly p x = 0} \ {s. z I s = 1}. consistent_sign_vec_copr qs x = s} = {s. z I s = 1 \ consistent_sign_vec_copr qs x = s}" by auto from horiz_vector_helper_pos[OF assms(1) assms(5) rt] have 2: "... = {s. (0 < poly (prod_list (retrieve_polys qs I)) x) \ consistent_sign_vec_copr qs x = s}" using welldefined by blast have 3: "... = (if (0 < poly (prod_list (retrieve_polys qs I)) x) then {consistent_sign_vec_copr qs x} else {})" by auto thus "card {s \ consistent_sign_vec_copr qs ` {x. poly p x = 0} \ {s. z I s = 1}. consistent_sign_vec_copr qs x = s} = (if 0 < poly (prod_list (retrieve_polys qs I)) x then 1 else 0) " using 1 2 3 by auto qed have e1: "(\s\consistent_sign_vec_copr qs ` {x. poly p x = 0} \ {s. z I s = 1}. card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}) = (sum (\x. if (poly (prod_list (retrieve_polys qs I)) x) > 0 then 1 else 0) {x. poly p x = 0})" unfolding * apply (rule sum_multicount_gen) using \finite (set (map (consistent_sign_vec_copr qs) (characterize_root_list_p p)))\ setroots apply auto[1] apply (metis List.finite_set setroots) using lem_e1 by auto have gtchr: "(\s\?gt. z I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) = rat_of_int (card {x. poly p x = 0 \ 0 < poly (prod_list (retrieve_polys qs I)) x})" apply (auto simp add: setroots) apply (subst of_nat_sum[symmetric]) apply (subst of_nat_eq_iff) apply (subst e1) apply (subst card_eq_sum) apply (rule sum.mono_neutral_cong_right) apply (metis List.finite_set setroots) by auto have lem_e2: "\x. x \ {x. poly p x = 0} \ card {s \ consistent_sign_vec_copr qs ` {x. poly p x = 0} \ {s. z I s = -1}. consistent_sign_vec_copr qs x = s} = (if poly (prod_list (retrieve_polys qs I)) x < 0 then 1 else 0)" proof - fix x assume rt: "x \ {x. poly p x = 0}" then have 1: "{s \ consistent_sign_vec_copr qs ` {x. poly p x = 0} \ {s. z I s = -1}. consistent_sign_vec_copr qs x = s} = {s. z I s = -1 \ consistent_sign_vec_copr qs x = s}" by auto from horiz_vector_helper_neg[OF assms(1) assms(5) rt] have 2: "... = {s. (0 > poly (prod_list (retrieve_polys qs I)) x) \ consistent_sign_vec_copr qs x = s}" using welldefined by blast have 3: "... = (if (0 > poly (prod_list (retrieve_polys qs I)) x) then {consistent_sign_vec_copr qs x} else {})" by auto thus "card {s \ consistent_sign_vec_copr qs ` {x. poly p x = 0} \ {s. z I s = -1}. consistent_sign_vec_copr qs x = s} = (if poly (prod_list (retrieve_polys qs I)) x < 0 then 1 else 0)" using 1 2 3 by auto qed have e2: " (\s\consistent_sign_vec_copr qs ` {x. poly p x = 0} \ {s. z I s = - 1}. card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}) = (sum (\x. if (poly (prod_list (retrieve_polys qs I)) x) < 0 then 1 else 0) {x. poly p x = 0})" unfolding * apply (rule sum_multicount_gen) using \finite (set (map (consistent_sign_vec_copr qs) (characterize_root_list_p p)))\ setroots apply auto[1] apply (metis List.finite_set setroots) using lem_e2 by auto have ltchr: "(\s\?lt. z I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) = - rat_of_int (card {x. poly p x = 0 \ 0 > poly (prod_list (retrieve_polys qs I)) x})" apply (auto simp add: setroots sum_negf) apply (subst of_nat_sum[symmetric]) apply (subst of_nat_eq_iff) apply (subst e2) apply (subst card_eq_sum) apply (rule sum.mono_neutral_cong_right) apply (metis List.finite_set setroots) by auto show ?thesis using gtchr ltchr using \(\s\set (map (consistent_sign_vec_copr qs) (characterize_root_list_p p)). z I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) = (\s\set (map (consistent_sign_vec_copr qs) (characterize_root_list_p p)) \ {s. z I s = 1}. z I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) + (\s\set (map (consistent_sign_vec_copr qs) (characterize_root_list_p p)) \ {s. z I s = - 1}. z I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}))\ \(\s\characterize_consistent_signs_at_roots_copr p qs. z I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) = (\s\set (map (consistent_sign_vec_copr qs) (characterize_root_list_p p)). z I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}))\ \vec_of_list (mtx_row (characterize_consistent_signs_at_roots_copr p qs) I) \ construct_lhs_vector p qs (characterize_consistent_signs_at_roots_copr p qs) = (\s\characterize_consistent_signs_at_roots_copr p qs. z I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}))\ \vec_of_list (mtx_row signs I) \ construct_lhs_vector p qs signs = vec_of_list (mtx_row (characterize_consistent_signs_at_roots_copr p qs) I) \ construct_lhs_vector p qs (characterize_consistent_signs_at_roots_copr p qs)\ by linarith qed (* A clean restatement of the helper lemma *) lemma matrix_equation_main_step: fixes p:: "real poly" fixes qs:: "real poly list" fixes I:: "nat list" fixes signs:: "rat list list" assumes nonzero: "p\0" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" assumes welldefined: "list_constr I (length qs)" assumes pairwise_rel_prime_1: "\q. ((List.member qs q) \ (coprime p q))" shows "(vec_of_list (mtx_row signs I) \ (construct_lhs_vector p qs signs)) = construct_NofI p (retrieve_polys qs I)" unfolding construct_NofI_prop[OF nonzero] using matrix_equation_helper_step[OF assms] by linarith lemma map_vec_vec_of_list_eq_intro: assumes "map f xs = map g ys" shows "map_vec f (vec_of_list xs) = map_vec g (vec_of_list ys)" by (metis assms vec_of_list_map) (* Shows that as long as we have a "basis" of sign assignments (see assumptions all_info, welldefined), and some other mild assumptions on our inputs (given in nonzero, distinct_signs, pairwise_rel_prime), the construction will be satisfied *) theorem matrix_equation: fixes p:: "real poly" fixes qs:: "real poly list" fixes subsets:: "nat list list" fixes signs:: "rat list list" assumes nonzero: "p\0" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" assumes pairwise_rel_prime: "\q. ((List.member qs q) \ (coprime p q))" assumes welldefined: "all_list_constr (subsets) (length qs)" shows "satisfy_equation p qs subsets signs" unfolding satisfy_equation_def matrix_A_def construct_lhs_vector_def construct_rhs_vector_def all_list_constr_def apply (subst mult_mat_vec_of_list) apply (auto simp add: mtx_row_length intro!: map_vec_vec_of_list_eq_intro) using matrix_equation_main_step[OF assms(1-3) _ assms(4), unfolded construct_lhs_vector_def] using all_list_constr_def in_set_member welldefined by fastforce (* Prettifying some theorems*) definition roots:: "real poly \ real set" where "roots p = {x. poly p x = 0}" definition sgn::"'a::linordered_field \ rat" where "sgn x = (if x > 0 then 1 else if x < 0 then -1 else 0)" definition sgn_vec::"real poly list \ real \ rat list" where "sgn_vec qs x \ map (sgn \ (\q. poly q x)) qs" definition consistent_signs_at_roots:: "real poly \ real poly list \ rat list set" where "consistent_signs_at_roots p qs = (sgn_vec qs) ` (roots p)" lemma consistent_signs_at_roots_eq: assumes "p \ 0" shows "consistent_signs_at_roots p qs = set (characterize_consistent_signs_at_roots p qs)" unfolding consistent_signs_at_roots_def characterize_consistent_signs_at_roots_def characterize_root_list_p_def apply auto apply (subst set_sorted_list_of_set) using assms poly_roots_finite apply blast unfolding sgn_vec_def sgn_def signs_at_def squash_def o_def using roots_def apply auto[1] by (smt Collect_cong assms image_iff poly_roots_finite roots_def sorted_list_of_set(1)) abbreviation w_vec:: "real poly \ real poly list \ rat list list \ rat vec" where "w_vec \ construct_lhs_vector" abbreviation v_vec:: "real poly \ real poly list \ nat list list \ rat vec" where "v_vec \ construct_rhs_vector" abbreviation M_mat:: "rat list list \ nat list list \ rat mat" where "M_mat \ matrix_A" theorem matrix_equation_pretty: assumes "p\0" assumes "\q. q \ set qs \ coprime p q" assumes "distinct signs" assumes "consistent_signs_at_roots p qs \ set signs" assumes "\l i. l \ set subsets \ i \ set l \ i < length qs" shows "M_mat signs subsets *\<^sub>v w_vec p qs signs = v_vec p qs subsets" unfolding satisfy_equation_def[symmetric] apply (rule matrix_equation[OF assms(1) assms(3)]) apply (metis assms(1) assms(2) assms(4) consistent_signs_at_roots_eq csa_list_copr_rel member_def) apply (simp add: assms(2) in_set_member) using Ball_set all_list_constr_def assms(5) list_constr_def member_def by fastforce end \ No newline at end of file diff --git a/thys/BenOr_Kozen_Reif/More_Matrix.thy b/thys/BenOr_Kozen_Reif/More_Matrix.thy --- a/thys/BenOr_Kozen_Reif/More_Matrix.thy +++ b/thys/BenOr_Kozen_Reif/More_Matrix.thy @@ -1,2007 +1,2007 @@ theory More_Matrix imports "Jordan_Normal_Form.Matrix" "Jordan_Normal_Form.DL_Rank" "Jordan_Normal_Form.VS_Connect" "Jordan_Normal_Form.Gauss_Jordan_Elimination" begin section "Kronecker Product" definition kronecker_product :: "'a :: ring mat \ 'a mat \ 'a mat" where "kronecker_product A B = (let ra = dim_row A; ca = dim_col A; rb = dim_row B; cb = dim_col B in mat (ra*rb) (ca*cb) (\(i,j). A $$ (i div rb, j div cb) * B $$ (i mod rb, j mod cb) ))" lemma arith: assumes "d < a" assumes "c < b" shows "b*d+c < a*(b::nat)" proof - have "b*d+c < b*(d+1)" by (simp add: assms(2)) thus ?thesis by (metis One_nat_def Suc_leI add.right_neutral add_Suc_right assms(1) less_le_trans mult.commute mult_le_cancel2) qed lemma dim_kronecker[simp]: "dim_row (kronecker_product A B) = dim_row A * dim_row B" "dim_col (kronecker_product A B) = dim_col A * dim_col B" unfolding kronecker_product_def Let_def by auto lemma kronecker_inverse_index: assumes "r < dim_row A" "s < dim_col A" assumes "v < dim_row B" "w < dim_col B" shows "kronecker_product A B $$ (dim_row B*r+v, dim_col B*s+w) = A $$ (r,s) * B $$ (v,w)" proof - from arith[OF assms(1) assms(3)] have "dim_row B*r+v < dim_row A * dim_row B" . moreover from arith[OF assms(2) assms(4)] have "dim_col B * s + w < dim_col A * dim_col B" . ultimately show ?thesis unfolding kronecker_product_def Let_def using assms by auto qed lemma kronecker_distr_left: assumes "dim_row B = dim_row C" "dim_col B = dim_col C" shows "kronecker_product A (B+C) = kronecker_product A B + kronecker_product A C" unfolding kronecker_product_def Let_def using assms apply (auto simp add: mat_eq_iff) by (metis (no_types, lifting) distrib_left index_add_mat(1) mod_less_divisor mult_eq_0_iff neq0_conv not_less_zero) lemma kronecker_distr_right: assumes "dim_row B = dim_row C" "dim_col B = dim_col C" shows "kronecker_product (B+C) A = kronecker_product B A + kronecker_product C A" unfolding kronecker_product_def Let_def using assms by (auto simp add: mat_eq_iff less_mult_imp_div_less distrib_right) lemma index_mat_mod[simp]: "nr > 0 & nc > 0 \ mat nr nc f $$ (i mod nr,j mod nc) = f (i mod nr,j mod nc)" by auto lemma kronecker_assoc: shows "kronecker_product A (kronecker_product B C) = kronecker_product (kronecker_product A B) C" unfolding kronecker_product_def Let_def apply (case_tac "dim_row B * dim_row C > 0 & dim_col B * dim_col C > 0") apply (auto simp add: mat_eq_iff less_mult_imp_div_less) by (smt (verit, best) div_less_iff_less_mult div_mult2_eq kronecker_inverse_index linordered_semiring_strict_class.mult_pos_pos mod_less_divisor mod_mult2_eq mult.assoc mult.commute mult_div_mod_eq) lemma sum_sum_mod_div: "(\ia = 0::nat..ja = 0..ia = 0..ia. (ia div y, ia mod y)) {0.. {0.. (\ia. (ia div y, ia mod y)) ` {0.. {0.. (\ia. (ia div y, ia mod y)) ` {0..a * y + b \ {0.. add.commute add.right_neutral div_less div_mult_self1 less_zeroE mod_eq_self_iff_div_eq_0 mod_mult_self1) qed have 22:"(\ia. (ia div y, ia mod y)) ` {0.. {0.. {0.. {0..ia. (ia div y, ia mod y)) ` {0..ia = 0::nat..ja = 0..(x, y)\{0.. {0..ia. (ia div y, ia mod y)"]) using 1 2 by auto qed (* Kronecker product distributes over matrix multiplication *) lemma kronecker_of_mult: assumes "dim_col (A :: 'a :: comm_ring mat) = dim_row C" assumes "dim_col B = dim_row D" shows "kronecker_product A B * kronecker_product C D = kronecker_product (A * C) (B * D)" unfolding kronecker_product_def Let_def mat_eq_iff proof clarsimp fix i j assume ij: "i < dim_row A * dim_row B" "j < dim_col C * dim_col D" have 1: "(A * C) $$ (i div dim_row B, j div dim_col D) = row A (i div dim_row B) \ col C (j div dim_col D)" using ij less_mult_imp_div_less by (auto intro!: index_mult_mat) have 2: "(B * D) $$ (i mod dim_row B, j mod dim_col D) = row B (i mod dim_row B) \ col D (j mod dim_col D)" using ij apply (auto intro!: index_mult_mat) using gr_implies_not0 apply fastforce using gr_implies_not0 by fastforce have 3: "\x. x < dim_row C * dim_row D \ A $$ (i div dim_row B, x div dim_row D) * B $$ (i mod dim_row B, x mod dim_row D) * (C $$ (x div dim_row D, j div dim_col D) * D $$ (x mod dim_row D, j mod dim_col D)) = row A (i div dim_row B) $ (x div dim_row D) * col C (j div dim_col D) $ (x div dim_row D) * (row B (i mod dim_row B) $ (x mod dim_row D) * col D (j mod dim_col D) $ (x mod dim_row D))" proof - fix x assume *:"x < dim_row C * dim_row D" have 1: "row A (i div dim_row B) $ (x div dim_row D) = A $$ (i div dim_row B, x div dim_row D)" by (simp add: * assms(1) less_mult_imp_div_less row_def) have 2: "row B (i mod dim_row B) $ (x mod dim_row D) = B $$ (i mod dim_row B, x mod dim_row D)" by (metis "*" assms(2) ij(1) index_row(1) mod_less_divisor nat_0_less_mult_iff neq0_conv not_less_zero) have 3: "col C (j div dim_col D) $ (x div dim_row D) = C $$ (x div dim_row D, j div dim_col D)" by (simp add: "*" ij(2) less_mult_imp_div_less) have 4: "col D (j mod dim_col D) $ (x mod dim_row D) = D $$ (x mod dim_row D, j mod dim_col D)" by (metis "*" bot_nat_0.not_eq_extremum ij(2) index_col mod_less_divisor mult_zero_right not_less_zero) show "A $$ (i div dim_row B, x div dim_row D) * B $$ (i mod dim_row B, x mod dim_row D) * (C $$ (x div dim_row D, j div dim_col D) * D $$ (x mod dim_row D, j mod dim_col D)) = row A (i div dim_row B) $ (x div dim_row D) * col C (j div dim_col D) $ (x div dim_row D) * (row B (i mod dim_row B) $ (x mod dim_row D) * col D (j mod dim_col D) $ (x mod dim_row D))" unfolding 1 2 3 4 by (simp add: mult.assoc mult.left_commute) qed have *: "(A * C) $$ (i div dim_row B, j div dim_col D) * (B * D) $$ (i mod dim_row B, j mod dim_col D) = (\ia = 0..j. A $$ (i div dim_row B, j div dim_col B) * B $$ (i mod dim_row B, j mod dim_col B)) \ vec (dim_row C * dim_row D) (\i. C $$ (i div dim_row D, j div dim_col D) * D $$ (i mod dim_row D, j mod dim_col D)) = (A * C) $$ (i div dim_row B, j div dim_col D) * (B * D) $$ (i mod dim_row B, j mod dim_col D)" unfolding * scalar_prod_def by (auto simp add: assms sum_product sum_sum_mod_div intro!: sum.cong) qed lemma inverts_mat_length: assumes "square_mat A" "inverts_mat A B" "inverts_mat B A" shows "dim_row B = dim_row A" "dim_col B = dim_col A" apply (metis assms(1) assms(3) index_mult_mat(3) index_one_mat(3) inverts_mat_def square_mat.simps) by (metis assms(1) assms(2) index_mult_mat(3) index_one_mat(3) inverts_mat_def square_mat.simps) lemma less_mult_imp_mod_less: "m mod i < i" if "m < n * i" for m n i :: nat using gr_implies_not_zero that by fastforce lemma kronecker_one: shows "kronecker_product ((1\<^sub>m x)::'a :: ring_1 mat) (1\<^sub>m y) = 1\<^sub>m (x*y)" unfolding kronecker_product_def Let_def apply (auto simp add:mat_eq_iff less_mult_imp_div_less less_mult_imp_mod_less) by (metis div_mult_mod_eq) lemma kronecker_invertible: assumes "invertible_mat (A :: 'a :: comm_ring_1 mat)" "invertible_mat B" shows "invertible_mat (kronecker_product A B)" proof - obtain Ai where Ai: "inverts_mat A Ai" "inverts_mat Ai A" using assms invertible_mat_def by blast obtain Bi where Bi: "inverts_mat B Bi" "inverts_mat Bi B" using assms invertible_mat_def by blast have "square_mat (kronecker_product A B)" by (metis (no_types, lifting) assms(1) assms(2) dim_col_mat(1) dim_row_mat(1) invertible_mat_def kronecker_product_def square_mat.simps) moreover have "inverts_mat (kronecker_product A B) (kronecker_product Ai Bi)" using Ai Bi unfolding inverts_mat_def by (metis (no_types, lifting) dim_kronecker(1) index_mult_mat(3) index_one_mat(3) kronecker_of_mult kronecker_one) moreover have "inverts_mat (kronecker_product Ai Bi) (kronecker_product A B)" using Ai Bi unfolding inverts_mat_def by (metis (no_types, lifting) dim_kronecker(1) index_mult_mat(3) index_one_mat(3) kronecker_of_mult kronecker_one) ultimately show ?thesis unfolding invertible_mat_def by blast qed section "More DL Rank" (* conjugate matrices *) instantiation mat :: (conjugate) conjugate begin definition conjugate_mat :: "'a :: conjugate mat \ 'a mat" where "conjugate m = mat (dim_row m) (dim_col m) (\(i,j). conjugate (m $$ (i,j)))" lemma dim_row_conjugate[simp]: "dim_row (conjugate m) = dim_row m" unfolding conjugate_mat_def by auto lemma dim_col_conjugate[simp]: "dim_col (conjugate m) = dim_col m" unfolding conjugate_mat_def by auto lemma carrier_vec_conjugate[simp]: "m \ carrier_mat nr nc \ conjugate m \ carrier_mat nr nc" by (auto) lemma mat_index_conjugate[simp]: shows "i < dim_row m \ j < dim_col m \ conjugate m $$ (i,j) = conjugate (m $$ (i,j))" unfolding conjugate_mat_def by auto lemma row_conjugate[simp]: "i < dim_row m \ row (conjugate m) i = conjugate (row m i)" by (auto) lemma col_conjugate[simp]: "i < dim_col m \ col (conjugate m) i = conjugate (col m i)" by (auto) lemma rows_conjugate: "rows (conjugate m) = map conjugate (rows m)" by (simp add: list_eq_iff_nth_eq) lemma cols_conjugate: "cols (conjugate m) = map conjugate (cols m)" by (simp add: list_eq_iff_nth_eq) instance proof fix a b :: "'a mat" show "conjugate (conjugate a) = a" unfolding mat_eq_iff by auto let ?a = "conjugate a" let ?b = "conjugate b" show "conjugate a = conjugate b \ a = b" by (metis dim_col_conjugate dim_row_conjugate mat_index_conjugate conjugate_cancel_iff mat_eq_iff) qed end abbreviation conjugate_transpose :: "'a::conjugate mat \ 'a mat" where "conjugate_transpose A \ conjugate (A\<^sup>T)" notation conjugate_transpose ("(_\<^sup>H)" [1000]) lemma transpose_conjugate: shows "(conjugate A)\<^sup>T = A\<^sup>H" unfolding conjugate_mat_def by auto lemma vec_module_col_helper: fixes A:: "('a :: field) mat" shows "(0\<^sub>v (dim_row A) \ LinearCombinations.module.span class_ring \carrier = carrier_vec (dim_row A), mult = undefined, one = undefined, zero = 0\<^sub>v (dim_row A), add = (+), smult = (\\<^sub>v)\ (set (cols A)))" proof - have "\v. (0::'a) \\<^sub>v v + v = v" by auto then show "0\<^sub>v (dim_row A) \ LinearCombinations.module.span class_ring \carrier = carrier_vec (dim_row A), mult = undefined, one = undefined, zero = 0\<^sub>v (dim_row A), add = (+), smult = (\\<^sub>v)\ (set (cols A))" by (metis cols_dim module_vec_def right_zero_vec smult_carrier_vec vec_space.prod_in_span zero_carrier_vec) qed lemma vec_module_col_helper2: fixes A:: "('a :: field) mat" shows "\a x. x \ LinearCombinations.module.span class_ring \carrier = carrier_vec (dim_row A), mult = undefined, one = undefined, zero = 0\<^sub>v (dim_row A), add = (+), smult = (\\<^sub>v)\ (set (cols A)) \ (\a b v. (a + b) \\<^sub>v v = a \\<^sub>v v + b \\<^sub>v v) \ a \\<^sub>v x \ LinearCombinations.module.span class_ring \carrier = carrier_vec (dim_row A), mult = undefined, one = undefined, zero = 0\<^sub>v (dim_row A), add = (+), smult = (\\<^sub>v)\ (set (cols A))" proof - fix a :: 'a and x :: "'a vec" assume "x \ LinearCombinations.module.span class_ring \carrier = carrier_vec (dim_row A), mult = undefined, one = undefined, zero = 0\<^sub>v (dim_row A), add = (+), smult = (\\<^sub>v)\ (set (cols A))" then show "a \\<^sub>v x \ LinearCombinations.module.span class_ring \carrier = carrier_vec (dim_row A), mult = undefined, one = undefined, zero = 0\<^sub>v (dim_row A), add = (+), smult = (\\<^sub>v)\ (set (cols A))" by (metis (full_types) cols_dim idom_vec.smult_in_span module_vec_def) qed lemma vec_module_col: "module (class_ring :: 'a :: field ring) (module_vec TYPE('a) (dim_row A) \carrier := LinearCombinations.module.span class_ring (module_vec TYPE('a) (dim_row A)) (set (cols A))\)" proof - interpret abelian_group "module_vec TYPE('a) (dim_row A) \carrier := LinearCombinations.module.span class_ring (module_vec TYPE('a) (dim_row A)) (set (cols A))\" apply (unfold_locales) apply (auto simp add:module_vec_def) apply (metis cols_dim module_vec_def partial_object.select_convs(1) ring.simps(2) vec_vs vectorspace.span_add1) apply (metis assoc_add_vec cols_dim module_vec_def vec_space.cV vec_vs vectorspace.span_closed) using vec_module_col_helper[of A] apply (auto) apply (metis cols_dim left_zero_vec module_vec_def partial_object.select_convs(1) vec_vs vectorspace.span_closed) apply (metis cols_dim module_vec_def partial_object.select_convs(1) right_zero_vec vec_vs vectorspace.span_closed) apply (metis cols_dim comm_add_vec module_vec_def vec_space.cV vec_vs vectorspace.span_closed) unfolding Units_def apply auto by (metis (no_types, opaque_lifting) cols_dim comm_add_vec module_vec_def partial_object.select_convs(1) uminus_l_inv_vec vec_space.vec_neg vec_vs vectorspace.span_closed vectorspace.span_neg) show ?thesis apply (unfold_locales) unfolding class_ring_simps apply auto unfolding module_vec_simps using add_smult_distrib_vec apply auto apply (auto simp add:module_vec_def) using vec_module_col_helper2 apply blast by (smt (verit) cols_dim module_vec_def smult_add_distrib_vec vec_space.cV vec_vs vectorspace.span_closed) qed (* The columns of a matrix form a vectorspace *) lemma vec_vs_col: "vectorspace (class_ring :: 'a :: field ring) (module_vec TYPE('a) (dim_row A) \carrier := LinearCombinations.module.span class_ring (module_vec TYPE('a) (dim_row A)) (set (cols A))\)" unfolding vectorspace_def using vec_module_col class_field by (auto simp: class_field_def) lemma cols_mat_mul_map: shows "cols (A * B) = map ((*\<^sub>v) A) (cols B)" unfolding list_eq_iff_nth_eq by auto lemma cols_mat_mul: shows "set (cols (A * B)) = (*\<^sub>v) A ` set (cols B)" by (simp add: cols_mat_mul_map) lemma set_obtain_sublist: assumes "S \ set ls" obtains ss where "distinct ss" "S = set ss" using assms finite_distinct_list infinite_super by blast lemma mul_mat_of_cols: assumes "A \ carrier_mat nr n" assumes "\j. j < length cs \ cs ! j \ carrier_vec n" shows "A * (mat_of_cols n cs) = mat_of_cols nr (map ((*\<^sub>v) A) cs)" unfolding mat_eq_iff using assms apply auto apply (subst mat_of_cols_index) by auto lemma helper: fixes x y z ::"'a :: {conjugatable_ring, comm_ring}" shows "x * (y * z) = y * x * z" by (simp add: mult.assoc mult.left_commute) lemma cscalar_prod_conjugate_transpose: fixes x y ::"'a :: {conjugatable_ring, comm_ring} vec" assumes "A \ carrier_mat nr nc" assumes "x \ carrier_vec nr" assumes "y \ carrier_vec nc" shows "x \c (A *\<^sub>v y) = (A\<^sup>H *\<^sub>v x) \c y" unfolding mult_mat_vec_def scalar_prod_def using assms apply (auto simp add: sum_distrib_left sum_distrib_right sum_conjugate conjugate_dist_mul) apply (subst sum.swap) by (meson helper mult.assoc mult.left_commute sum.cong) lemma mat_mul_conjugate_transpose_vec_eq_0: fixes v ::"'a :: {conjugatable_ordered_ring,semiring_no_zero_divisors,comm_ring} vec" assumes "A \ carrier_mat nr nc" assumes "v \ carrier_vec nr" assumes "A *\<^sub>v (A\<^sup>H *\<^sub>v v) = 0\<^sub>v nr" shows "A\<^sup>H *\<^sub>v v = 0\<^sub>v nc" proof - have "(A\<^sup>H *\<^sub>v v) \c (A\<^sup>H *\<^sub>v v) = (A *\<^sub>v (A\<^sup>H *\<^sub>v v)) \c v" by (metis (mono_tags, lifting) Matrix.carrier_vec_conjugate assms(1) assms(2) assms(3) carrier_matD(2) conjugate_zero_vec cscalar_prod_conjugate_transpose dim_row_conjugate index_transpose_mat(2) mult_mat_vec_def scalar_prod_left_zero scalar_prod_right_zero vec_carrier) also have "... = 0" by (simp add: assms(2) assms(3)) (* this step requires real entries *) ultimately have "(A\<^sup>H *\<^sub>v v) \c (A\<^sup>H *\<^sub>v v) = 0" by auto thus ?thesis apply (subst conjugate_square_eq_0_vec[symmetric]) using assms(1) carrier_dim_vec apply fastforce by auto qed lemma row_mat_of_cols: assumes "i < nr" shows "row (mat_of_cols nr ls) i = vec (length ls) (\j. (ls ! j) $i)" by (simp add: assms mat_of_cols_index vec_eq_iff) lemma mat_of_cols_cons_mat_vec: fixes v ::"'a::comm_ring vec" assumes "v \ carrier_vec (length ls)" assumes "dim_vec a = nr" shows "mat_of_cols nr (a # ls) *\<^sub>v (vCons m v) = m \\<^sub>v a + mat_of_cols nr ls *\<^sub>v v" unfolding mult_mat_vec_def vec_eq_iff using assms by (auto simp add: row_mat_of_cols vec_Suc o_def mult.commute) lemma smult_vec_zero: fixes v ::"'a::ring vec" shows "0 \\<^sub>v v = 0\<^sub>v (dim_vec v)" unfolding smult_vec_def vec_eq_iff by (auto) lemma helper2: fixes A ::"'a::comm_ring mat" fixes v ::"'a vec" assumes "v \ carrier_vec (length ss)" assumes "\x. x \ set ls \ dim_vec x = nr" shows "mat_of_cols nr ss *\<^sub>v v = mat_of_cols nr (ls @ ss) *\<^sub>v (0\<^sub>v (length ls) @\<^sub>v v)" using assms(2) proof (induction ls) case Nil then show ?case by auto next case (Cons a ls) then show ?case apply (auto simp add:zero_vec_Suc) apply (subst mat_of_cols_cons_mat_vec) by (auto simp add:assms smult_vec_zero) qed lemma mat_of_cols_mult_mat_vec_permute_list: fixes v ::"'a::comm_ring list" assumes "f permutes {..v vec_of_list (permute_list f v) = mat_of_cols nr ss *\<^sub>v vec_of_list v" unfolding mat_of_cols_def mult_mat_vec_def vec_eq_iff scalar_prod_def proof clarsimp fix i assume "i < nr" from sum.permute[OF assms(1)] have "(\iaia. ss ! f ia $ i * v ! f ia) \ f) {..ia = 0..ia = 0..ia = 0..\g. sum g {.. f) {.. assms(2) comp_apply lessThan_atLeast0 sum.cong) show "(\ia = 0..j. permute_list f ss ! j $ i) $ ia * vec_of_list (permute_list f v) $ ia) = (\ia = 0..j. ss ! j $ i) $ ia * vec_of_list v $ ia)" using assms * by (auto simp add: permute_list_nth vec_of_list_index) qed (* permute everything in a subset of the indices to the back *) lemma subindex_permutation: assumes "distinct ss" "set ss \ {..i. i \ set ss) [0..i. i \ set ss) [0..i. i \ set ss) [0..i. i \ set ss) [0.. {..i. i \ set ss) [0.. set ls" obtains ids where "distinct ids" "set ids \ {.. {..ids. distinct ids \ set ids \ {.. ss = map ((!) ls) ids \ thesis) \ thesis" using 1 2 3 by blast qed lemma helper3: fixes A ::"'a::comm_ring mat" assumes A: "A \ carrier_mat nr nc" assumes ss:"distinct ss" "set ss \ set (cols A)" assumes "v \ carrier_vec (length ss)" obtains c where "mat_of_cols nr ss *\<^sub>v v = A *\<^sub>v c" "dim_vec c = nc" proof - from distinct_list_subset_nths[OF ss] obtain ids where ids: "distinct ids" "set ids \ {..x. x \ set ?ls \ dim_vec x = nr" using A by auto let ?cs1 = "(list_of_vec (0\<^sub>v (length ?ls) @\<^sub>v v))" from helper2[OF assms(4) ] have "mat_of_cols nr ss *\<^sub>v v = mat_of_cols nr (?ls @ ss) *\<^sub>v vec_of_list (?cs1)" using * by (metis vec_list) also have "... = mat_of_cols nr (permute_list f (?ls @ ss)) *\<^sub>v vec_of_list (permute_list f ?cs1)" apply (auto intro!: mat_of_cols_mult_mat_vec_permute_list[symmetric]) apply (metis cols_length f(1) f(2) length_append length_map length_permute_list) using assms(4) by auto also have "... = A *\<^sub>v vec_of_list (permute_list f ?cs1)" using f(2) assms by auto ultimately show "(\c. mat_of_cols nr ss *\<^sub>v v = A *\<^sub>v c \ dim_vec c = nc \ thesis) \ thesis" by (metis A assms(4) carrier_matD(2) carrier_vecD cols_length dim_vec_of_list f(2) index_append_vec(2) index_zero_vec(2) length_append length_list_of_vec length_permute_list) qed lemma mat_mul_conjugate_transpose_sub_vec_eq_0: fixes A ::"'a :: {conjugatable_ordered_ring,semiring_no_zero_divisors,comm_ring} mat" assumes "A \ carrier_mat nr nc" assumes "distinct ss" "set ss \ set (cols (A\<^sup>H))" assumes "v \ carrier_vec (length ss)" assumes "A *\<^sub>v (mat_of_cols nc ss *\<^sub>v v) = 0\<^sub>v nr" shows "(mat_of_cols nc ss *\<^sub>v v) = 0\<^sub>v nc" proof - have "A\<^sup>H \ carrier_mat nc nr" using assms(1) by auto from helper3[OF this assms(2-4)] obtain c where c: "mat_of_cols nc ss *\<^sub>v v = A\<^sup>H *\<^sub>v c" "dim_vec c = nr" by blast have 1: "c \ carrier_vec nr" using c carrier_vec_dim_vec by blast have 2: "A *\<^sub>v (A\<^sup>H *\<^sub>v c) = 0\<^sub>v nr" using c assms(5) by auto from mat_mul_conjugate_transpose_vec_eq_0[OF assms(1) 1 2] have "A\<^sup>H *\<^sub>v c = 0\<^sub>v nc" . thus ?thesis unfolding c(1)[symmetric] . qed lemma Units_invertible: fixes A:: "'a::semiring_1 mat" assumes "A \ Units (ring_mat TYPE('a) n b)" shows "invertible_mat A" using assms unfolding Units_def invertible_mat_def apply (auto simp add: ring_mat_def) using inverts_mat_def by blast lemma invertible_Units: fixes A:: "'a::semiring_1 mat" assumes "invertible_mat A" shows "A \ Units (ring_mat TYPE('a) (dim_row A) b)" using assms unfolding Units_def invertible_mat_def apply (auto simp add: ring_mat_def) by (metis assms carrier_mat_triv invertible_mat_def inverts_mat_def inverts_mat_length(1) inverts_mat_length(2)) lemma invertible_det: fixes A:: "'a::field mat" assumes "A \ carrier_mat n n" shows "invertible_mat A \ det A \ 0" apply auto using invertible_Units unit_imp_det_non_zero apply fastforce using assms by (auto intro!: Units_invertible det_non_zero_imp_unit) context vec_space begin lemma find_indices_distinct: assumes "distinct ss" assumes "i < length ss" shows "find_indices (ss ! i) ss = [i]" proof - have "set (find_indices (ss ! i) ss) = {i}" using assms apply auto by (simp add: assms(1) assms(2) nth_eq_iff_index_eq) thus ?thesis by (metis distinct.simps(2) distinct_find_indices empty_iff empty_set insert_iff list.exhaust list.simps(15)) qed lemma lin_indpt_lin_comb_list: assumes "distinct ss" assumes "lin_indpt (set ss)" assumes "set ss \ carrier_vec n" assumes "lincomb_list f ss = 0\<^sub>v n" assumes "i < length ss" shows "f i = 0" proof - from lincomb_list_as_lincomb[OF assms(3)] have "lincomb_list f ss = lincomb (mk_coeff ss f) (set ss)" . also have "... = lincomb (\v. sum f (set (find_indices v ss))) (set ss)" unfolding mk_coeff_def apply (subst R.sumlist_map_as_finsum) by (auto simp add: distinct_find_indices) ultimately have "lincomb_list f ss = lincomb (\v. sum f (set (find_indices v ss))) (set ss)" by auto then have *:"lincomb (\v. sum f (set (find_indices v ss))) (set ss) = 0\<^sub>v n" using assms(4) by auto have "finite (set ss)" by simp from not_lindepD[OF assms(2) this _ _ *] have "(\v. sum f (set (find_indices v ss))) \ set ss \ {0}" by auto from funcset_mem[OF this] have "sum f (set (find_indices (nth ss i) ss)) \ {0}" using assms(5) by auto thus ?thesis unfolding find_indices_distinct[OF assms(1) assms(5)] by auto qed (* Note: in this locale dim_row A = n, e.g.: lemma foo: assumes "dim_row A = n" shows "rank A = vec_space.rank (dim_row A) A" by (simp add: assms) *) lemma span_mat_mul_subset: assumes "A \ carrier_mat n d" assumes "B \ carrier_mat d nc" shows "span (set (cols (A * B))) \ span (set (cols A))" proof - have *: "\v. \ca. lincomb_list v (cols (A * B)) = lincomb_list ca (cols A)" proof - fix v have "lincomb_list v (cols (A * B)) = (A * B) *\<^sub>v vec nc v" apply (subst lincomb_list_as_mat_mult) apply (metis assms(1) carrier_dim_vec carrier_matD(1) cols_dim index_mult_mat(2) subset_code(1)) by (metis assms(1) assms(2) carrier_matD(1) carrier_matD(2) cols_length index_mult_mat(2) index_mult_mat(3) mat_of_cols_cols) also have "... = A *\<^sub>v (B *\<^sub>v vec nc v)" using assms(1) assms(2) by auto also have "... = lincomb_list (\i. (B *\<^sub>v vec nc v) $ i) (cols A)" apply (subst lincomb_list_as_mat_mult) using assms(1) carrier_dim_vec cols_dim apply blast by (metis assms(1) assms(2) carrier_matD(1) carrier_matD(2) cols_length dim_mult_mat_vec dim_vec eq_vecI index_vec mat_of_cols_cols) ultimately have "lincomb_list v (cols (A * B)) = lincomb_list (\i. (B *\<^sub>v vec nc v) $ i) (cols A)" by auto thus "\ca. lincomb_list v (cols (A * B)) = lincomb_list ca (cols A)" by auto qed show ?thesis apply (subst span_list_as_span[symmetric]) apply (metis assms(1) carrier_matD(1) cols_dim index_mult_mat(2)) apply (subst span_list_as_span[symmetric]) using assms(1) cols_dim apply blast by (auto simp add:span_list_def *) qed lemma rank_mat_mul_right: assumes "A \ carrier_mat n d" assumes "B \ carrier_mat d nc" shows "rank (A * B) \ rank A" proof - have "subspace class_ring (local.span (set (cols (A*B)))) (vs (local.span (set (cols A))))" unfolding subspace_def by (metis assms(1) assms(2) carrier_matD(1) cols_dim index_mult_mat(2) nested_submodules span_is_submodule vec_space.span_mat_mul_subset vec_vs_col) from vectorspace.subspace_dim[OF _ this] have "vectorspace.dim class_ring (vs (local.span (set (cols A))) \carrier := local.span (set (cols (A * B)))\) \ vectorspace.dim class_ring (vs (local.span (set (cols A))))" apply auto by (metis (no_types) assms(1) carrier_matD(1) fin_dim_span_cols index_mult_mat(2) mat_of_cols_carrier(1) mat_of_cols_cols vec_vs_col) thus ?thesis unfolding rank_def by auto qed lemma sumlist_drop: assumes "\v. v \ set ls \ dim_vec v = n" shows "sumlist ls = sumlist (filter (\v. v \ 0\<^sub>v n) ls)" using assms proof (induction ls) case Nil then show ?case by auto next case (Cons a ls) then show ?case using dim_sumlist by auto qed lemma lincomb_list_alt: shows "lincomb_list c s = sumlist (map2 (\i j. i \\<^sub>v s ! j) (map (\i. c i) [0..v. v \ set s \ dim_vec v = n" assumes "\i. i \ set ls \ i < length s" shows " sumlist (map2 (\i j. i \\<^sub>v s ! j) (map (\i. c i) ls) ls) = sumlist (map2 (\i j. i \\<^sub>v s ! j) (map (\i. c i) (filter (\i. c i \ 0) ls)) (filter (\i. c i \ 0) ls))" using assms(2) proof (induction ls) case Nil then show ?case by auto next case (Cons a s) then show ?case apply auto apply (subst smult_l_null) apply (simp add: assms(1) carrier_vecI) apply (subst left_zero_vec) apply (subst sumlist_carrier) apply auto by (metis (no_types, lifting) assms(1) carrier_dim_vec mem_Collect_eq nth_mem set_filter set_zip_rightD) qed lemma two_set: assumes "distinct ls" assumes "set ls = set [a,b]" assumes "a \ b" shows "ls = [a,b] \ ls = [b,a]" apply (cases ls) using assms(2) apply auto[1] proof - fix x xs assume ls:"ls = x # xs" obtain y ys where xs:"xs = y # ys" by (metis (no_types) \ls = x # xs\ assms(2) assms(3) list.set_cases list.set_intros(1) list.set_intros(2) set_ConsD) have 1:"x = a \ x = b" using \ls = x # xs\ assms(2) by auto have 2:"y = a \ y = b" using \ls = x # xs\ \xs = y # ys\ assms(2) by auto have 3:"ys = []" by (metis (no_types) \ls = x # xs\ \xs = y # ys\ assms(1) assms(2) distinct.simps(2) distinct_length_2_or_more in_set_member member_rec(2) neq_Nil_conv set_ConsD) show "ls = [a, b] \ ls = [b, a]" using ls xs 1 2 3 assms by auto qed lemma filter_disj_inds: assumes "i < length ls" "j < length ls" "i \ j" shows "filter (\ia. ia \ j \ ia = i) [0.. filter (\ia. ia \ j \ ia = i) [0..ia. ia = i \ ia = j) [0..ia. ia = i \ ia = j) [0..v. v \ set ls \ dim_vec v = n" assumes "\c. lincomb_list c ls = 0\<^sub>v n \ (\i. i < (length ls) \ c i = 0)" shows "distinct ls" unfolding distinct_conv_nth proof clarsimp fix i j assume ij: "i < length ls" "j < length ls" "i \ j" assume lsij: "ls ! i = ls ! j" have "lincomb_list (\k. if k = i then 1 else if k = j then -1 else 0) ls = (ls ! i) - (ls ! j)" unfolding lincomb_list_alt apply (subst lincomb_list_alt2[OF assms(1)]) apply auto using filter_disj_inds[OF ij] apply auto using ij(3) apply force using assms(1) ij(2) apply auto[1] using ij(3) apply blast using assms(1) ij(2) by auto also have "... = 0\<^sub>v n" unfolding lsij apply (rule minus_cancel_vec) using \j < length ls\ assms(1) using carrier_vec_dim_vec nth_mem by blast ultimately have "lincomb_list (\k. if k = i then 1 else if k = j then -1 else 0) ls = 0\<^sub>v n" by auto from assms(2)[OF this] show False using \i < length ls\ by auto qed end locale conjugatable_vec_space = vec_space f_ty n for f_ty::"'a::conjugatable_ordered_field itself" and n begin lemma transpose_rank_mul_conjugate_transpose: fixes A :: "'a mat" assumes "A \ carrier_mat n nc" shows "vec_space.rank nc A\<^sup>H \ rank (A * A\<^sup>H)" proof - have 1: "A\<^sup>H \ carrier_mat nc n" using assms by auto have 2: "A * A\<^sup>H \ carrier_mat n n" using assms by auto (* S is a maximal linearly independent set of rows A (or cols A\<^sup>T) *) let ?P = "(\T. T \ set (cols A\<^sup>H) \ module.lin_indpt class_ring (module_vec TYPE('a) nc) T)" have *:"\A. ?P A \ finite A \ card A \ n" proof clarsimp fix S assume S: "S \ set (cols A\<^sup>H)" have "card S \ card (set (cols A\<^sup>H))" using S using card_mono by blast also have "... \ length (cols A\<^sup>H)" using card_length by blast also have "... \ n" using assms by auto ultimately show "finite S \ card S \ n" by (meson List.finite_set S dual_order.trans finite_subset) qed have **:"?P {}" apply (subst module.lin_dep_def) by (auto simp add: vec_module) from maximal_exists[OF *] obtain S where S: "maximal S ?P" using ** by (metis (no_types, lifting)) (* Some properties of S *) from vec_space.rank_card_indpt[OF 1 S] have rankeq: "vec_space.rank nc A\<^sup>H = card S" . have s_hyp: "S \ set (cols A\<^sup>H)" using S unfolding maximal_def by simp have modhyp: "module.lin_indpt class_ring (module_vec TYPE('a) nc) S" using S unfolding maximal_def by simp (* switch to a list representation *) obtain ss where ss: "set ss = S" "distinct ss" by (metis (mono_tags) S maximal_def set_obtain_sublist) have ss2: "set (map ((*\<^sub>v) A) ss) = (*\<^sub>v) A ` S" by (simp add: ss(1)) have rw_hyp: "cols (mat_of_cols n (map ((*\<^sub>v) A) ss)) = cols (A * mat_of_cols nc ss)" unfolding cols_def apply (auto) using mat_vec_as_mat_mat_mult[of A n nc] by (metis (no_types, lifting) "1" assms carrier_matD(1) cols_dim mul_mat_of_cols nth_mem s_hyp ss(1) subset_code(1)) then have rw: "mat_of_cols n (map ((*\<^sub>v) A) ss) = A * mat_of_cols nc ss" by (metis assms carrier_matD(1) index_mult_mat(2) mat_of_cols_carrier(2) mat_of_cols_cols) have indpt: "\c. lincomb_list c (map ((*\<^sub>v) A) ss) = 0\<^sub>v n \ \i. (i < (length ss) \ c i = 0)" proof clarsimp fix c i assume *: "lincomb_list c (map ((*\<^sub>v) A) ss) = 0\<^sub>v n" assume i: "i < length ss" have "\w\set (map ((*\<^sub>v) A) ss). dim_vec w = n" using assms by auto from lincomb_list_as_mat_mult[OF this] have "A * mat_of_cols nc ss *\<^sub>v vec (length ss) c = 0\<^sub>v n" using * rw by auto then have hq: "A *\<^sub>v (mat_of_cols nc ss *\<^sub>v vec (length ss) c) = 0\<^sub>v n" by (metis assms assoc_mult_mat_vec mat_of_cols_carrier(1) vec_carrier) then have eq1: "(mat_of_cols nc ss *\<^sub>v vec (length ss) c) = 0\<^sub>v nc" apply (intro mat_mul_conjugate_transpose_sub_vec_eq_0) using assms ss s_hyp by auto (* Rewrite the inner vector back to a lincomb_list *) have dim_hyp2: "\w\set ss. dim_vec w = nc" using ss(1) s_hyp by (metis "1" carrier_matD(1) carrier_vecD cols_dim subsetD) from vec_module.lincomb_list_as_mat_mult[OF this, symmetric] have "mat_of_cols nc ss *\<^sub>v vec (length ss) c = module.lincomb_list (module_vec TYPE('a) nc) c ss" . then have "module.lincomb_list (module_vec TYPE('a) nc) c ss = 0\<^sub>v nc" using eq1 by auto from vec_space.lin_indpt_lin_comb_list[OF ss(2) _ _ this i] show "c i = 0" using modhyp ss s_hyp using "1" cols_dim by blast qed have distinct: "distinct (map ((*\<^sub>v) A) ss)" by (metis (no_types, lifting) assms carrier_matD(1) dim_mult_mat_vec imageE indpt length_map lincomb_list_indpt_distinct ss2) then have 3: "card S = card ((*\<^sub>v) A ` S)" by (metis ss distinct_card image_set length_map) then have 4: "(*\<^sub>v) A ` S \ set (cols (A * A\<^sup>H))" using cols_mat_mul \S \ set (cols A\<^sup>H)\ by blast have 5: "lin_indpt ((*\<^sub>v) A ` S)" proof clarsimp assume ld:"lin_dep ((*\<^sub>v) A ` S)" have *: "finite ((*\<^sub>v) A ` S)" by (metis List.finite_set ss2) have **: "(*\<^sub>v) A ` S \ carrier_vec n" using "2" "4" cols_dim by blast from finite_lin_dep[OF * ld **] obtain a v where a: "lincomb a ((*\<^sub>v) A ` S) = 0\<^sub>v n" and v: "v \ (*\<^sub>v) A ` S" "a v \ 0" by blast obtain i where i:"v = map ((*\<^sub>v) A) ss ! i" "i < length ss" using v unfolding ss2[symmetric] using find_first_le nth_find_first by force from ss2[symmetric] have "set (map ((*\<^sub>v) A) ss)\ carrier_vec n" using ** ss2 by auto from lincomb_as_lincomb_list_distinct[OF this distinct] have "lincomb_list (\i. a (map ((*\<^sub>v) A) ss ! i)) (map ((*\<^sub>v) A) ss) = 0\<^sub>v n" using a ss2 by auto from indpt[OF this] show False using v i by simp qed from rank_ge_card_indpt[OF 2 4 5] have "card ((*\<^sub>v) A ` S) \ rank (A * A\<^sup>H)" . thus ?thesis using rankeq 3 by linarith qed lemma conjugate_transpose_rank_le: assumes "A \ carrier_mat n nc" shows "vec_space.rank nc (A\<^sup>H) \ rank A" by (metis assms carrier_matD(2) carrier_mat_triv dim_row_conjugate dual_order.trans index_transpose_mat(2) rank_mat_mul_right transpose_rank_mul_conjugate_transpose) lemma conjugate_finsum: assumes f: "f : U \ carrier_vec n" shows "conjugate (finsum V f U) = finsum V (conjugate \ f) U" using f proof (induct U rule: infinite_finite_induct) case (infinite A) then show ?case by auto next case empty then show ?case by auto next case (insert u U) hence f: "f : U \ carrier_vec n" "f u : carrier_vec n" by auto then have cf: "conjugate \ f : U \ carrier_vec n" "(conjugate \ f) u : carrier_vec n" apply (simp add: Pi_iff) by (simp add: f(2)) then show ?case unfolding finsum_insert[OF insert(1) insert(2) f] unfolding finsum_insert[OF insert(1) insert(2) cf ] apply (subst conjugate_add_vec[of _ n]) using f(2) apply blast using M.finsum_closed f(1) apply blast by (simp add: comp_def f(1) insert.hyps(3)) qed lemma rank_conjugate_le: assumes A:"A \ carrier_mat n d" shows "rank (conjugate (A)) \ rank A" proof - (* S is a maximal linearly independent set of (conjugate A) *) let ?P = "(\T. T \ set (cols (conjugate A)) \ lin_indpt T)" have *:"\A. ?P A \ finite A \ card A \ d" by (metis List.finite_set assms card_length card_mono carrier_matD(2) cols_length dim_col_conjugate dual_order.trans rev_finite_subset) have **:"?P {}" by (simp add: finite_lin_indpt2) from maximal_exists[OF *] obtain S where S: "maximal S ?P" using ** by (metis (no_types, lifting)) have s_hyp: "S \ set (cols (conjugate A))" "lin_indpt S" using S unfolding maximal_def apply blast by (metis (no_types, lifting) S maximal_def) from rank_card_indpt[OF _ S, of d] have rankeq: "rank (conjugate A) = card S" using assms by auto have 1:"conjugate ` S \ set (cols A)" using S apply auto by (metis (no_types, lifting) cols_conjugate conjugate_id image_eqI in_mono list.set_map s_hyp(1)) have 2: "lin_indpt (conjugate ` S)" apply (rule ccontr) apply (auto simp add: lin_dep_def) proof - fix T c v assume T: "T \ conjugate ` S" "finite T" and lc:"lincomb c T = 0\<^sub>v n" and "v \ T" "c v \ 0" let ?T = "conjugate ` T" let ?c = "conjugate \ c \ conjugate" have 1: "finite ?T" using T by auto have 2: "?T \ S" using T by auto have 3: "?c \ ?T \ UNIV" by auto have "lincomb ?c ?T = (\\<^bsub>V\<^esub>x\T. conjugate (c x) \\<^sub>v conjugate x)" unfolding lincomb_def apply (subst finsum_reindex) apply auto apply (metis "2" carrier_vec_conjugate assms carrier_matD(1) cols_dim image_eqI s_hyp(1) subsetD) by (meson conjugate_cancel_iff inj_onI) also have "... = (\\<^bsub>V\<^esub>x\T. conjugate (c x \\<^sub>v x)) " by (simp add: conjugate_smult_vec) also have "... = conjugate (\\<^bsub>V\<^esub>x\T. (c x \\<^sub>v x))" apply(subst conjugate_finsum[of "\x.(c x \\<^sub>v x)" T]) apply (auto simp add:o_def) by (smt (verit, ccfv_SIG) Matrix.carrier_vec_conjugate Pi_I' T(1) assms carrier_matD(1) cols_dim dim_row_conjugate imageE s_hyp(1) smult_carrier_vec subset_eq) also have "... = conjugate (lincomb c T)" using lincomb_def by presburger ultimately have "lincomb ?c ?T = conjugate (lincomb c T)" by auto then have 4:"lincomb ?c ?T = 0\<^sub>v n" using lc by auto from not_lindepD[OF s_hyp(2) 1 2 3 4] have "conjugate \ c \ conjugate \ conjugate ` T \ {0}" . then have "c v = 0" by (simp add: Pi_iff \v \ T\) thus False using \c v \ 0\ by auto qed from rank_ge_card_indpt[OF A 1 2] have 3:"card (conjugate ` S) \ rank A" . have 4: "card (conjugate ` S) = card S" apply (auto intro!: card_image) by (meson conjugate_cancel_iff inj_onI) show ?thesis using rankeq 3 4 by auto qed lemma rank_conjugate: assumes "A \ carrier_mat n d" shows "rank (conjugate A) = rank A" using rank_conjugate_le by (metis carrier_vec_conjugate assms conjugate_id dual_order.antisym) end (* exit the context *) lemma conjugate_transpose_rank: fixes A::"'a::{conjugatable_ordered_field} mat" shows "vec_space.rank (dim_row A) A = vec_space.rank (dim_col A) (A\<^sup>H)" using conjugatable_vec_space.conjugate_transpose_rank_le by (metis (no_types, lifting) Matrix.transpose_transpose carrier_matI conjugate_id dim_col_conjugate dual_order.antisym index_transpose_mat(2) transpose_conjugate) lemma transpose_rank: fixes A::"'a::{conjugatable_ordered_field} mat" shows "vec_space.rank (dim_row A) A = vec_space.rank (dim_col A) (A\<^sup>T)" by (metis carrier_mat_triv conjugatable_vec_space.rank_conjugate conjugate_transpose_rank index_transpose_mat(2)) lemma rank_mat_mul_left: fixes A::"'a::{conjugatable_ordered_field} mat" assumes "A \ carrier_mat n d" assumes "B \ carrier_mat d nc" shows "vec_space.rank n (A * B) \ vec_space.rank d B" by (metis (no_types, lifting) Matrix.transpose_transpose assms(1) assms(2) carrier_matD(1) carrier_matD(2) carrier_mat_triv conjugatable_vec_space.rank_conjugate conjugate_transpose_rank index_mult_mat(3) index_transpose_mat(3) transpose_mult vec_space.rank_mat_mul_right) section "Results on Invertibility" (* Extract specific columns of a matrix *) definition take_cols :: "'a mat \ nat list \ 'a mat" where "take_cols A inds = mat_of_cols (dim_row A) (map ((!) (cols A)) (filter ((>) (dim_col A)) inds))" definition take_cols_var :: "'a mat \ nat list \ 'a mat" where "take_cols_var A inds = mat_of_cols (dim_row A) (map ((!) (cols A)) (inds))" definition take_rows :: "'a mat \ nat list \ 'a mat" where "take_rows A inds = mat_of_rows (dim_col A) (map ((!) (rows A)) (filter ((>) (dim_row A)) inds))" lemma cong1: "x = y \ mat_of_cols n x = mat_of_cols n y" by auto lemma nth_filter: assumes "j < length (filter P ls)" shows "P ((filter P ls) ! j)" by (simp add: assms list_ball_nth) lemma take_cols_mat_mul: assumes "A \ carrier_mat nr n" assumes "B \ carrier_mat n nc" shows "A * take_cols B inds = take_cols (A * B) inds" proof - have "\j. j < length (map ((!) (cols B)) (filter ((>) nc) inds)) \ (map ((!) (cols B)) (filter ((>) nc) inds)) ! j \ carrier_vec n" using assms apply auto apply (subst cols_nth) using nth_filter by auto from mul_mat_of_cols[OF assms(1) this] have "A * take_cols B inds = mat_of_cols nr (map (\x. A *\<^sub>v cols B ! x) (filter ((>) (dim_col B)) inds))" unfolding take_cols_def using assms by (auto simp add: o_def) also have "... = take_cols (A * B) inds" unfolding take_cols_def using assms by (auto intro!: cong1) ultimately show ?thesis by auto qed lemma take_cols_carrier_mat: assumes "A \ carrier_mat nr nc" obtains n where "take_cols A inds \ carrier_mat nr n" unfolding take_cols_def using assms by fastforce lemma take_cols_carrier_mat_strict: assumes "A \ carrier_mat nr nc" assumes "\i. i \ set inds \ i < nc" shows "take_cols A inds \ carrier_mat nr (length inds)" unfolding take_cols_def using assms by auto lemma gauss_jordan_take_cols: assumes "gauss_jordan A (take_cols A inds) = (C,D)" shows "D = take_cols C inds" proof - obtain nr nc where A: "A \ carrier_mat nr nc" by auto from take_cols_carrier_mat[OF this] obtain n where B: "take_cols A inds \ carrier_mat nr n" by auto from gauss_jordan_transform[OF A B assms, of undefined] obtain P where PP:"P\Units (ring_mat TYPE('a) nr undefined)" and CD: "C = P * A" "D = P * take_cols A inds" by blast have P: "P \ carrier_mat nr nr" by (metis (no_types, lifting) Units_def PP mem_Collect_eq partial_object.select_convs(1) ring_mat_def) from take_cols_mat_mul[OF P A] have "P * take_cols A inds = take_cols (P * A) inds" . thus ?thesis using CD by blast qed lemma dim_col_take_cols: assumes "\j. j \ set inds \ j < dim_col A" shows "dim_col (take_cols A inds) = length inds" unfolding take_cols_def using assms by auto lemma dim_col_take_rows[simp]: shows "dim_col (take_rows A inds) = dim_col A" unfolding take_rows_def by auto lemma cols_take_cols_subset: shows "set (cols (take_cols A inds)) \ set (cols A)" unfolding take_cols_def apply (subst cols_mat_of_cols) apply auto using in_set_conv_nth by fastforce lemma dim_row_take_cols[simp]: shows "dim_row (take_cols A ls) = dim_row A" by (simp add: take_cols_def) lemma dim_row_append_rows[simp]: shows "dim_row (A @\<^sub>r B) = dim_row A + dim_row B" by (simp add: append_rows_def) lemma rows_inj: assumes "dim_col A = dim_col B" assumes "rows A = rows B" shows "A = B" unfolding mat_eq_iff apply auto apply (metis assms(2) length_rows) using assms(1) apply blast by (metis assms(1) assms(2) mat_of_rows_rows) lemma append_rows_index: assumes "dim_col A = dim_col B" assumes "i < dim_row A + dim_row B" assumes "j < dim_col A" shows "(A @\<^sub>r B) $$ (i,j) = (if i < dim_row A then A $$ (i,j) else B $$ (i-dim_row A,j))" unfolding append_rows_def apply (subst index_mat_four_block) using assms by auto lemma row_append_rows: assumes "dim_col A = dim_col B" assumes "i < dim_row A + dim_row B" shows "row (A @\<^sub>r B) i = (if i < dim_row A then row A i else row B (i-dim_row A))" unfolding vec_eq_iff using assms by (auto simp add: append_rows_def) lemma append_rows_mat_mul: assumes "dim_col A = dim_col B" shows "(A @\<^sub>r B) * C = A * C @\<^sub>r B * C" unfolding mat_eq_iff apply auto apply (simp add: append_rows_def) apply (subst index_mult_mat) apply auto apply (simp add: append_rows_def) apply (subst append_rows_index) apply auto apply (simp add: append_rows_def) apply (metis add.right_neutral append_rows_def assms index_mat_four_block(3) index_mult_mat(1) index_mult_mat(3) index_zero_mat(3) row_append_rows trans_less_add1) by (metis add_cancel_right_right add_diff_inverse_nat append_rows_def assms index_mat_four_block(3) index_mult_mat(1) index_mult_mat(3) index_zero_mat(3) nat_add_left_cancel_less row_append_rows) lemma cardlt: shows "card {i. i < (n::nat)} \ n" by simp lemma row_echelon_form_zero_rows: assumes row_ech: "row_echelon_form A" assumes dim_asm: "dim_col A \ dim_row A" shows "take_rows A [0..r 0\<^sub>m (dim_row A - length (pivot_positions A)) (dim_col A) = A" proof - have ex_pivot_fun: "\ f. pivot_fun A f (dim_col A)" using row_ech unfolding row_echelon_form_def by auto have len_help: "length (pivot_positions A) = card {i. i < dim_row A \ row A i \ 0\<^sub>v (dim_col A)}" using ex_pivot_fun pivot_positions[where A = "A",where nr = "dim_row A", where nc = "dim_col A"] by auto then have len_help2: "length (pivot_positions A) \ dim_row A" by (metis (no_types, lifting) card_mono cardlt finite_Collect_less_nat le_trans mem_Collect_eq subsetI) have fileq: "filter (\y. y < dim_row A) [0..< length (pivot_positions A)] = [0..n. card {i. i < n \ row A i \ 0\<^sub>v (dim_col A)} \ n" proof clarsimp fix n have h: "\x. x \ {i. i < n \ row A i \ 0\<^sub>v (dim_col A)} \ x\{.. row A i \ 0\<^sub>v (dim_col A)} \ {.. row A i \ 0\<^sub>v (dim_col A)}::nat) \ (card {.. row A i \ 0\<^sub>v (dim_col A)}::nat) \ (n::nat)" using h2 card_lessThan[of n] by auto qed then have pivot_len: "length (pivot_positions A) \ dim_row A " using len_help by simp have alt_char: "mat_of_rows (dim_col A) (map ((!) (rows A)) (filter (\y. y < dim_col A) [0..i j. i < dim_row A \ j < dim_col A \ i < dim_row (take_rows A [0.. take_rows A [0.. dim_row A" using pivot_len by auto have h1: "take_rows A [0..i < dim_row A\ j_lt) qed let ?nc = "dim_col A" let ?nr = "dim_row A" have h2: "\i j. i < dim_row A \ j < dim_col A \ \ i < dim_row (take_rows A [0.. 0\<^sub>m (dim_row A - length (pivot_positions A)) (dim_col A) $$ (i - dim_row (take_rows A [0.. i < dim_row (take_rows A [0..f. pivot_fun A f (dim_col A) \ f i = ?nc" proof - have half1: "\f. pivot_fun A f (dim_col A)" using assms unfolding row_echelon_form_def by blast have half2: "\f. pivot_fun A f (dim_col A) \ f i = ?nc " proof clarsimp fix f assume is_piv: "pivot_fun A f (dim_col A)" have len_pp: "length (pivot_positions A) = card {i. i < ?nr \ row A i \ 0\<^sub>v ?nc}" using is_piv pivot_positions[of A ?nr ?nc f] by auto have "\i. (i < ?nr \ row A i \ 0\<^sub>v ?nc) \ (i < ?nr \ f i \ ?nc)" using is_piv pivot_fun_zero_row_iff[of A f ?nc ?nr] by blast then have len_pp_var: "length (pivot_positions A) = card {i. i < ?nr \ f i \ ?nc}" using len_pp by auto have allj_hyp: "\j < ?nr. f j = ?nc \ ((Suc j) < ?nr \ f (Suc j) = ?nc)" using is_piv unfolding pivot_fun_def using lt_i by (metis le_antisym le_less) have if_then_bad: "f i \ ?nc \ (\j. j \ i \ f j \ ?nc)" proof clarsimp fix j assume not_i: "f i \ ?nc" assume j_leq: "j \ i" assume bad_asm: "f j = ?nc" have "\k. k \ j \ k < ?nr \ f k = ?nc" proof - fix k :: nat assume a1: "j \ k" assume a2: "k < dim_row A" have f3: "\n. \ n < dim_row A \ f n \ f j \ \ Suc n < dim_row A \ f (Suc n) = f j" using allj_hyp bad_asm by presburger obtain nn :: "nat \ nat \ (nat \ bool) \ nat" where f4: "\n na p nb nc. (\ n \ na \ Suc n \ Suc na) \ (\ p nb \ \ nc \ nb \ \ p (nn nc nb p) \ p nc) \ (\ p nb \ \ nc \ nb \ p nc \ p (Suc (nn nc nb p)))" - using inc_induct order_refl by moura + using inc_induct by (metis Suc_le_mono) then have f5: "\p. \ p k \ p j \ p (Suc (nn j k p))" using a1 by presburger have f6: "\p. \ p k \ \ p (nn j k p) \ p j" using f4 a1 by meson { assume "nn j k (\n. n < dim_row A \ f n \ dim_col A) < dim_row A \ f (nn j k (\n. n < dim_row A \ f n \ dim_col A)) \ dim_col A" moreover { assume "(nn j k (\n. n < dim_row A \ f n \ dim_col A) < dim_row A \ f (nn j k (\n. n < dim_row A \ f n \ dim_col A)) \ dim_col A) \ (\ j < dim_row A \ f j = dim_col A)" then have "\ k < dim_row A \ f k = dim_col A" using f6 by (metis (mono_tags, lifting)) } ultimately have "(\ j < dim_row A \ f j = dim_col A) \ (\ Suc (nn j k (\n. n < dim_row A \ f n \ dim_col A)) < dim_row A \ f (Suc (nn j k (\n. n < dim_row A \ f n \ dim_col A))) = dim_col A) \ \ k < dim_row A \ f k = dim_col A" using bad_asm by blast } moreover { assume "(\ j < dim_row A \ f j = dim_col A) \ (\ Suc (nn j k (\n. n < dim_row A \ f n \ dim_col A)) < dim_row A \ f (Suc (nn j k (\n. n < dim_row A \ f n \ dim_col A))) = dim_col A)" then have "\ k < dim_row A \ f k = dim_col A" using f5 proof - have "\ (Suc (nn j k (\n. n < dim_row A \ f n \ dim_col A)) < dim_row A \ f (Suc (nn j k (\n. n < dim_row A \ f n \ dim_col A))) \ dim_col A) \ \ (j < dim_row A \ f j \ dim_col A)" using \(\ j < dim_row A \ f j = dim_col A) \ (\ Suc (nn j k (\n. n < dim_row A \ f n \ dim_col A)) < dim_row A \ f (Suc (nn j k (\n. n < dim_row A \ f n \ dim_col A))) = dim_col A)\ by linarith then have "\ (k < dim_row A \ f k \ dim_col A)" by (metis (mono_tags, lifting) a2 bad_asm f5 le_less) then show ?thesis by meson qed } ultimately show "f k = dim_col A" using f3 a2 by (metis (lifting) Suc_lessD bad_asm) qed then show "False" using lt_i not_i using j_leq by blast qed have "f i \ ?nc \ ({0.. {y. y < ?nr \ f y \ dim_col A})" proof - have h1: "f i \ dim_col A \ (\j\i. j < ?nr \ f j \ dim_col A)" using if_then_bad lt_i by auto then show ?thesis by auto qed then have gteq: "f i \ ?nc \ (card {i. i < ?nr \ f i \ dim_col A} \ (i+1))" using card_lessThan[of ?ip] card_mono[where B = "{i. i < ?nr \ f i \ dim_col A} ", where A = "{0.. length (pivot_positions A)" using not_lt clear by auto then show "f i = ?nc" using gteq len_pp_var by auto qed show ?thesis using half1 half2 by blast qed then have h1a: "row A i = 0\<^sub>v (dim_col A)" using pivot_fun_zero_row_iff[of A _ ?nc ?nr] using lt_i by blast then have h1: "A $$ (i, j) = 0" using index_row(1) lt_i lt_j by fastforce have h2a: "i - dim_row (take_rows A [0..m (dim_row A - length (pivot_positions A)) (dim_col A) $$ (i - dim_row (take_rows A [0..m (dim_row A - length (pivot_positions A)) (dim_col A) $$ (i - dim_row (take_rows A [0..i j. i < dim_row A \ j < dim_col A \ i < dim_row (take_rows A [0.. dim_row A" proof - have 1: "A \ carrier_mat (dim_row A) (dim_col A)" by auto obtain f where 2: "pivot_fun A f (dim_col A)" using assms row_echelon_form_def by blast from pivot_positions(4)[OF 1 2] have "length (pivot_positions A) = card {i. i < dim_row A \ row A i \ 0\<^sub>v (dim_col A)}" . also have "... \ card {i. i < dim_row A}" apply (rule card_mono) by auto ultimately show ?thesis by auto qed lemma rref_pivot_positions: assumes "row_echelon_form R" assumes R: "R \ carrier_mat nr nc" shows "\i j. (i,j) \ set (pivot_positions R) \ i < nr \ j < nc" proof - obtain f where f: "pivot_fun R f nc" using assms(1) assms(2) row_echelon_form_def by blast have *: "\i. i < nr \ f i \ nc" using f using R pivot_funD(1) by blast from pivot_positions[OF R f] have "set (pivot_positions R) = {(i, f i) |i. i < nr \ f i \ nc}" by auto then have **: "set (pivot_positions R) = {(i, f i) |i. i < nr \ f i < nc}" using * by fastforce fix i j assume "(i, j) \ set (pivot_positions R)" thus "i < nr \ j < nc" using ** by simp qed lemma pivot_fun_monoton: assumes pf: "pivot_fun A f (dim_col A)" assumes dr: "dim_row A = nr" shows "\ i. i < nr \ (\ k. ((k < nr \ i < k) \ f i \ f k))" proof - fix i assume "i < nr" show "(\ k. ((k < nr \ i < k) \ f i \ f k))" proof - fix k show "((k < nr \ i < k) \ f i \ f k)" proof (induct k) case 0 then show ?case by blast next case (Suc k) then show ?case by (smt (verit, ccfv_SIG) dr le_less_trans less_Suc_eq less_imp_le_nat pf pivot_funD(1) pivot_funD(3)) qed qed qed lemma pivot_positions_contains: assumes row_ech: "row_echelon_form A" assumes dim_h: "dim_col A \ dim_row A" assumes "pivot_fun A f (dim_col A)" shows "\i < (length (pivot_positions A)). (i, f i) \ set (pivot_positions A)" proof - let ?nr = "dim_row A" let ?nc = "dim_col A" let ?pp = "pivot_positions A" have i_nr: "\i < (length ?pp). i < ?nr" using rref_pivot_positions assms using length_pivot_positions_dim_row less_le_trans by blast have i_nc: "\i < (length ?pp). f i < ?nc" proof clarsimp fix i assume i_lt: "i < length ?pp" have fis_nc: "f i = ?nc \ (\ k > i. k < ?nr \ f k = ?nc)" proof - assume is_nc: "f i = ?nc" show "(\ k > i. k < ?nr \ f k = ?nc)" proof clarsimp fix k assume k_gt: "k > i" assume k_lt: "k < ?nr" have fk_lt: "f k \ ?nc" using pivot_funD(1)[of A ?nr f ?nc k] k_lt apply (auto) using \pivot_fun A f (dim_col A)\ by blast show "f k = ?nc" using fk_lt is_nc k_gt k_lt assms pivot_fun_monoton[of A f ?nr i k] using \pivot_fun A f (dim_col A)\ by auto qed qed have ncimp: "f i = ?nc \ (\ k \i. k \ { i. i < ?nr \ row A i \ 0\<^sub>v ?nc})" proof - assume nchyp: "f i = ?nc" show "(\ k \i. k \ { i. i < ?nr \ row A i \ 0\<^sub>v ?nc})" proof clarsimp fix k assume i_lt: "i \ k" assume k_lt: "k < dim_row A" show "row A k = 0\<^sub>v (dim_col A) " using i_lt k_lt fis_nc using pivot_fun_zero_row_iff[of A f ?nc ?nr] using \pivot_fun A f (dim_col A)\ le_neq_implies_less nchyp by blast qed qed then have "f i = ?nc \ card { i. i < ?nr \ row A i \ 0\<^sub>v ?nc} \ i" proof - assume nchyp: "f i = ?nc" have h: "{ i. i < ?nr \ row A i \ 0\<^sub>v ?nc} \ {0.. row A i \ 0\<^sub>v ?nc} \ i" using card_lessThan using subset_eq_atLeast0_lessThan_card by blast qed then show "f i < ?nc" using i_lt pivot_positions(4)[of A ?nr ?nc f] apply (auto) by (metis \pivot_fun A f (dim_col A)\ i_nr le_neq_implies_less not_less pivot_funD(1)) qed then show ?thesis using pivot_positions(1) by (smt (verit, ccfv_SIG) \pivot_fun A f (dim_col A)\ carrier_matI i_nr less_not_refl mem_Collect_eq) qed lemma pivot_positions_form_helper_1: shows "(a, b) \ set (pivot_positions_main_gen z A nr nc i j) \ i \ a" proof (induct i j rule: pivot_positions_main_gen.induct[of nr nc A z]) case (1 i j) then show ?case using pivot_positions_main_gen.simps[of z A nr nc i j] by (metis Pair_inject Suc_leD emptyE list.set(1) nle_le set_ConsD) qed lemma pivot_positions_form_helper_2: shows "sorted_wrt (<) (map fst (pivot_positions_main_gen z A nr nc i j))" proof (induct i j rule: pivot_positions_main_gen.induct[of nr nc A z]) case (1 i j) then show ?case using pivot_positions_main_gen.simps[of z A nr nc i j] by (auto simp: pivot_positions_form_helper_1 Suc_le_lessD) qed lemma sorted_pivot_positions: shows "sorted_wrt (<) (map fst (pivot_positions A))" using pivot_positions_form_helper_2 by (simp add: pivot_positions_form_helper_2 pivot_positions_gen_def) lemma pivot_positions_form: assumes row_ech: "row_echelon_form A" assumes dim_h: "dim_col A \ dim_row A" shows "\ i < (length (pivot_positions A)). fst ((pivot_positions A) ! i) = i" proof clarsimp let ?nr = "dim_row A" let ?nc = "dim_col A" let ?pp = "pivot_positions A :: (nat \ nat) list" fix i assume i_lt: "i < length (pivot_positions A)" have "\f. pivot_fun A f ?nc" using row_ech unfolding row_echelon_form_def by blast then obtain f where pf:"pivot_fun A f ?nc" by blast have all_f_in: "\i < (length ?pp). (i, f i) \ set ?pp" using pivot_positions_contains pf assms by blast have sorted_hyp: "\ (p::nat) (q::nat). p < (length ?pp) \ q < (length ?pp) \ p < q \ (fst (?pp ! p) < fst (?pp ! q))" proof - fix p::nat fix q::nat assume p_lt: "p < q" assume p_welldef: "p < (length ?pp)" assume q_welldef: "q < (length ?pp)" show "fst (?pp ! p) < fst (?pp ! q)" using sorted_pivot_positions p_lt p_welldef q_welldef sorted_wrt_nth_less by fastforce qed have h: "i < (length ?pp) \ fst (pivot_positions A ! i) = i" proof (induct i) case 0 have "\j. fst (pivot_positions A ! j) = 0" by (metis all_f_in fst_conv i_lt in_set_conv_nth length_greater_0_conv list.size(3) not_less0) then obtain j where jth:" fst (pivot_positions A ! j) = 0" by blast have "j \ 0 \ (fst (pivot_positions A ! 0) > 0 \ j \ 0)" by (smt (verit, ccfv_SIG) all_f_in fst_conv i_lt in_set_conv_nth less_nat_zero_code not_gr_zero sorted_hyp) then show ?case using jth neq0_conv by blast next case (Suc i) have ind_h: "i < length (pivot_positions A) \ fst (pivot_positions A ! i) = i" using Suc.hyps by blast have thesis_h: "(Suc i) < length (pivot_positions A) \ fst (pivot_positions A ! (Suc i)) = (Suc i)" proof - assume suc_i_lt: "(Suc i) < length (pivot_positions A)" have fst_i_is: "fst (pivot_positions A ! i) = i" using suc_i_lt ind_h using Suc_lessD by blast have "(\j < (length ?pp). fst (pivot_positions A ! j) = (Suc i))" by (metis suc_i_lt all_f_in fst_conv in_set_conv_nth) then obtain j where jth: "j < (length ?pp) \ fst (pivot_positions A ! j) = (Suc i)" by blast have "j > i" using sorted_hyp apply (auto) by (metis Suc_lessD \fst (pivot_positions A ! i) = i\ jth less_not_refl linorder_neqE_nat n_not_Suc_n suc_i_lt) have "j > (Suc i) \ False" proof - assume j_gt: "j > (Suc i)" then have h1: "fst (pivot_positions A ! (Suc i)) > i" using fst_i_is sorted_pivot_positions using sorted_hyp suc_i_lt by force have "fst (pivot_positions A ! j) > fst (pivot_positions A ! (Suc i))" using jth j_gt sorted_hyp apply (auto) by fastforce then have h2: "fst (pivot_positions A ! (Suc i)) < (Suc i)" using jth by simp show "False" using h1 h2 using not_less_eq by blast qed show "fst (pivot_positions A ! (Suc i)) = (Suc i)" using Suc_lessI \Suc i < j \ False\ \i < j\ jth by blast qed then show ?case by blast qed then show "fst (pivot_positions A ! i) = i" using i_lt by auto qed lemma take_cols_pivot_eq: assumes row_ech: "row_echelon_form A" assumes dim_h: "dim_col A \ dim_row A" shows "take_cols A (map snd (pivot_positions A)) = 1\<^sub>m (length (pivot_positions A)) @\<^sub>r 0\<^sub>m (dim_row A - length (pivot_positions A)) (length (pivot_positions A))" proof - let ?nr = "dim_row A" let ?nc = "dim_col A" have h1: " dim_col (1\<^sub>m (length (pivot_positions A)) @\<^sub>r 0\<^sub>m (dim_row A - length (pivot_positions A)) (length (pivot_positions A))) = (length (pivot_positions A))" by (simp add: append_rows_def) have len_pivot: "length (pivot_positions A) = card {i. i < ?nr \ row A i \ 0\<^sub>v ?nc}" using row_ech pivot_positions(4) row_echelon_form_def by blast have pp_leq_nc: "\f. pivot_fun A f ?nc \ (\i < ?nr. f i \ ?nc)" unfolding pivot_fun_def by meson have pivot_set: "\f. pivot_fun A f ?nc \ set (pivot_positions A) = {(i, f i) | i. i < ?nr \ f i \ ?nc}" using row_ech row_echelon_form_def pivot_positions(1) by (smt (verit) Collect_cong carrier_matI) then have pivot_set_alt: "\f. pivot_fun A f ?nc \ set (pivot_positions A) = {(i, f i) | i. i < ?nr \ row A i \ 0\<^sub>v ?nc}" using pivot_positions pivot_fun_zero_row_iff Collect_cong carrier_mat_triv by (smt (verit, best)) have "\f. pivot_fun A f ?nc \ set (pivot_positions A) = {(i, f i) | i. f i \ ?nc \ i < ?nr \ f i \ ?nc}" using pivot_set pp_leq_nc by auto then have pivot_set_var: "\f. pivot_fun A f ?nc \ set (pivot_positions A) = {(i, f i) | i. i < ?nr \ f i < ?nc}" by auto have "length (map snd (pivot_positions A)) = card (set (map snd (pivot_positions A)))" using row_ech row_echelon_form_def pivot_positions(3) distinct_card[where xs = "map snd (pivot_positions A)"] by (metis carrier_mat_triv) then have "length (map snd (pivot_positions A)) = card (set (pivot_positions A))" by (metis card_distinct distinct_card distinct_map length_map) then have "length (map snd (pivot_positions A)) = card {i. i < ?nr \ row A i \ 0\<^sub>v ?nc}" using pivot_set_alt by (simp add: len_pivot) then have length_asm: "length (map snd (pivot_positions A)) = length (pivot_positions A)" using len_pivot by linarith then have "\a. List.member (map snd (pivot_positions A)) a \ a < dim_col A" proof clarsimp fix a assume a_in: "List.member (map snd (pivot_positions A)) a" have "\v \ set (pivot_positions A). a = snd v" using a_in in_set_member[where xs = "(pivot_positions A)"] apply (auto) by (metis in_set_impl_in_set_zip2 in_set_member length_map snd_conv zip_map_fst_snd) then show "a < dim_col A" using pivot_set_var in_set_member by auto qed then have h2b: "(filter (\y. y < dim_col A) (map snd (pivot_positions A))) = (map snd (pivot_positions A))" by (meson filter_True in_set_member) then have h2a: "length (map ((!) (cols A)) (filter (\y. y < dim_col A) (map snd (pivot_positions A)))) = length (pivot_positions A)" using length_asm by (simp add: h2b) then have h2: "length (pivot_positions A) \ dim_row A \ dim_col (take_cols A (map snd (pivot_positions A))) = (length (pivot_positions A))" unfolding take_cols_def using mat_of_cols_carrier by auto have h_len: "length (pivot_positions A) \ dim_row A \ dim_col (take_cols A (map snd (pivot_positions A))) = dim_col (1\<^sub>m (length (pivot_positions A)) @\<^sub>r 0\<^sub>m (dim_row A - length (pivot_positions A)) (length (pivot_positions A)))" using h1 h2 by (simp add: h1 assms length_pivot_positions_dim_row) have h2: "\i j. length (pivot_positions A) \ dim_row A \ i < dim_row A \ j < dim_col (1\<^sub>m (length (pivot_positions A)) @\<^sub>r 0\<^sub>m (dim_row A - length (pivot_positions A)) (length (pivot_positions A))) \ take_cols A (map snd (pivot_positions A)) $$ (i, j) = (1\<^sub>m (length (pivot_positions A)) @\<^sub>r 0\<^sub>m (dim_row A - length (pivot_positions A)) (length (pivot_positions A))) $$ (i, j)" proof - fix i fix j let ?pp = "(pivot_positions A)" assume len_lt: "length (pivot_positions A) \ dim_row A" assume i_lt: " i < dim_row A" assume j_lt: "j < dim_col (1\<^sub>m (length (pivot_positions A)) @\<^sub>r 0\<^sub>m (dim_row A - length (pivot_positions A)) (length (pivot_positions A)))" let ?w = "((map snd (pivot_positions A)) ! j)" have breaking_it_down: "mat_of_cols (dim_row A) (map ((!) (cols A)) (map snd (pivot_positions A))) $$ (i, j) = ((cols A) ! ?w) $ i" apply (auto) by (metis comp_apply h1 i_lt j_lt length_map mat_of_cols_index nth_map) have h1a: "i < (length ?pp) \ (mat_of_cols (dim_row A) (map ((!) (cols A)) (map snd (pivot_positions A))) $$ (i, j) = (1\<^sub>m (length (pivot_positions A))) $$ (i, j))" proof - (* need to, using row_ech, rely heavily on pivot_fun_def, that num_cols \ num_rows, and row_echelon form*) assume "i < (length ?pp)" have "\f. pivot_fun A f ?nc" using row_ech unfolding row_echelon_form_def by blast then obtain f where "pivot_fun A f ?nc" by blast have j_nc: "j < (length ?pp)" using j_lt by (simp add: h1) then have j_lt_nr: "j < ?nr" using dim_h using len_lt by linarith then have is_this_true: "(pivot_positions A) ! j = (j, f j)" using pivot_positions_form pivot_positions(1)[of A ?nr ?nc f] proof - have "pivot_positions A ! j \ set (pivot_positions A)" using j_nc nth_mem by blast then have "\n. pivot_positions A ! j = (n, f n) \ n < dim_row A \ f n \ dim_col A" using \\A \ carrier_mat (dim_row A) (dim_col A); pivot_fun A f (dim_col A)\ \ set (pivot_positions A) = {(i, f i) |i. i < dim_row A \ f i \ dim_col A}\ \pivot_fun A f (dim_col A)\ by blast then show ?thesis by (metis (no_types) \\A. \row_echelon_form A; dim_row A \ dim_col A\ \ \i dim_h fst_conv j_nc row_ech) qed then have w_is: "?w = f j" by (metis h1 j_lt nth_map snd_conv) have h0: "i = j \ ((cols A) ! ?w) $ i = 1" using w_is pivot_funD(4)[of A ?nr f ?nc i] by (metis \\a. List.member (map snd (pivot_positions A)) a \ a < dim_col A\ \i < length (pivot_positions A)\ \pivot_fun A f (dim_col A)\ cols_length i_lt in_set_member length_asm mat_of_cols_cols mat_of_cols_index nth_mem) have h1: "i \ j \ ((cols A) ! ?w) $ i = 0" using w_is pivot_funD(5) by (metis \\a. List.member (map snd (pivot_positions A)) a \ a < dim_col A\ \pivot_fun A f (dim_col A)\ cols_length h1 i_lt in_set_member j_lt len_lt length_asm less_le_trans mat_of_cols_cols mat_of_cols_index nth_mem) show "(mat_of_cols (dim_row A) (map ((!) (cols A)) (map snd (pivot_positions A))) $$ (i, j) = (1\<^sub>m (length (pivot_positions A))) $$ (i, j))" using h0 h1 breaking_it_down by (metis \i < length (pivot_positions A)\ h2 h_len index_one_mat(1) j_lt len_lt) qed have h1b: "i \ (length ?pp) \ (mat_of_cols (dim_row A) (map ((!) (cols A)) (map snd (pivot_positions A))) $$ (i, j) = 0)" proof - assume i_gt: "i \ (length ?pp)" have h0a: "((cols A) ! ((map snd (pivot_positions A)) ! j)) $ i = (row A i) $ ?w" by (metis \\a. List.member (map snd (pivot_positions A)) a \ a < dim_col A\ cols_length h1 i_lt in_set_member index_row(1) j_lt length_asm mat_of_cols_cols mat_of_cols_index nth_mem) have h0b: "take_rows A [0..r 0\<^sub>m (dim_row A - length (pivot_positions A)) (dim_col A) = A" using assms row_echelon_form_zero_rows[of A] by blast then have h0c: "(row A i) = 0\<^sub>v (dim_col A)" using i_gt by (smt (verit, best) add_diff_cancel_left' add_diff_cancel_right' add_less_cancel_left dim_col_take_rows dim_row_append_rows i_lt index_zero_mat(2) index_zero_mat(3) le_Suc_ex len_lt nat_less_le nle_le row_append_rows row_zero) then show ?thesis using h0a breaking_it_down apply (auto) by (metis \\a. List.member (map snd (pivot_positions A)) a \ a < dim_col A\ h1 in_set_member index_zero_vec(1) j_lt length_asm nth_mem) qed have h1: " mat_of_cols (dim_row A) (map ((!) (cols A)) (map snd (pivot_positions A))) $$ (i, j) = (1\<^sub>m (length (pivot_positions A)) @\<^sub>r 0\<^sub>m (dim_row A - length (pivot_positions A)) (length (pivot_positions A))) $$ (i, j) " using h1a h1b by (smt (verit) add_diff_inverse_nat append_rows_index diff_less_mono h1 i_lt index_one_mat(2) index_one_mat(3) index_zero_mat(1) index_zero_mat(2) index_zero_mat(3) j_lt leD len_lt not_le_imp_less) then show "take_cols A (map snd (pivot_positions A)) $$ (i, j) = (1\<^sub>m (length (pivot_positions A)) @\<^sub>r 0\<^sub>m (dim_row A - length (pivot_positions A)) (length (pivot_positions A))) $$ (i, j)" unfolding take_cols_def by (simp add: h2b) qed show ?thesis unfolding mat_eq_iff using length_pivot_positions_dim_row[OF assms(1)] h_len h2 by auto qed lemma rref_right_mul: assumes "row_echelon_form A" assumes "dim_col A \ dim_row A" shows "take_cols A (map snd (pivot_positions A)) * take_rows A [0..m (length (pivot_positions A)) @\<^sub>r 0\<^sub>m (dim_row A - length (pivot_positions A)) (length (pivot_positions A))" . have 2: "take_cols A (map snd (pivot_positions A)) * take_rows A [0..r 0\<^sub>m (dim_row A - length (pivot_positions A)) (dim_col A)" unfolding 1 apply (simp add: append_rows_mat_mul) by (metis (no_types, lifting) "1" add_right_imp_eq assms dim_col_take_rows dim_row_append_rows dim_row_take_cols index_one_mat(2) index_zero_mat(2) left_mult_one_mat' left_mult_zero_mat' row_echelon_form_zero_rows) from row_echelon_form_zero_rows[OF assms] have "... = A" . thus ?thesis by (simp add: "2") qed context conjugatable_vec_space begin lemma lin_indpt_id: shows "lin_indpt (set (cols (1\<^sub>m n)::'a vec list))" proof - have *: "set (cols (1\<^sub>m n)) = set (rows (1\<^sub>m n))" by (metis cols_transpose transpose_one) have "det (1\<^sub>m n) \ 0" using det_one by auto from det_not_0_imp_lin_indpt_rows[OF _ this] have "lin_indpt (set (rows (1\<^sub>m n)))" using one_carrier_mat by blast thus ?thesis by (simp add: *) qed lemma lin_indpt_take_cols_id: shows "lin_indpt (set (cols (take_cols (1\<^sub>m n) inds)))" proof - have subset_h: "set (cols (take_cols (1\<^sub>m n) inds)) \ set (cols (1\<^sub>m n)::'a vec list)" using cols_take_cols_subset by blast then show ?thesis using lin_indpt_id subset_li_is_li by auto qed lemma cols_id_unit_vecs: shows "cols (1\<^sub>m d) = unit_vecs d" unfolding unit_vecs_def list_eq_iff_nth_eq by auto lemma distinct_cols_id: shows "distinct (cols (1\<^sub>m d)::'a vec list)" by (simp add: conjugatable_vec_space.cols_id_unit_vecs vec_space.unit_vecs_distinct) lemma distinct_map_nth: assumes "distinct ls" assumes "distinct inds" assumes "\j. j \ set inds \ j < length ls" shows "distinct (map ((!) ls) inds)" by (simp add: assms(1) assms(2) assms(3) distinct_map inj_on_nth) lemma distinct_take_cols_id: assumes "distinct inds" shows "distinct (cols (take_cols (1\<^sub>m n) inds) :: 'a vec list)" unfolding take_cols_def apply (subst cols_mat_of_cols) apply (auto intro!: distinct_map_nth simp add: distinct_cols_id) using assms distinct_filter by blast lemma rank_take_cols: assumes "distinct inds" shows "rank (take_cols (1\<^sub>m n) inds) = length (filter ((>) n) inds)" apply (subst lin_indpt_full_rank[of _ "length (filter ((>) n) inds)"]) apply (auto simp add: lin_indpt_take_cols_id) apply (metis (full_types) index_one_mat(2) index_one_mat(3) length_map mat_of_cols_carrier(1) take_cols_def) by (simp add: assms distinct_take_cols_id) lemma rank_mul_left_invertible_mat: fixes A::"'a mat" assumes "invertible_mat A" assumes "A \ carrier_mat n n" assumes "B \ carrier_mat n nc" shows "rank (A * B) = rank B" proof - obtain C where C: "inverts_mat A C" "inverts_mat C A" using assms invertible_mat_def by blast from C have ceq: "C * A = 1\<^sub>m n" by (metis assms(2) carrier_matD(2) index_mult_mat(3) index_one_mat(3) inverts_mat_def) then have *:"B = C*A*B" using assms(3) by auto from rank_mat_mul_left[OF assms(2-3)] have **: "rank (A*B) \ rank B" . have 1: "C \ carrier_mat n n" using C ceq by (metis assms(2) carrier_matD(1) carrier_matI index_mult_mat(3) index_one_mat(3) inverts_mat_def) have 2: "A * B \ carrier_mat n nc" using assms by auto have "rank B = rank (C* A * B)" using * by auto also have "... \ rank (A*B)" using rank_mat_mul_left[OF 1 2] using "1" assms(2) assms(3) by auto ultimately show ?thesis using ** by auto qed lemma invertible_take_cols_rank: fixes A::"'a mat" assumes "invertible_mat A" assumes "A \ carrier_mat n n" assumes "distinct inds" shows "rank (take_cols A inds) = length (filter ((>) n) inds)" proof - have " A = A * 1\<^sub>m n" using assms(2) by auto then have "take_cols A inds = A * take_cols (1\<^sub>m n) inds" by (metis assms(2) one_carrier_mat take_cols_mat_mul) then have "rank (take_cols A inds) = rank (take_cols (1\<^sub>m n) inds)" by (metis assms(1) assms(2) conjugatable_vec_space.rank_mul_left_invertible_mat one_carrier_mat take_cols_carrier_mat) thus ?thesis by (simp add: assms(3) conjugatable_vec_space.rank_take_cols) qed lemma rank_take_cols_leq: assumes R:"R \ carrier_mat n nc" shows "rank (take_cols R ls) \ rank R" proof - from take_cols_mat_mul[OF R] have "take_cols R ls = R * take_cols (1\<^sub>m nc) ls" by (metis assms one_carrier_mat right_mult_one_mat) thus ?thesis by (metis assms one_carrier_mat take_cols_carrier_mat vec_space.rank_mat_mul_right) qed lemma rank_take_cols_geq: assumes R:"R \ carrier_mat n nc" assumes t:"take_cols R ls \ carrier_mat n r" assumes B:"B \ carrier_mat r nc" assumes "R = (take_cols R ls) * B" shows "rank (take_cols R ls) \ rank R" by (metis B assms(4) t vec_space.rank_mat_mul_right) lemma rref_drop_pivots: assumes row_ech: "row_echelon_form R" assumes dims: "R \ carrier_mat n nc" assumes order: "nc \ n" shows "rank (take_cols R (map snd (pivot_positions R))) = rank R" proof - let ?B = "take_rows R [0..r. take_cols R (map snd (pivot_positions R)) \ carrier_mat n r \ ?B \ carrier_mat r nc" proof - have h1: "take_cols R (map snd (pivot_positions R)) \ carrier_mat n (length (pivot_positions R))" using assms by (metis in_set_impl_in_set_zip2 length_map rref_pivot_positions take_cols_carrier_mat_strict zip_map_fst_snd) have "\ f. pivot_fun R f nc" using row_ech unfolding row_echelon_form_def using dims by blast then have "length (pivot_positions R) = card {i. i < n \ row R i \ 0\<^sub>v nc}" using pivot_positions[of R n nc] using dims by auto then have "nc \ length (pivot_positions R)" using order using carrier_matD(1) dims dual_order.trans length_pivot_positions_dim_row row_ech by blast then have "dim_col R \ length (pivot_positions R)" using dims by auto then have h2: "?B \ carrier_mat (length (pivot_positions R)) nc" unfolding take_rows_def using dims by (smt (verit) atLeastLessThan_iff carrier_matD(2) filter_True le_eq_less_or_eq length_map length_pivot_positions_dim_row less_trans map_nth mat_of_cols_carrier(1) row_ech set_upt transpose_carrier_mat transpose_mat_of_rows) show ?thesis using h1 h2 by blast qed (* prove the other two dimensionality assumptions *) have "rank R \ rank (take_cols R (map snd (pivot_positions R)))" using dims ex_r rank_take_cols_geq[where R = "R", where B = "?B", where ls = "(map snd (pivot_positions R))", where nc = "nc"] using equa by blast thus ?thesis using assms(2) conjugatable_vec_space.rank_take_cols_leq le_antisym by blast qed lemma gjs_and_take_cols_var: fixes A::"'a mat" assumes A:"A \ carrier_mat n nc" assumes order: "nc \ n" shows "(take_cols A (map snd (pivot_positions (gauss_jordan_single A)))) = (take_cols_var A (map snd (pivot_positions (gauss_jordan_single A))))" proof - let ?gjs = "(gauss_jordan_single A)" have "\x. List.member (map snd (pivot_positions (gauss_jordan_single A))) x \ x \ dim_col A" using rref_pivot_positions gauss_jordan_single(3) carrier_matD(2) gauss_jordan_single(2) in_set_impl_in_set_zip2 in_set_member length_map less_irrefl less_trans not_le_imp_less zip_map_fst_snd by (metis (no_types, lifting) carrier_mat_triv) then have "(filter (\y. y < dim_col A) (map snd (pivot_positions (gauss_jordan_single A)))) = (map snd (pivot_positions (gauss_jordan_single A)))" by (metis (no_types, lifting) A carrier_matD(2) filter_True gauss_jordan_single(2) gauss_jordan_single(3) in_set_impl_in_set_zip2 length_map rref_pivot_positions zip_map_fst_snd) then show ?thesis unfolding take_cols_def take_cols_var_def by simp qed lemma gauss_jordan_single_rank: fixes A::"'a mat" assumes A:"A \ carrier_mat n nc" assumes order: "nc \ n" shows "rank (take_cols A (map snd (pivot_positions (gauss_jordan_single A)))) = rank A" proof - let ?R = "gauss_jordan_single A" obtain P where P:"P\Units (ring_mat TYPE('a) n undefined)" and i: "?R = P * A" using gauss_jordan_transform[OF A] by (metis A carrier_matD(1) fst_eqD gauss_jordan_single_def surj_pair zero_carrier_mat) have pcarrier: "P \ carrier_mat n n" using P unfolding Units_def by (auto simp add: ring_mat_def) have "invertible_mat P" using P unfolding invertible_mat_def Units_def inverts_mat_def apply auto apply (simp add: ring_mat_simps(5)) by (metis index_mult_mat(2) index_one_mat(2) ring_mat_simps(1) ring_mat_simps(3)) then obtain Pi where Pi: "invertible_mat Pi" "Pi * P = 1\<^sub>m n" proof - assume a1: "\Pi. \invertible_mat Pi; Pi * P = 1\<^sub>m n\ \ thesis" have "dim_row P = n" by (metis (no_types) A assms(1) carrier_matD(1) gauss_jordan_single(2) i index_mult_mat(2)) then show ?thesis using a1 by (metis (no_types) \invertible_mat P\ index_mult_mat(3) index_one_mat(3) invertible_mat_def inverts_mat_def square_mat.simps) qed then have pi_carrier:"Pi \ carrier_mat n n" by (metis carrier_mat_triv index_mult_mat(2) index_one_mat(2) invertible_mat_def square_mat.simps) have R1:"row_echelon_form ?R" using assms(2) gauss_jordan_single(3) by blast have R2: "?R \ carrier_mat n nc" using A assms(2) gauss_jordan_single(2) by auto have Rcm: "take_cols ?R (map snd (pivot_positions ?R)) \ carrier_mat n (length (map snd (pivot_positions ?R)))" apply (rule take_cols_carrier_mat_strict[OF R2]) using rref_pivot_positions[OF R1 R2] by auto have "Pi * ?R = A" using i Pi by (smt (verit, best) A assoc_mult_mat left_mult_one_mat pcarrier pi_carrier) then have "rank (take_cols A (map snd (pivot_positions ?R))) = rank (take_cols (Pi * ?R) (map snd (pivot_positions ?R)))" by auto also have "... = rank ( Pi * take_cols ?R (map snd (pivot_positions ?R)))" by (metis A gauss_jordan_single(2) pi_carrier take_cols_mat_mul) also have "... = rank (take_cols ?R (map snd (pivot_positions ?R)))" by (intro rank_mul_left_invertible_mat[OF Pi(1) pi_carrier Rcm]) also have "... = rank ?R" using assms(2) conjugatable_vec_space.rref_drop_pivots gauss_jordan_single(3) using R1 R2 by blast ultimately show ?thesis using A \P \ carrier_mat n n\ \invertible_mat P\ conjugatable_vec_space.rank_mul_left_invertible_mat i by auto qed lemma lin_indpt_subset_cols: fixes A:: "'a mat" fixes B:: "'a vec set" assumes "A \ carrier_mat n n" assumes inv: "invertible_mat A" assumes "B \ set (cols A)" shows "lin_indpt B" proof - have "det A \ 0" using assms(1) inv invertible_det by blast then have "lin_indpt (set (rows A\<^sup>T))" using assms(1) idom_vec.lin_dep_cols_imp_det_0 by auto thus ?thesis using subset_li_is_li assms(3) by auto qed lemma rank_invertible_subset_cols: fixes A:: "'a mat" fixes B:: "'a vec list" assumes inv: "invertible_mat A" assumes A_square: "A \ carrier_mat n n" assumes set_sub: "set (B) \ set (cols A)" assumes dist_B: "distinct B" shows "rank (mat_of_cols n B) = length B" proof - let ?B_mat = "(mat_of_cols n B)" have h1: "lin_indpt (set(B))" using assms lin_indpt_subset_cols[of A] by auto have "set B \ carrier_vec n" using set_sub A_square cols_dim[of A] by auto then have cols_B: "cols (mat_of_cols n B) = B" using cols_mat_of_cols by auto then have "maximal (set B) (\T. T \ set (B) \ lin_indpt T)" using h1 by (simp add: maximal_def subset_antisym) then have h2: "maximal (set B) (\T. T \ set (cols (mat_of_cols n B)) \ lin_indpt T)" using cols_B by auto have h3: "rank (mat_of_cols n B) = card (set B)" using h1 h2 rank_card_indpt[of ?B_mat] using mat_of_cols_carrier(1) by blast then show ?thesis using assms distinct_card by auto qed end end \ No newline at end of file diff --git a/thys/BenOr_Kozen_Reif/Renegar_Proofs.thy b/thys/BenOr_Kozen_Reif/Renegar_Proofs.thy --- a/thys/BenOr_Kozen_Reif/Renegar_Proofs.thy +++ b/thys/BenOr_Kozen_Reif/Renegar_Proofs.thy @@ -1,2779 +1,2779 @@ theory Renegar_Proofs imports Renegar_Algorithm BKR_Proofs begin (* Note that there is significant overlap between Renegar and BKR in general, so there is some similarity between this file and BKR_Proofs.thy The main difference is that the RHS vector in Renegar is different from the RHS vector in BKR In BKR, all of the qs are assumed to be relatively prime to p. Renegar removes this assumption. In general, the _R's on definition and lemma names in this file are to emphasize that we are working with Renegar style. *) section "Tarski Queries Changed" lemma construct_NofI_R_relation: fixes p:: "real poly" fixes I1:: "real poly list" fixes I2:: "real poly list" shows "construct_NofI_R p I1 I2 = construct_NofI (sum_list (map power2 (p # I1))) I2" unfolding construct_NofI_R_def construct_NofI_def by metis lemma sum_list_map_power2: shows "sum_list (map power2 ls) \ (0::real poly)" apply (induct ls) by auto lemma sum_list_map_power2_poly: shows "poly (sum_list (map power2 (ls::real poly list))) x \ (0::real)" apply (induct ls) by auto lemma construct_NofI_R_prop_helper: fixes p:: "real poly" fixes I1:: "real poly list" fixes I2:: "real poly list" assumes nonzero: "p\0" shows "construct_NofI_R p I1 I2 = rat_of_int (int (card {x. poly (sum_list (map (\x. x^2) (p # I1))) x = 0 \ poly (prod_list I2) x > 0}) - int (card {x. poly (sum_list (map (\x. x^2) (p # I1))) x = 0 \ poly (prod_list I2) x < 0}))" proof - show ?thesis unfolding construct_NofI_R_relation[of p I1 I2] apply (subst construct_NofI_prop[of _ I2]) apply auto using assms sum_list_map_power2 by (metis le_add_same_cancel1 power2_less_eq_zero_iff) qed lemma zer_iff: fixes p:: "real poly" fixes ls:: "real poly list" shows "poly (sum_list (map (\x. x^2) ls)) x = 0 \ (\i \ set ls. poly i x = 0)" proof (induct ls) case Nil then show ?case by auto next case (Cons a I1) then show ?case apply simp apply (subst add_nonneg_eq_0_iff) using zero_le_power2 apply blast using sum_list_map_power2_poly apply presburger by simp qed lemma construct_NofI_prop_R: fixes p:: "real poly" fixes I1:: "real poly list" fixes I2:: "real poly list" assumes nonzero: "p\0" shows "construct_NofI_R p I1 I2 = rat_of_int (int (card {x. poly p x = 0 \ (\q \ set I1. poly q x = 0) \ poly (prod_list I2) x > 0}) - int (card {x. poly p x = 0 \ (\q \ set I1. poly q x = 0) \ poly (prod_list I2) x < 0}))" unfolding construct_NofI_R_prop_helper[OF nonzero] using zer_iff apply auto by (smt (verit, del_insts) Collect_cong sum_list_map_power2_poly zero_le_power2 zero_less_power2) section "Matrix Equation" definition map_sgas:: "rat list \ rat list" where "map_sgas l = map (\r. (1 - r^2)) l" definition z_R:: "(nat list*nat list) \ rat list \ rat" where "z_R index_list sign_asg \ (prod_list (map (nth (map_sgas sign_asg)) (fst(index_list))))*(prod_list (map (nth sign_asg) (snd(index_list))))" definition mtx_row_R:: "rat list list \ (nat list * nat list) \ rat list" where "mtx_row_R sign_list index_list \ (map ((z_R index_list)) sign_list)" definition matrix_A_R:: "rat list list \ (nat list * nat list) list \ rat mat" where "matrix_A_R sign_list subset_list = (mat_of_rows_list (length sign_list) (map (\i .(mtx_row_R sign_list i)) subset_list))" definition all_list_constr_R:: "(nat list*nat list) list \ nat \ bool" where "all_list_constr_R L n \ (\x. List.member L x \ (list_constr (fst x) n \ list_constr (snd x) n))" definition alt_matrix_A_R:: "rat list list \ (nat list*nat list) list \ rat mat" where "alt_matrix_A_R signs subsets = (mat (length subsets) (length signs) (\(i, j). z_R (subsets ! i) (signs ! j)))" lemma alt_matrix_char_R: "alt_matrix_A_R signs subsets = matrix_A_R signs subsets" proof - have h0: "(\i j. i < length subsets \ j < length signs \ map (\index_list. map (z_R index_list) signs) subsets ! i ! j = z_R (subsets ! i) (signs ! j))" proof - fix i fix j assume i_lt: "i < length subsets" assume j_lt: "j < length signs" show "((map (\index_list. map (z_R index_list) signs) subsets) ! i) ! j = z_R (subsets ! i) (signs ! j)" proof - have h0: "(map (\index_list. map (z_R index_list) signs) subsets) ! i = map (z_R (subsets ! i)) signs" using nth_map i_lt by blast then show ?thesis using nth_map j_lt by simp qed qed have h: " mat (length subsets) (length signs) (\(i, j). z_R (subsets ! i) (signs ! j)) = mat (length subsets) (length signs) (\(i, y). map (\index_list. map (z_R index_list) signs) subsets ! i ! y)" using h0 eq_matI[where A = "mat (length subsets) (length signs) (\(i, j). z_R (subsets ! i) (signs ! j))", where B = "mat (length subsets) (length signs) (\(i, y). map (\index_list. map (z_R index_list) signs) subsets ! i ! y)"] by auto show ?thesis unfolding alt_matrix_A_R_def matrix_A_R_def mat_of_rows_list_def apply (auto) unfolding mtx_row_R_def using h by blast qed lemma subsets_are_rows_R: "\i < (length subsets). row (alt_matrix_A_R signs subsets) i = vec (length signs) (\j. z_R (subsets ! i) (signs ! j))" unfolding row_def unfolding alt_matrix_A_R_def by auto lemma signs_are_cols_R: "\i < (length signs). col (alt_matrix_A_R signs subsets) i = vec (length subsets) (\j. z_R (subsets ! j) (signs ! i))" unfolding col_def unfolding alt_matrix_A_R_def by auto definition consistent_sign_vec::"real poly list \ real \ rat list" where "consistent_sign_vec qs x \ map (\ q. if (poly q x > 0) then (1::rat) else (if (poly q x = 0) then (0::rat) else (-1::rat))) qs" definition construct_lhs_vector_R:: "real poly \ real poly list \ rat list list \ rat vec" where "construct_lhs_vector_R p qs signs \ vec_of_list (map (\w. rat_of_int (int (length (filter (\v. v = w) (map (consistent_sign_vec qs) (characterize_root_list_p p)))))) signs)" (* The ith entry of LHS vector is the number of (distinct) real zeros of p where the sign vector of the qs is the ith entry of signs.*) (* Putting all of the pieces of the construction together *) definition satisfy_equation_R:: "real poly \ real poly list \ (nat list*nat list) list \ rat list list \ bool" where "satisfy_equation_R p qs subset_list sign_list = (mult_mat_vec (matrix_A_R sign_list subset_list) (construct_lhs_vector_R p qs sign_list) = (construct_rhs_vector_R p qs subset_list))" (* Recharacterize the LHS vector *) lemma construct_lhs_vector_clean_R: assumes "p \ 0" assumes "i < length signs" shows "(construct_lhs_vector_R p qs signs) $ i = card {x. poly p x = 0 \ ((consistent_sign_vec qs x) = (nth signs i))}" proof - from poly_roots_finite[OF assms(1)] have "finite {x. poly p x = 0}" . then have eq: "(Collect ((\v. v = signs ! i) \ consistent_sign_vec qs) \ set (sorted_list_of_set {x. poly p x = 0})) = {x. poly p x = 0 \ consistent_sign_vec qs x = signs ! i}" by auto show ?thesis unfolding construct_lhs_vector_R_def vec_of_list_index characterize_root_list_p_def apply auto apply (subst nth_map[OF assms(2)]) apply auto apply (subst distinct_length_filter) apply (auto) using eq by auto qed lemma construct_lhs_vector_cleaner_R: assumes "p \ 0" shows "(construct_lhs_vector_R p qs signs) = vec_of_list (map (\s. rat_of_int (card {x. poly p x = 0 \ ((consistent_sign_vec qs x) = s)})) signs)" apply (rule eq_vecI) apply (auto simp add: construct_lhs_vector_clean_R[OF assms] ) apply (simp add: vec_of_list_index) unfolding construct_lhs_vector_R_def using assms construct_lhs_vector_clean_R construct_lhs_vector_def apply auto[1] apply (simp add: construct_lhs_vector_R_def) by auto (* Show that because our consistent sign vectors consist of 1, 0's, and -1's, z returns 1, 0, or -1 when applied to a consistent sign vector *) lemma z_signs_R2: fixes I:: "nat list" fixes signs:: "rat list" assumes lf: "list_all (\i. i < length signs) I" assumes la: "list_all (\s. s = 1 \ s = 0 \ s = -1) signs" shows "(prod_list (map (nth signs) I)) = 1 \ (prod_list (map (nth signs) I)) = 0 \ (prod_list (map (nth signs) I)) = -1" using assms proof (induct I) case Nil then show ?case by auto next case (Cons a I) moreover have eo: "signs ! a = 1 \ signs ! a = 0 \ signs ! a = -1" using assms by (smt (verit, del_insts) calculation(2) list_all_length list_all_simps(1)) have "prod_list (map ((!) (map_sgas signs)) (a # I)) = (1 - (signs ! a)^2)*prod_list (map ((!) (map_sgas signs)) (I))" unfolding map_sgas_def apply (auto) using calculation(2) by auto then show ?case using eo using Cons.hyps calculation(2) la by auto qed lemma z_signs_R1: fixes I:: "nat list" fixes signs:: "rat list" assumes lf: "list_all (\i. i < length signs) I" assumes la: "list_all (\s. s = 1 \ s = 0 \ s = -1) signs" shows "(prod_list (map (nth (map_sgas signs)) I)) = 1 \ (prod_list (map (nth (map_sgas signs)) I)) = 0" using assms proof (induct I) case Nil then show ?case by auto next case (Cons a I) moreover have "signs ! a = 1 \ signs ! a = 0 \ signs ! a = -1" using assms by (smt (verit, best) calculation(2) list_all_length list_all_simps(1)) then have eo: "1 - (signs ! a)^2 = 1 \ 1 - (signs !a)^2 = 0" using cancel_comm_monoid_add_class.diff_cancel diff_zero by fastforce have "prod_list (map ((!) (map_sgas signs)) (a # I)) = (1 - (signs ! a)^2)*prod_list (map ((!) (map_sgas signs)) (I))" unfolding map_sgas_def apply (auto) using calculation(2) by auto then show ?case using eo using Cons.hyps calculation(2) la by auto qed lemma z_signs_R: fixes I:: "(nat list * nat list)" fixes signs:: "rat list" assumes lf: "list_all (\i. i < length signs) (fst(I))" assumes ls: "list_all (\i. i < length signs) (snd(I))" assumes la: "list_all (\s. s = 1 \ s = 0 \ s = -1) signs" shows "(z_R I signs = 1) \ (z_R I signs = 0) \ (z_R I signs = -1)" using assms z_signs_R2 z_signs_R1 unfolding z_R_def apply (auto) by (metis (no_types, lifting) mult_cancel_left1 mult_minus1_right) lemma z_lemma_R: fixes I:: "nat list * nat list" fixes sign:: "rat list" assumes consistent: "sign \ set (characterize_consistent_signs_at_roots p qs)" assumes welldefined1: "list_constr (fst I) (length qs)" assumes welldefined2: "list_constr (snd I) (length qs)" shows "(z_R I sign = 1) \ (z_R I sign = 0) \ (z_R I sign = -1)" proof (rule z_signs_R) have same: "length sign = length qs" using consistent using characterize_consistent_signs_at_roots_def signs_at_def by fastforce thus "(list_all (\i. i < length sign) (fst I))" using welldefined1 by (auto simp add: list_constr_def characterize_consistent_signs_at_roots_def consistent_sign_vec_copr_def) thus "(list_all (\i. i < length sign) (snd I))" using same welldefined2 by (auto simp add: list_constr_def characterize_consistent_signs_at_roots_def consistent_sign_vec_copr_def) show "list_all (\s. s = 1 \ s = 0 \ s = - 1) sign" using consistent apply (auto simp add: list.pred_map characterize_consistent_signs_at_roots_def consistent_sign_vec_def) using Ball_set by (simp add: list_all_length signs_at_def squash_def) qed (* Show that all consistent sign vectors on roots of polynomials are in characterize_consistent_signs_at_roots *) lemma in_set_R: fixes p:: "real poly" assumes nonzero: "p\0" fixes qs:: "real poly list" fixes sign:: "rat list" fixes x:: "real" assumes root_p: "x \ {x. poly p x = 0}" assumes sign_fix: "sign = consistent_sign_vec qs x" shows "sign \ set (characterize_consistent_signs_at_roots p qs)" proof - have h1: "consistent_sign_vec qs x \ set (remdups (map (signs_at qs) (sorted_list_of_set {x. poly p x = 0})))" unfolding consistent_sign_vec_def signs_at_def squash_def using root_p nonzero poly_roots_finite set_sorted_list_of_set apply (auto) by (smt (verit, ccfv_SIG) Collect_cong comp_def image_eqI map_eq_conv mem_Collect_eq poly_roots_finite set_sorted_list_of_set) thus ?thesis unfolding characterize_consistent_signs_at_roots_def characterize_root_list_p_def using sign_fix by blast qed lemma consistent_signs_prop_R: fixes p:: "real poly" assumes nonzero: "p\0" fixes qs:: "real poly list" fixes sign:: "rat list" fixes x:: "real" assumes root_p: "x \ {x. poly p x = 0}" assumes sign_fix: "sign = consistent_sign_vec qs x" shows "list_all (\s. s = 1 \ s = 0 \ s = -1) sign" using assms unfolding consistent_sign_vec_def squash_def apply (auto) by (smt (z3) length_map list_all_length nth_map) (* The next few lemmas relate z_R to the signs of the product of subsets of polynomials of qs *) lemma horiz_vector_helper_pos_ind_R1: fixes p:: "real poly" assumes nonzero: "p\0" fixes qs:: "real poly list" fixes I:: "nat list" fixes sign:: "rat list" fixes x:: "real" assumes root_p: "x \ {x. poly p x = 0}" assumes sign_fix: "sign = consistent_sign_vec qs x" assumes asm: "list_constr I (length qs)" shows "(prod_list (map (nth (map_sgas sign)) I)) = 1 \ (\p \ set (retrieve_polys qs I). poly p x = 0)" using asm proof (induction "I") case Nil then show ?case unfolding map_sgas_def apply (auto) by (simp add: retrieve_polys_def) next case (Cons a xa) then have same0: "(prod_list (map (nth (map_sgas sign)) xa)) = 1 \ (\p \ set (retrieve_polys qs xa). poly p x = 0)" unfolding list_constr_def by auto have welldef: "a < length qs" using Cons.prems assms unfolding list_constr_def by auto then have h: "prod_list (map ((!) (map_sgas sign)) (a#xa)) = (1 - (sign ! a)^2)*(prod_list (map ((!) (map_sgas sign)) (xa)))" unfolding map_sgas_def using assms apply (auto) by (metis (no_types, lifting) consistent_sign_vec_def length_map nth_map) have "list_all (\s. s = 1 \ s = 0 \ s = -1) sign" using sign_fix unfolding consistent_sign_vec_def squash_def apply (auto) by (smt (z3) length_map list_all_length nth_map) then have eo: "(prod_list (map ((!) (map_sgas sign)) (xa))) = 0 \ (prod_list (map ((!) (map_sgas sign)) (xa))) = 1" using z_signs_R1 assms Cons.prems consistent_sign_vec_def length_map list_all_simps(1) length_map list_all_length list_constr_def by (smt (verit, best)) have "(sign ! a)^2 = 1 \ (sign ! a)^2 = 0" using sign_fix welldef unfolding consistent_sign_vec_def by auto then have s1: "(prod_list (map (nth (map_sgas sign)) (a#xa))) = 1 \ ((sign ! a)^2 = 0 \ (prod_list (map ((!) (map_sgas sign)) (xa))) = 1)" using eo h by (metis (mono_tags, opaque_lifting) cancel_comm_monoid_add_class.diff_cancel diff_zero mult.left_neutral mult_zero_left) have "(sign ! a)^2 = 0 \ (poly (qs ! a) x = 0)" using sign_fix unfolding consistent_sign_vec_def welldef apply (auto) apply (smt (z3) class_field.neg_1_not_0 class_field.zero_not_one nth_map welldef) by (smt (z3) nth_map welldef) then have same1:"(prod_list (map (nth (map_sgas sign)) (a#xa))) = 1 \ ((poly (qs ! a) x = 0) \ (prod_list (map ((!) (map_sgas sign)) (xa))) = 1)" using s1 by auto have "set (retrieve_polys qs (a#xa)) = {(qs ! a)} \ set (retrieve_polys qs xa)" by (metis (no_types, lifting) insert_is_Un list.simps(15) list.simps(9) retrieve_polys_def) then have same2:"(\p \ set (retrieve_polys qs (a#xa)). poly p x = 0) = ((poly (qs ! a) x = 0) \ (\p \ set (retrieve_polys qs (xa)). poly p x = 0))" by auto then show ?case using same0 same1 same2 assms by auto qed lemma csv_length_same_as_qlist: fixes p:: "real poly" assumes nonzero: "p\0" fixes qs:: "real poly list" fixes sign:: "rat list" fixes x:: "real" assumes root_p: "x \ {x. poly p x = 0}" assumes sign_fix: "sign = consistent_sign_vec qs x" shows "length sign = length qs" using assms unfolding consistent_sign_vec_def by auto lemma horiz_vector_helper_zer_ind_R2: fixes p:: "real poly" assumes nonzero: "p\0" fixes qs:: "real poly list" fixes I:: "nat list" fixes sign:: "rat list" fixes x:: "real" assumes root_p: "x \ {x. poly p x = 0}" assumes sign_fix: "sign = consistent_sign_vec qs x" assumes asm: "list_constr I (length qs)" shows "(prod_list (map (nth sign) I)) = 0 \ (poly (prod_list (retrieve_polys qs I)) x = 0)" using asm proof (induction "I") case Nil then show ?case unfolding map_sgas_def apply (auto) unfolding retrieve_polys_def by auto next case (Cons a xa) have "poly (prod_list (retrieve_polys qs (a # xa))) x = (poly (qs ! a) x)*poly (prod_list (retrieve_polys qs (xa))) x" by (simp add: retrieve_polys_def) then show ?case using Cons.prems by (smt (z3) Cons.IH class_field.neg_1_not_0 class_field.zero_not_one consistent_sign_vec_def list.simps(9) list_all_simps(1) list_constr_def mult_eq_0_iff nth_map prod_list.Cons sign_fix) qed lemma horiz_vector_helper_pos_ind_R2: fixes p:: "real poly" assumes nonzero: "p\0" fixes qs:: "real poly list" fixes I:: "nat list" fixes sign:: "rat list" fixes x:: "real" assumes root_p: "x \ {x. poly p x = 0}" assumes sign_fix: "sign = consistent_sign_vec qs x" assumes asm: "list_constr I (length qs)" shows "(prod_list (map (nth sign) I)) = 1 \ (poly (prod_list (retrieve_polys qs I)) x > 0)" using asm proof (induction "I") case Nil then show ?case unfolding map_sgas_def apply (auto) unfolding retrieve_polys_def by auto next case (Cons a xa) then have ih: "(prod_list (map ((!) sign) xa) = 1) = (0 < poly (prod_list (retrieve_polys qs xa)) x)" unfolding list_constr_def by auto have lensame: "length sign = length qs" using sign_fix csv_length_same_as_qlist[of p x sign qs] nonzero root_p by auto have "poly (prod_list (retrieve_polys qs (a # xa))) x = (poly (qs ! a) x)*poly (prod_list (retrieve_polys qs (xa))) x" by (simp add: retrieve_polys_def) then have iff1: "(poly (prod_list (retrieve_polys qs (a#xa))) x > 0) \ (((poly (qs ! a) x) > 0 \ poly (prod_list (retrieve_polys qs (xa))) x > 0) \ ((poly (qs ! a) x) < 0 \ poly (prod_list (retrieve_polys qs (xa))) x < 0))" by (metis zero_less_mult_iff) have prodsame: "(prod_list (map (nth sign) (a#xa))) = (sign ! a)* (prod_list (map (nth sign) (xa)))" using lensame Cons.prems unfolding list_constr_def by auto have sagt: "sign ! a = 1 \ (poly (qs ! a) x) > 0" using assms unfolding consistent_sign_vec_def apply (auto) apply (smt (verit, best) Cons.prems list_all_simps(1) list_constr_def neg_equal_zero nth_map zero_neq_one) by (smt (verit, ccfv_threshold) Cons.prems list_all_simps(1) list_constr_def nth_map) have salt: "sign ! a = -1 \ (poly (qs ! a) x) < 0" using assms unfolding consistent_sign_vec_def apply (auto) apply (smt (verit, ccfv_SIG) Cons.prems less_minus_one_simps(1) less_minus_one_simps(3) list_all_simps(1) list_constr_def neg_0_less_iff_less nth_map) by (smt (verit, best) Cons.prems list_all_simps(1) list_constr_def nth_map) have h1: "((poly (qs ! a) x) > 0 \ poly (prod_list (retrieve_polys qs (xa))) x > 0) \ (prod_list (map (nth sign) (a#xa))) = 1" using prodsame sagt ih by auto have eo: "(prod_list (map ((!) sign) xa) = 1) \ (prod_list (map ((!) sign) xa) = 0) \ (prod_list (map ((!) sign) xa) = -1)" using Cons.prems consistent_signs_prop_R lensame list_constr_def nonzero root_p sign_fix z_signs_R2 by auto have d1:"(prod_list (map ((!) sign) xa) = -1) \ (0 > poly (prod_list (retrieve_polys qs xa)) x)" proof - assume "(prod_list (map ((!) sign) xa) = -1) " then show "(0 > poly (prod_list (retrieve_polys qs xa)) x)" using prodsame salt ih assms Cons.prems class_field.neg_1_not_0 equal_neg_zero horiz_vector_helper_zer_ind_R2 linorder_neqE_linordered_idom list_all_simps(1) list_constr_def apply (auto) apply (smt (verit, ccfv_threshold) class_field.neg_1_not_0 list.set_map list_all_length semidom_class.prod_list_zero_iff) by (smt (verit, ccfv_threshold) class_field.neg_1_not_0 list.set_map list_all_length semidom_class.prod_list_zero_iff) qed have d2: "(0 > poly (prod_list (retrieve_polys qs xa)) x) \ (prod_list (map ((!) sign) xa) = -1)" using eo assms horiz_vector_helper_zer_ind_R2[where p = "p", where x = "x", where sign = "sign", where I ="I"] apply (auto) using ih apply force by (metis (full_types, lifting) Cons.prems class_field.neg_1_not_0 horiz_vector_helper_zer_ind_R2 ih imageI list.set_map list_all_simps(1) list_constr_def mem_Collect_eq neg_equal_0_iff_equal semidom_class.prod_list_zero_iff) have "(prod_list (map ((!) sign) xa) = -1) = (0 > poly (prod_list (retrieve_polys qs xa)) x)" using d1 d2 by blast then have h2: "((poly (qs ! a) x) < 0 \ poly (prod_list (retrieve_polys qs (xa))) x < 0) \ (prod_list (map (nth sign) (a#xa))) = 1" using prodsame salt ih by auto have h3: "(prod_list (map (nth sign) (a#xa))) = 1 \ (((poly (qs ! a) x) > 0 \ poly (prod_list (retrieve_polys qs (xa))) x > 0) \ ((poly (qs ! a) x) < 0 \ poly (prod_list (retrieve_polys qs (xa))) x < 0))" using prodsame salt ih assms horiz_vector_helper_zer_ind_R2[where p = "p", where x = "x", where sign = "sign", where I ="I"] by (smt (verit, ccfv_threshold) Cons.prems \poly (prod_list (retrieve_polys qs (a # xa))) x = poly (qs ! a) x * poly (prod_list (retrieve_polys qs xa)) x\ horiz_vector_helper_zer_ind_R2 mem_Collect_eq mult_cancel_left1 mult_not_zero sagt) then show ?case using h1 h2 h3 iff1 Cons.prems by auto qed lemma horiz_vector_helper_pos_ind_R: fixes p:: "real poly" assumes nonzero: "p\0" fixes qs:: "real poly list" fixes I:: "nat list * nat list" fixes sign:: "rat list" fixes x:: "real" assumes root_p: "x \ {x. poly p x = 0}" assumes sign_fix: "sign = consistent_sign_vec qs x" assumes asm1: "list_constr (fst I) (length qs)" assumes asm2: "list_constr (snd I) (length qs)" shows "((\p \ set (retrieve_polys qs (fst I)). poly p x = 0) \ (poly (prod_list (retrieve_polys qs (snd I))) x > 0)) \ (z_R I sign = 1)" proof - have len: "length sign = length qs" using sign_fix csv_length_same_as_qlist[of p x sign qs] nonzero root_p by auto have d1: "((\p \ set (retrieve_polys qs (fst I)). poly p x = 0) \ (poly (prod_list (retrieve_polys qs (snd I))) x > 0)) \ (z_R I sign = 1)" using assms horiz_vector_helper_pos_ind_R1[where p = "p", where qs = "qs", where sign = "sign", where x = "x", where I = "fst I"] horiz_vector_helper_pos_ind_R2[where p = "p", where qs = "qs", where sign = "sign", where x = "x", where I = "snd I"] unfolding z_R_def by auto have d2: "(z_R I sign = 1) \ ((\p \ set (retrieve_polys qs (fst I)). poly p x = 0) \ (poly (prod_list (retrieve_polys qs (snd I))) x > 0))" proof - have h1: "(z_R I sign = 1) \ (\p \ set (retrieve_polys qs (fst I)). poly p x = 0)" proof - have "(prod_list (map (nth (map_sgas sign)) (fst I))) = 1 \ (prod_list (map (nth (map_sgas sign)) (fst I))) = 0" using len consistent_signs_prop_R[where p = "p", where qs = "qs", where x = "x", where sign = "sign"] z_signs_R1[where signs = "sign", where I = "fst I"] assms unfolding list_constr_def by auto then show ?thesis using z_signs_R1[where signs = "sign", where I = "fst I"] horiz_vector_helper_pos_ind_R1[where sign = "sign", where I = "fst I", where p = "p", where x = "x"] assms apply (auto) by (metis (mono_tags, opaque_lifting) \prod_list (map ((!) (map_sgas sign)) (fst I)) = 1 \ prod_list (map ((!) (map_sgas sign)) (fst I)) = 0\ mult_zero_left z_R_def) qed then have h2: "(z_R I sign = 1) \ (poly (prod_list (retrieve_polys qs (snd I))) x > 0)" unfolding z_R_def using assms horiz_vector_helper_pos_ind_R2[where p = "p", where x = "x", where sign = "sign", where qs = "qs", where I ="snd I"] by (metis horiz_vector_helper_pos_ind_R1 mult.left_neutral) show ?thesis using h1 h2 by auto qed show ?thesis using d1 d2 by blast qed lemma horiz_vector_helper_pos_R: fixes p:: "real poly" assumes nonzero: "p\0" fixes qs:: "real poly list" fixes I:: "nat list*nat list" fixes sign:: "rat list" fixes x:: "real" assumes root_p: "x \ {x. poly p x = 0}" assumes sign_fix: "sign = consistent_sign_vec qs x" assumes welldefined1: "list_constr (fst I) (length qs)" assumes welldefined2: "list_constr (snd I) (length qs)" shows "((\p \ set (retrieve_polys qs (fst I)). poly p x = 0) \ (poly (prod_list (retrieve_polys qs (snd I))) x > 0)) \ (z_R I sign = 1)" using horiz_vector_helper_pos_ind_R using nonzero root_p sign_fix welldefined1 welldefined2 by blast lemma horiz_vector_helper_neg_R: fixes p:: "real poly" assumes nonzero: "p\0" fixes qs:: "real poly list" fixes I:: "nat list*nat list" fixes sign:: "rat list" fixes x:: "real" assumes root_p: "x \ {x. poly p x = 0}" assumes sign_fix: "sign = consistent_sign_vec qs x" assumes welldefined1: "list_constr (fst I) (length qs)" assumes welldefined2: "list_constr (snd I) (length qs)" shows "((\p \ set (retrieve_polys qs (fst I)). poly p x = 0) \ (poly (prod_list (retrieve_polys qs (snd I))) x < 0)) \ (z_R I sign = -1)" proof - have set_hyp: "sign \ set (characterize_consistent_signs_at_roots p qs)" using in_set_R[of p x sign qs] nonzero root_p sign_fix by blast have z_hyp: "((z_R I sign = 1) \ (z_R I sign = 0) \ (z_R I sign = -1))" using welldefined1 welldefined2 set_hyp z_lemma_R[where sign="sign", where I = "I", where p="p", where qs="qs"] by blast have d1: "((\p \ set (retrieve_polys qs (fst I)). poly p x = 0) \ (poly (prod_list (retrieve_polys qs (snd I))) x < 0)) \ (z_R I sign = -1)" using horiz_vector_helper_pos_R using nonzero root_p sign_fix welldefined1 welldefined2 by (smt (verit, ccfv_threshold) horiz_vector_helper_pos_ind_R1 horiz_vector_helper_zer_ind_R2 mem_Collect_eq mult_eq_0_iff z_R_def z_hyp) have d2: "(z_R I sign = -1) \ ((\p \ set (retrieve_polys qs (fst I)). poly p x = 0) \ (poly (prod_list (retrieve_polys qs (snd I))) x < 0))" using horiz_vector_helper_pos_ind_R1 horiz_vector_helper_zer_ind_R2 nonzero root_p sign_fix welldefined1 welldefined2 by (smt (verit, ccfv_threshold) class_field.neg_1_not_0 consistent_sign_vec_def consistent_signs_prop_R equal_neg_zero horiz_vector_helper_pos_ind_R2 length_map list_all_length list_constr_def mem_Collect_eq mem_Collect_eq mult_cancel_left1 mult_not_zero retrieve_polys_def z_R_def z_signs_R1 zero_neq_one) then show ?thesis using d1 d2 by linarith qed lemma lhs_dot_rewrite: fixes p:: "real poly" fixes qs:: "real poly list" fixes I:: "nat list*nat list" fixes signs:: "rat list list" assumes nonzero: "p\0" shows "(vec_of_list (mtx_row_R signs I) \ (construct_lhs_vector_R p qs signs)) = sum_list (map (\s. (z_R I s) * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) signs)" proof - have "p \ 0" using nonzero by auto from construct_lhs_vector_cleaner[OF this] have rhseq: "construct_lhs_vector_R p qs signs = vec_of_list (map (\s. rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) signs)" using construct_lhs_vector_cleaner_R nonzero by presburger have "(vec_of_list (mtx_row_R signs I) \ (construct_lhs_vector_R p qs signs)) = sum_list (map2 (*) (mtx_row_R signs I) (map (\s. rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) signs))" unfolding rhseq apply (intro vec_of_list_dot_rewrite) by (auto simp add: mtx_row_R_def) thus ?thesis unfolding mtx_row_R_def using map2_map_map by (auto simp add: map2_map_map) qed (* If we have a superset of the signs, we can drop to just the consistent ones *) lemma construct_lhs_vector_drop_consistent_R: fixes p:: "real poly" fixes qs:: "real poly list" fixes I:: "nat list*nat list" fixes signs:: "rat list list" assumes nonzero: "p\0" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots p qs) \ set(signs)" assumes welldefined1: "list_constr (fst I) (length qs)" assumes welldefined2: "list_constr (snd I) (length qs)" shows "(vec_of_list (mtx_row_R signs I) \ (construct_lhs_vector_R p qs signs)) = (vec_of_list (mtx_row_R (characterize_consistent_signs_at_roots p qs) I) \ (construct_lhs_vector_R p qs (characterize_consistent_signs_at_roots p qs)))" proof - have h0: "\ sgn. sgn \ set signs \ sgn \ consistent_sign_vec qs ` set (characterize_root_list_p p) \ 0 < rat_of_nat (card {xa. poly p xa = 0 \ consistent_sign_vec qs xa = sgn}) \ z_R I sgn = 0" proof - have "\ sgn. sgn \ set signs \ sgn \ consistent_sign_vec qs ` set (characterize_root_list_p p) \ 0 < rat_of_int (card {xa. poly p xa = 0 \ consistent_sign_vec qs xa = sgn}) \ {xa. poly p xa = 0 \ consistent_sign_vec qs xa = sgn} \ {}" by fastforce then show ?thesis proof - { fix iis :: "rat list" have ff1: "0 \ p" using nonzero rsquarefree_def by blast obtain rr :: "(real \ bool) \ real" where ff2: "\p. p (rr p) \ Collect p = {}" - by moura + by (metis Collect_empty_eq) { assume "\is. is = iis \ {r. poly p r = 0 \ consistent_sign_vec qs r = is} \ {}" then have "\is. consistent_sign_vec qs (rr (\r. poly p r = 0 \ consistent_sign_vec qs r = is)) = iis \ {r. poly p r = 0 \ consistent_sign_vec qs r = is} \ {}" using ff2 by (metis (mono_tags, lifting)) then have "\r. poly p r = 0 \ consistent_sign_vec qs r = iis" using ff2 by smt then have "iis \ consistent_sign_vec qs ` set (sorted_list_of_set {r. poly p r = 0})" using ff1 poly_roots_finite by (metis (mono_tags) imageI mem_Collect_eq set_sorted_list_of_set) } then have "iis \ set signs \ iis \ consistent_sign_vec qs ` set (characterize_root_list_p p) \ \ 0 < rat_of_int (int (card {r. poly p r = 0 \ consistent_sign_vec qs r = iis}))" by (metis (no_types) \\sgn. sgn \ set signs \ sgn \ consistent_sign_vec qs ` set (characterize_root_list_p p) \ 0 < rat_of_int (int (card {xa. poly p xa = 0 \ consistent_sign_vec qs xa = sgn})) \ {xa. poly p xa = 0 \ consistent_sign_vec qs xa = sgn} \ {}\ characterize_root_list_p_def) } then show ?thesis by fastforce qed qed then have "\ sgn. sgn \ set signs \ sgn \ consistent_sign_vec qs ` set (characterize_root_list_p p) \ ((0 = rat_of_nat (card {xa. poly p xa = 0 \ consistent_sign_vec qs xa = sgn}) \ z_R I sgn = 0))" by auto then have hyp: "\ s. s \ set signs \ s \ consistent_sign_vec qs ` set (characterize_root_list_p p) \ (z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s}) = 0)" by auto then have "(\s\ set(signs). z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) = (\s\(set (signs) \ (consistent_sign_vec qs ` set (characterize_root_list_p p))). z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s}))" proof - have "set(signs) =(set (signs) \ (consistent_sign_vec qs ` set (characterize_root_list_p p))) \ (set(signs)-(consistent_sign_vec qs ` set (characterize_root_list_p p)))" by blast then have sum_rewrite: "(\s\ set(signs). z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) = (\s\ (set (signs) \ (consistent_sign_vec qs ` set (characterize_root_list_p p))) \ (set(signs)-(consistent_sign_vec qs ` set (characterize_root_list_p p))). z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s}))" by auto then have sum_split: "(\s\ (set (signs) \ (consistent_sign_vec qs ` set (characterize_root_list_p p))) \ (set(signs)-(consistent_sign_vec qs ` set (characterize_root_list_p p))). z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) = (\s\ (set (signs) \ (consistent_sign_vec qs ` set (characterize_root_list_p p))). z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) + (\s\ (set(signs)-(consistent_sign_vec qs ` set (characterize_root_list_p p))). z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s}))" by (metis (no_types, lifting) List.finite_set sum.Int_Diff) have sum_zero: "(\s\ (set(signs)-(consistent_sign_vec qs ` set (characterize_root_list_p p))). z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) = 0" using hyp by (simp add: hyp) show ?thesis using sum_rewrite sum_split sum_zero by linarith qed then have set_eq: "set (remdups (map (consistent_sign_vec qs) (characterize_root_list_p p))) = set (signs) \ (consistent_sign_vec qs ` set (characterize_root_list_p p))" using all_info apply (simp add: characterize_consistent_signs_at_roots_def subset_antisym) by (smt (z3) Int_subset_iff consistent_sign_vec_def list.set_map map_eq_conv o_apply signs_at_def squash_def subsetI subset_antisym) have hyp1: "(\s\signs. z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) = (\s\set (signs). z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s}))" using distinct_signs sum_list_distinct_conv_sum_set by blast have hyp2: "(\s\remdups (map (consistent_sign_vec qs) (characterize_root_list_p p)). z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) = (\s\ set (remdups (map (consistent_sign_vec qs) (characterize_root_list_p p))). z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s}))" using sum_list_distinct_conv_sum_set by blast have set_sum_eq: "(\s\(set (signs) \ (consistent_sign_vec qs ` set (characterize_root_list_p p))). z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) = (\s\ set (remdups (map (consistent_sign_vec qs) (characterize_root_list_p p))). z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s}))" using set_eq by auto then have "(\s\signs. z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) = (\s\remdups (map (consistent_sign_vec qs) (characterize_root_list_p p)). z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s}))" using set_sum_eq hyp1 hyp2 using \(\s\set signs. z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) = (\s\set signs \ consistent_sign_vec qs ` set (characterize_root_list_p p). z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s}))\ by linarith then have "consistent_sign_vec qs ` set (characterize_root_list_p p) \ set signs \ (\p qss. characterize_consistent_signs_at_roots p qss = remdups (map (consistent_sign_vec qss) (characterize_root_list_p p))) \ (\s\signs. z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) = (\s\remdups (map (consistent_sign_vec qs) (characterize_root_list_p p)). z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s}))" by linarith then show ?thesis unfolding lhs_dot_rewrite[OF nonzero] apply (auto intro!: sum_list_distinct_filter simp add: distinct_signs consistent_sign_vec_def characterize_consistent_signs_at_roots_def) using all_info consistent_sign_vec_def characterize_consistent_signs_at_roots_def by (smt (z3) list.set_map map_eq_conv o_apply set_remdups signs_at_def squash_def) qed lemma matrix_equation_helper_step_R: fixes p:: "real poly" fixes qs:: "real poly list" fixes I:: "nat list*nat list" fixes signs:: "rat list list" assumes nonzero: "p\0" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots p qs) \ set(signs)" assumes welldefined1: "list_constr (fst I) (length qs)" assumes welldefined2: "list_constr (snd I) (length qs)" shows "(vec_of_list (mtx_row_R signs I) \ (construct_lhs_vector_R p qs signs)) = rat_of_int (card {x. poly p x = 0 \ (\p \ set (retrieve_polys qs (fst I)). poly p x = 0) \ poly (prod_list (retrieve_polys qs (snd I))) x > 0}) - rat_of_int (card {x. poly p x = 0 \ (\p \ set (retrieve_polys qs (fst I)). poly p x = 0) \ poly (prod_list (retrieve_polys qs (snd I))) x < 0})" proof - have finset: "finite (set (map (consistent_sign_vec qs) (characterize_root_list_p p)))" by auto let ?gt = "(set (map (consistent_sign_vec qs) (characterize_root_list_p p)) \ {s. z_R I s = 1})" let ?lt = " (set (map (consistent_sign_vec qs) (characterize_root_list_p p)) \ {s. z_R I s = -1})" let ?zer = "(set (map (consistent_sign_vec qs) (characterize_root_list_p p)) \ {s. z_R I s = 0})" have eq: "set (map (consistent_sign_vec qs) (characterize_root_list_p p)) = (?gt \ ?lt) \ ?zer" proof safe fix x assume h:"x \ set (map (consistent_sign_vec qs) (characterize_root_list_p p))" "z_R I x \ 0" "z_R I x \ - 1" then have "x \ set (characterize_consistent_signs_at_roots p qs)" unfolding characterize_consistent_signs_at_roots_def by (smt (verit, del_insts) characterize_consistent_signs_at_roots_def characterize_root_list_p_def imageE in_set_R nonzero poly_roots_finite set_map sorted_list_of_set(1)) thus "z_R I x = 1" using h welldefined1 welldefined2 z_lemma_R by blast qed have sumeq: "(\s\(?gt\?lt). z_R I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) = (\s\?gt. z_R I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) + (\s\?lt. z_R I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s}))" apply (rule sum.union_disjoint) by auto (* First, drop the signs that are irrelevant *) from construct_lhs_vector_drop_consistent_R[OF assms(1-5)] have "vec_of_list (mtx_row_R signs I) \ construct_lhs_vector_R p qs signs = vec_of_list (mtx_row_R (characterize_consistent_signs_at_roots p qs) I) \ construct_lhs_vector_R p qs (characterize_consistent_signs_at_roots p qs)" . (* Now we split the sum *) from lhs_dot_rewrite[OF assms(1)] moreover have "... = (\s\characterize_consistent_signs_at_roots p qs. z_R I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s}))" . moreover have "... = (\s\set (map (consistent_sign_vec qs) (characterize_root_list_p p)). z_R I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s}))" unfolding characterize_consistent_signs_at_roots_def sum_code[symmetric] apply (auto) by (smt (verit, best) consistent_sign_vec_def list.set_map map_eq_conv o_apply signs_at_def squash_def sum.set_conv_list) moreover have setc1:"... = (\s\(?gt\?lt). z_R I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) + (\s\?zer. z_R I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) " apply (subst eq) apply (rule sum.union_disjoint) by auto ultimately have setc: "... = (\s\?gt. z_R I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) + (\s\?lt. z_R I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) + (\s\?zer. z_R I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s}))" using sumeq by auto have "\s \ ?zer. (z_R I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) = 0" by auto then have obvzer: "(\s\?zer. z_R I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) = 0" by auto (* Now recharacterize lt, gt*) have setroots: "set (characterize_root_list_p p) = {x. poly p x = 0}" unfolding characterize_root_list_p_def using poly_roots_finite nonzero rsquarefree_def set_sorted_list_of_set by blast have *: "\s. {x. poly p x = 0 \ consistent_sign_vec qs x = s} = {x \{x. poly p x = 0}. consistent_sign_vec qs x = s}" by auto have e1: "(\s\consistent_sign_vec qs ` {x. poly p x = 0} \ {s. z_R I s = 1}. card {x. poly p x = 0 \ consistent_sign_vec qs x = s}) = (sum (\x. if (\p \ set (retrieve_polys qs (fst I)). poly p x = 0) \ (poly (prod_list (retrieve_polys qs (snd I))) x) > 0 then 1 else 0) {x. poly p x = 0})" unfolding * apply (rule sum_multicount_gen) using \finite (set (map (consistent_sign_vec qs) (characterize_root_list_p p)))\ setroots apply auto[1] apply (metis List.finite_set setroots) proof safe fix x assume rt: "poly p x = 0" then have 1: "{s \ consistent_sign_vec qs ` {x. poly p x = 0} \ {s. z_R I s = 1}. consistent_sign_vec qs x = s} = {s. z_R I s = 1 \ consistent_sign_vec qs x = s}" by auto have 2: "... = {s. (\p \ set (retrieve_polys qs (fst I)). poly p x = 0) \ (0 < poly (prod_list (retrieve_polys qs (snd I))) x) \ consistent_sign_vec qs x = s}" using horiz_vector_helper_pos_R assms welldefined1 welldefined2 rt by blast have 3: "... = (if (\p \ set (retrieve_polys qs (fst I)). poly p x = 0) \ (0 < poly (prod_list (retrieve_polys qs (snd I))) x) then {consistent_sign_vec qs x} else {})" by auto then have "card {s \ consistent_sign_vec qs ` {x. poly p x = 0} \ {s. z_R I s = 1}. consistent_sign_vec qs x = s} = (if ((\p \ set (retrieve_polys qs (fst I)). poly p x = 0) \ 0 < poly (prod_list (retrieve_polys qs (snd I))) x) then 1 else 0)" using 1 2 3 by auto thus " card {s \ consistent_sign_vec qs ` {x. poly p x = 0} \ {s. z_R I s = 1}. consistent_sign_vec qs x = s} = (if (\p\set (retrieve_polys qs (fst I)). poly p x = 0) \ 0 < poly (prod_list (retrieve_polys qs (snd I))) x then 1 else 0)" by auto qed have gtchr: "(\s\?gt. z_R I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) = rat_of_int (card {x. poly p x = 0 \ (\p\set (retrieve_polys qs (fst I)). poly p x = 0) \ 0 < poly (prod_list (retrieve_polys qs (snd I))) x})" apply (auto simp add: setroots) apply (subst of_nat_sum[symmetric]) apply (subst of_nat_eq_iff) apply (subst e1) apply (subst card_eq_sum) apply (rule sum.mono_neutral_cong_right) apply (metis List.finite_set setroots) by auto have e2: " (\s\consistent_sign_vec qs ` {x. poly p x = 0} \ {s. z_R I s = - 1}. card {x. poly p x = 0 \ consistent_sign_vec qs x = s}) = (sum (\x. if (\p\set (retrieve_polys qs (fst I)). poly p x = 0) \ (poly (prod_list (retrieve_polys qs (snd I))) x) < 0 then 1 else 0) {x. poly p x = 0})" unfolding * apply (rule sum_multicount_gen) using \finite (set (map (consistent_sign_vec qs) (characterize_root_list_p p)))\ setroots apply auto[1] apply (metis List.finite_set setroots) proof safe fix x assume rt: "poly p x = 0" then have 1: "{s \ consistent_sign_vec qs ` {x. poly p x = 0} \ {s. z_R I s = -1}. consistent_sign_vec qs x = s} = {s. z_R I s = -1 \ consistent_sign_vec qs x = s}" by auto have 2: "... = {s. ((\p\set (retrieve_polys qs (fst I)). poly p x = 0) \ 0 > poly (prod_list (retrieve_polys qs (snd I))) x) \ consistent_sign_vec qs x = s}" using horiz_vector_helper_neg_R assms rt welldefined1 welldefined2 by blast have 3: "... = (if (\p\set (retrieve_polys qs (fst I)). poly p x = 0) \ (0 > poly (prod_list (retrieve_polys qs (snd I))) x) then {consistent_sign_vec qs x} else {})" by auto thus "card {s \ consistent_sign_vec qs ` {x. poly p x = 0} \ {s. z_R I s = -1}. consistent_sign_vec qs x = s} = (if (\p\set (retrieve_polys qs (fst I)). poly p x = 0) \ 0 > poly (prod_list (retrieve_polys qs (snd I))) x then 1 else 0)" using 1 2 3 by auto qed have ltchr: "(\s\?lt. z_R I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) = - rat_of_int (card {x. poly p x = 0 \ (\p\set (retrieve_polys qs (fst I)). poly p x = 0) \ 0 > poly (prod_list (retrieve_polys qs (snd I))) x})" apply (auto simp add: setroots sum_negf) apply (subst of_nat_sum[symmetric]) apply (subst of_nat_eq_iff) apply (subst e2) apply (subst card_eq_sum) apply (rule sum.mono_neutral_cong_right) apply (metis List.finite_set setroots) by auto show ?thesis using gtchr ltchr obvzer setc using \(\s\characterize_consistent_signs_at_roots p qs. z_R I s * rat_of_int (int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s}))) = (\s\set (map (consistent_sign_vec qs) (characterize_root_list_p p)). z_R I s * rat_of_int (int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})))\ \vec_of_list (mtx_row_R (characterize_consistent_signs_at_roots p qs) I) \ construct_lhs_vector_R p qs (characterize_consistent_signs_at_roots p qs) = (\s\characterize_consistent_signs_at_roots p qs. z_R I s * rat_of_int (int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})))\ \vec_of_list (mtx_row_R signs I) \ construct_lhs_vector_R p qs signs = vec_of_list (mtx_row_R (characterize_consistent_signs_at_roots p qs) I) \ construct_lhs_vector_R p qs (characterize_consistent_signs_at_roots p qs)\ setc1 by linarith qed lemma matrix_equation_main_step_R: fixes p:: "real poly" fixes qs:: "real poly list" fixes I:: "nat list*nat list" fixes signs:: "rat list list" assumes nonzero: "p\0" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots p qs) \ set(signs)" assumes welldefined1: "list_constr (fst I) (length qs)" assumes welldefined2: "list_constr (snd I) (length qs)" shows "(vec_of_list (mtx_row_R signs I) \ (construct_lhs_vector_R p qs signs)) = construct_NofI_R p (retrieve_polys qs (fst I)) (retrieve_polys qs (snd I))" proof - show ?thesis unfolding construct_NofI_prop_R[OF nonzero] using matrix_equation_helper_step_R[OF assms] by linarith qed lemma mtx_row_length_R: "list_all (\r. length r = length signs) (map (mtx_row_R signs) ls)" apply (induction ls) by (auto simp add: mtx_row_R_def) (* Shows that as long as we have a "basis" of sign assignments (see assumptions all_info, welldefined), and some other mild assumptions on our inputs (given in nonzero, distinct_signs), the construction will be satisfied *) theorem matrix_equation_R: fixes p:: "real poly" fixes qs:: "real poly list" fixes subsets:: "(nat list*nat list) list" fixes signs:: "rat list list" assumes nonzero: "p\0" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots p qs) \ set(signs)" assumes welldefined: "all_list_constr_R (subsets) (length qs)" shows "satisfy_equation_R p qs subsets signs" unfolding satisfy_equation_R_def matrix_A_R_def construct_lhs_vector_R_def construct_rhs_vector_R_def all_list_constr_R_def apply (subst mult_mat_vec_of_list) apply (auto simp add: mtx_row_length_R intro!: map_vec_vec_of_list_eq_intro) using matrix_equation_main_step_R[OF assms(1-3), unfolded construct_lhs_vector_R_def] using all_list_constr_R_def in_set_member welldefined by fastforce (* Prettifying some theorems*) lemma consistent_signs_at_roots_eq: assumes "p \ 0" shows "consistent_signs_at_roots p qs = set (characterize_consistent_signs_at_roots p qs)" unfolding consistent_signs_at_roots_def characterize_consistent_signs_at_roots_def characterize_root_list_p_def apply auto apply (subst set_sorted_list_of_set) using assms poly_roots_finite apply blast unfolding sgn_vec_def sgn_def signs_at_def squash_def o_def using roots_def apply auto[1] by (smt Collect_cong assms image_iff poly_roots_finite roots_def sorted_list_of_set(1)) abbreviation w_vec_R:: "real poly \ real poly list \ rat list list \ rat vec" where "w_vec_R \ construct_lhs_vector_R" abbreviation v_vec_R:: "real poly \ real poly list \ (nat list*nat list) list \ rat vec" where "v_vec_R \ construct_rhs_vector_R" abbreviation M_mat_R:: "rat list list \ (nat list*nat list) list \ rat mat" where "M_mat_R \ matrix_A_R" theorem matrix_equation_pretty: fixes subsets:: "(nat list*nat list) list" assumes "p\0" assumes "distinct signs" assumes "consistent_signs_at_roots p qs \ set signs" assumes "\a b i. (a, b) \ set ( subsets) \ (i \ set a \ i \ set b) \ i < length qs" shows "M_mat_R signs subsets *\<^sub>v w_vec_R p qs signs = v_vec_R p qs subsets" unfolding satisfy_equation_R_def[symmetric] using matrix_equation_R[of p signs qs subsets] assms using consistent_signs_at_roots_eq unfolding all_list_constr_R_def list_constr_def apply (auto) by (metis (no_types, lifting) Ball_set in_set_member) section "Base Case" definition satisfies_properties_R:: "real poly \ real poly list \ (nat list*nat list) list \ rat list list \ rat mat \ bool" where "satisfies_properties_R p qs subsets signs matrix = ( all_list_constr_R subsets (length qs) \ well_def_signs (length qs) signs \ distinct signs \ satisfy_equation_R p qs subsets signs \ invertible_mat matrix \ matrix = matrix_A_R signs subsets \ set (characterize_consistent_signs_at_roots p qs) \ set(signs) )" lemma mat_base_case_R: shows "matrix_A_R [[1],[0],[-1]] [([], []),([0], []),([], [0])] = (mat_of_rows_list 3 [[1,1,1], [0,1,0], [1,0,-1]])" unfolding matrix_A_R_def mtx_row_R_def z_R_def map_sgas_def apply (auto) by (simp add: numeral_3_eq_3) lemma base_case_sgas_R: fixes q p:: "real poly" assumes nonzero: "p \ 0" shows "set (characterize_consistent_signs_at_roots p [q]) \ {[1],[0], [- 1]}" unfolding characterize_consistent_signs_at_roots_def signs_at_def apply (auto) by (meson squash_def) lemma base_case_sgas_alt_R: fixes p:: "real poly" fixes qs:: "real poly list" assumes len1: "length qs = 1" assumes nonzero: "p \ 0" shows "set (characterize_consistent_signs_at_roots p qs) \ {[1], [0], [- 1]}" proof - have ex_q: "\(q::real poly). qs = [q]" using len1 using length_Suc_conv[of qs 0] by auto then show ?thesis using base_case_sgas_R nonzero by auto qed lemma base_case_satisfy_equation_R: fixes q p:: "real poly" assumes nonzero: "p \ 0" shows "satisfy_equation_R p [q] [([], []),([0], []),([], [0])] [[1],[0],[-1]] " proof - have h1: "set (characterize_consistent_signs_at_roots p [q]) \ {[1], [0],[- 1]}" using base_case_sgas_R assms by auto have h2: "all_list_constr_R [([], []),([0], []),([], [0])] (Suc 0)" unfolding all_list_constr_R_def by (simp add: list_constr_def member_def) then show ?thesis using matrix_equation_R[where p = "p", where qs = "[q]", where signs = "[[1],[0],[-1]] ", where subsets = "[([], []),([0], []),([], [0])]"] nonzero h1 h2 by auto qed lemma base_case_satisfy_equation_alt_R: fixes p:: "real poly" fixes qs:: "real poly list" assumes len1: "length qs = 1" assumes nonzero: "p \ 0" shows "satisfy_equation_R p qs [([], []),([0], []),([], [0])] [[1],[0],[-1]]" proof - have ex_q: "\(q::real poly). qs = [q]" using len1 using length_Suc_conv[of qs 0] by auto then show ?thesis using base_case_satisfy_equation_R nonzero by auto qed lemma base_case_matrix_eq: fixes q p:: "real poly" assumes nonzero: "p \ 0" shows "(mult_mat_vec (mat_of_rows_list 3 [[1,1,1], [0,1,0], [1,0,-1]]) (construct_lhs_vector_R p [q] [[1],[0],[-1]]) = (construct_rhs_vector_R p [q] [([], []),([0], []),([], [0])]))" using mat_base_case_R base_case_satisfy_equation_R unfolding satisfy_equation_R_def by (simp add: nonzero) lemma less_three: "(n::nat) < Suc (Suc (Suc 0)) \ n = 0 \ n = 1 \ n = 2" by auto lemma inverse_mat_base_case_R: shows "inverts_mat (mat_of_rows_list 3 [[1/2, -1/2, 1/2], [0, 1, 0], [1/2, -1/2, -1/2]]::rat mat) (mat_of_rows_list 3 [[1,1,1], [0,1,0], [1,0,-1]]::rat mat)" unfolding inverts_mat_def mat_of_rows_list_def mat_eq_iff apply auto unfolding less_three by (auto simp add: scalar_prod_def) lemma inverse_mat_base_case_2_R: shows "inverts_mat (mat_of_rows_list 3 [[1,1,1], [0,1,0], [1,0,-1]]::rat mat) (mat_of_rows_list 3 [[1/2, -1/2, 1/2], [0, 1, 0], [1/2, -1/2, -1/2]]:: rat mat)" unfolding inverts_mat_def mat_of_rows_list_def mat_eq_iff apply auto unfolding less_three by (auto simp add: scalar_prod_def) lemma base_case_invertible_mat_R: shows "invertible_mat (matrix_A_R [[1],[0], [- 1]] [([], []),([0], []),([], [0])])" unfolding invertible_mat_def using inverse_mat_base_case_R inverse_mat_base_case_2_R apply (auto) apply (simp add: mat_base_case mat_of_rows_list_def) using mat_base_case_R by auto section "Inductive Step" (***** Need some properties of smashing; smashing signs is just as it was in BKR *****) subsection "Lemmas on smashing subsets " definition subsets_first_component_list::"(nat list*nat list) list \ nat list list" where "subsets_first_component_list I = map (\I. (fst I)) I" definition subsets_second_component_list::"(nat list*nat list) list \ nat list list" where "subsets_second_component_list I = map (\I. (snd I)) I" definition smash_list_list::"('a list*'a list) list \ ('a list*'a list) list \ ('a list*'a list) list" where "smash_list_list s1 s2 = concat (map (\l1. map (\l2. (fst l1 @ fst l2, snd l1 @ snd l2)) s2) s1)" lemma smash_list_list_property_set: fixes l1 l2 :: "('a list*'a list) list" fixes a b:: "nat" shows "\ (elem :: ('a list*'a list)). (elem \ (set (smash_list_list l1 l2)) \ (\ (elem1:: ('a list*'a list)). \ (elem2:: ('a list*'a list)). (elem1 \ set(l1) \ elem2 \ set(l2) \ elem = (fst elem1@ fst elem2, snd elem1 @ snd elem2))))" proof clarsimp fix a b assume assum: "(a, b) \ set (smash_list_list l1 l2)" show "\aa ba. (aa, ba) \ set l1 \ (\ab bb. (ab, bb) \ set l2 \ a = aa @ ab \ b = ba @ bb)" using assum unfolding smash_list_list_def apply (auto) by blast qed lemma subsets_smash_property_R: fixes subsets1 subsets2 :: "(nat list*nat list) list" fixes n:: "nat" shows "\ (elem :: nat list*nat list). (List.member (subsets_smash_R n subsets1 subsets2) elem) \ (\ (elem1::nat list*nat list). \ (elem2::nat list*nat list). (elem1 \ set(subsets1) \ elem2 \ set(subsets2) \ elem = ((fst elem1) @ (map ((+) n) (fst elem2)), (snd elem1) @ (map ((+) n) (snd elem2)))))" proof - let ?fst_component2 = "subsets_first_component_list subsets2" let ?snd_component2 = "subsets_second_component_list subsets2" let ?new_subsets = "map (\I. ((map ((+) n)) (fst I), (map ((+) n)) (snd I))) subsets2" (* a slightly unorthodox use of signs_smash, but it closes the proof *) have "subsets_smash_R n subsets1 subsets2 = smash_list_list subsets1 ?new_subsets" unfolding subsets_smash_R_def smash_list_list_def apply (auto) by (simp add: comp_def) then show ?thesis using smash_list_list_property_set[of subsets1 ?new_subsets] apply (auto) using imageE in_set_member set_map smash_list_list_property_set by (smt (z3) prod.exhaust_sel) qed (* If subsets for smaller systems are well-defined, then subsets for combined systems are, too *) subsection "Well-defined subsets preserved when smashing" lemma well_def_step_R: fixes qs1 qs2 :: "real poly list" fixes subsets1 subsets2 :: "(nat list*nat list) list" assumes well_def_subsets1: "all_list_constr_R (subsets1) (length qs1)" assumes well_def_subsets2: "all_list_constr_R (subsets2) (length qs2)" shows "all_list_constr_R ((subsets_smash_R (length qs1) subsets1 subsets2)) (length (qs1 @ qs2))" proof - let ?n = "(length qs1)" have h1: "\elem. List.member (subsets_smash_R ?n subsets1 subsets2) elem \ (\ (elem1::nat list*nat list). \ (elem2::nat list*nat list). (elem1 \ set(subsets1) \ elem2 \ set(subsets2) \ elem = ((fst elem1) @ (map ((+) ?n) (fst elem2)), (snd elem1) @ (map ((+) ?n) (snd elem2)))))" using subsets_smash_property_R by blast have h2: "\elem1 elem2. (elem1 \ set subsets1 \ elem2 \ set subsets2) \ list_constr ((fst elem1) @ map ((+) (length qs1)) (fst elem2)) (length (qs1 @ qs2))" proof clarsimp fix elem1 b elem2 ba assume e1: "(elem1, b) \ set subsets1" assume e2: "(elem2, ba) \ set subsets2" have h1: "list_constr elem1 (length qs1 + length qs2) " proof - have h1: "list_constr elem1 (length qs1)" using e1 well_def_subsets1 unfolding all_list_constr_R_def apply (auto) by (simp add: in_set_member) then show ?thesis unfolding list_constr_def by (simp add: list.pred_mono_strong) qed have h2: "list_constr (map ((+) (length qs1)) elem2) (length qs1 + length qs2)" proof - have h1: "list_constr elem2 (length qs2)" using e2 well_def_subsets2 unfolding all_list_constr_R_def by (simp add: in_set_member) then show ?thesis unfolding list_constr_def by (simp add: list_all_length) qed show "list_constr (elem1 @ map ((+) (length qs1)) elem2) (length qs1 + length qs2)" using h1 h2 list_constr_append by blast qed have h3: "\elem1 elem2. (elem1 \ set subsets1 \ elem2 \ set subsets2) \ list_constr ((snd elem1) @ map ((+) (length qs1)) (snd elem2)) (length (qs1 @ qs2))" proof clarsimp fix elem1 b elem2 ba assume e1: "(b, elem1) \ set subsets1" assume e2: "(ba, elem2) \ set subsets2" have h1: "list_constr elem1 (length qs1 + length qs2) " proof - have h1: "list_constr elem1 (length qs1)" using e1 well_def_subsets1 unfolding all_list_constr_R_def apply (auto) by (simp add: in_set_member) then show ?thesis unfolding list_constr_def by (simp add: list.pred_mono_strong) qed have h2: "list_constr (map ((+) (length qs1)) elem2) (length qs1 + length qs2)" proof - have h1: "list_constr elem2 (length qs2)" using e2 well_def_subsets2 unfolding all_list_constr_R_def by (simp add: in_set_member) then show ?thesis unfolding list_constr_def by (simp add: list_all_length) qed show "list_constr (elem1 @ map ((+) (length qs1)) elem2) (length qs1 + length qs2)" using h1 h2 list_constr_append by blast qed show ?thesis using h1 h2 h3 unfolding all_list_constr_def by (metis all_list_constr_R_def fst_conv snd_conv) qed subsection "Consistent Sign Assignments Preserved When Smashing" lemma subset_helper_R: fixes p:: "real poly" fixes qs1 qs2 :: "real poly list" fixes signs1 signs2 :: "rat list list" shows "\ x \ set (characterize_consistent_signs_at_roots p (qs1 @ qs2)). \ x1 \ set (characterize_consistent_signs_at_roots p qs1). \ x2 \ set (characterize_consistent_signs_at_roots p qs2). x = x1@x2" proof clarsimp fix x assume x_in: "x \ set (characterize_consistent_signs_at_roots p (qs1 @ qs2))" have x_in_csv: "x \ set(map (consistent_sign_vec (qs1 @ qs2)) (characterize_root_list_p p))" using x_in unfolding characterize_consistent_signs_at_roots_def by (smt (z3) consistent_sign_vec_def map_eq_conv o_apply set_remdups signs_at_def squash_def) have csv_hyp: "\r. consistent_sign_vec (qs1 @ qs2) r = (consistent_sign_vec qs1 r)@(consistent_sign_vec qs2 r)" unfolding consistent_sign_vec_def by auto have exr_iff: "(\r \ set (characterize_root_list_p p). x = consistent_sign_vec (qs1 @ qs2) r) \ (\r \ set (characterize_root_list_p p). x = (consistent_sign_vec qs1 r)@(consistent_sign_vec qs2 r))" using csv_hyp by auto have exr: "x \ set(map (consistent_sign_vec (qs1 @ qs2)) (characterize_root_list_p p)) \ (\r \ set (characterize_root_list_p p). x = consistent_sign_vec (qs1 @ qs2) r)" by auto have mem_list1: "\ r \ set (characterize_root_list_p p). (consistent_sign_vec qs1 r) \ set(map (consistent_sign_vec qs1) (characterize_root_list_p p))" by simp have mem_list2: "\ r \ set (characterize_root_list_p p). (consistent_sign_vec qs2 r) \ set(map (consistent_sign_vec qs2) (characterize_root_list_p p))" by simp then show "\x1\set (characterize_consistent_signs_at_roots p qs1). \x2\set (characterize_consistent_signs_at_roots p qs2). x = x1 @ x2" using x_in_csv exr mem_list1 mem_list2 characterize_consistent_signs_at_roots_def exr_iff using imageE image_eqI map_append set_map set_remdups signs_at_def x_in by auto qed lemma subset_step_R: fixes p:: "real poly" fixes qs1 qs2 :: "real poly list" fixes signs1 signs2 :: "rat list list" assumes csa1: "set (characterize_consistent_signs_at_roots p qs1) \ set (signs1)" assumes csa2: "set (characterize_consistent_signs_at_roots p qs2) \ set (signs2)" shows "set (characterize_consistent_signs_at_roots p (qs1 @ qs2)) \ set (signs_smash signs1 signs2)" proof - have h0: "\ x \ set (characterize_consistent_signs_at_roots p (qs1 @ qs2)). \ x1 \ set (characterize_consistent_signs_at_roots p qs1). \ x2 \ set (characterize_consistent_signs_at_roots p qs2). (x = x1@x2)" using subset_helper_R[of p qs1 qs2] by blast then have "\x \ set (characterize_consistent_signs_at_roots p (qs1 @ qs2)). x \ set (signs_smash (characterize_consistent_signs_at_roots p qs1) (characterize_consistent_signs_at_roots p qs2))" unfolding signs_smash_def apply (auto) by fastforce then have "\x \ set (characterize_consistent_signs_at_roots p (qs1 @ qs2)). x \ set (signs_smash signs1 signs2)" using csa1 csa2 subset_smash_signs[of _ signs1 _ signs2] apply (auto) by blast thus ?thesis by blast qed subsection "Main Results" lemma dim_row_matrix_A_R[simp]: shows "dim_row (matrix_A_R signs subsets) = length subsets" unfolding matrix_A_R_def by auto lemma dim_col_matrix_A_R[simp]: shows "dim_col (matrix_A_R signs subsets) = length signs" unfolding matrix_A_R_def by auto lemma length_subsets_smash_R: shows "length (subsets_smash_R n subs1 subs2) = length subs1 * length subs2" unfolding subsets_smash_R_def length_concat by (auto simp add: o_def sum_list_triv) lemma z_append_R: fixes xs:: "(nat list * nat list)" assumes "\i. i \ set (fst xs) \ i < length as" assumes "\i. i \ set (snd xs) \ i < length as" shows "z_R ((fst xs) @ (map ((+) (length as)) (fst ys)), (snd xs) @ (map ((+) (length as)) (snd ys))) (as @ bs) = z_R xs as * z_R ys bs" proof - have 1: "map ((!) (as @ bs)) (fst xs) = map ((!) as) (fst xs)" unfolding list_eq_iff_nth_eq using assms by (auto simp add:nth_append) have 2: "map ((!) (as @ bs) \ (+) (length as)) (fst ys) = map ((!) bs) (fst ys)" unfolding list_eq_iff_nth_eq using assms by auto have 3: "map ((!) (as @ bs)) (snd xs) = map ((!) as) (snd xs)" unfolding list_eq_iff_nth_eq using assms by (auto simp add:nth_append) have 4: "map ((!) (as @ bs) \ (+) (length as)) (snd ys) = map ((!) bs) (snd ys)" unfolding list_eq_iff_nth_eq using assms by auto show ?thesis unfolding z_R_def apply auto unfolding 1 2 3 4 apply (auto) by (smt (z3) assms(1) comp_apply length_map map_append map_eq_conv map_sgas_def nth_append nth_append_length_plus) qed (* Shows that the matrix of a smashed system is the Kronecker product of the matrices of the two systems being combined *) lemma matrix_construction_is_kronecker_product_R: fixes qs1 :: "real poly list" fixes subs1 subs2 :: "(nat list*nat list) list" fixes signs1 signs2 :: "rat list list" (* n1 is the number of polynomials in the "1" sets *) assumes "\l i. l \ set subs1 \ (i \ set (fst l) \ i \ set (snd l)) \ i < n1" assumes "\j. j \ set signs1 \ length j = n1" shows "(matrix_A_R (signs_smash signs1 signs2) (subsets_smash_R n1 subs1 subs2)) = kronecker_product (matrix_A_R signs1 subs1) (matrix_A_R signs2 subs2)" unfolding mat_eq_iff dim_row_matrix_A_R dim_col_matrix_A_R length_subsets_smash_R length_signs_smash dim_kronecker proof safe fix i j assume i: "i < length subs1 * length subs2" assume j: "j < length signs1 * length signs2" have ld: "i div length subs2 < length subs1" "j div length signs2 < length signs1" using i j less_mult_imp_div_less by auto have lm: "i mod length subs2 < length subs2" "j mod length signs2 < length signs2" using i j less_mult_imp_mod_less by auto have n1: "n1 = length (signs1 ! (j div length signs2))" using assms(2) ld(2) nth_mem by blast have 1: "matrix_A_R (signs_smash signs1 signs2) (subsets_smash_R n1 subs1 subs2) $$ (i, j) = z_R (subsets_smash_R n1 subs1 subs2 ! i) (signs_smash signs1 signs2 ! j)" unfolding mat_of_rows_list_def matrix_A_R_def mtx_row_R_def using i j by (auto simp add: length_signs_smash length_subsets_smash_R) have 2: " ... = z_R ((fst (subs1 ! (i div length subs2)) @ map ((+) n1) (fst(subs2 ! (i mod length subs2)))), (snd (subs1 ! (i div length subs2)) @ map ((+) n1) (snd (subs2 ! (i mod length subs2))))) (signs1 ! (j div length signs2) @ signs2 ! (j mod length signs2))" unfolding signs_smash_def subsets_smash_R_def apply (subst length_eq_concat) using i apply (auto simp add: mult.commute) apply (subst length_eq_concat) using j apply (auto simp add: mult.commute) using ld lm by auto have 3: "... = z_R (subs1 ! (i div length subs2)) (signs1 ! (j div length signs2)) * z_R (subs2 ! (i mod length subs2)) (signs2 ! (j mod length signs2))" unfolding n1 apply (subst z_append_R) apply (auto simp add: n1[symmetric]) using assms(1) ld(1) nth_mem apply blast using assms(1) ld(1) nth_mem by blast have 4: "kronecker_product (matrix_A_R signs1 subs1) (matrix_A_R signs2 subs2) $$ (i,j) = z_R (subs1 ! (i div length subs2)) (signs1 ! (j div length signs2)) * z_R (subs2 ! (i mod length subs2)) (signs2 ! (j mod length signs2))" unfolding kronecker_product_def matrix_A_R_def mat_of_rows_list_def mtx_row_R_def using i j apply (auto simp add: Let_def) apply (subst index_mat(1)[OF ld]) apply (subst index_mat(1)[OF lm]) using ld lm by auto show "matrix_A_R (signs_smash signs1 signs2) (subsets_smash_R n1 subs1 subs2) $$ (i, j) = kronecker_product (matrix_A_R signs1 subs1) (matrix_A_R signs2 subs2) $$ (i, j)" using 1 2 3 4 by auto qed (* Given that two smaller systems satisfy some mild constraints, show that their combined system (after smashing the signs and subsets) satisfies the matrix equation, and that the matrix of the combined system is invertible. *) lemma inductive_step_R: fixes p:: "real poly" fixes qs1 qs2 :: "real poly list" fixes subsets1 subsets2 :: "(nat list*nat list) list" fixes signs1 signs2 :: "rat list list" assumes nonzero: "p \ 0" assumes nontriv1: "length qs1 > 0" assumes nontriv2: "length qs2 > 0" assumes welldefined_signs1: "well_def_signs (length qs1) signs1" assumes welldefined_signs2: "well_def_signs (length qs2) signs2" assumes distinct_signs1: "distinct signs1" assumes distinct_signs2: "distinct signs2" assumes all_info1: "set (characterize_consistent_signs_at_roots p qs1) \ set(signs1)" assumes all_info2: "set (characterize_consistent_signs_at_roots p qs2) \ set(signs2)" assumes welldefined_subsets1: "all_list_constr_R (subsets1) (length qs1)" assumes welldefined_subsets2: "all_list_constr_R (subsets2) (length qs2)" assumes invertibleMtx1: "invertible_mat (matrix_A_R signs1 subsets1)" assumes invertibleMtx2: "invertible_mat (matrix_A_R signs2 subsets2)" shows "satisfy_equation_R p (qs1@qs2) (subsets_smash_R (length qs1) subsets1 subsets2) (signs_smash signs1 signs2) \ invertible_mat (matrix_A_R (signs_smash signs1 signs2) (subsets_smash_R (length qs1) subsets1 subsets2))" proof - have h1: "invertible_mat (matrix_A_R (signs_smash signs1 signs2) (subsets_smash_R (length qs1) subsets1 subsets2))" using matrix_construction_is_kronecker_product_R kronecker_invertible invertibleMtx1 invertibleMtx2 welldefined_subsets1 welldefined_subsets2 unfolding all_list_constr_R_def list_constr_def using Ball_set in_set_member well_def_signs_def welldefined_signs1 in_set_conv_nth list_all_length apply (auto) by (smt (z3) Ball_set kronecker_invertible member_def) have h2a: "distinct (signs_smash signs1 signs2)" using distinct_signs1 distinct_signs2 welldefined_signs1 welldefined_signs2 nontriv1 nontriv2 distinct_step by auto have h2c: "all_list_constr_R ((subsets_smash_R (length qs1) subsets1 subsets2)) (length (qs1@qs2))" using well_def_step_R welldefined_subsets1 welldefined_subsets2 by blast have h2d: "set (characterize_consistent_signs_at_roots p (qs1@qs2)) \ set(signs_smash signs1 signs2)" using subset_step_R all_info1 all_info2 by simp have h2: "satisfy_equation_R p (qs1@qs2) (subsets_smash_R (length qs1) subsets1 subsets2) (signs_smash signs1 signs2)" using matrix_equation_R[where p="p", where qs="qs1@qs2", where subsets = "subsets_smash_R (length qs1) subsets1 subsets2", where signs = "signs_smash signs1 signs2"] h2a h2c h2d apply (auto) using nonzero by blast show ?thesis using h1 h2 by blast qed section "Reduction Step Proofs" (* Some definitions *) definition get_matrix_R:: "(rat mat \ ((nat list*nat list) list \ rat list list)) \ rat mat" where "get_matrix_R data = fst(data)" definition get_subsets_R:: "(rat mat \ ((nat list*nat list) list \ rat list list)) \ (nat list*nat list) list" where "get_subsets_R data = fst(snd(data))" definition get_signs_R:: "(rat mat \ ((nat list*nat list) list \ rat list list)) \ rat list list" where "get_signs_R data = snd(snd(data))" definition reduction_signs_R:: "real poly \ real poly list \ rat list list \ (nat list*nat list) list \ rat mat \ rat list list" where "reduction_signs_R p qs signs subsets matr = (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets matr)))" definition reduction_subsets_R:: "real poly \ real poly list \ rat list list \ (nat list*nat list) list \ rat mat \ (nat list*nat list) list" where "reduction_subsets_R p qs signs subsets matr = (take_indices subsets (rows_to_keep (reduce_mat_cols matr (solve_for_lhs_R p qs subsets matr))))" (* Some basic lemmas *) lemma reduction_signs_is_get_signs_R: "reduction_signs_R p qs signs subsets m = get_signs_R (reduce_system_R p (qs, (m, (subsets, signs))))" unfolding reduction_signs_R_def get_signs_R_def apply (simp) using reduction_step_R.elims snd_conv by metis lemma reduction_subsets_is_get_subsets_R: "reduction_subsets_R p qs signs subsets m = get_subsets_R (reduce_system_R p (qs, (m, (subsets, signs))))" unfolding reduction_subsets_R_def get_subsets_R_def using reduce_system.simps reduction_step.elims fst_conv snd_conv by (metis reduce_system_R.simps reduction_step_R.simps) subsection "Showing sign conditions preserved when reducing" lemma take_indices_lem_R: fixes p:: "real poly" fixes qs :: "real poly list" fixes arb_list :: "('a list*'a list) list" fixes index_list :: "nat list" fixes n:: "nat" assumes lest: "n < length (take_indices arb_list index_list)" assumes well_def: "\q.(List.member index_list q \ q < length arb_list)" shows "\k carrier_mat (length subsets) (length signs)" unfolding matrix_A_R_def carrier_mat_def by auto lemma size_of_lhs_R: fixes p:: "real poly" fixes qs :: "real poly list" fixes signs :: "rat list list" shows "dim_vec (construct_lhs_vector_R p qs signs) = length signs" unfolding construct_lhs_vector_R_def by simp lemma size_of_rhs_R: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "(nat list*nat list) list" shows "dim_vec (construct_rhs_vector_R p qs subsets) = length subsets" unfolding construct_rhs_vector_R_def by simp lemma same_size_R: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "(nat list*nat list) list" fixes signs :: "rat list list" assumes invertible_mat: "invertible_mat (matrix_A_R signs subsets)" shows "length subsets = length signs" using invertible_mat unfolding invertible_mat_def using size_of_mat_R[of signs subsets] size_of_lhs_R[of p qs signs] size_of_rhs_R[of p qs subsets] by simp lemma construct_lhs_matches_solve_for_lhs_R: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "(nat list*nat list) list" fixes signs :: "rat list list" assumes match: "satisfy_equation_R p qs subsets signs" assumes invertible_mat: "invertible_mat (matrix_A_R signs subsets)" shows "(construct_lhs_vector_R p qs signs) = solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)" proof - have matrix_equation_hyp: "(mult_mat_vec (matrix_A_R signs subsets) (construct_lhs_vector_R p qs signs) = (construct_rhs_vector_R p qs subsets))" using match unfolding satisfy_equation_R_def by blast then have eqn_hyp: " ((matr_option (dim_row (matrix_A_R signs subsets)) (mat_inverse (matrix_A_R signs subsets))) *\<^sub>v (mult_mat_vec (matrix_A_R signs subsets) (construct_lhs_vector_R p qs signs)) = mult_mat_vec (matr_option (dim_row (matrix_A_R signs subsets)) (mat_inverse (matrix_A_R signs subsets))) (construct_rhs_vector_R p qs subsets))" using invertible_mat by (simp add: matrix_equation_hyp) have match_hyp: "length subsets = length signs" using same_size_R invertible_mat by auto then have dim_hyp1: "matrix_A_R signs subsets \ carrier_mat (length signs) (length signs)" using size_of_mat by auto then have dim_hyp2: "matr_option (dim_row (matrix_A_R signs subsets)) (mat_inverse (matrix_A_R signs subsets)) \ carrier_mat (length signs) (length signs)" using invertible_mat dim_invertible by blast have mult_assoc_hyp: "((matr_option (dim_row (matrix_A_R signs subsets)) (mat_inverse (matrix_A_R signs subsets))) *\<^sub>v (mult_mat_vec (matrix_A_R signs subsets) (construct_lhs_vector_R p qs signs))) = (((matr_option (dim_row (matrix_A_R signs subsets)) (mat_inverse (matrix_A_R signs subsets))) * (matrix_A_R signs subsets)) *\<^sub>v (construct_lhs_vector_R p qs signs))" using mult_assoc dim_hyp1 dim_hyp2 size_of_lhs_R by auto have cancel_helper: "(((matr_option (dim_row (matrix_A_R signs subsets)) (mat_inverse (matrix_A_R signs subsets))) * (matrix_A_R signs subsets)) *\<^sub>v (construct_lhs_vector_R p qs signs)) = (1\<^sub>m (length signs)) *\<^sub>v (construct_lhs_vector_R p qs signs)" using invertible_means_mult_id[where A= "matrix_A_R signs subsets"] dim_hyp1 by (simp add: invertible_mat match_hyp) then have cancel_hyp: "(((matr_option (dim_row (matrix_A_R signs subsets)) (mat_inverse (matrix_A_R signs subsets))) * (matrix_A_R signs subsets)) *\<^sub>v (construct_lhs_vector_R p qs signs)) = (construct_lhs_vector_R p qs signs)" apply (auto) by (metis carrier_vec_dim_vec one_mult_mat_vec size_of_lhs_R) then have mult_hyp: "((matr_option (dim_row (matrix_A_R signs subsets)) (mat_inverse (matrix_A_R signs subsets))) *\<^sub>v (mult_mat_vec (matrix_A_R signs subsets) (construct_lhs_vector_R p qs signs))) = (construct_lhs_vector_R p qs signs)" using mult_assoc_hyp cancel_hyp by simp then have "(construct_lhs_vector_R p qs signs) = (mult_mat_vec (matr_option (dim_row (matrix_A_R signs subsets)) (mat_inverse (matrix_A_R signs subsets))) (construct_rhs_vector_R p qs subsets)) " using eqn_hyp by simp then show ?thesis unfolding solve_for_lhs_R_def by (simp add: mat_inverse_same) qed (* Then show that dropping columns doesn't affect the consistent sign assignments *) lemma reduction_doesnt_break_things_signs_R: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "(nat list*nat list) list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes welldefined_signs1: "well_def_signs (length qs) signs" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots p qs) \ set(signs)" assumes match: "satisfy_equation_R p qs subsets signs" assumes invertible_mat: "invertible_mat (matrix_A_R signs subsets)" shows "set (characterize_consistent_signs_at_roots p qs) \ set(reduction_signs_R p qs signs subsets (matrix_A_R signs subsets))" proof - have dim_hyp2: "matr_option (dim_row (matrix_A_R signs subsets)) (mat_inverse (matrix_A_R signs subsets)) \ carrier_mat (length signs) (length signs)" using invertible_mat dim_invertible using same_size_R by fastforce have "(construct_lhs_vector_R p qs signs) = solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)" using construct_lhs_matches_solve_for_lhs_R assms by auto then have "(solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)) = vec_of_list (map rat_of_nat (map (\s. card {x. poly p x = 0 \ consistent_sign_vec qs x = s}) signs))" using construct_lhs_vector_cleaner_R assms by (metis (mono_tags, lifting) list.map_cong map_map o_apply of_int_of_nat_eq) then have "\ n < (dim_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))). (((solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)) $ n = 0) \ (nth signs n) \ set (characterize_consistent_signs_at_roots p qs))" proof - have h0: "\ n < (dim_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))). (((solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)) $ n = 0) \ rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = (nth signs n)}) = 0)" by (metis (mono_tags, lifting) \construct_lhs_vector_R p qs signs = solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)\ construct_lhs_vector_clean_R nonzero of_nat_0_eq_iff of_rat_of_nat_eq size_of_lhs_R) have h1: "\ w. (rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = w}) = 0 \ (\ x. poly p x = 0 \ consistent_sign_vec qs x = w))" proof clarsimp fix x assume card_asm: "card {xa. poly p xa = 0 \ consistent_sign_vec qs xa = consistent_sign_vec qs x} = 0" assume zero_asm: "poly p x = 0" have card_hyp: "{xa. poly p xa = 0 \ consistent_sign_vec qs xa = consistent_sign_vec qs x} = {}" using card_eq_0_iff using card_asm nonzero poly_roots_finite by (metis (full_types) finite_Collect_conjI) have "x \ {xa. poly p xa = 0 \ consistent_sign_vec qs xa = consistent_sign_vec qs x}" using zero_asm by auto thus "False" using card_hyp by blast qed have h2: "\ w. (rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = w}) = 0 \ (\List.member (characterize_consistent_signs_at_roots p qs) w))" proof clarsimp fix w assume card_asm: "card {x. poly p x = 0 \ consistent_sign_vec qs x = w} = 0" assume mem_asm: "List.member (characterize_consistent_signs_at_roots p qs) w" have h0: "\ x. poly p x = 0 \ consistent_sign_vec qs x = w" using h1 card_asm by (simp add: h1) have h1: "\ x. poly p x = 0 \ consistent_sign_vec qs x = w" using mem_asm unfolding characterize_consistent_signs_at_roots_def characterize_root_list_p_def proof - have "w \ Collect (List.member (remdups (map (consistent_sign_vec qs) (sorted_list_of_set {r. poly p r = 0}))))" using characterize_consistent_signs_at_roots_def mem_asm characterize_root_list_p_def by (smt (verit, ccfv_SIG) consistent_sign_vec_def h0 imageE in_set_member list.set_map map_cong mem_Collect_eq nonzero o_apply poly_roots_finite set_remdups set_sorted_list_of_set signs_at_def squash_def) then have f1: "w \ set (map (consistent_sign_vec qs) (sorted_list_of_set {r. poly p r = 0}))" by (metis (no_types) in_set_member mem_Collect_eq set_remdups) have "finite {r. poly p r = 0}" using nonzero poly_roots_finite by blast then show ?thesis using f1 by auto qed show "False" using h0 h1 by auto qed then have h3: "\ w. rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = w}) = 0 \ w \ set (characterize_consistent_signs_at_roots p qs)" by (simp add: in_set_member) show ?thesis using h0 h3 by blast qed then have "set (characterize_consistent_signs_at_roots p qs) \ set (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))))" using all_info reduction_signs_set_helper_lemma[where A = "set (characterize_consistent_signs_at_roots p qs)", where B = "signs", where C = "(solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))"] using dim_hyp2 solve_for_lhs_R_def by (simp add: mat_inverse_same) then show ?thesis unfolding reduction_signs_R_def by auto qed lemma reduction_deletes_bad_sign_conds_R: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "(nat list*nat list) list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes welldefined_signs1: "well_def_signs (length qs) signs" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots p qs) \ set(signs)" assumes match: "satisfy_equation_R p qs subsets signs" assumes invertible_mat: "invertible_mat (matrix_A_R signs subsets)" shows "set (characterize_consistent_signs_at_roots p qs) = set(reduction_signs_R p qs signs subsets (matrix_A_R signs subsets))" proof - have dim_hyp2: "matr_option (dim_row (matrix_A_R signs subsets)) (mat_inverse (matrix_A_R signs subsets)) \ carrier_mat (length signs) (length signs)" using invertible_mat dim_invertible using same_size_R by fastforce have supset: "set (characterize_consistent_signs_at_roots p qs) \ set(reduction_signs_R p qs signs subsets (matrix_A_R signs subsets))" proof - have "(construct_lhs_vector_R p qs signs) = solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)" using construct_lhs_matches_solve_for_lhs_R assms by auto then have "(solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)) = vec_of_list (map rat_of_nat (map (\s. card {x. poly p x = 0 \ consistent_sign_vec qs x = s}) signs))" using construct_lhs_vector_cleaner_R assms by (metis (mono_tags, lifting) list.map_cong map_map o_apply of_int_of_nat_eq) then have "\ n < (dim_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))). (((solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)) $ n \ 0) \ (nth signs n) \ set (characterize_consistent_signs_at_roots p qs))" proof - have h0: "\ n < (dim_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))). (((solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)) $ n = 0) \ rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = (nth signs n)}) = 0)" by (simp add: \solve_for_lhs_R p qs subsets (M_mat_R signs subsets) = vec_of_list (map rat_of_nat (map (\s. card {x. poly p x = 0 \ consistent_sign_vec qs x = s}) signs))\ vec_of_list_index) have h1: "\ w. (rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = w}) \ 0 \ (\ x. poly p x = 0 \ consistent_sign_vec qs x = w))" proof clarsimp fix w assume card_asm: "0 < card {x. poly p x = 0 \ consistent_sign_vec qs x = w}" show "\x. poly p x = 0 \ consistent_sign_vec qs x = w" by (metis (mono_tags, lifting) Collect_empty_eq card_asm card_eq_0_iff gr_implies_not0) qed have h2: "\ w. (rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = w}) \ 0 \ (List.member (characterize_consistent_signs_at_roots p qs) w))" proof clarsimp fix w assume card_asm: " 0 < card {x. poly p x = 0 \ consistent_sign_vec qs x = w}" have h0: "\x. poly p x = 0 \ consistent_sign_vec qs x = w" using card_asm by (simp add: h1) then show "List.member (characterize_consistent_signs_at_roots p qs) w" unfolding characterize_consistent_signs_at_roots_def using in_set_member nonzero poly_roots_finite characterize_root_list_p_def by (smt (verit) characterize_consistent_signs_at_roots_def in_set_R mem_Collect_eq) qed then have h3: "\ w. rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = w}) \ 0 \ w \ set (characterize_consistent_signs_at_roots p qs)" by (simp add: in_set_member) show ?thesis using h0 h3 by (metis (no_types, lifting) \solve_for_lhs_R p qs subsets (matrix_A_R signs subsets) = vec_of_list (map rat_of_nat (map (\s. card {x. poly p x = 0 \ consistent_sign_vec qs x = s}) signs))\ dim_vec_of_list length_map nth_map vec_of_list_index) qed then have "set (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))) \ set (characterize_consistent_signs_at_roots p qs)" using all_info reduction_signs_set_helper_lemma2[where A = "set (characterize_consistent_signs_at_roots p qs)", where B = "signs", where C = "(solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))"] using distinct_signs dim_hyp2 solve_for_lhs_R_def by (simp add: mat_inverse_same) then show ?thesis unfolding reduction_signs_R_def by auto qed have subset: "set (characterize_consistent_signs_at_roots p qs) \ set(reduction_signs_R p qs signs subsets (matrix_A_R signs subsets))" using reduction_doesnt_break_things_signs_R[of p qs signs subsets] assms by blast then show ?thesis using supset subset by auto qed theorem reduce_system_sign_conditions_R: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "(nat list*nat list) list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes welldefined_signs1: "well_def_signs (length qs) signs" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots p qs) \ set(signs)" assumes match: "satisfy_equation_R p qs subsets signs" assumes invertible_mat: "invertible_mat (matrix_A_R signs subsets)" shows "set (get_signs_R (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs))))) = set (characterize_consistent_signs_at_roots p qs)" unfolding get_signs_R_def using reduction_deletes_bad_sign_conds_R[of p qs signs subsets] apply (auto) apply (simp add: all_info distinct_signs match nonzero reduction_signs_def welldefined_signs1) using nonzero invertible_mat snd_conv apply (metis reduction_signs_R_def) by (metis all_info distinct_signs invertible_mat match nonzero reduction_signs_R_def snd_conv welldefined_signs1) subsection "Showing matrix equation preserved when reducing" lemma reduce_system_matrix_equation_preserved_R: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "(nat list*nat list) list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes welldefined_signs: "well_def_signs (length qs) signs" assumes welldefined_subsets: "all_list_constr_R (subsets) (length qs)" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots p qs) \ set(signs)" assumes match: "satisfy_equation_R p qs subsets signs" assumes invertible_mat: "invertible_mat (matrix_A_R signs subsets)" shows "satisfy_equation_R p qs (get_subsets_R (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs))))) (get_signs_R (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs)))))" proof - have poly_type_hyp: "p \ 0" using nonzero by auto have distinct_signs_hyp: "distinct (snd (snd (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs))))))" proof - let ?sym = "(find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))" have h1: "\ i < length (take_indices signs ?sym). \j < length(take_indices signs ?sym). i \ j \ nth (take_indices signs ?sym) i \ nth (take_indices signs ?sym) j" using distinct_signs unfolding take_indices_def proof clarsimp fix i fix j assume "distinct signs" assume "i < length (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))" assume "j < length (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))" assume neq_hyp: "i \ j" assume "signs ! (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)) ! i) = signs ! (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)) ! j)" have h1: "find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)) ! i \ find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)) ! j" unfolding find_nonzeros_from_input_vec_def using neq_hyp by (metis \i < length (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))\ \j < length (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))\ distinct_conv_nth distinct_filter distinct_upt find_nonzeros_from_input_vec_def) then show "False" using distinct_signs proof - have f1: "\p ns n. ((n::nat) \ {n \ set ns. p n}) = (n \ set ns \ n \ Collect p)" by simp then have f2: "filter (\n. solve_for_lhs_R p qs subsets (matrix_A_R signs subsets) $ n \ 0) [0.. set [0..i < length (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))\ construct_lhs_matches_solve_for_lhs_R find_nonzeros_from_input_vec_def invertible_mat match nth_mem set_filter size_of_lhs_R) have "filter (\n. solve_for_lhs_R p qs subsets (matrix_A_R signs subsets) $ n \ 0) [0.. set [0..j < length (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))\ construct_lhs_matches_solve_for_lhs_R find_nonzeros_from_input_vec_def invertible_mat match nth_mem set_filter size_of_lhs_R) then show ?thesis using f2 by (metis \signs ! (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)) ! i) = signs ! (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)) ! j)\ atLeastLessThan_iff distinct_conv_nth distinct_signs find_nonzeros_from_input_vec_def h1 set_upt) qed qed then have "distinct (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))))" using distinct_conv_nth by blast then show ?thesis using get_signs_R_def reduction_signs_R_def reduction_signs_is_get_signs_R by auto qed have all_info_hyp: "set (characterize_consistent_signs_at_roots p qs) \ set(snd (snd (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs))))))" using reduce_system_sign_conditions_R assms unfolding get_signs_R_def by auto have welldefined_hyp: "all_list_constr_R (fst (snd (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs)))))) (length qs)" using welldefined_subsets rows_to_keep_lem unfolding all_list_constr_R_def List.member_def list_constr_def list_all_def apply (auto simp add: Let_def take_indices_def take_cols_from_matrix_def) using nth_mem apply (smt (z3) mat_of_cols_carrier(2) rows_to_keep_lem) by (smt (z3) mat_of_cols_carrier(2) nth_mem rows_to_keep_lem) then show ?thesis using poly_type_hyp distinct_signs_hyp all_info_hyp welldefined_hyp using matrix_equation_R unfolding get_subsets_R_def get_signs_R_def by blast qed (* Show that we are tracking the correct matrix in the algorithm *) subsection "Showing matrix preserved" lemma reduce_system_matrix_signs_helper_aux_R: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "(nat list*nat list) list" fixes signs :: "rat list list" fixes S:: "nat list" assumes well_def_h: "\x. List.member S x \ x < length signs" assumes nonzero: "p \ 0" shows "alt_matrix_A_R (take_indices signs S) subsets = take_cols_from_matrix (alt_matrix_A_R signs subsets) S" proof - have h0a: "dim_col (take_cols_from_matrix (alt_matrix_A_R signs subsets) S) = length (take_indices signs S)" unfolding take_cols_from_matrix_def apply (auto) unfolding take_indices_def by auto have h0: "\i < length (take_indices signs S). (col (alt_matrix_A_R (take_indices signs S) subsets ) i = col (take_cols_from_matrix (alt_matrix_A_R signs subsets) S) i)" proof clarsimp fix i assume asm: "i < length (take_indices signs S)" have i_lt: "i < length (map ((!) (cols (alt_matrix_A_R signs subsets))) S)" using asm apply (auto) unfolding take_indices_def by auto have h0: " vec (length subsets) (\j. z_R (subsets ! j) (map ((!) signs) S ! i)) = vec (length subsets) (\j. z_R (subsets ! j) (signs ! (S ! i)))" using nth_map by (metis \i < length (take_indices signs S)\ length_map take_indices_def) have dim: "(map ((!) (cols (alt_matrix_A_R signs subsets))) S) ! i \ carrier_vec (dim_row (alt_matrix_A_R signs subsets))" proof - have "dim_col (alt_matrix_A_R signs subsets) = length (signs)" by (simp add: alt_matrix_A_R_def) have well_d: "S ! i < length (signs)" using well_def_h using i_lt in_set_member by fastforce have map_eq: "(map ((!) (cols (alt_matrix_A_R signs subsets))) S) ! i = nth (cols (alt_matrix_A_R signs subsets)) (S ! i)" using i_lt by auto have "nth (cols (alt_matrix_A_R signs subsets)) (S ! i) \ carrier_vec (dim_row (alt_matrix_A_R signs subsets))" using col_dim unfolding cols_def using nth_map well_d by (simp add: \dim_col (alt_matrix_A_R signs subsets) = length signs\) then show ?thesis using map_eq by auto qed have h1: "col (take_cols_from_matrix (alt_matrix_A_R signs subsets) S) i = col (mat_of_cols (dim_row (alt_matrix_A_R signs subsets)) (map ((!) (cols (alt_matrix_A_R signs subsets))) S)) i " by (simp add: take_cols_from_matrix_def take_indices_def) have h2: "col (mat_of_cols (dim_row (alt_matrix_A_R signs subsets)) (map ((!) (cols (alt_matrix_A_R signs subsets))) S)) i = nth (map ((!) (cols (alt_matrix_A_R signs subsets))) S) i " using dim i_lt asm col_mat_of_cols[where j = "i", where n = "(dim_row (alt_matrix_A_R signs subsets))", where vs = "(map ((!) (cols (alt_matrix_A_R signs subsets))) S)"] by blast have h3: "col (take_cols_from_matrix (alt_matrix_A_R signs subsets) S) i = (col (alt_matrix_A_R signs subsets) (S !i))" using h1 h2 apply (auto) by (metis alt_matrix_char_R asm cols_nth dim_col_mat(1) in_set_member length_map mat_of_rows_list_def matrix_A_R_def nth_map nth_mem take_indices_def well_def_h) have "vec (length subsets) (\j. z_R (subsets ! j) (signs ! (S ! i))) = (col (alt_matrix_A_R signs subsets) (S !i))" by (metis asm in_set_member length_map nth_mem signs_are_cols_R take_indices_def well_def_h) then have "vec (length subsets) (\j. z_R (subsets ! j) (take_indices signs S ! i)) = col (take_cols_from_matrix (alt_matrix_A_R signs subsets) S) i " using h0 h3 by (simp add: take_indices_def) then show "col (alt_matrix_A_R (take_indices signs S) subsets) i = col (take_cols_from_matrix (alt_matrix_A_R signs subsets) S) i " using asm signs_are_cols_R[where signs = "(take_indices signs S)", where subsets = "subsets"] by auto qed then show ?thesis unfolding alt_matrix_A_R_def take_cols_from_matrix_def apply (auto) using h0a mat_col_eqI by (metis alt_matrix_A_R_def dim_col_mat(1) dim_row_mat(1) h0 mat_of_cols_def take_cols_from_matrix_def) qed lemma reduce_system_matrix_signs_helper_R: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "(nat list*nat list) list" fixes signs :: "rat list list" fixes S:: "nat list" assumes well_def_h: "\x. List.member S x \ x < length signs" assumes nonzero: "p \ 0" shows "matrix_A_R (take_indices signs S) subsets = take_cols_from_matrix (matrix_A_R signs subsets) S" using reduce_system_matrix_signs_helper_aux_R alt_matrix_char_R assms by auto lemma reduce_system_matrix_subsets_helper_aux_R: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "(nat list* nat list) list" fixes signs :: "rat list list" fixes S:: "nat list" assumes inv: "length subsets \ length signs" assumes well_def_h: "\x. List.member S x \ x < length subsets" assumes nonzero: "p \ 0" shows "alt_matrix_A_R signs (take_indices subsets S) = take_rows_from_matrix (alt_matrix_A_R signs subsets) S" proof - have h0a: "dim_row (take_rows_from_matrix (alt_matrix_A_R signs subsets) S) = length (take_indices subsets S)" unfolding take_rows_from_matrix_def apply (auto) unfolding take_indices_def by auto have h0: "\i < length (take_indices subsets S). (row (alt_matrix_A_R signs (take_indices subsets S) ) i = row (take_rows_from_matrix (alt_matrix_A_R signs subsets) S) i)" proof clarsimp fix i assume asm: "i < length (take_indices subsets S)" have i_lt: "i < length (map ((!) (rows (alt_matrix_A_R signs subsets))) S)" using asm apply (auto) unfolding take_indices_def by auto have h0: "row (take_rows_from_matrix (alt_matrix_A_R signs subsets) S) i = row (mat_of_rows (dim_col (alt_matrix_A_R signs subsets)) (map ((!) (rows (alt_matrix_A_R signs subsets))) S)) i " unfolding take_rows_from_matrix_def take_indices_def by auto have dim: "(map ((!) (rows (alt_matrix_A_R signs subsets))) S) ! i \ carrier_vec (dim_col (alt_matrix_A_R signs subsets))" proof - have "dim_col (alt_matrix_A_R signs subsets) = length (signs)" by (simp add: alt_matrix_A_R_def) then have lenh: "dim_col (alt_matrix_A_R signs subsets) \ length (subsets)" using assms by auto have well_d: "S ! i < length (subsets)" using well_def_h using i_lt in_set_member by fastforce have map_eq: "(map ((!) (rows (alt_matrix_A_R signs subsets))) S) ! i = nth (rows (alt_matrix_A_R signs subsets)) (S ! i)" using i_lt by auto have "nth (rows (alt_matrix_A_R signs subsets)) (S ! i) \ carrier_vec (dim_col (alt_matrix_A_R signs subsets))" using col_dim unfolding rows_def using nth_map well_d using lenh by (simp add: alt_matrix_A_R_def) then show ?thesis using map_eq unfolding alt_matrix_A_R_def by auto qed have h1: " row (mat_of_rows (dim_col (alt_matrix_A_R signs subsets)) (map ((!) (rows (alt_matrix_A_R signs subsets))) S)) i = (row (alt_matrix_A_R signs subsets) (S ! i)) " using dim i_lt mat_of_rows_row[where i = "i", where n = "(dim_col (alt_matrix_A_R signs subsets))", where vs = "(map ((!) (rows (alt_matrix_A_R signs subsets))) S)"] using alt_matrix_char_R in_set_member nth_mem well_def_h by fastforce have "row (alt_matrix_A_R signs (take_indices subsets S) ) i = (row (alt_matrix_A_R signs subsets) (S ! i))" using subsets_are_rows_R proof - have f1: "i < length S" by (metis (no_types) asm length_map take_indices_def) then have "List.member S (S ! i)" by (meson in_set_member nth_mem) then show ?thesis using f1 by (simp add: \\subsets signs. \ij. z_R (subsets ! i) (signs ! j))\ take_indices_def well_def_h) qed then show "(row (alt_matrix_A_R signs (take_indices subsets S) ) i = row (take_rows_from_matrix (alt_matrix_A_R signs subsets) S) i)" using h0 h1 unfolding take_indices_def by auto qed then show ?thesis unfolding alt_matrix_A_R_def take_rows_from_matrix_def apply (auto) using eq_rowI by (metis alt_matrix_A_R_def dim_col_mat(1) dim_row_mat(1) h0 h0a mat_of_rows_def take_rows_from_matrix_def) qed lemma reduce_system_matrix_subsets_helper_R: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "(nat list*nat list) list" fixes signs :: "rat list list" fixes S:: "nat list" assumes nonzero: "p \ 0" assumes inv: "length subsets \ length signs" assumes well_def_h: "\x. List.member S x \ x < length subsets" shows "matrix_A_R signs (take_indices subsets S) = take_rows_from_matrix (matrix_A_R signs subsets) S" using assms reduce_system_matrix_subsets_helper_aux_R alt_matrix_char_R by auto lemma reduce_system_matrix_match_R: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "(nat list*nat list) list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes welldefined_signs1: "well_def_signs (length qs) signs" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots p qs) \ set(signs)" assumes match: "satisfy_equation_R p qs subsets signs" assumes inv: "invertible_mat (matrix_A_R signs subsets)" shows "matrix_A_R (get_signs_R (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs))))) (get_subsets_R (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs))))) = (get_matrix_R (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs)))))" proof - let ?A = "(matrix_A_R signs subsets)" let ?lhs_vec = "(solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))" let ?red_mtx = "(take_rows_from_matrix (reduce_mat_cols (matrix_A_R signs subsets) ?lhs_vec) (rows_to_keep (reduce_mat_cols (matrix_A_R signs subsets) ?lhs_vec)))" have h1: "matrix_A_R (get_signs_R (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs))))) (get_subsets_R (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs))))) = matrix_A_R (reduction_signs_R p qs signs subsets (matrix_A_R signs subsets)) (reduction_subsets_R p qs signs subsets (matrix_A_R signs subsets))" using reduction_signs_is_get_signs_R reduction_subsets_is_get_subsets_R by auto have h1_var: "matrix_A_R (get_signs_R (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs))))) (get_subsets_R (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs))))) = matrix_A_R (take_indices signs (find_nonzeros_from_input_vec ?lhs_vec)) (take_indices subsets (rows_to_keep (reduce_mat_cols ?A ?lhs_vec)))" using h1 reduction_signs_R_def reduction_subsets_R_def by auto have h2: "?red_mtx = (take_rows_from_matrix (take_cols_from_matrix ?A (find_nonzeros_from_input_vec ?lhs_vec)) (rows_to_keep (take_cols_from_matrix ?A (find_nonzeros_from_input_vec ?lhs_vec))))" by simp have h30: "(construct_lhs_vector_R p qs signs) = ?lhs_vec" using assms construct_lhs_matches_solve_for_lhs_R by simp have h3a: "\x. List.member (find_nonzeros_from_input_vec ?lhs_vec) x \x < length (signs)" using h30 size_of_lhs_R[of p qs signs] unfolding find_nonzeros_from_input_vec_def apply (auto) using in_set_member by force have h3b_a: "\x. List.member (find_nonzeros_from_input_vec ?lhs_vec) x \x < length (subsets)" using assms h30 size_of_lhs_R same_size_R unfolding find_nonzeros_from_input_vec_def apply (auto) by (simp add: find_nonzeros_from_input_vec_def h3a) have h3ba: "length (filter (\i. solve_for_lhs_R p qs subsets (matrix_A_R signs subsets) $ i \ 0) [0.. length subsets" using length_filter_le[where P = "(\i. solve_for_lhs_R p qs subsets (matrix_A_R signs subsets) $ i \ 0)", where xs = "[0..i. solve_for_lhs_R p qs subsets (matrix_A_R signs subsets) $ i \ 0) [0.. length subsets" using h3ba by auto then have h3b: "length subsets \ length (take_indices signs (find_nonzeros_from_input_vec ?lhs_vec))" unfolding take_indices_def find_nonzeros_from_input_vec_def by auto have h3c: "\x. List.member (rows_to_keep (reduce_mat_cols ?A ?lhs_vec)) x \ x < length (subsets)" proof clarsimp fix x assume x_mem: "List.member (rows_to_keep (take_cols_from_matrix (matrix_A_R signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))))) x" obtain nn :: "rat list list \ nat list \ nat" where "\x2 x3. (\v4. v4 \ set x3 \ \ v4 < length x2) = (nn x2 x3 \ set x3 \ \ nn x2 x3 < length x2)" by moura then have f4: "nn signs (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))) \ set (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))) \ \ nn signs (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))) < length signs \ matrix_A_R (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))) subsets = take_cols_from_matrix (matrix_A_R signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))" using nonzero using h3a reduce_system_matrix_signs_helper_R by auto then have "matrix_A_R (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))) subsets = take_cols_from_matrix (matrix_A_R signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))) \ x \ set (map snd (pivot_positions (gauss_jordan_single (take_cols_from_matrix (matrix_A_R signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))))\<^sup>T)))" using f4 by (metis h3a in_set_member rows_to_keep_def x_mem) thus "x < length (subsets)" using x_mem unfolding rows_to_keep_def by (metis dim_row_matrix_A_R rows_to_keep_def rows_to_keep_lem) qed have h3: "matrix_A_R (take_indices signs (find_nonzeros_from_input_vec ?lhs_vec)) (take_indices subsets (rows_to_keep (reduce_mat_cols ?A ?lhs_vec))) = (take_rows_from_matrix (take_cols_from_matrix ?A (find_nonzeros_from_input_vec ?lhs_vec)) (rows_to_keep (take_cols_from_matrix ?A (find_nonzeros_from_input_vec ?lhs_vec))))" using inv h3a h3b h3c reduce_system_matrix_subsets_helper_R reduce_system_matrix_signs_helper_R assms by auto show ?thesis using h1 h2 h3 by (metis fst_conv get_matrix_R_def h1_var reduce_system_R.simps reduction_step_R.simps) qed subsection "Showing invertibility preserved when reducing" (* gauss_jordan_single_rank is critically helpful in this subsection *) thm conjugatable_vec_space.gauss_jordan_single_rank thm vec_space.full_rank_lin_indpt (* Overall: Start with a matrix equation. Input a matrix, subsets, and signs. Drop columns of the matrix based on the 0's on the LHS---so extract a list of 0's. Reduce signs accordingly. Then find a list of rows to delete based on using rank (use the transpose result, pivot positions!), and delete those rows. Reduce subsets accordingly. End with a reduced system! *) lemma well_def_find_zeros_from_lhs_vec_R: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "(nat list*nat list) list" fixes signs :: "rat list list" assumes len_eq: "length subsets = length signs" assumes inv: "invertible_mat (matrix_A_R signs subsets)" assumes nonzero: "p \ 0" assumes welldefined_signs1: "well_def_signs (length qs) signs" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots p qs) \ set(signs)" assumes match: "satisfy_equation_R p qs subsets signs" shows "(\j. j \ set (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))) \ j < length (cols (matrix_A_R signs subsets)))" proof - fix i fix j assume j_in: " j \ set (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))) " let ?og_mat = "(matrix_A_R signs subsets)" let ?lhs = "(solve_for_lhs_R p qs subsets ?og_mat)" let ?new_mat = "(take_rows_from_matrix (reduce_mat_cols ?og_mat ?lhs) (rows_to_keep (reduce_mat_cols ?og_mat ?lhs)))" have "square_mat (matrix_A_R signs subsets)" using inv using invertible_mat_def by blast then have mat_size: "?og_mat \ carrier_mat (length signs) (length signs)" using size_of_mat by auto have "dim_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)) = (length signs)" using size_of_lhs_R construct_lhs_matches_solve_for_lhs_R assms by (metis (full_types)) then have h: "j < (length signs)" using j_in unfolding find_nonzeros_from_input_vec_def by simp then show "j < length (cols (matrix_A_R signs subsets))" using mat_size by auto qed lemma take_cols_subsets_og_cols_R: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "(nat list*nat list) list" fixes signs :: "rat list list" assumes len_eq: "length subsets = length signs" assumes inv: "invertible_mat (matrix_A_R signs subsets)" assumes nonzero: "p \ 0" assumes welldefined_signs1: "well_def_signs (length qs) signs" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots p qs) \ set(signs)" assumes match: "satisfy_equation_R p qs subsets signs" shows "set (take_indices (cols (matrix_A_R signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))) \ set (cols (matrix_A_R signs subsets))" proof - have well_def: "(\j. j \ set (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))) \ j < length (cols (matrix_A_R signs subsets)))" using assms well_def_find_zeros_from_lhs_vec_R by auto have "\x. x \ set (take_indices (cols (matrix_A_R signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))) \ x \ set (cols (matrix_A_R signs subsets))" proof clarsimp fix x let ?og_list = "(cols (matrix_A_R signs subsets))" let ?ind_list = "(find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))" assume x_in: "x \ set (take_indices ?og_list ?ind_list)" show "x \ set (cols (matrix_A_R signs subsets))" using x_in unfolding take_indices_def using in_set_member apply (auto) using in_set_conv_nth well_def by fastforce qed then show ?thesis by blast qed lemma reduction_doesnt_break_things_invertibility_step1_R: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "(nat list*nat list) list" fixes signs :: "rat list list" assumes len_eq: "length subsets = length signs" assumes inv: "invertible_mat (matrix_A_R signs subsets)" assumes nonzero: "p \ 0" assumes welldefined_signs1: "well_def_signs (length qs) signs" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots p qs) \ set(signs)" assumes match: "satisfy_equation_R p qs subsets signs" shows "vec_space.rank (length signs) (reduce_mat_cols (matrix_A_R signs subsets) (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))) = (length (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))))" proof - let ?og_mat = "(matrix_A_R signs subsets)" let ?lhs = "(solve_for_lhs_R p qs subsets ?og_mat)" let ?new_mat = "(take_rows_from_matrix (reduce_mat_cols ?og_mat ?lhs) (rows_to_keep (reduce_mat_cols ?og_mat ?lhs)))" have "square_mat (matrix_A_R signs subsets)" using inv using invertible_mat_def by blast then have mat_size: "?og_mat \ carrier_mat (length signs) (length signs)" using size_of_mat by auto then have mat_size_alt: "?og_mat \ carrier_mat (length subsets) (length subsets)" using size_of_mat same_size assms by auto have det_h: "det ?og_mat \ 0" using invertible_det[where A = "matrix_A_R signs subsets"] mat_size using inv by blast then have rank_h: "vec_space.rank (length signs) ?og_mat = (length signs)" using vec_space.det_rank_iff mat_size by auto then have dist_cols: "distinct (cols ?og_mat)" using mat_size vec_space.non_distinct_low_rank[where A = ?og_mat, where n = "length signs"] by auto have well_def: "(\j. j \ set (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))) \ j < length (cols (matrix_A_R signs subsets)))" using assms well_def_find_zeros_from_lhs_vec_R by auto have dist1: "distinct (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))" unfolding find_nonzeros_from_input_vec_def by auto have clear: "set (take_indices (cols (matrix_A_R signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))) \ set (cols (matrix_A_R signs subsets))" using assms take_cols_subsets_og_cols_R by auto then have "distinct (take_indices (cols (matrix_A_R signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))))" unfolding take_indices_def using dist1 dist_cols well_def conjugatable_vec_space.distinct_map_nth[where ls = "cols (matrix_A_R signs subsets)", where inds = "(find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))"] by auto then have unfold_thesis: "vec_space.rank (length signs) (mat_of_cols (dim_row ?og_mat) (take_indices (cols ?og_mat) (find_nonzeros_from_input_vec ?lhs))) = (length (find_nonzeros_from_input_vec ?lhs))" using clear inv conjugatable_vec_space.rank_invertible_subset_cols[where A= "matrix_A_R signs subsets", where B = "(take_indices (cols (matrix_A_R signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))))"] by (simp add: len_eq mat_size take_indices_def) then show ?thesis apply (simp) unfolding take_cols_from_matrix_def by auto qed lemma reduction_doesnt_break_things_invertibility_R: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "(nat list*nat list) list" fixes signs :: "rat list list" assumes len_eq: "length subsets = length signs" assumes inv: "invertible_mat (matrix_A_R signs subsets)" assumes nonzero: "p \ 0" assumes welldefined_signs1: "well_def_signs (length qs) signs" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots p qs) \ set(signs)" assumes match: "satisfy_equation_R p qs subsets signs" shows "invertible_mat (get_matrix_R (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs)))))" proof - let ?og_mat = "(matrix_A_R signs subsets)" let ?lhs = "(solve_for_lhs_R p qs subsets ?og_mat)" let ?step1_mat = "(reduce_mat_cols ?og_mat ?lhs)" let ?new_mat = "take_rows_from_matrix ?step1_mat (rows_to_keep ?step1_mat)" let ?ind_list = "(find_nonzeros_from_input_vec ?lhs)" have "square_mat (matrix_A_R signs subsets)" using inv using invertible_mat_def by blast then have og_mat_size: "?og_mat \ carrier_mat (length signs) (length signs)" using size_of_mat by auto have "dim_col (mat_of_cols (dim_row ?og_mat) (take_indices (cols ?og_mat) ?ind_list)) = (length (find_nonzeros_from_input_vec ?lhs))" by (simp add: take_indices_def) then have "mat_of_cols (dim_row ?og_mat) (take_indices (cols ?og_mat) ?ind_list) \ carrier_mat (length signs) (length (find_nonzeros_from_input_vec ?lhs))" by (simp add: len_eq mat_of_cols_def) then have step1_mat_size: "?step1_mat \ carrier_mat (length signs) (length (find_nonzeros_from_input_vec ?lhs))" by (simp add: take_cols_from_matrix_def) then have "dim_row ?step1_mat \ dim_col ?step1_mat" by (metis carrier_matD(1) carrier_matD(2) construct_lhs_matches_solve_for_lhs_R diff_zero find_nonzeros_from_input_vec_def inv length_filter_le length_upt match size_of_lhs_R) then have gt_eq_assm: "dim_col ?step1_mat\<^sup>T \ dim_row ?step1_mat\<^sup>T" by simp have det_h: "det ?og_mat \ 0" using invertible_det[where A = "matrix_A_R signs subsets"] og_mat_size using inv by blast then have rank_h: "vec_space.rank (length signs) ?og_mat = (length signs)" using vec_space.det_rank_iff og_mat_size by auto have rank_drop_cols: "vec_space.rank (length signs) ?step1_mat = (dim_col ?step1_mat)" using assms reduction_doesnt_break_things_invertibility_step1_R step1_mat_size by auto let ?step1_T = "?step1_mat\<^sup>T" have rank_transpose: "vec_space.rank (length signs) ?step1_mat = vec_space.rank (length (find_nonzeros_from_input_vec ?lhs)) ?step1_T" using transpose_rank[of ?step1_mat] using step1_mat_size by auto have obv: "?step1_T \ carrier_mat (dim_row ?step1_T) (dim_col ?step1_T)" by auto have should_have_this:"vec_space.rank (length (find_nonzeros_from_input_vec ?lhs)) (take_cols ?step1_T (map snd (pivot_positions (gauss_jordan_single (?step1_T))))) = vec_space.rank (length (find_nonzeros_from_input_vec ?lhs)) ?step1_T" using obv gt_eq_assm conjugatable_vec_space.gauss_jordan_single_rank[where A = "?step1_T", where n = "dim_row ?step1_T", where nc = "dim_col ?step1_T"] by (simp add: take_cols_from_matrix_def take_indices_def) then have "vec_space.rank (length (find_nonzeros_from_input_vec ?lhs)) (take_cols ?step1_T (map snd (pivot_positions (gauss_jordan_single (?step1_T))))) = dim_col ?step1_mat" using rank_drop_cols rank_transpose by auto then have with_take_cols_from_matrix: "vec_space.rank (length (find_nonzeros_from_input_vec ?lhs)) (take_cols_from_matrix ?step1_T (map snd (pivot_positions (gauss_jordan_single (?step1_T))))) = dim_col ?step1_mat" using rank_transpose rechar_take_cols conjugatable_vec_space.gjs_and_take_cols_var apply (auto) by (smt conjugatable_vec_space.gjs_and_take_cols_var gt_eq_assm obv rechar_take_cols reduce_mat_cols.simps) have "(take_rows_from_matrix ?step1_mat (rows_to_keep ?step1_mat)) = (take_cols_from_matrix ?step1_T (rows_to_keep ?step1_mat))\<^sup>T" using take_rows_and_take_cols by blast then have rank_new_mat: "vec_space.rank (dim_row ?new_mat) ?new_mat = dim_col ?step1_mat" using with_take_cols_from_matrix transpose_rank apply (auto) by (smt (verit, ccfv_threshold) carrier_matD(2) index_transpose_mat(2) mat_of_cols_carrier(2) reduce_mat_cols.simps rows_to_keep_def step1_mat_size take_cols_from_matrix_def transpose_rank) have "length (pivot_positions (gauss_jordan_single (?step1_mat\<^sup>T))) \ (length (find_nonzeros_from_input_vec ?lhs))" using obv length_pivot_positions_dim_row[where A = "(gauss_jordan_single (?step1_mat\<^sup>T))"] by (metis carrier_matD(1) carrier_matD(2) gauss_jordan_single(2) gauss_jordan_single(3) index_transpose_mat(2) step1_mat_size) then have len_lt_eq: "length (pivot_positions (gauss_jordan_single (?step1_mat\<^sup>T))) \ dim_col ?step1_mat" using step1_mat_size by simp have len_gt_false: "length (rows_to_keep ?step1_mat) < (dim_col ?step1_mat) \ False" proof - assume length_lt: "length (rows_to_keep ?step1_mat) < (dim_col ?step1_mat)" have h: "dim_row ?new_mat < (dim_col ?step1_mat)" by (metis Matrix.transpose_transpose index_transpose_mat(3) length_lt length_map mat_of_cols_carrier(3) take_cols_from_matrix_def take_indices_def take_rows_and_take_cols) then show "False" using rank_new_mat by (metis Matrix.transpose_transpose carrier_matI index_transpose_mat(2) nat_less_le not_less_iff_gr_or_eq transpose_rank vec_space.rank_le_nc) qed then have len_gt_eq: "length (rows_to_keep ?step1_mat) \ (dim_col ?step1_mat)" using not_less by blast have len_match: "length (rows_to_keep ?step1_mat) = (dim_col ?step1_mat)" using len_lt_eq len_gt_eq by (simp add: rows_to_keep_def) have mat_match: "matrix_A_R (get_signs_R (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs))))) (get_subsets_R (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs))))) = (get_matrix_R (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs)))))" using reduce_system_matrix_match_R[of p qs signs subsets] assms by blast have f2: "length (get_subsets_R (take_rows_from_matrix (mat_of_cols (dim_row (matrix_A_R signs subsets)) (map ((!) (cols (matrix_A_R signs subsets))) (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))))) (rows_to_keep (mat_of_cols (dim_row (matrix_A_R signs subsets)) (map ((!) (cols (matrix_A_R signs subsets))) (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))))), map ((!) subsets) (rows_to_keep (mat_of_cols (dim_row (matrix_A_R signs subsets)) (map ((!) (cols (matrix_A_R signs subsets))) (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))))), map ((!) signs) (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))))) = length (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))" by (metis (no_types) \dim_col (mat_of_cols (dim_row (matrix_A_R signs subsets)) (take_indices (cols (matrix_A_R signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))))) = length (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))\ \length (rows_to_keep (reduce_mat_cols (matrix_A_R signs subsets) (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))) = dim_col (reduce_mat_cols (matrix_A_R signs subsets) (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))\ length_map reduce_mat_cols.simps reduce_system_R.simps reduction_step_R.simps reduction_subsets_R_def reduction_subsets_is_get_subsets_R take_cols_from_matrix_def take_indices_def) have f3: "\p ps rss nss m. map ((!) rss) (find_nonzeros_from_input_vec (solve_for_lhs_R p ps nss m)) = get_signs_R (reduce_system_R p (ps, m, nss, rss))" by (metis (no_types) reduction_signs_R_def reduction_signs_is_get_signs_R take_indices_def) have square_final_mat: "square_mat (get_matrix_R (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs)))))" using mat_match assms size_of_mat_R same_size_R apply (auto) using f2 f3 by (metis (no_types, lifting) Matrix.transpose_transpose fst_conv get_matrix_R_def index_transpose_mat(2) len_match length_map mat_of_cols_carrier(2) mat_of_cols_carrier(3) reduce_mat_cols.simps take_cols_from_matrix_def take_indices_def take_rows_and_take_cols) have rank_match: "vec_space.rank (dim_row ?new_mat) ?new_mat = dim_row ?new_mat" using len_match rank_new_mat by (simp add: take_cols_from_matrix_def take_indices_def take_rows_and_take_cols) have "invertible_mat ?new_mat" using invertible_det og_mat_size using vec_space.det_rank_iff square_final_mat by (metis dim_col_matrix_A_R dim_row_matrix_A_R fst_conv get_matrix_R_def mat_match rank_match reduce_system_R.simps reduction_step_R.simps size_of_mat_R square_mat.elims(2)) then show ?thesis apply (simp) by (metis fst_conv get_matrix_R_def) qed subsection "Well def signs preserved when reducing" lemma reduction_doesnt_break_length_signs_R: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "(nat list*nat list) list" fixes signs :: "rat list list" assumes length_init : "\ x\ set(signs). length x = length qs" assumes sat_eq: "satisfy_equation_R p qs subsets signs" assumes inv_mat: "invertible_mat (matrix_A_R signs subsets)" shows "\x \ set(reduction_signs_R p qs signs subsets (matrix_A_R signs subsets)). length x = length qs" unfolding reduction_signs_def using take_indices_lem by (smt (verit) atLeastLessThan_iff construct_lhs_matches_solve_for_lhs_R filter_is_subset find_nonzeros_from_input_vec_def in_set_conv_nth in_set_member inv_mat length_init reduction_signs_R_def sat_eq set_upt size_of_lhs_R subset_code(1)) subsection "Distinct signs preserved when reducing" lemma reduction_signs_are_distinct_R: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "(nat list*nat list) list" fixes signs :: "rat list list" assumes sat_eq: "satisfy_equation_R p qs subsets signs" assumes inv_mat: "invertible_mat (matrix_A_R signs subsets)" assumes distinct_init: "distinct signs" shows "distinct (reduction_signs_R p qs signs subsets (matrix_A_R signs subsets))" proof - have solve_construct: "construct_lhs_vector_R p qs signs = solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)" using construct_lhs_matches_solve_for_lhs_R assms by simp have h1: "distinct (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))" unfolding find_nonzeros_from_input_vec_def using distinct_filter using distinct_upt by blast have h2: "(\j. j \ set (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))) \ j < length signs)" proof - fix j assume "j \ set (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))" show "j < length signs" using solve_construct size_of_lhs_R by (metis \j \ set (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))\ atLeastLessThan_iff filter_is_subset find_nonzeros_from_input_vec_def set_upt subset_iff) qed then show ?thesis unfolding reduction_signs_R_def unfolding take_indices_def using distinct_init h1 h2 conjugatable_vec_space.distinct_map_nth[where ls = "signs", where inds = "(find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))"] by blast qed subsection "Well def subsets preserved when reducing" lemma reduction_doesnt_break_subsets_R: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "(nat list* nat list) list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes length_init : "all_list_constr_R subsets (length qs)" assumes sat_eq: "satisfy_equation_R p qs subsets signs" assumes inv_mat: "invertible_mat (matrix_A_R signs subsets)" shows "all_list_constr_R (reduction_subsets_R p qs signs subsets (matrix_A_R signs subsets)) (length qs)" unfolding all_list_constr_R_def proof clarsimp fix a b assume in_red_subsets: "List.member (reduction_subsets_R p qs signs subsets (matrix_A_R signs subsets)) (a, b) " have solve_construct: "construct_lhs_vector_R p qs signs = solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)" using construct_lhs_matches_solve_for_lhs_R assms by simp have rows_to_keep_hyp: "\y. y \ set (rows_to_keep (reduce_mat_cols (matrix_A_R signs subsets) (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))) \ y < length subsets" proof clarsimp fix y assume in_set: "y \ set (rows_to_keep (take_cols_from_matrix (matrix_A_R signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))))" have in_set_2: "y \ set (rows_to_keep (take_cols_from_matrix (matrix_A_R signs subsets) (find_nonzeros_from_input_vec (construct_lhs_vector_R p qs signs))))" using in_set solve_construct by simp let ?lhs_vec = "(solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))" have h30: "(construct_lhs_vector_R p qs signs) = ?lhs_vec" using assms construct_lhs_matches_solve_for_lhs_R by simp have h3a: "\x. List.member (find_nonzeros_from_input_vec ?lhs_vec) x \x < length (signs)" using h30 size_of_lhs_R unfolding find_nonzeros_from_input_vec_def using atLeastLessThan_iff filter_is_subset member_def set_upt subset_eq apply (auto) by (smt (verit, best) atLeastLessThan_iff in_set_member mem_Collect_eq set_filter set_upt) have h3c: "\x. List.member (rows_to_keep (reduce_mat_cols (matrix_A_R signs subsets) (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))) x \ x < length (subsets)" proof clarsimp fix x assume x_mem: "List.member (rows_to_keep (take_cols_from_matrix (matrix_A_R signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))))) x" show "x < length (subsets)" using x_mem unfolding rows_to_keep_def using pivot_positions using dim_row_matrix_A h3a in_set_member nonzero reduce_system_matrix_signs_helper_R rows_to_keep_def rows_to_keep_lem apply (auto) by (smt (verit, best) List.member_def dim_row_matrix_A_R rows_to_keep_def rows_to_keep_lem) qed then show "y < length subsets" using in_set_member apply (auto) using in_set_2 solve_construct by fastforce qed show "list_constr a (length qs) \ list_constr b (length qs)" using in_red_subsets unfolding reduction_subsets_def using take_indices_lem_R[of _ subsets] rows_to_keep_hyp using all_list_constr_R_def in_set_conv_nth in_set_member length_init by (metis fst_conv reduction_subsets_R_def snd_conv) qed section "Overall Lemmas" lemma combining_to_smash_R: "combine_systems_R p (qs1, m1, (sub1, sgn1)) (qs2, m2, (sub2, sgn2)) = smash_systems_R p qs1 qs2 sub1 sub2 sgn1 sgn2 m1 m2" by simp lemma getter_functions_R: "calculate_data_R p qs = (get_matrix_R (calculate_data_R p qs), (get_subsets_R (calculate_data_R p qs), get_signs_R (calculate_data_R p qs))) " unfolding get_matrix_R_def get_subsets_R_def get_signs_R_def by auto subsection "Key properties preserved" subsubsection "Properties preserved when combining and reducing systems" lemma combining_sys_satisfies_properties_helper_R: fixes p:: "real poly" fixes qs1 :: "real poly list" fixes qs2 :: "real poly list" fixes subsets1 subsets2 :: "(nat list * nat list) list" fixes signs1 signs2 :: "rat list list" fixes matrix1 matrix2:: "rat mat" assumes nonzero: "p \ 0" assumes nontriv1: "length qs1 > 0" assumes nontriv2: "length qs2 > 0" assumes satisfies_properties_sys1: "satisfies_properties_R p qs1 subsets1 signs1 matrix1" assumes satisfies_properties_sys2: "satisfies_properties_R p qs2 subsets2 signs2 matrix2" shows "satisfies_properties_R p (qs1@qs2) (get_subsets_R (snd ((combine_systems_R p (qs1,(matrix1, (subsets1, signs1))) (qs2,(matrix2, (subsets2, signs2))))))) (get_signs_R (snd ((combine_systems_R p (qs1,(matrix1, (subsets1, signs1))) (qs2,(matrix2, (subsets2, signs2))))))) (get_matrix_R (snd ((combine_systems_R p (qs1,(matrix1, (subsets1, signs1))) (qs2,(matrix2, (subsets2, signs2)))))))" proof - let ?subsets = "(get_subsets_R (snd (combine_systems_R p (qs1, matrix1, subsets1, signs1) (qs2, matrix2, subsets2, signs2))))" let ?signs = "(get_signs_R (snd (combine_systems_R p (qs1, matrix1, subsets1, signs1) (qs2, matrix2, subsets2, signs2))))" let ?matrix = "(get_matrix_R (snd (combine_systems_R p (qs1, matrix1, subsets1, signs1) (qs2, matrix2, subsets2, signs2))))" have h1: "all_list_constr_R ?subsets (length (qs1 @ qs2))" using well_def_step_R[of subsets1 qs1 subsets2 qs2] assms by (simp add: nontriv2 get_subsets_R_def satisfies_properties_R_def smash_systems_R_def) have h2: "well_def_signs (length (qs1 @ qs2)) ?signs" using well_def_signs_step[of qs1 qs2 signs1 signs2] using get_signs_R_def nontriv1 nontriv2 satisfies_properties_R_def satisfies_properties_sys1 satisfies_properties_sys2 smash_systems_R_def by (metis combining_to_smash_R snd_conv) have h3: "distinct ?signs" using distinct_step[of _ signs1 _ signs2] assms using combine_systems.simps get_signs_R_def satisfies_properties_R_def smash_systems_R_def snd_conv by (metis combining_to_smash_R) have h4: "satisfy_equation_R p (qs1 @ qs2) ?subsets ?signs" using assms inductive_step_R[of p qs1 qs2 signs1 signs2 subsets1 subsets2] using get_signs_R_def get_subsets_R_def satisfies_properties_R_def smash_systems_R_def by (metis (no_types, opaque_lifting) combining_to_smash_R h1 h3 matrix_equation_R snd_conv subset_step_R) have h5: " invertible_mat ?matrix" using assms inductive_step_R[of p qs1 qs2 signs1 signs2 subsets1 subsets2] by (metis combining_to_smash_R fst_conv get_matrix_R_def kronecker_invertible satisfies_properties_R_def smash_systems_R_def snd_conv) have h6: "?matrix = matrix_A_R ?signs ?subsets" unfolding get_matrix_R_def combine_systems_R.simps smash_systems_R_def get_signs_R_def get_subsets_R_def apply simp apply (subst matrix_construction_is_kronecker_product_R[of subsets1 _ signs1 signs2 subsets2]) apply (metis Ball_set all_list_constr_R_def in_set_member list_constr_def satisfies_properties_R_def satisfies_properties_sys1) using satisfies_properties_R_def satisfies_properties_sys1 well_def_signs_def apply blast using satisfies_properties_R_def satisfies_properties_sys1 satisfies_properties_sys2 by auto have h7: "set (characterize_consistent_signs_at_roots p (qs1 @ qs2)) \ set (?signs)" using subset_step_R[of p qs1 signs1 qs2 signs2] assms by (simp add: nonzero get_signs_R_def satisfies_properties_R_def smash_systems_R_def) then show ?thesis unfolding satisfies_properties_R_def using h1 h2 h3 h4 h5 h6 h7 by blast qed lemma combining_sys_satisfies_properties_R: fixes p:: "real poly" fixes qs1 :: "real poly list" fixes qs2 :: "real poly list" assumes nonzero: "p \ 0" assumes nontriv1: "length qs1 > 0" assumes nontriv2: "length qs2 > 0" assumes satisfies_properties_sys1: "satisfies_properties_R p qs1 (get_subsets_R (calculate_data_R p qs1)) (get_signs_R (calculate_data_R p qs1)) (get_matrix_R (calculate_data_R p qs1))" assumes satisfies_properties_sys2: "satisfies_properties_R p qs2 (get_subsets_R (calculate_data_R p qs2)) (get_signs_R (calculate_data_R p qs2)) (get_matrix_R (calculate_data_R p qs2))" shows "satisfies_properties_R p (qs1@qs2) (get_subsets_R (snd ((combine_systems_R p (qs1,calculate_data_R p qs1) (qs2,calculate_data_R p qs2))))) (get_signs_R (snd ((combine_systems_R p (qs1,calculate_data_R p qs1) (qs2,calculate_data_R p qs2))))) (get_matrix_R (snd ((combine_systems_R p (qs1,calculate_data_R p qs1) (qs2,calculate_data_R p qs2)))))" using combining_sys_satisfies_properties_helper_R[of p qs1 qs2] by (metis getter_functions_R nontriv1 nontriv2 nonzero satisfies_properties_sys1 satisfies_properties_sys2) lemma reducing_sys_satisfies_properties_R: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "(nat list*nat list) list" fixes signs :: "rat list list" fixes matrix:: "rat mat" assumes nonzero: "p \ 0" assumes nontriv: "length qs > 0" assumes satisfies_properties_sys: "satisfies_properties_R p qs subsets signs matrix" shows "satisfies_properties_R p qs (get_subsets_R (reduce_system_R p (qs,matrix,subsets,signs))) (get_signs_R (reduce_system_R p (qs,matrix,subsets,signs))) (get_matrix_R (reduce_system_R p (qs,matrix,subsets,signs)))" proof - have h1: " all_list_constr_R (get_subsets_R (reduce_system_R p (qs, matrix, subsets, signs))) (length qs)" using reduction_doesnt_break_subsets_R assms reduction_subsets_is_get_subsets_R satisfies_properties_R_def satisfies_properties_sys by auto have h2: "well_def_signs (length qs) (get_signs_R (reduce_system_R p (qs, matrix, subsets, signs)))" using reduction_doesnt_break_length_signs_R[of signs qs p subsets] assms reduction_signs_is_get_signs_R satisfies_properties_R_def well_def_signs_def by auto have h3: "distinct (get_signs_R (reduce_system_R p (qs, matrix, subsets, signs)))" using reduction_signs_are_distinct_R[of p qs subsets signs] assms reduction_signs_is_get_signs_R satisfies_properties_R_def by auto have h4: "satisfy_equation_R p qs (get_subsets_R (reduce_system_R p (qs, matrix, subsets, signs))) (get_signs_R (reduce_system_R p (qs, matrix, subsets, signs)))" using reduce_system_matrix_equation_preserved_R[of p qs signs subsets] assms satisfies_properties_R_def by auto have h5: "invertible_mat (get_matrix_R (reduce_system_R p (qs, matrix, subsets, signs)))" using reduction_doesnt_break_things_invertibility_R assms same_size_R satisfies_properties_R_def by auto have h6: "get_matrix_R (reduce_system_R p (qs, matrix, subsets, signs)) = matrix_A_R (get_signs_R (reduce_system_R p (qs, matrix, subsets, signs))) (get_subsets_R (reduce_system_R p (qs, matrix, subsets, signs)))" using reduce_system_matrix_match_R[of p qs signs subsets] assms satisfies_properties_R_def by auto have h7: "set (characterize_consistent_signs_at_roots p qs) \ set (get_signs_R (reduce_system_R p (qs, matrix, subsets, signs)))" using reduction_doesnt_break_things_signs_R[of p qs signs subsets] assms reduction_signs_is_get_signs_R satisfies_properties_R_def by auto then show ?thesis unfolding satisfies_properties_R_def using h1 h2 h3 h4 h5 h6 h7 by blast qed subsubsection "For length 1 qs" lemma length_1_calculate_data_satisfies_properties_R: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "(nat list*nat list) list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes len1: "length qs = 1" shows "satisfies_properties_R p qs (get_subsets_R (calculate_data_R p qs)) (get_signs_R (calculate_data_R p qs)) (get_matrix_R (calculate_data_R p qs))" proof - have h1: "all_list_constr_R [([], []),([0], []),([], [0])] (length qs)" using len1 unfolding all_list_constr_R_def list_constr_def apply (auto) apply (smt (verit, best) Ball_set in_set_member member_rec(1) member_rec(2) prod.inject) by (smt (verit, ccfv_threshold) Ball_set in_set_member member_rec(1) member_rec(2) prod.inject) have h2: "well_def_signs (length qs) [[1],[-1]]" unfolding well_def_signs_def using len1 in_set_member by auto have h3: "distinct ([[1],[0],[-1]]::rat list list)" unfolding distinct_def using in_set_member by auto have h4: "satisfy_equation_R p qs [([], []),([0], []),([], [0])] [[1],[0],[-1]]" using assms base_case_satisfy_equation_alt_R[of qs p] by auto have h6: "(mat_of_rows_list 3 [[1,1,1], [0,1,0], [1,0,-1]]::rat mat) = (matrix_A_R ([[1],[0],[-1]]) ([([], []),([0], []),([], [0])]) :: rat mat)" using mat_base_case_R by auto then have h5: "invertible_mat (mat_of_rows_list 3 [[1,1,1], [0,1,0], [1,0,-1]]::rat mat)" using base_case_invertible_mat_R by simp have h7: "set (characterize_consistent_signs_at_roots p qs) \ set ([[1],[0],[-1]])" using assms base_case_sgas_alt_R[of qs p] by simp have base_case_hyp: "satisfies_properties_R p qs [([], []),([0], []),([], [0])] [[1],[0],[-1]] (mat_of_rows_list 3 [[1,1,1], [0,1,0], [1,0,-1]]::rat mat)" using h1 h2 h3 h4 h5 h6 h7 using satisfies_properties_R_def apply (auto) by (simp add: well_def_signs_def) then have key_hyp: "satisfies_properties_R p qs (get_subsets_R (reduce_system_R p (qs,base_case_info_R))) (get_signs_R (reduce_system_R p (qs,base_case_info_R))) (get_matrix_R (reduce_system_R p (qs,base_case_info_R)))" using reducing_sys_satisfies_properties_R by (metis base_case_info_R_def len1 nonzero nonzero zero_less_one_class.zero_less_one) show ?thesis by (simp add: key_hyp len1) qed subsubsection "For arbitrary qs" lemma calculate_data_satisfies_properties_R: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "(nat list*nat list) list" fixes signs :: "rat list list" shows "(p \ 0 \ (length qs > 0)) \ satisfies_properties_R p qs (get_subsets_R (calculate_data_R p qs)) (get_signs_R (calculate_data_R p qs)) (get_matrix_R (calculate_data_R p qs))" proof (induction "length qs" arbitrary: qs rule: less_induct) case less have len1_h: "length qs = 1 \ ( p \ 0 \ (length qs > 0)) \ satisfies_properties_R p qs (get_subsets_R (calculate_data_R p qs)) (get_signs_R (calculate_data_R p qs)) (get_matrix_R (calculate_data_R p qs))" using length_1_calculate_data_satisfies_properties_R by blast let ?len = "length qs" let ?q1 = "take (?len div 2) qs" let ?left = "calculate_data_R p ?q1" let ?q2 = "drop (?len div 2) qs" let ?right = "calculate_data_R p ?q2" let ?comb = "combine_systems_R p (?q1,?left) (?q2,?right)" let ?red = "reduce_system_R p ?comb" have h_q1_len: "length qs > 1 \ (length ?q1 > 0)" by auto have h_q2_len: "length qs > 1 \ (length ?q2 > 0)" by auto have q1_sat_props: "length qs > 1 \ (p \ 0 \ (length qs > 0)) \ satisfies_properties_R p ?q1 (get_subsets_R (calculate_data_R p ?q1)) (get_signs_R (calculate_data_R p ?q1)) (get_matrix_R (calculate_data_R p ?q1))" using less.hyps[of ?q1] h_q1_len by (metis div_le_dividend div_less_dividend length_take min.absorb2 one_less_numeral_iff semiring_norm(76)) have q2_help: "length qs > 1 \ length (drop (length qs div 2) qs) < length qs" using h_q1_len by auto then have q2_sat_props: "length qs > 1 \ (p \ 0 \ (length qs > 0)) \ satisfies_properties_R p ?q2 (get_subsets_R (calculate_data_R p ?q2)) (get_signs_R (calculate_data_R p ?q2)) (get_matrix_R (calculate_data_R p ?q2))" using less.hyps[of ?q2] h_q2_len by blast have put_tog: "?q1@?q2 = qs" using append_take_drop_id by blast then have comb_sat_props: "length qs > 1 \ (p \ 0 \ (length qs > 0)) \ (satisfies_properties_R p (qs) (get_subsets_R (snd ((combine_systems_R p (?q1,calculate_data_R p ?q1) (?q2,calculate_data_R p ?q2))))) (get_signs_R (snd ((combine_systems_R p (?q1,calculate_data_R p ?q1) (?q2,calculate_data_R p ?q2))))) (get_matrix_R (snd ((combine_systems_R p (?q1,calculate_data_R p ?q1) (?q2,calculate_data_R p ?q2))))))" using q1_sat_props q2_sat_props combining_sys_satisfies_properties_R using h_q1_len h_q2_len put_tog by metis then have comb_sat: "length qs > 1 \ (p \ 0 \ (length qs > 0)) \ (satisfies_properties_R p (qs) (get_subsets_R (snd ?comb)) (get_signs_R (snd ?comb)) (get_matrix_R (snd ?comb)))" by blast have red_char: "?red = (reduce_system_R p (qs,(get_matrix_R (snd ?comb)),(get_subsets_R (snd ?comb)),(get_signs_R (snd ?comb))))" using getter_functions by (smt (z3) combine_systems_R.simps find_consistent_signs_at_roots_R_def find_consistent_signs_at_roots_thm_R fst_conv get_matrix_R_def get_signs_R_def get_subsets_R_def prod.collapse put_tog smash_systems_R_def) then have "length qs > 1 \ (p \ 0 \ (length qs > 0)) \ (satisfies_properties_R p qs (get_subsets_R ?red) (get_signs_R ?red) (get_matrix_R ?red))" using reducing_sys_satisfies_properties_R comb_sat by presburger then have len_gt1: "length qs > 1 \ (p \ 0 \ (length qs > 0) ) \ satisfies_properties_R p qs (get_subsets_R (calculate_data_R p qs)) (get_signs_R (calculate_data_R p qs)) (get_matrix_R (calculate_data_R p qs))" apply (auto) by (smt (z3) div_le_dividend min.absorb2) then show ?case using len1_h len_gt1 by (metis One_nat_def Suc_lessI) qed subsection "Some key results on consistent sign assignments" lemma find_consistent_signs_at_roots_len1_R: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "(nat list*nat list) list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes len1: "length qs = 1" shows "set (find_consistent_signs_at_roots_R p qs) = set (characterize_consistent_signs_at_roots p qs)" proof - let ?signs = "[[1],[0],[-1]]::rat list list" let ?subsets = "[([], []),([0], []),([], [0])]::(nat list*nat list) list" let ?mat = "(mat_of_rows_list 3 [[1,1,1], [0,1,0], [1,0,-1]])" have mat_help: "matrix_A_R ?signs ?subsets = (mat_of_rows_list 3 [[1,1,1], [0,1,0], [1,0,-1]])" using mat_base_case_R by auto have well_def_signs: "well_def_signs (length qs) ?signs" unfolding well_def_signs_def using len1 by auto have distinct_signs: "distinct ?signs" unfolding distinct_def by auto have ex_q: "\(q::real poly). qs = [q]" using len1 using length_Suc_conv[of qs 0] by auto then have all_info: "set (characterize_consistent_signs_at_roots p qs) \ set(?signs)" using assms base_case_sgas_R by auto have match: "satisfy_equation_R p qs ?subsets ?signs" using ex_q base_case_satisfy_equation_R nonzero by auto have invertible_mat: "invertible_mat (matrix_A_R ?signs ?subsets)" using inverse_mat_base_case_R inverse_mat_base_case_2_R unfolding invertible_mat_def using mat_base_case_R by auto have h: "set (get_signs_R (reduce_system_R p (qs, ((matrix_A_R ?signs ?subsets), (?subsets, ?signs))))) = set (characterize_consistent_signs_at_roots p qs)" using nonzero nonzero well_def_signs distinct_signs all_info match invertible_mat reduce_system_sign_conditions_R[where p = "p", where qs = "qs", where signs = ?signs, where subsets = ?subsets] by blast then have "set (snd (snd (reduce_system_R p (qs, (?mat, (?subsets, ?signs)))))) = set (characterize_consistent_signs_at_roots p qs)" unfolding get_signs_R_def using mat_help by auto then have "set (snd (snd (reduce_system_R p (qs, base_case_info_R)))) = set (characterize_consistent_signs_at_roots p qs)" unfolding base_case_info_R_def by auto then show ?thesis using len1 by (simp add: find_consistent_signs_at_roots_thm_R) qed lemma smaller_sys_are_good_R: fixes p:: "real poly" fixes qs1 :: "real poly list" fixes qs2 :: "real poly list" fixes subsets :: "(nat list*nat list) list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes nontriv1: "length qs1 > 0" assumes nontriv2: "length qs2 > 0" assumes "set(find_consistent_signs_at_roots_R p qs1) = set(characterize_consistent_signs_at_roots p qs1)" assumes "set(find_consistent_signs_at_roots_R p qs2) = set(characterize_consistent_signs_at_roots p qs2)" shows "set(snd(snd(reduce_system_R p (combine_systems_R p (qs1,calculate_data_R p qs1) (qs2,calculate_data_R p qs2))))) = set(characterize_consistent_signs_at_roots p (qs1@qs2))" proof - let ?signs = "(get_signs_R (snd ((combine_systems_R p (qs1,calculate_data_R p qs1) (qs2,calculate_data_R p qs2))))) " let ?subsets = "(get_subsets_R (snd ((combine_systems_R p (qs1,calculate_data_R p qs1) (qs2,calculate_data_R p qs2))))) " have h0: "satisfies_properties_R p (qs1@qs2) ?subsets ?signs (get_matrix_R (snd ((combine_systems_R p (qs1,calculate_data_R p qs1) (qs2,calculate_data_R p qs2)))))" using calculate_data_satisfies_properties_R combining_sys_satisfies_properties_R using nontriv1 nontriv2 nonzero nonzero by simp then have h1: "set(characterize_consistent_signs_at_roots p (qs1@qs2)) \ set ?signs" unfolding satisfies_properties_R_def by linarith have h2: "well_def_signs (length (qs1@qs2)) ?signs" using calculate_data_satisfies_properties_R using h0 satisfies_properties_R_def by blast have h3: "distinct ?signs" using calculate_data_satisfies_properties_R using h0 satisfies_properties_R_def by blast have h4: "satisfy_equation_R p (qs1@qs2) ?subsets ?signs" using calculate_data_satisfies_properties_R using h0 satisfies_properties_R_def by blast have h5: "invertible_mat (matrix_A_R ?signs ?subsets) " using calculate_data_satisfies_properties_R using h0 satisfies_properties_R_def by auto have h: "set (take_indices ?signs (find_nonzeros_from_input_vec (solve_for_lhs_R p (qs1@qs2) ?subsets (matrix_A_R ?signs ?subsets)))) = set(characterize_consistent_signs_at_roots p (qs1@qs2))" using h1 h2 h3 h4 h5 reduction_deletes_bad_sign_conds_R using nonzero reduction_signs_R_def by auto then have h: "set (characterize_consistent_signs_at_roots p (qs1@qs2)) = set (reduction_signs_R p (qs1@qs2) ?signs ?subsets (matrix_A_R ?signs ?subsets ))" unfolding reduction_signs_R_def get_signs_R_def by blast have help_h: "reduction_signs_R p (qs1@qs2) ?signs ?subsets (matrix_A_R ?signs ?subsets) = (take_indices ?signs (find_nonzeros_from_input_vec (solve_for_lhs_R p (qs1@qs2) ?subsets (matrix_A_R?signs ?subsets))))" unfolding reduction_signs_R_def by auto have clear_signs: "(signs_smash (get_signs_R (calculate_data_R p qs1)) (get_signs_R (calculate_data_R p qs2))) = (get_signs_R (snd ((combine_systems_R p (qs1,calculate_data_R p qs1) (qs2,calculate_data_R p qs2)))))" using combining_to_smash get_signs_R_def getter_functions_R smash_systems_R_def snd_conv proof - have "combine_systems_R p (qs1, calculate_data_R p qs1) (qs2, calculate_data_R p qs2) = (qs1 @ qs2, kronecker_product (get_matrix_R (calculate_data_R p qs1)) (get_matrix_R (calculate_data_R p qs2)), subsets_smash_R (length qs1) (get_subsets_R (calculate_data_R p qs1)) (get_subsets_R (calculate_data_R p qs2)), signs_smash (snd (snd (calculate_data_R p qs1))) (snd (snd (calculate_data_R p qs2))))" by (metis (no_types) combine_systems_R.simps get_signs_R_def getter_functions_R smash_systems_R_def) then show ?thesis by (simp add: get_signs_R_def) qed have clear_subsets: "(subsets_smash_R (length qs1) (get_subsets_R(calculate_data_R p qs1)) (get_subsets_R (calculate_data_R p qs2))) = (get_subsets_R (snd ((combine_systems_R p (qs1,calculate_data_R p qs1) (qs2,calculate_data_R p qs2)))))" using Pair_inject combining_to_smash get_subsets_R_def prod.collapse smash_systems_R_def by (smt (z3) combine_systems_R.simps) have "well_def_signs (length qs1) (get_signs_R (calculate_data_R p qs1))" using calculate_data_satisfies_properties_R using nontriv1 nonzero nonzero satisfies_properties_R_def by auto then have well_def_signs1: "(\j. j \ set (get_signs_R (calculate_data_R p qs1)) \ length j = (length qs1))" using well_def_signs_def by blast have "all_list_constr_R (get_subsets_R(calculate_data_R p qs1)) (length qs1) " using calculate_data_satisfies_properties_R using nontriv1 nonzero nonzero satisfies_properties_R_def by auto then have well_def_subsets1: "(\l i. l \ set (get_subsets_R(calculate_data_R p qs1)) \ (i \ set (fst l) \ i < (length qs1)) \ (i \ set (snd l) \ i < (length qs1)))" unfolding all_list_constr_R_def list_constr_def using in_set_member by (metis in_set_conv_nth list_all_length) have extra_matrix_same: "matrix_A_R (signs_smash (get_signs_R (calculate_data_R p qs1)) (get_signs_R (calculate_data_R p qs2))) (subsets_smash_R (length qs1) (get_subsets_R(calculate_data_R p qs1)) (get_subsets_R (calculate_data_R p qs2))) = kronecker_product (get_matrix_R (calculate_data_R p qs1)) (get_matrix_R (calculate_data_R p qs2))" using well_def_signs1 well_def_subsets1 using matrix_construction_is_kronecker_product_R using calculate_data_satisfies_properties_R nontriv1 nontriv2 nonzero nonzero satisfies_properties_R_def by fastforce then have matrix_same: "matrix_A_R ?signs ?subsets = kronecker_product (get_matrix_R (calculate_data_R p qs1)) (get_matrix_R (calculate_data_R p qs2))" using clear_signs clear_subsets by simp have comb_sys_h: "snd(snd(reduce_system_R p (combine_systems_R p (qs1,calculate_data_R p qs1) (qs2,calculate_data_R p qs2)))) = snd(snd(reduce_system_R p (qs1@qs2, (matrix_A_R ?signs ?subsets, (?subsets, ?signs)))))" unfolding get_signs_R_def get_subsets_R_def using matrix_same by (metis (full_types) clear_signs clear_subsets combine_systems_R.simps get_signs_R_def get_subsets_R_def getter_functions_R smash_systems_R_def) then have extra_h: " snd(snd(reduce_system_R p (qs1@qs2, (matrix_A_R ?signs ?subsets, (?subsets, ?signs))))) = snd(snd(reduction_step_R (matrix_A_R ?signs ?subsets) ?signs ?subsets (solve_for_lhs_R p (qs1@qs2) ?subsets (matrix_A_R ?signs ?subsets)))) " by simp then have same_h: "set(snd(snd(reduce_system_R p (combine_systems_R p (qs1,calculate_data_R p qs1) (qs2,calculate_data_R p qs2))))) = set (reduction_signs_R p (qs1@qs2) ?signs ?subsets (matrix_A_R ?signs ?subsets ))" using comb_sys_h unfolding reduction_signs_R_def by (metis get_signs_R_def help_h reduction_signs_is_get_signs_R) then show ?thesis using h by blast qed lemma find_consistent_signs_at_roots_1_R: fixes p:: "real poly" fixes qs :: "real poly list" shows "(p \ 0 \ length qs > 0) \ set(find_consistent_signs_at_roots_R p qs) = set(characterize_consistent_signs_at_roots p qs)" proof (induction "length qs" arbitrary: qs rule: less_induct) case less then show ?case proof clarsimp assume ind_hyp: "(\qsa. length qsa < length qs \ qsa \ [] \ set (find_consistent_signs_at_roots_R p qsa) = set (characterize_consistent_signs_at_roots p qsa))" assume nonzero: "p \ 0" assume nontriv: "qs \ []" let ?len = "length qs" let ?q1 = "take ((?len) div 2) qs" let ?left = "calculate_data_R p ?q1" let ?q2 = "drop ((?len) div 2) qs" let ?right = "calculate_data_R p ?q2" have nontriv_q1: "length qs>1 \ length ?q1 > 0" by auto have qs_more_q1: "length qs>1 \ length qs > length ?q1" by auto have nontriv_q2: "length qs>1 \length ?q2 > 0" by auto have qs_more_q2: "length qs>1 \ length qs > length ?q2" by auto have key_h: "set (snd (snd (if ?len \ Suc 0 then reduce_system_R p (qs, base_case_info_R) else Let (combine_systems_R p (?q1, ?left) (?q2, ?right)) (reduce_system_R p)))) = set (characterize_consistent_signs_at_roots p qs)" proof - have h_len1 : "?len = 1 \ set (snd (snd (if ?len \ Suc 0 then reduce_system_R p (qs, base_case_info_R) else Let (combine_systems_R p (?q1, ?left) (?q2, ?right)) (reduce_system_R p)))) = set (characterize_consistent_signs_at_roots p qs)" using find_consistent_signs_at_roots_len1_R[of p qs] nonzero nontriv by (simp add: find_consistent_signs_at_roots_thm_R) have h_len_gt1 : "?len > 1 \ set (snd (snd (if ?len \ Suc 0 then reduce_system_R p (qs, base_case_info_R) else Let (combine_systems_R p (?q1, ?left) (?q2, ?right)) (reduce_system_R p)))) = set (characterize_consistent_signs_at_roots p qs)" proof - have h_imp_a: "?len > 1 \ set (snd (snd (reduce_system_R p (combine_systems_R p (?q1, ?left) (?q2, ?right))))) = set (characterize_consistent_signs_at_roots p qs)" proof - have h1: "?len > 1 \ set(snd(snd(?left))) = set (characterize_consistent_signs_at_roots p ?q1)" using nontriv_q1 ind_hyp[of ?q1] qs_more_q1 by (metis find_consistent_signs_at_roots_thm_R less_numeral_extra(3) list.size(3)) have h2: "?len > 1 \ set(snd(snd(?right))) = set (characterize_consistent_signs_at_roots p ?q2)" using nontriv_q2 ind_hyp[of ?q2] qs_more_q2 by (metis (full_types) find_consistent_signs_at_roots_thm_R list.size(3) not_less_zero) show ?thesis using nonzero nontriv_q1 nontriv_q2 h1 h2 smaller_sys_are_good_R by (metis append_take_drop_id find_consistent_signs_at_roots_thm_R) qed then have h_imp: "?len > 1 \ set (snd (snd (Let (combine_systems_R p (?q1, ?left) (?q2, ?right)) (reduce_system_R p)))) = set (characterize_consistent_signs_at_roots p qs)" by auto then show ?thesis by auto qed show ?thesis using h_len1 h_len_gt1 by (meson \qs \ []\ length_0_conv less_one nat_neq_iff) qed then show "set (find_consistent_signs_at_roots_R p qs) = set (characterize_consistent_signs_at_roots p qs)" using One_nat_def calculate_data.simps find_consistent_signs_at_roots_thm length_0_conv nontriv by (smt (z3) calculate_data_R.simps find_consistent_signs_at_roots_thm_R) qed qed lemma find_consistent_signs_at_roots_0_R: fixes p:: "real poly" assumes "p \ 0" shows "set(find_consistent_signs_at_roots_R p []) = set(characterize_consistent_signs_at_roots p [])" proof - obtain a b c where abc: "reduce_system_R p ([1], base_case_info_R) = (a,b,c)" using prod_cases3 by blast have "find_consistent_signs_at_roots_R p [1] = c" using abc by (simp add: find_consistent_signs_at_roots_thm_R) have *: "set (find_consistent_signs_at_roots_R p [1]) = set (characterize_consistent_signs_at_roots p [1])" apply (subst find_consistent_signs_at_roots_1_R) using assms by auto have "set(characterize_consistent_signs_at_roots p []) = drop 1 ` set(characterize_consistent_signs_at_roots p [1])" unfolding characterize_consistent_signs_at_roots_def consistent_sign_vec_def signs_at_def squash_def apply simp using drop0 drop_Suc_Cons image_cong image_image proof - have "(\r. []) ` set (characterize_root_list_p p) = (\r. drop (Suc 0) [1::rat]) ` set (characterize_root_list_p p)" by force then show "(\r. []) ` set (characterize_root_list_p p) = drop (Suc 0) ` (\r. [1::rat]) ` set (characterize_root_list_p p)" by blast qed thus ?thesis using abc * apply (auto) apply (simp add: find_consistent_signs_at_roots_thm_R) by (simp add: find_consistent_signs_at_roots_thm_R) qed lemma find_consistent_signs_at_roots_R: fixes p:: "real poly" fixes qs :: "real poly list" assumes "p \ 0" shows "set(find_consistent_signs_at_roots_R p qs) = set(characterize_consistent_signs_at_roots p qs)" by (metis assms find_consistent_signs_at_roots_0_R find_consistent_signs_at_roots_1_R length_greater_0_conv) end \ No newline at end of file diff --git a/thys/Differential_Dynamic_Logic/Lib.thy b/thys/Differential_Dynamic_Logic/Lib.thy --- a/thys/Differential_Dynamic_Logic/Lib.thy +++ b/thys/Differential_Dynamic_Logic/Lib.thy @@ -1,469 +1,469 @@ theory Lib imports Ordinary_Differential_Equations.ODE_Analysis begin section \Generic Mathematical Lemmas\ text\General lemmas that don't have anything to do with dL specifically and could be fit for general-purpose libraries, mostly dealing with derivatives, ODEs and vectors.\ lemma vec_extensionality:"(\i. v$i = w$i) \ (v = w)" by (simp add: vec_eq_iff) lemma norm_axis: "norm (axis i x) = norm x" unfolding axis_def norm_vec_def unfolding L2_set_def by(clarsimp simp add: if_distrib[where f=norm] if_distrib[where f="\x. x\<^sup>2"] sum.If_cases) lemma bounded_linear_axis: "bounded_linear (axis i)" proof show "axis i (x + y) = axis i x + axis i y" "axis i (r *\<^sub>R x) = r *\<^sub>R axis i x" for x y :: "'a" and r by (auto simp: vec_eq_iff axis_def) show "\K. \x::'a. norm (axis i x) \ norm x * K" by (auto simp add: norm_axis intro!: exI[of _ 1]) qed lemma bounded_linear_vec: fixes f::"('a::finite) \ 'b::real_normed_vector \ 'c::real_normed_vector" assumes bounds:"\i. bounded_linear (f i)" shows "bounded_linear (\x. \ i. f i x)" proof unfold_locales fix r x y interpret bounded_linear "f i" for i by fact show "(\ i. f i (x + y)) = (\ i. f i x) + (\ i. f i y)" by (vector add) show "(\ i. f i (r *\<^sub>R x)) = r *\<^sub>R (\ i. f i x)" by (vector scaleR) obtain K where "norm (f i x) \ norm x * K i" for x i using bounded by metis then have "norm (\ i. f i x) \ norm x * (\i\UNIV. K i)" (is "?lhs \ ?rhs") for x unfolding sum_distrib_left unfolding norm_vec_def by (auto intro!: L2_set_le_sum_abs[THEN order_trans] sum_mono simp: abs_mult) then show "\K. \x. norm (\ i. f i x) \ norm x * K" by blast qed lift_definition blinfun_vec::"('a::finite \ 'b::real_normed_vector \\<^sub>L real) \ 'b \\<^sub>L (real ^ 'a)" is "(\(f::('a \ 'b \ real)) (x::'b). \ (i::'a). f i x)" by(rule bounded_linear_vec, simp) lemmas blinfun_vec_simps[simp] = blinfun_vec.rep_eq lemma continuous_blinfun_vec:"(\i. continuous_on UNIV (blinfun_apply (g i))) \ continuous_on UNIV (blinfun_vec g)" by (simp add: continuous_on_vec_lambda) lemma blinfun_elim:"\g. (blinfun_apply (blinfun_vec g)) = (\x. \ i. g i x)" using blinfun_vec.rep_eq by auto lemma sup_plus: fixes f g::"('b::metric_space) \ real" assumes nonempty:"R \ {}" assumes bddf:"bdd_above (f ` R)" assumes bddg:"bdd_above (g ` R)" shows "(SUP x\R. f x + g x) \ (SUP x\R. f x) + (SUP x\R. g x)" proof - have bddfg:"bdd_above((\x. f x + g x ) ` R)" using bddf bddg apply (auto simp add: bdd_above_def) using add_mono_thms_linordered_semiring(1) by blast have eq:"(SUP x\R. f x + g x) \ (SUP x\R. f x) + (SUP x\R. g x) \ (\x\R. (f x + g x) \ (SUP x\R. f x) + (SUP x\R. g x))" apply(rule cSUP_le_iff) subgoal by (rule nonempty) subgoal by (rule bddfg) done have fs:"\x. x \ R \ f x \ (SUP x\R. f x)" using bddf by (simp add: cSUP_upper) have gs:"\x. x \ R \ g x \ (SUP x\R. g x)" using bddg by (simp add: cSUP_upper) have "(\x\R. (f x + g x) \ (SUP x\R. f x) + (SUP x\R. g x))" apply auto subgoal for x using fs[of x] gs[of x] by auto done then show ?thesis by (auto simp add: eq) qed lemma continuous_blinfun_vec': fixes f::"'a::{finite,linorder} \ 'b::{metric_space, real_normed_vector,abs} \ 'b \\<^sub>L real" fixes S::"'b set" assumes conts:"\i. continuous_on UNIV (f i)" shows "continuous_on UNIV (\x. blinfun_vec (\ i. f i x))" proof (auto simp add: LIM_def continuous_on_def) fix x1 and \::real assume \:"0 < \" let ?n = "card (UNIV::'a set)" have conts':" \i x1 \. 0 < \ \ \\>0. \x2. x2 \ x1 \ dist x2 x1 < \ \ dist (f i x2) (f i x1) < \" using conts by(auto simp add: LIM_def continuous_on_def) have conts'':"\i. \\>0. \x2. x2 \ x1 \ dist x2 x1 < \ \ dist (f i x2) (f i x1) < (\/?n)" subgoal for i using conts'[of "\ / ?n" x1 i] \ by auto done let ?f = "(\x. blinfun_vec (\ i. f i x))" let ?P\ = "(\ i \. (\>0 \ (\x2. x2 \ x1 \ dist x2 x1 < \ \ dist (f i x2) (f i x1) < (\/?n))))" let ?\i = "(\i. SOME \. ?P\ i \)" have Ps:"\i. \\. ?P\ i \" using conts'' by auto have P\i:"\i. ?P\ i (?\i i)" subgoal for i using someI[of "?P\ i" ] Ps[of i] by auto done have finU:"finite (UNIV::'a set)" by auto let ?\ = "linorder_class.Min (?\i ` UNIV)" have \0s:"\i. ?\i i > 0" using P\i by blast then have \0s':"\i. 0 < ?\i i" by auto have bounds:"bdd_below (?\i ` UNIV)" unfolding bdd_below_def using \0s less_eq_real_def by blast have \s:"\i. ?\ \ ?\i i" using bounds cINF_lower[of ?\i] by auto have finite:"finite ((?\i ` UNIV))" by auto have nonempty:"((?\i ` UNIV)) \ {}" by auto have \:"?\ > 0 " using Min_gr_iff[OF finite nonempty] \0s' by blast have conts''':"\i x2. x2 \ x1 \ dist x2 x1 < ?\i i \ dist (f i x2) (f i x1) < (\/?n)" subgoal for i x2 using conts''[of i] apply auto subgoal for \ apply(erule allE[where x=x2]) using P\i \s[of i] apply (auto simp add: \s[of i]) done done done have "\x2. x2 \ x1 \ dist x2 x1 < ?\ \ dist (blinfun_vec (\i. f i x2)) (blinfun_vec (\i. f i x1)) < \" proof (auto) fix x2 assume ne:"x2 \ x1" assume dist:"\i. dist x2 x1 < ?\i i" have dists:"\i. dist x2 x1 < ?\i i" subgoal for i using dist \s[of i] by auto done have euclid:"\y. norm(?f x1 y - ?f x2 y) = (L2_set (\i. norm(f i x1 y - f i x2 y)) UNIV)" by (simp add: norm_vec_def) have finite:"finite (UNIV::'a set)" by auto have nonempty: "(UNIV::'a set) \ {}" by auto have nonemptyB: "(UNIV::'b set) \ {}" by auto have nonemptyR: "(UNIV::real set) \ {}" by auto have SUP_leq:"\f::('b \ real). \ g::('b \ real). \ S::'b set. S \ {} \ bdd_above (g ` S) \ (\x. x \ (S::'b set) \ ((f x)::real) \ ((g x)::real)) \ (SUP x\S. f x) \ (SUP x\S. g x)" by(rule cSup_mono, auto) have SUP_sum_comm':"\R S f . finite (S::'a set) \ (R::'d::metric_space set) \ {} \ (\i x. ((f i x)::real) \ 0) \ (\i. bdd_above (f i ` R)) \ (SUP x\R . (\i \ S. f i x)) \ (\i \ S. (SUP x\R. f i x))" proof - fix R::"'d set" and S ::"('a)set" and f ::"'a \ 'd \ real" assume non:"R \ {} " assume fin:"finite S" assume every:"(\i x. 0 \ f i x)" assume bddF:"\i. bdd_above (f i ` R)" then have bddF':"\i. \M. \x \R. f i x \ M " unfolding bdd_above_def by auto let ?boundP = "(\i M. \x \R. f i x \ M)" let ?bound = "(\i::'a. SOME M. \x \R. f i x \ M)" have "\i. \M. ?boundP i M" using bddF' by auto then have each_bound:"\i. ?boundP i (?bound i)" subgoal for i using someI[of "?boundP i"] by blast done let ?bigBound = "(\F. \i\F. (?bound i))" have bddG:"\i::'a. \F. bdd_above ((\x. \i\F. f i x) ` R)" subgoal for i F using bddF[of i] unfolding bdd_above_def apply auto apply(rule exI[where x="?bigBound F"]) subgoal for M apply auto subgoal for x using each_bound by (simp add: sum_mono) done done done show "?thesis R S f" using fin assms proof (induct) case empty have "((SUP x\R. \i\{}. f i x)::real) \ (\i\{}. SUP a\R. f i a)" by (simp add: non) then show ?case by auto next case (insert x F) have "((SUP xa\R. \i\insert x F. f i xa)::real) \ (SUP xa\R. f x xa + (\i\F. f i xa))" using insert.hyps(2) by auto moreover have "... \ (SUP xa\ R. f x xa) + (SUP xa\R. (\i\F. f i xa))" by(rule sup_plus, rule non, rule bddF, rule bddG) moreover have "... \ (SUP xa\ R. f x xa) + (\i\F. SUP a\R. f i a)" using add_le_cancel_left conts insert.hyps(3) by blast moreover have "... \ (\i\(insert x F). SUP a\R. f i a)" by (simp add: insert.hyps(2)) ultimately have "((SUP xa\R. \i\insert x F. f i xa)::real) \ (\i\(insert x F). SUP a\R. f i a)" by linarith then show ?case by auto qed qed have SUP_sum_comm:"\R S y1 y2 . finite (S::'a set) \ (R::'b set) \ {} \ (SUP x\R . (\i \ S. norm(f i y1 x - f i y2 x)/norm(x))) \ (\i \ S. (SUP x\R. norm(f i y1 x - f i y2 x)/norm(x)))" apply(rule SUP_sum_comm') apply(auto)[3] proof (unfold bdd_above_def) fix R S y1 y2 i { fix rr :: "real \ real" obtain bb :: "real \ ('b \ real) \ 'b set \ 'b" where ff1: "\r f B. r \ f ` B \ f (bb r f B) = r" - by moura + unfolding image_iff by moura { assume "\r. \ rr r \ norm (f i y1 - f i y2)" then have "\r. norm (blinfun_apply (f i y1) (bb (rr r) (\b. norm (blinfun_apply (f i y1) b - blinfun_apply (f i y2) b) / norm b) R) - blinfun_apply (f i y2) (bb (rr r) (\b. norm (blinfun_apply (f i y1) b - blinfun_apply (f i y2) b) / norm b) R)) / norm (bb (rr r) (\b. norm (blinfun_apply (f i y1) b - blinfun_apply (f i y2) b) / norm b) R) \ rr r" by (metis (no_types) le_norm_blinfun minus_blinfun.rep_eq) then have "\r. rr r \ r \ rr r \ (\b. norm (blinfun_apply (f i y1) b - blinfun_apply (f i y2) b) / norm b) ` R" using ff1 by meson } then have "\r. rr r \ r \ rr r \ (\b. norm (blinfun_apply (f i y1) b - blinfun_apply (f i y2) b) / norm b) ` R" by blast } then show "\r. \ra\(\b. norm (blinfun_apply (f i y1) b - blinfun_apply (f i y2) b) / norm b) ` R. ra \ r" by meson qed have SUM_leq:"\S::('a) set. \ f g ::('a \ real). S \ {} \ finite S \ (\x. x \ S \ f x < g x) \ (\x\S. f x) < (\x\S. g x)" by(rule sum_strict_mono, auto) have L2:"\f S. L2_set (\x. norm(f x)) S \ (\x \ S. norm(f x))" using L2_set_le_sum norm_ge_zero by metis have L2':"\y. (L2_set (\i. norm(f i x1 y - f i x2 y)) UNIV)/norm(y) \ (\i\UNIV. norm(f i x1 y - f i x2 y))/norm(y)" subgoal for y using L2[of "(\ x. f x x1 y - f x x2 y)" UNIV] by (auto simp add: divide_right_mono) done have "\i. (SUP y\UNIV. norm((f i x1 - f i x2) y)/norm(y)) = norm(f i x1 - f i x2)" by (simp add: onorm_def norm_blinfun.rep_eq) then have each_norm:"\i. (SUP y\UNIV. norm(f i x1 y - f i x2 y)/norm(y)) = norm(f i x1 - f i x2)" by (metis (no_types, lifting) SUP_cong blinfun.diff_left) have bounded_linear:"\i. bounded_linear (\y. f i x1 y - f i x2 y)" by (simp add: blinfun.bounded_linear_right bounded_linear_sub) have each_bound:"\i. bdd_above ((\y. norm(f i x1 y - f i x2 y)/norm(y)) ` UNIV)" using bounded_linear unfolding bdd_above_def proof - fix i :: 'a { fix rr :: "real \ real" have "\a r. norm (blinfun_apply (f a x1) r - blinfun_apply (f a x2) r) / norm r \ norm (f a x1 - f a x2)" by (metis le_norm_blinfun minus_blinfun.rep_eq) then have "\r R. r \ (\r. norm (blinfun_apply (f i x1) r - blinfun_apply (f i x2) r) / norm r) ` R \ r \ norm (f i x1 - f i x2)" by blast then have "\r. rr r \ r \ rr r \ range (\r. norm (blinfun_apply (f i x1) r - blinfun_apply (f i x2) r) / norm r)" by blast } then show "\r. \ra\range (\r. norm (blinfun_apply (f i x1) r - blinfun_apply (f i x2) r) / norm r). ra \ r" by meson qed have bdd_above:"(bdd_above ((\y. (\i\UNIV. norm(f i x1 y - f i x2 y)/norm(y))) ` UNIV))" using each_bound unfolding bdd_above_def apply auto proof - assume each:"(\i. \M. \x. \blinfun_apply (f i x1) x - blinfun_apply (f i x2) x\ / norm x \ M)" let ?boundP = "(\i M. \x. \blinfun_apply (f i x1) x - blinfun_apply (f i x2) x\ / norm x \ M)" let ?bound = "(\i. SOME x. ?boundP i x)" have bounds:"\i. ?boundP i (?bound i)" subgoal for i using each someI[of "?boundP i"] by blast done let ?bigBound = "\i\(UNIV::'a set). ?bound i" show "\M. \x. (\i\UNIV. \blinfun_apply (f i x1) x - blinfun_apply (f i x2) x\ / norm x) \ M" apply(rule exI[where x= ?bigBound]) by(auto simp add: bounds sum_mono) qed have bdd_above:"(bdd_above ((\y. (\i\UNIV. norm(f i x1 y - f i x2 y))/norm(y)) ` UNIV))" using bdd_above unfolding bdd_above_def apply auto proof - fix M :: real assume a1: "\x. (\i\UNIV. \blinfun_apply (f i x1) x - blinfun_apply (f i x2) x\ / norm x) \ M" { fix bb :: "real \ 'b" have "\b. (\a\UNIV. \blinfun_apply (f a x1) b - blinfun_apply (f a x2) b\) / norm b \ M" using a1 by (simp add: sum_divide_distrib) then have "\r. (\a\UNIV. \blinfun_apply (f a x1) (bb r) - blinfun_apply (f a x2) (bb r)\) / norm (bb r) \ r" by blast } then show "\r. \b. (\a\UNIV. \blinfun_apply (f a x1) b - blinfun_apply (f a x2) b\) / norm b \ r" by meson qed have "dist (?f x2) (?f x1) = norm((?f x2) - (?f x1))" by (simp add: dist_blinfun_def) moreover have "... = (SUP y\UNIV. norm(?f x1 y - ?f x2 y)/norm(y))" by (metis (no_types, lifting) SUP_cong blinfun.diff_left norm_blinfun.rep_eq norm_minus_commute onorm_def) moreover have "... = (SUP y\UNIV. (L2_set (\i. norm(f i x1 y - f i x2 y)) UNIV)/norm(y))" using euclid by auto moreover have "... \ (SUP y\UNIV. (\i\UNIV. norm(f i x1 y - f i x2 y))/norm(y))" using L2' SUP_cong SUP_leq bdd_above by auto moreover have "... = (SUP y\UNIV. (\i\UNIV. norm(f i x1 y - f i x2 y)/norm(y)))" by (simp add: sum_divide_distrib) moreover have "... \ (\i\UNIV. (SUP y\UNIV. norm(f i x1 y - f i x2 y)/norm(y)))" by(rule SUP_sum_comm[OF finite nonemptyB, of x1 x2]) moreover have "... = (\i\UNIV. norm(f i x1 - f i x2))" using each_norm by simp moreover have "... = (\i\UNIV. dist(f i x1) (f i x2))" by (simp add: dist_blinfun_def) moreover have "... < (\i\(UNIV::'a set). \ / ?n)" using conts'''[OF ne dists] using SUM_leq[OF nonempty, of "(\i. dist (f i x1) (f i x2))" "(\i. \ / ?n)"] by (simp add: dist_commute) moreover have "... = \" by(auto) ultimately show "dist (?f x2) (?f x1) < \" by linarith qed then show "\s>0. \x2. x2 \ x1 \ dist x2 x1 < s \ dist (blinfun_vec (\i. f i x2)) (blinfun_vec (\i. f i x1)) < \" using \ by blast qed lemma has_derivative_vec[derivative_intros]: assumes "\i. ((\x. f i x) has_derivative (\h. f' i h)) F" shows "((\x. \ i. f i x) has_derivative (\h. \ i. f' i h)) F" proof - have *: "(\ i. f i x) = (\i\UNIV. axis i (f i x))" "(\ i. f' i x) = (\i\UNIV. axis i (f' i x))" for x by (simp_all add: axis_def sum.If_cases vec_eq_iff) show ?thesis unfolding * by (intro has_derivative_sum bounded_linear.has_derivative[OF bounded_linear_axis] assms) qed lemma has_derivative_proj: fixes j::"('a::finite)" fixes f::"'a \ real \ real" assumes assm:"((\x. \ i. f i x) has_derivative (\h. \ i. f' i h)) F" shows "((\x. f j x) has_derivative (\h. f' j h)) F" proof - have bounded_proj:"bounded_linear (\ x::(real^'a). x $ j)" by (simp add: bounded_linear_vec_nth) show "?thesis" using bounded_linear.has_derivative[OF bounded_proj, of "(\x. \ i. f i x)" "(\h. \ i. f' i h)", OF assm] by auto qed lemma has_derivative_proj': fixes i::"'a::finite" shows "\x. ((\ x. x $ i) has_derivative (\x::(real^'a). x $ i)) (at x)" proof - have bounded_proj:"bounded_linear (\ x::(real^'a). x $ i)" by (simp add: bounded_linear_vec_nth) show "?thesis" using bounded_proj unfolding has_derivative_def by auto qed lemma constant_when_zero: fixes v::"real \ (real, 'i::finite) vec" assumes x0: "(v t0) $ i = x0" assumes sol: "(v solves_ode f) T S" assumes f0: "\s x. s \ T \ f s x $ i = 0" assumes t0:"t0 \ T" assumes t:"t \ T" assumes convex:"convex T" shows "v t $ i = x0" proof - from solves_odeD[OF sol] have deriv: "(v has_vderiv_on (\t. f t (v t))) T" by simp then have "((\t. v t $ i) has_vderiv_on (\t. 0)) T" using f0 by (auto simp: has_vderiv_on_def has_vector_derivative_def cart_eq_inner_axis intro!: derivative_eq_intros) from has_vderiv_on_zero_constant[OF convex this] obtain c where c:"\x. x \ T \ v x $ i = c" by blast with x0 have "c = x0" "v t $ i = c" using t t0 c x0 by blast+ then show ?thesis by simp qed lemma solves_ode_subset: assumes x: "(x solves_ode f) T X" assumes s: "S \ T" shows "(x solves_ode f) S X" apply(rule solves_odeI) using has_vderiv_on_subset s solves_ode_vderivD x apply force using assms by (auto intro!: solves_odeI dest!: solves_ode_domainD) lemma solves_ode_supset_range: assumes x: "(x solves_ode f) T X" assumes y: "X \ Y" shows "(x solves_ode f) T Y" apply(rule solves_odeI) using has_vderiv_on_subset y solves_ode_vderivD x apply force using assms by (auto intro!: solves_odeI dest!: solves_ode_domainD) lemma usolves_ode_subset: assumes x: "(x usolves_ode f from t0) T X" assumes s: "S \ T" assumes t0: "t0 \ S" assumes S: "is_interval S" shows "(x usolves_ode f from t0) S X" proof (rule usolves_odeI) note usolves_odeD[OF x] show "(x solves_ode f) S X" by (rule solves_ode_subset; fact) show "t0 \ S" "is_interval S" by(fact+) fix z t assume s: "{t0 -- t} \ S" and z: "(z solves_ode f) {t0 -- t} X" and z0: "z t0 = x t0" then have "t0 \ {t0 -- t}" "is_interval {t0 -- t}" by auto moreover note s moreover have "(z solves_ode f) {t0--t} X" using solves_odeD[OF z] \S \ T\ by (intro solves_ode_subset_range[OF z]) force moreover note z0 moreover have "t \ {t0 -- t}" by simp ultimately show "z t = x t" by (meson \\z ta T'. \t0 \ T'; is_interval T'; T' \ T; (z solves_ode f) T' X; z t0 = x t0; ta \ T'\ \ z ta = x ta\ assms(2) dual_order.trans) qed \ \Example of using lemmas above to show a lemma that could be useful for dL: The constant ODE\ \ \0 does not change the state.\ lemma example: fixes x t::real and i::"('sz::finite)" assumes "t > 0" shows "x = (ll_on_open.flow UNIV (\t. \x. \ (i::('sz::finite)). 0) UNIV 0 (\ i. x) t) $ i" proof - let ?T = UNIV let ?f = "(\t. \x. \ i::('sz::finite). 0)" let ?X = UNIV let ?t0.0 = 0 let ?x0.0 = "\ i::('sz::finite). x" interpret ll: ll_on_open "UNIV" "(\t x. \ i::('sz::finite). 0)" UNIV using gt_ex by unfold_locales (auto simp: interval_def continuous_on_def local_lipschitz_def intro!: lipschitz_intros) have foo1:"?t0.0 \ ?T" by auto have foo2:"?x0.0 \ ?X" by auto let ?v = "ll.flow ?t0.0 ?x0.0" from ll.flow_solves_ode[OF foo1 foo2] have solves:"(ll.flow ?t0.0 ?x0.0 solves_ode ?f) (ll.existence_ivl ?t0.0 ?x0.0) ?X" by (auto) then have solves:"(?v solves_ode ?f) (ll.existence_ivl ?t0.0 ?x0.0) ?X" by auto have thex0: "(?v ?t0.0) $ (i::('sz::finite)) = x" by auto have sol_help: "(?v solves_ode ?f) (ll.existence_ivl ?t0.0 ?x0.0) ?X" using solves by auto have ivl:"ll.existence_ivl ?t0.0 ?x0.0 = UNIV" by (rule ll.existence_ivl_eq_domain) (auto intro!: exI[where x=0] simp: vec_eq_iff) have sol: "(?v solves_ode ?f) UNIV ?X" using solves ivl by auto have thef0: "\t x. ?f t x $ i = 0" by auto from constant_when_zero [OF thex0 sol thef0] have "?v t $ i = x" by auto thus ?thesis by auto qed lemma MVT_ivl: fixes f::"'a::ordered_euclidean_space\'b::ordered_euclidean_space" assumes fderiv: "\x. x \ D \ (f has_derivative J x) (at x within D)" assumes J_ivl: "\x. x \ D \ J x u \ J0" assumes line_in: "\x. x \ {0..1} \ a + x *\<^sub>R u \ D" shows "f (a + u) - f a \ J0" proof - from MVT_corrected[OF fderiv line_in] obtain t where t: "t\Basis \ {0<..<1}" and mvt: "f (a + u) - f a = (\i\Basis. (J (a + t i *\<^sub>R u) u \ i) *\<^sub>R i)" by auto note mvt also have "\ \ J0" proof - have J: "\i. i \ Basis \ J0 \ J (a + t i *\<^sub>R u) u" using J_ivl t line_in by (auto simp: Pi_iff) show ?thesis using J unfolding atLeastAtMost_iff eucl_le[where 'a='b] by auto qed finally show ?thesis . qed lemma MVT_ivl': fixes f::"'a::ordered_euclidean_space\'b::ordered_euclidean_space" assumes fderiv: "(\x. x \ D \ (f has_derivative J x) (at x within D))" assumes J_ivl: "\x. x \ D \ J x (a - b) \ J0" assumes line_in: "\x. x \ {0..1} \ b + x *\<^sub>R (a - b) \ D" shows "f a \ f b + J0" proof - have "f (b + (a - b)) - f b \ J0" apply (rule MVT_ivl[OF fderiv ]) apply assumption apply (rule J_ivl) apply assumption using line_in apply (auto simp: diff_le_eq le_diff_eq ac_simps) done thus ?thesis by (auto simp: diff_le_eq le_diff_eq ac_simps) qed end diff --git a/thys/Differential_Game_Logic/Coincidence.thy b/thys/Differential_Game_Logic/Coincidence.thy --- a/thys/Differential_Game_Logic/Coincidence.thy +++ b/thys/Differential_Game_Logic/Coincidence.thy @@ -1,787 +1,786 @@ theory "Coincidence" imports "Lib" "Syntax" "Denotational_Semantics" "Static_Semantics" "HOL.Finite_Set" begin section \Static Semantics Properties\ subsection Auxiliaries text \The state interpolating \stateinterpol \ \ S\ between \\\ and \\\ that is \\\ on \S\ and \\\ elsewhere\ definition stateinterpol:: "state \ state \ variable set \ state" where "stateinterpol \ \ S = (\x. if (x\S) then \(x) else \(x))" definition statediff:: "state \ state \ variable set" where "statediff \ \ = {x. \(x)\\(x)}" lemma nostatediff: "x\statediff \ \ \ \(x)=\(x)" by (simp add: statediff_def) lemma stateinterpol_empty: "stateinterpol \ \ {} = \" proof fix x have empty: "\x. \(x\{})" by auto show "\x. stateinterpol \ \ {} x = \ x" using empty by (simp add: stateinterpol_def) qed lemma stateinterpol_left [simp]: "x\S \ (stateinterpol \ \ S)(x)=\(x)" by (simp add: stateinterpol_def) lemma stateinterpol_right [simp]: "x\S \ (stateinterpol \ \ S)(x)=\(x)" by (simp add: stateinterpol_def) lemma Vagree_stateinterpol [simp]: "Vagree (stateinterpol \ \ S) \ S" and "Vagree (stateinterpol \ \ S) \ (-S)" unfolding Vagree_def by auto lemma Vagree_ror: "Vagree \ \' (V\W) \ (\\. (Vagree \ \ V \ Vagree \ \' W))" proof - assume "Vagree \ \' (V\W)" hence "\x. x\V\W \ \(x)=\'(x)" by (simp add: Vagree_def) let ?w="stateinterpol \ \' V" have l: "Vagree \ ?w V" by (simp add: Vagree_def) have r: "Vagree ?w \' W \ Vagree ?w \' W" by (simp add: Vagree_def stateinterpol_def \\x. x\V\W \ \ x = \' x\) have "Vagree \ ?w V \ Vagree ?w \' W" using l and r by blast thus ?thesis by auto qed text \Remark 8 \<^url>\https://doi.org/10.1007/978-3-319-94205-6_15\ about simple properties of projections\ lemma restrictto_extends [simp]: "restrictto X V \ X" by (auto simp add: restrictto_def) lemma restrictto_compose [simp]: "restrictto (restrictto X V) W = restrictto X (V\W)" proof show "restrictto (restrictto X V) W \ restrictto X (V\W)" by (auto simp add: restrictto_def Vagree_def) next show "restrictto X (V\W) \ restrictto (restrictto X V) W" (*by (smt Vagree_ror mem_Collect_eq restrictto_def subsetI)*) proof - (* sledgehammer *) obtain rr :: "(variable \ real) set \ (variable \ real) set \ variable \ real" where "\x0 x1. (\v2. v2 \ x1 \ v2 \ x0) = (rr x0 x1 \ x1 \ rr x0 x1 \ x0)" - by moura + by moura then have f1: "\F Fa. rr Fa F \ F \ rr Fa F \ Fa \ F \ Fa" by (meson subsetI) obtain rra :: "(variable \ real) \ variable set \ (variable \ real) set \ variable \ real" where "\x0 x1 x2. (\v3. v3 \ x2 \ Vagree v3 x0 x1) = (rra x0 x1 x2 \ x2 \ Vagree (rra x0 x1 x2) x0 x1)" by moura then have f2: "\F V f. (f \ {f. \fa. fa \ F \ Vagree fa f V} \ rra f V F \ F \ Vagree (rra f V F) f V) \ (f \ {f. \fa. fa \ F \ Vagree fa f V} \ (\fa. fa \ F \ \ Vagree fa f V))" by blast moreover { assume "\f. f \ X \ Vagree f (v4_1 W V (rr {f. \fa. fa \ {f. \fa. fa \ X \ Vagree fa f V} \ Vagree fa f W} {f. \fa. fa \ X \ Vagree fa f (V \ W)}) (rra (rr {f. \fa. fa \ {f. \fa. fa \ X \ Vagree fa f V} \ Vagree fa f W} {f. \fa. fa \ X \ Vagree fa f (V \ W)}) (V \ W) X)) V" moreover { assume "\f. f \ {f. \fa. fa \ X \ Vagree fa f V} \ Vagree f (rr {f. \fa. fa \ {f. \fa. fa \ X \ Vagree fa f V} \ Vagree fa f W} {f. \fa. fa \ X \ Vagree fa f (V \ W)}) W" then have "rr {f. \fa. fa \ {f. \fa. fa \ X \ Vagree fa f V} \ Vagree fa f W} {f. \fa. fa \ X \ Vagree fa f (V \ W)} \ {f. \fa. fa \ X \ Vagree fa f (V \ W)} \ rr {f. \fa. fa \ {f. \fa. fa \ X \ Vagree fa f V} \ Vagree fa f W} {f. \fa. fa \ X \ Vagree fa f (V \ W)} \ {f. \fa. fa \ {f. \fa. fa \ X \ Vagree fa f V} \ Vagree fa f W}" by blast then have "{f. \fa. fa \ X \ Vagree fa f (V \ W)} \ {f. \fa. fa \ {f. \fa. fa \ X \ Vagree fa f V} \ Vagree fa f W}" using f1 by meson } ultimately have "(\ Vagree (rra (rr {f. \fa. fa \ {f. \fa. fa \ X \ Vagree fa f V} \ Vagree fa f W} {f. \fa. fa \ X \ Vagree fa f (V \ W)}) (V \ W) X) (v4_1 W V (rr {f. \fa. fa \ {f. \fa. fa \ X \ Vagree fa f V} \ Vagree fa f W} {f. \fa. fa \ X \ Vagree fa f (V \ W)}) (rra (rr {f. \fa. fa \ {f. \fa. fa \ X \ Vagree fa f V} \ Vagree fa f W} {f. \fa. fa \ X \ Vagree fa f (V \ W)}) (V \ W) X)) V \ \ Vagree (v4_1 W V (rr {f. \fa. fa \ {f. \fa. fa \ X \ Vagree fa f V} \ Vagree fa f W} {f. \fa. fa \ X \ Vagree fa f (V \ W)}) (rra (rr {f. \fa. fa \ {f. \fa. fa \ X \ Vagree fa f V} \ Vagree fa f W} {f. \fa. fa \ X \ Vagree fa f (V \ W)}) (V \ W) X)) (rr {f. \fa. fa \ {f. \fa. fa \ X \ Vagree fa f V} \ Vagree fa f W} {f. \fa. fa \ X \ Vagree fa f (V \ W)}) W) \ {f. \fa. fa \ X \ Vagree fa f (V \ W)} \ {f. \fa. fa \ {f. \fa. fa \ X \ Vagree fa f V} \ Vagree fa f W}" using f2 by meson } ultimately have "{f. \fa. fa \ X \ Vagree fa f (V \ W)} \ {f. \fa. fa \ {f. \fa. fa \ X \ Vagree fa f V} \ Vagree fa f W}" using f1 by (meson Vagree_ror) then show ?thesis using restrictto_def by presburger qed qed lemma restrictto_antimon [simp]: "W\V \ restrictto X W \ restrictto X V" proof - assume "W\V" then have "\U. V=W\U" by auto then obtain U where "V=W\U" by auto hence "restrictto X V = restrictto (restrictto X W) U" by simp hence "restrictto X V \ restrictto X W" using restrictto_extends by blast thus ?thesis by auto qed lemma restrictto_empty [simp]: "X\{} \ restrictto X {} = worlds" by (auto simp add: restrictto_def worlds_def) lemma selectlike_shrinks [simp]: "selectlike X \ V \ X" by (auto simp add: selectlike_def) lemma selectlike_compose [simp]: "selectlike (selectlike X \ V) \ W = selectlike X \ (V\W)" by (auto simp add: selectlike_def) lemma selectlike_antimon [simp]: "W\V \ selectlike X \ W \ selectlike X \ V" by (auto simp add: selectlike_def) lemma selectlike_empty [simp]: "selectlike X \ {} = X" by (auto simp add: selectlike_def) lemma selectlike_self [simp]: "(\ \ selectlike X \ V) = (\\X)" by (auto simp add: selectlike_def) lemma selectlike_complement [simp]: "selectlike (-X) \ V \ -selectlike X \ V" by (auto simp add: selectlike_def) lemma selectlike_union: "selectlike (X\Y) \ V = selectlike X \ V \ selectlike Y \ V" by (auto simp add: selectlike_def) lemma selectlike_Sup: "selectlike (Sup M) \ V = Sup {selectlike X \ V | X. X\M}" using selectlike_def by auto lemma selectlike_equal_cond: "(selectlike X \ V = selectlike Y \ V) = (\\. Uvariation \ \ (-V) \ (\\X) = (\\Y))" unfolding selectlike_def using Uvariation_Vagree by auto lemma selectlike_equal_cocond: "(selectlike X \ (-V) = selectlike Y \ (-V)) = (\\. Uvariation \ \ V \ (\\X) = (\\Y))" using selectlike_equal_cond[where V=\-V\] by simp lemma selectlike_equal_cocond_rule: "(\\. Uvariation \ \ (-V) \ (\\X) = (\\Y)) \ (selectlike X \ V = selectlike Y \ V)" using selectlike_equal_cond[where V=\V\] by simp lemma selectlike_equal_cocond_corule: "(\\. Uvariation \ \ V \ (\\X) = (\\Y)) \ (selectlike X \ (-V) = selectlike Y \ (-V))" using selectlike_equal_cond[where V=\-V\] by simp lemma co_selectlike: "-(selectlike X \ V) = (-X) \ {\. \Vagree \ \ V}" unfolding selectlike_def by auto lemma selectlike_co_selectlike: "selectlike (-(selectlike X \ V)) \ V = selectlike (-X) \ V" unfolding selectlike_def by auto lemma selectlike_Vagree: "Vagree \ \ V \ selectlike X \ V = selectlike X \ V" using Vagree_def selectlike_def by auto lemma similar_selectlike_mem: "Vagree \ \ V \ (\\selectlike X \ V) = (\\X)" unfolding selectlike_def using Vagree_sym_rel by blast (* also see nonBVG_rule *) lemma BVG_nonelem [simp] :"(x\BVG \) = (\I \ X. (\ \ game_sem I \ X) = (\ \ game_sem I \ (selectlike X \ {x})))" using BVG_elem monotone selectlike_shrinks by (metis subset_iff) text \\statediff\ interoperability\ lemma Vagree_statediff [simp]: "Vagree \ \' S \ statediff \ \' \ -S" by (auto simp add: Vagree_def statediff_def) lemma stateinterpol_diff [simp]: "stateinterpol \ \ (statediff \ \) = \" proof fix x show sp: "(stateinterpol \ \ (statediff \ \))(x) = \(x)" proof (cases "x\statediff \ \") case True then show ?thesis by simp next case False then show ?thesis by (simp add: stateinterpol_def nostatediff) qed qed lemma stateinterpol_insert: "Vagree (stateinterpol v w S) (stateinterpol v w (insert z S)) (-{z})" by (simp add: Vagree_def stateinterpol_def) lemma stateinterpol_FVT [simp]: "z\FVT(t) \ term_sem I t (stateinterpol \ \' S) = term_sem I t (stateinterpol \ \' (insert z S))" proof - assume a: "z\FVT(t)" have fvc: "\v. \w. Vagree v w (-{z}) \ (term_sem I t v = term_sem I t w)" using a by (simp add: FVT_def) then show "term_sem I t (stateinterpol \ \' S) = term_sem I t (stateinterpol \ \' (insert z S))" using fvc and stateinterpol_insert by blast qed subsection \Coincidence Lemmas\ paragraph \Coincidence for Terms\ text \Lemma 10 \<^url>\https://doi.org/10.1007/978-3-319-94205-6_15\\ theorem coincidence_term: "Vagree \ \' (FVT \) \ term_sem I \ \ = term_sem I \ \'" proof - assume a: "Vagree \ \' (FVT \)" have isS: "statediff \ \' \ -FVT(\)" using a and Vagree_statediff by simp have gen: "S\-FVT(\) \ (term_sem I \ \' = term_sem I \ (stateinterpol \ \' S))" if "finite S" for S using that proof (induction S) case empty show ?case by (simp add: stateinterpol_empty) next case (insert z S) thus ?case by auto qed from isS have finS: "finite (statediff \ \')" using allvars_finite by (metis FVT_finite UNIV_def finite_compl rev_finite_subset) show ?thesis using gen[where S=\statediff \ \'\, OF finS, OF isS] by simp qed corollary coincidence_term_cor: "Uvariation \ \' U \ (FVT \)\U={} \ term_sem I \ \ = term_sem I \ \'" using coincidence_term Uvariation_Vagree by (metis Vagree_antimon disjoint_eq_subset_Compl double_compl) lemma stateinterpol_FVF [simp]: "z\FVF(e) \ ((stateinterpol \ \' S) \ fml_sem I e \ (stateinterpol \ \' (insert z S)) \ fml_sem I e)" proof - assume a: "z\FVF(e)" have agr: "Vagree (stateinterpol \ \' S) (stateinterpol \ \' (insert z S)) (-{z})" by (simp add: Vagree_def stateinterpol_def) have fvc: "\v. \w. (Vagree v w (-{z}) \ (v\fml_sem I e \ w\fml_sem I e))" using a by (simp add: FVF_def) then have fvce: "\v. \w. (Vagree v w (-{z}) \ ((v\fml_sem I e) = (w\fml_sem I e)))" using Vagree_sym_rel by blast then show "(stateinterpol \ \' S) \ fml_sem I e \ (stateinterpol \ \' (insert z S)) \ fml_sem I e" using agr by simp qed paragraph \Coincidence for Formulas\ text \Lemma 11 \<^url>\https://doi.org/10.1007/978-3-319-94205-6_15\\ theorem coincidence_formula: "Vagree \ \' (FVF \) \ (\ \ fml_sem I \ \ \' \ fml_sem I \)" proof - assume a: "Vagree \ \' (FVF \)" have isS: "statediff \ \' \ -FVF(\)" using a and Vagree_statediff by simp have gen: "S\-FVF(\) \ (\' \ fml_sem I \ \ (stateinterpol \ \' S) \ fml_sem I \)" if "finite S" for S using that proof (induction S) case empty show ?case by (simp add: stateinterpol_empty) next case (insert z S) thus ?case by auto qed from isS have finS: "finite (statediff \ \')" using allvars_finite by (metis FVF_finite UNIV_def finite_compl rev_finite_subset) show ?thesis using gen[where S=\statediff \ \'\, OF finS, OF isS] by simp qed corollary coincidence_formula_cor: "Uvariation \ \' U \ (FVF \)\U={} \ (\ \ fml_sem I \ \ \' \ fml_sem I \)" using coincidence_formula Uvariation_Vagree by (metis Uvariation_def disjoint_eq_subset_Compl inf.commute subsetCE) paragraph \Coincidence for Games\ text \\Cignorabimus \ V\ is the set of all sets of variables that can be ignored for the coincidence game lemma\ definition Cignorabimus:: "game \ variable set \ variable set set" where "Cignorabimus \ V = {M. \I.\\.\\'.\X. (Vagree \ \' (-M) \ (\\game_sem I \ (restrictto X V)) \ (\'\game_sem I \ (restrictto X V)))}" lemma Cignorabimus_finite [simp]: "finite (Cignorabimus \ V)" unfolding Cignorabimus_def using finite_powerset[OF allvars_finite] finite_subset using Finite_Set.finite_subset by fastforce lemma Cignorabimus_equiv [simp]: "Cignorabimus \ V = {M. \I.\\.\\'.\X. (Vagree \ \' (-M) \ (\\game_sem I \ (restrictto X V)) = (\'\game_sem I \ (restrictto X V)))}" unfolding Cignorabimus_def by (metis (no_types, lifting) Vagree_sym_rel) lemma Cignorabimus_antimon [simp]: "M \ Cignorabimus \ V \ N\M \ N \ Cignorabimus \ V" unfolding Cignorabimus_def using Vagree_antimon by blast lemma coempty: "-{}=allvars" by simp lemma Cignorabimus_empty [simp]: "{} \ Cignorabimus \ V" unfolding Cignorabimus_def using coempty Vagree_univ by simp text \Cignorabimus contains nonfree variables\ lemma Cignorabimus_init: "V\FVG(\) \ x\V \ {x}\Cignorabimus \ V" proof- assume "V\FVG(\)" assume a0: "x\V" hence a1: "x\FVG(\)" using \FVG \ \ V\ by blast hence "\I v w. Vagree v w (-{x}) \ (v \ game_sem I \ (restrictto X (-{x})) \ w \ game_sem I \ (restrictto X (-{x})))" by (metis (mono_tags, lifting) CollectI FVG_def Vagree_sym_rel) show "{x}\Cignorabimus \ V" proof- { fix I \ \' X have "Vagree \ \' (-{x}) \ (\\game_sem I \ (restrictto X V)) \ (\'\game_sem I \ (restrictto X V))" proof assume a2: "Vagree \ \' (-{x})" show "(\\game_sem I \ (restrictto X V)) \ (\'\game_sem I \ (restrictto X V))" proof assume "\\game_sem I \ (restrictto X V)" hence "\\game_sem I \ (restrictto (restrictto X V) (-{x}))" by (simp add: Int_absorb2 \x\V\) hence "\'\game_sem I \ (restrictto (restrictto X V) (-{x}))" using FVG_def a1 a2 by blast hence "\'\game_sem I \ (restrictto X (V\-{x}))" by simp show "\'\game_sem I \ (restrictto X V)" using a0 by (metis Int_absorb2 \\' \ game_sem I \ (restrictto X (V \ - {x}))\ subset_Compl_singleton) qed qed } thus ?thesis unfolding Cignorabimus_def by auto qed qed text \Cignorabimus is closed under union\ lemma Cignorabimus_union: "M\Cignorabimus \ V \ N\Cignorabimus \ V \ (M\N)\Cignorabimus \ V" proof- assume a1: "M\Cignorabimus \ V" assume a2: "N\Cignorabimus \ V" show "(M\N)\Cignorabimus \ V" (*using a1 a2 unfolding Cignorabimus_def *) proof- { fix I \ \' X assume a3: "Vagree \ \' (-(M\N))" have h1: "\I \ \'.\X. (Vagree \ \' (-M) \ (\\game_sem I \ (restrictto X V)) \ (\'\game_sem I \ (restrictto X V)))" using a1 by simp have h2: "\I \ \'.\X. (Vagree \ \' (-N) \ (\\game_sem I \ (restrictto X V)) \ (\'\game_sem I \ (restrictto X V)))" using a2 by simp let ?s = "stateinterpol \' \ M" have v1: "Vagree \ ?s (-(M\N))" using a3 by (simp add: Vagree_def) have v2: "Vagree ?s \' (-(M\N))" using a3 by (simp add: Vagree_def) have r1: "\\game_sem I \ (restrictto X V) \ ?s\game_sem I \ (restrictto X V)" by (metis ComplD Vagree_def h1 stateinterpol_right) have r2: "?s\game_sem I \ (restrictto X V) \ \'\game_sem I \ (restrictto X V)" by (metis Vagree_ror compl_sup h1 h2 v2) have res: "\\game_sem I \ (restrictto X V) \ \'\game_sem I \ (restrictto X V)" using r1 r2 by blast } thus ?thesis unfolding Cignorabimus_def by auto qed qed lemma powersetup_induct [case_names Base Cup]: "\C. (\M. M\C \ P M) \ (\S. (\M. M\S \ P M) \ P (\S)) \ P (\C)" by simp lemma Union_insert: "\(insert x S) = x\\S" by simp lemma powerset2up_induct [case_names Finite Nonempty Base Cup]: "(finite C) \ (C\{}) \ (\M. M\C \ P M) \ (\M N. P M \ P N \ P (M\N)) \ P (\C)" proof (induction rule: finite_induct) case empty then show ?case by simp next case (insert x F) then show ?case by force qed lemma Cignorabimus_step: "(\M. M\S \ M\Cignorabimus \ V) \ (\S)\Cignorabimus \ V" proof (cases "S={}") case True then show ?thesis using Cignorabimus_empty by simp next case nonem: False then show "\S\Cignorabimus \ V" if "\M. M\S \ M\Cignorabimus \ V" and nonemp:"S\{}" for S proof (induction rule: powerset2up_induct) case Finite then show ?case using Cignorabimus_finite by (meson infinite_super subset_eq that(1)) next case Nonempty then show ?case using nonemp by simp next case (Base M) then show ?case using that by simp next case (Cup S) then show ?case using that Cignorabimus_union by blast qed qed text \Lemma 12 \<^url>\https://doi.org/10.1007/978-3-319-94205-6_15\\ theorem coincidence_game: "Vagree \ \' V \ V\FVG(\) \ (\ \ game_sem I \ (restrictto X V)) = (\' \ game_sem I \ (restrictto X V))" proof- assume a1: "Vagree \ \' V" assume a2: "V \ FVG \" have base: "{x}\Cignorabimus \ V" if a3: "x\V" and a4: "V \ FVG \" for x V using a3 and a4 and Cignorabimus_init by simp have h: "-V = \{xx. \x. xx={x} \ x\V}" by auto (* finite *) have "(-V)\Cignorabimus \ V" using a2 base h using Cignorabimus_step - (*by (smt Cignorabimus_step mem_Collect_eq)*) proof - (*sledgehammer*) have f1: "\v V. v \ V \ \ FVG \ \ V \ {v} \ Cignorabimus \ V" using base by satx obtain VV :: "variable set \ game \ variable set set \ variable set" where f2: "\x0 x1 x2. (\v3. v3 \ x2 \ v3 \ Cignorabimus x1 x0) = (VV x0 x1 x2 \ x2 \ VV x0 x1 x2 \ Cignorabimus x1 x0)" by moura obtain vv :: "variable set \ variable" where f3: "((\v. VV V \ {{v} |v. v \ V} = {v} \ v \ V) \ VV V \ {{v} |v. v \ V} = {vv (VV V \ {{v} |v. v \ V})} \ vv (VV V \ {{v} |v. v \ V}) \ V) \ ((\v. VV V \ {{v} |v. v \ V} = {v} \ v \ V) \ (\v. VV V \ {{v} |v. v \ V} \ {v} \ v \ V))" - by moura + by fastforce moreover { assume "{vv (VV V \ {{v} |v. v \ V})} \ Cignorabimus \ V" then have "(VV V \ {{v} |v. v \ V} \ {vv (VV V \ {{v} |v. v \ V})} \ vv (VV V \ {{v} |v. v \ V}) \ V) \ VV V \ {{v} |v. v \ V} \ {{v} |v. v \ V} \ VV V \ {{v} |v. v \ V} \ Cignorabimus \ V" by metis then have "(\v. VV V \ {{v} |v. v \ V} = {v} \ v \ V) \ VV V \ {{v} |v. v \ V} \ {{v} |v. v \ V} \ VV V \ {{v} |v. v \ V} \ Cignorabimus \ V" using f3 by blast } ultimately have "VV V \ {{v} |v. v \ V} \ {{v} |v. v \ V} \ VV V \ {{v} |v. v \ V} \ Cignorabimus \ V" using f1 a2 by blast then have "\{{v} |v. v \ V} \ Cignorabimus \ V" using f2 by (meson Cignorabimus_step) then show ?thesis using h by presburger qed from this show ?thesis by (simp add: a1) qed corollary coincidence_game_cor: "Uvariation \ \' U \ U\FVG(\)={} \ (\ \ game_sem I \ (restrictto X (-U))) = (\' \ game_sem I \ (restrictto X (-U)))" using coincidence_game Uvariation_Vagree by (metis Uvariation_Vagree coincidence_game compl_le_swap1 disjoint_eq_subset_Compl double_compl) subsection \Bound Effect Lemmas\ text \\Bignorabimus \ V\ is the set of all sets of variables that can be ignored for boundeffect\ definition Bignorabimus:: "game \ variable set set" where "Bignorabimus \ = {M. \I.\\.\X. \\game_sem I \ X \ \\game_sem I \ (selectlike X \ M)}" lemma Bignorabimus_finite [simp]: "finite (Bignorabimus \)" unfolding Bignorabimus_def using finite_powerset[OF allvars_finite] finite_subset using Finite_Set.finite_subset by fastforce lemma Bignorabimus_single [simp]: "game_sem I \ (selectlike X \ M) \ game_sem I \ X" by (meson monotone selectlike_shrinks subsetCE) lemma Bignorabimus_equiv [simp]: "Bignorabimus \ = {M. \I.\\.\X. (\\game_sem I \ X \ \\game_sem I \ (selectlike X \ M))}" (*by (smt Bignorabimus_def Bignorabimus_single Collect_cong subsetCE)*) proof - (*sledgehammer transformed*) obtain VV :: "(variable set \ bool) \ (variable set \ bool) \ variable set" where f1: "\p pa. (\ p (VV pa p)) = pa (VV pa p) \ Collect p = Collect pa" by (metis (no_types) Collect_cong) obtain rr :: "variable set \ game \ variable \ real" and ii :: "variable set \ game \ interp" and FF :: "variable set \ game \ (variable \ real) set" where f2: "\x0 x1. (\v2 v3 v4. (v3 \ game_sem v2 x1 v4) \ (v3 \ game_sem v2 x1 (selectlike v4 v3 x0))) = ((rr x0 x1 \ game_sem (ii x0 x1) x1 (FF x0 x1)) \ (rr x0 x1 \ game_sem (ii x0 x1) x1 (selectlike (FF x0 x1) (rr x0 x1) x0)))" by moura have fact: "{V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))} = {V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V)}" using f1 by (metis (no_types, opaque_lifting) Bignorabimus_single subsetCE) have f3: "rr (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V))) \ \ game_sem (ii (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V))) \) \ (selectlike (FF (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V))) \) (rr (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V))) \) (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V)))) \ rr (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V))) \ \ game_sem (ii (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V))) \) \ (FF (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V))) \)" using Bignorabimus_single by blast have f4: "(\ (\i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V)))))) \ (\i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V))))))) \ ((\i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V)))))) \ (rr (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V))) \ \ game_sem (ii (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V))) \) \ (FF (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V))) \)) = (rr (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V))) \ \ game_sem (ii (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V))) \) \ (selectlike (FF (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V))) \) (rr (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V))) \) (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V))))))" using f2 by blast have "rr (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V))) \ \ game_sem (ii (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V))) \) \ (FF (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V))) \) \ (\i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V))))) \ rr (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V))) \ \ game_sem (ii (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V))) \) \ (selectlike (FF (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V))) \) (rr (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V))) \) (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V))))" by blast moreover { assume "\i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V))))" then have "\ (\i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V)))))" by blast moreover { assume "(\ (\i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V)))))) \ (\i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V))))))" then have "{V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))} = {V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V)}" using fact by simp } ultimately have "(\i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V)))))) \ {V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))} = {V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V)}" by satx } moreover { assume "\i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V)))))" then have "(\ (\i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V)))))) \ (\i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f (VV (\V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))) (\V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V))))))" by blast then have "{V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))} = {V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V)}" using fact by simp} ultimately have "{V. \i f F. (f \ game_sem i \ F) = (f \ game_sem i \ (selectlike F f V))} = {V. \i f F. f \ game_sem i \ F \ f \ game_sem i \ (selectlike F f V)}" using f4 f3 by satx then show ?thesis using Bignorabimus_def by presburger qed lemma Bignorabimus_empty [simp]: "{} \ Bignorabimus \" unfolding Bignorabimus_def using coempty selectlike_empty by simp lemma Bignorabimus_init: "x\BVG(\) \ {x}\Bignorabimus \" unfolding Bignorabimus_def BVG_def proof- assume "x \ {x. \I \ X. \ \ game_sem I \ X \ \ \ game_sem I \ (selectlike X \ {x})}" hence "\(\I \ X. \ \ game_sem I \ X \ \ \ game_sem I \ (selectlike X \ {x}))" by simp hence "\I \ X. (\ \ game_sem I \ X) = (\ \ game_sem I \ (selectlike X \ {x}))" using Bignorabimus_single by blast thus "{x} \ {M. \I \ X. (\ \ game_sem I \ X) = (\ \ game_sem I \ (selectlike X \ M))}" by simp qed text \Bignorabimus is closed under union\ lemma Bignorabimus_union: "M\Bignorabimus \ \ N\Bignorabimus \ \ (M\N)\Bignorabimus \" proof - assume a1: "M\Bignorabimus \" assume a2: "N\Bignorabimus \" have h1: "\I.\\.\X. (\\game_sem I \ X) \ (\\game_sem I \ (selectlike X \ M))" using a1 using Bignorabimus_equiv Bignorabimus_single by blast have h2: "\I.\\.\X. (\\game_sem I \ X) \ (\\game_sem I \ (selectlike X \ N))" using a2 using Bignorabimus_equiv Bignorabimus_single by blast have c: "\I.\\.\X. (\\game_sem I \ X) \ (\\game_sem I \ (selectlike X \ (M\N)))" by (metis h1 h2 selectlike_compose) then show "(M\N)\Bignorabimus \" unfolding Bignorabimus_def using c by fastforce qed lemma Bignorabimus_step: "(\M. M\S \ M\Bignorabimus \) \ (\S)\Bignorabimus \" proof (cases "S={}") case True then show ?thesis using Bignorabimus_empty by simp next case nonem: False then show "\S\Bignorabimus \" if "\M. M\S \ M\Bignorabimus \" and nonemp:"S\{}" for S proof (induction rule: powerset2up_induct) case Finite then show ?case using Bignorabimus_finite by (meson infinite_super subset_eq that(1)) next case Nonempty then show ?case using nonemp by simp next case (Base M) then show ?case using that by simp next case (Cup S) then show ?case using that Bignorabimus_union by blast qed qed text \Lemma 13 \<^url>\https://doi.org/10.1007/978-3-319-94205-6_15\\ theorem boundeffect: "(\ \ game_sem I \ X) = (\ \ game_sem I \ (selectlike X \ (-BVG(\))))" proof- have base: "{x}\Bignorabimus \" if a3: "x\BVG \" for x using a3 and Bignorabimus_init by simp have h: "-BVG \ = \{xx. \x. xx={x} \ x\BVG \}" by blast (* finite *) have "(-BVG \)\Bignorabimus \" using base h (*by (smt Bignorabimus_step mem_Collect_eq)*) proof - (*sledgehammer*) obtain VV :: "game \ variable set set \ variable set" where f1: "\x0 x1. (\v2. v2 \ x1 \ v2 \ Bignorabimus x0) = (VV x0 x1 \ x1 \ VV x0 x1 \ Bignorabimus x0)" by moura have "VV \ {{v} |v. v \ BVG \} \ {{v} |v. v \ BVG \} \ VV \ {{v} |v. v \ BVG \} \ Bignorabimus \" by fastforce then have "\{{v} |v. v \ BVG \} \ Bignorabimus \" using f1 by (meson Bignorabimus_step) then show ?thesis using h by presburger qed from this show ?thesis using Bignorabimus_def by blast qed corollary boundeffect_cor: "V\BVG(\)={} \ (\ \ game_sem I \ X) = (\ \ game_sem I \ (selectlike X \ V))" using boundeffect by (metis disjoint_eq_subset_Compl selectlike_compose sup.absorb_iff2) subsection \Static Analysis Observations\ lemma BVG_equiv: "game_equiv \ \ \ BVG(\) = BVG(\)" proof- assume a: "game_equiv \ \" show "BVG(\) = BVG(\)" unfolding BVG_def using game_equiv_subst_eq[OF a] by metis qed lemmas union_or = Set.Un_iff lemma not_union_or: "(x\A\B) = (x\A \ x\B)" by simp lemma repv_selectlike_self: "(repv \ x d \ selectlike X \ {x}) = (d=\(x) \ \ \ X)" unfolding selectlike_def using Vagree_repv_self Vagree_sym by (metis (no_types, lifting) mem_Collect_eq repv_self) lemma repv_selectlike_other: "x\y \ (repv \ x d \ selectlike X \ {y}) = (repv \ x d \ X)" proof assume a: "x\y" then have h: "{y}\-{x}" by simp show "(repv \ x d \ selectlike X \ {y}) \ (repv \ x d \ X)" using a selectlike_def Vagree_repv[of \ x d] by auto show "(repv \ x d \ X) \ (repv \ x d \ selectlike X \ {y})" using selectlike_def[where X=X and \=\ and V=\-{x}\] Vagree_repv[where \=\ and x=x and d=d] selectlike_antimon[where X=X and \=\ and V=\{y}\ and W=\-{x}\, OF h] Vagree_sym[where \=\repv \ x d\ and V=\-{x}\] by auto qed lemma repv_selectlike_other_converse: "x\y \ (repv \ x d \ X) = (repv \ x d \ selectlike X \ {y})" using repv_selectlike_other HOL.eq_commute by blast lemma BVG_assign_other: "x\y \ y\BVG(Assign x \)" using repv_selectlike_other_converse[where x=x and y=y] by simp lemma BVG_assign_meta: "(\I \. term_sem I \ \ = \(x)) \ BVG(Assign x \) = {}" and "term_sem I \ \ \ \(x) \ BVG(Assign x \) = {x}" (*using repv_selectlike_self BVG_assign_other BVG_def BVG_elem*) proof- have fact: "BVG(Assign x \) \ {x}" using BVG_assign_other by (metis singleton_iff subsetI) from fact show "(\I \. term_sem I \ \ = \(x)) \ BVG(Assign x \) = {}" using BVG_def by simp have h2: "\I \. term_sem I \ \ \ \(x) \ x \ BVG(Assign x \)" using repv_selectlike_self by auto from fact show "term_sem I \ \ \ \(x) \ BVG(Assign x \) = {x}" using BVG_elem h2 by blast qed lemma BVG_assign: "BVG(Assign x \) = (if (\I \. term_sem I \ \ = \(x)) then {} else {x})" using repv_selectlike_self repv_selectlike_other BVG_assign_other proof- have c0: "BVG(Assign x \) \ {x}" using BVG_assign_other by (metis singletonI subsetI) have c1: "\I \. term_sem I \ \ = \(x) \ BVG(Assign x \) = {}" using BVG_assign_other by auto have h2: "\I \. term_sem I \ \ \ \(x) \ x \ BVG(Assign x \)" using repv_selectlike_self by auto have c2: "\I \. term_sem I \ \ \ \(x) \ BVG(Assign x \) = {x}" using c0 h2 by blast from c1 and c2 show ?thesis by simp qed lemma BVG_ODE_other: "y\RVar x \ y\DVar x \ y\BVG(ODE x \)" (*using nonBVG_rule selectlike_equal_cocond_rule solves_ODE_def*) proof- assume yx: "y\RVar x" assume yxp: "y\DVar x" show "y\BVG(ODE x \)" proof (rule nonBVG_inc_rule) fix I \ X assume "\ \ game_sem I (ODE x \) X" then have "\F T. Vagree \ (F(0)) (-{DVar x}) \ F(T) \ X \ solves_ODE I F x \" by simp then obtain F T where "Vagree \ (F(0)) (-{DVar x}) \ F(T) \ X \ solves_ODE I F x \" by blast then have "Vagree \ (F(0)) (-{DVar x}) \ F(T) \ (selectlike X \ {y}) \ solves_ODE I F x \" using yx yxp solves_Vagree Vagree_def similar_selectlike_mem by auto then have "\F T. Vagree \ (F(0)) (-{DVar x}) \ F(T) \ (selectlike X \ {y}) \ solves_ODE I F x \" by blast then show "\ \ game_sem I (ODE x \) (selectlike X \ {y})" by simp qed qed text \This result could be strengthened to a conditional equality based on the RHS values\ lemma BVG_ODE: "BVG(ODE x \) \ {RVar x,DVar x}" using BVG_ODE_other by blast lemma BVG_test: "BVG(Test \) = {}" unfolding BVG_def game_sem.simps by auto lemma BVG_choice: "BVG(Choice \ \) \ BVG(\) \ BVG(\)" unfolding BVG_def game_sem.simps using not_union_or by auto (*proof- have f1: "\\ I X. (\ \ game_sem I \ X \ game_sem I \ X) = (\ \ game_sem I \ X \ \ \ game_sem I \ X)" by (rule union_or) have f2: "\\ I X. (\ \ game_sem I \ (selectlike X \ {x}) \ game_sem I \ (selectlike X \ {x})) = (\ \ game_sem I \ (selectlike X \ {x}) \ \ \ game_sem I \ (selectlike X \ {x}))" by (rule not_union_or) let ?lhs = "{x. \I \ X. \ \ game_sem I \ X \ game_sem I \ X \ \ \ game_sem I \ (selectlike X \ {x}) \ game_sem I \ (selectlike X \ {x})}" let ?rhs = "{x. \I \ X. \ \ game_sem I \ X \ \ \ game_sem I \ (selectlike X \ {x})} \ {x. \I \ X. \ \ game_sem I \ X \ \ \ game_sem I \ (selectlike X \ {x})}" show "?lhs\?rhs" using f1 f2 by auto qed*) lemma select_nonBV: "x\BVG(\) \ selectlike (game_sem I \ (selectlike X \ {x})) \ {x} = selectlike (game_sem I \ X) \ {x}" proof show "selectlike (game_sem I \ (selectlike X \ {x})) \ {x} \ selectlike (game_sem I \ X) \ {x}" using game_sem_mono selectlike_shrinks selectlike_antimon Bignorabimus_single by (metis selectlike_union sup.absorb_iff1) next assume nonbound: "x\BVG(\)" then have fact: "{x}\BVG(\)={}" by auto show "selectlike (game_sem I \ X) \ {x} \ selectlike (game_sem I \ (selectlike X \ {x})) \ {x}" proof fix \ assume "\ \ selectlike (game_sem I \ X) \ {x}" (* "V\BVG(\)={} \ (\ \ game_sem I \ X) = (\ \ game_sem I \ (selectlike X \ V))" *) then have "\ \ selectlike (game_sem I \ (selectlike X \ {x})) \ {x}" using boundeffect_cor[where \=\ and V=\{x}\ and \=\, OF fact] nonbound by (metis ComplD ComplI co_selectlike not_union_or) then show "\ \ selectlike (game_sem I \ (selectlike X \ {x})) \ {x}" using selectlike_Vagree selectlike_def by fastforce qed qed lemma BVG_compose: "BVG(Compose \ \) \ BVG(\) \ BVG(\)" (*unfolding BVG_def game_sem.simps using game_union union_or not_union_or selectlike_shrinks monotone selectlike_compose*) proof fix x assume a: "x\BVG(Compose \ \)" show "x \ BVG \ \ BVG \" proof (rule ccontr) assume "x\BVG \ \ BVG(\)" then have n\: "x\BVG(\)" and n\: "x\BVG(\)" by simp_all from a have "\I.\\.\X. \ \ game_sem I (Compose \ \) X \ \ \ game_sem I (Compose \ \) (selectlike X \ {x})" by simp then obtain I \ X where adef: "\ \ game_sem I (Compose \ \) X \ \ \ game_sem I (Compose \ \) (selectlike X \ {x})" by blast from adef have a1: "\ \ game_sem I \ (game_sem I \ X)" by simp from adef have a2: "\ \ game_sem I \ (game_sem I \ (selectlike X \ {x}))" by simp let ?Y = "selectlike X \ {x}" from n\ have n\c: "\I \ X. (\ \ game_sem I \ X) = (\ \ game_sem I \ (selectlike X \ {x}))" using BVG_nonelem by simp from n\ have n\c: "\I \ X. (\ \ game_sem I \ X) = (\ \ game_sem I \ (selectlike X \ {x}))" using BVG_nonelem by simp have c1: "\ \ game_sem I \ (selectlike (game_sem I \ X) \ {x})" using a1 n\c[where I=I and \=\ and X=\game_sem I \ X\] by blast have c2: "\ \ game_sem I \ (selectlike (game_sem I \ ?Y) \ {x})" using a2 n\c[where I=I and \=\ and X=\game_sem I \ ?Y\] by blast from c2 have c3: "\ \ game_sem I \ (selectlike (game_sem I \ X) \ {x})" using n\ selectlike_Vagree proof- have "selectlike (game_sem I \ ?Y) \ {x} = selectlike (game_sem I \ X) \ {x}" using n\ by (rule select_nonBV) thus ?thesis using c2 by simp qed show "False" using c1 c3 n\c[where I=I] by auto qed qed text \The converse inclusion does not hold generally, because \BVG(x := x+1; x:= x-1) = {} \ BVG(x := x+1) \ BVG(x := x-1) = {x}\\ lemma "BVG(Compose (Assign x (Plus (Var x) (Number 1))) (Assign x (Plus (Var x) (Number (-1))))) \ BVG(Assign x (Plus (Var x) (Number 1))) \ BVG(Assign x (Plus (Var x) (Number (-1))))" unfolding BVG_def selectlike_def repv_def Vagree_def by auto lemma BVG_loop: "BVG(Loop \) \ BVG(\)" proof fix x assume a: "x\BVG(Loop \)" show "x \ BVG(\)" proof (rule ccontr) assume "\ (x \ BVG(\))" then have n\: "x\BVG \" by simp from n\ have n\c: "\I \ X. (\ \ game_sem I \ X) = (\ \ game_sem I \ (selectlike X \ {x}))" using BVG_nonelem by simp have "x\BVG(Loop \)" (*using BVG_nonelem*) proof (rule nonBVG_rule) fix I \ X let ?f = "\Z. X \ game_sem I \ Z" let ?g = "\Y. (selectlike X \ {x}) \ game_sem I \ Y" let ?R = "\Z Y. selectlike Z \ {x} = selectlike Y \ {x}" have "?R (lfp ?f) (lfp ?g)" proof (induction rule: lfp_lockstep_induct[where f=\?f\ and g=\?g\ and R=\?R\]) case monof then show ?case using game_sem_loop_fixpoint_mono by simp next case monog then show ?case using game_sem_loop_fixpoint_mono by simp next case (step A B) then have IH: "selectlike A \ {x} = selectlike B \ {x}" by simp then show ?case (* using selectlike_union n\ select_nonBV IH by (smt insert_absorb2 insert_def selectlike_compose singleton_conv smt_solver=cvc4) *) proof- have "selectlike (X \ game_sem I \ A) \ {x} = selectlike X \ {x} \ selectlike (game_sem I \ A) \ {x}" using selectlike_union by simp also have "... = selectlike X \ {x} \ selectlike (game_sem I \ (selectlike A \ {x})) \ {x}" using n\ select_nonBV by blast also have "... = selectlike X \ {x} \ selectlike (game_sem I \ (selectlike B \ {x})) \ {x}" using IH by simp also have "... = selectlike (selectlike X \ {x} \ game_sem I \ B) \ {x}" using selectlike_union n\ select_nonBV by auto finally show "selectlike (X \ game_sem I \ A) \ {x} = selectlike (selectlike X \ {x} \ game_sem I \ B) \ {x}" . qed next case (union M) then have IH: "\(A,B)\M. selectlike A \ {x} = selectlike B \ {x}" . then show ?case using fst_proj_mem[where M=M] snd_proj_mem[where M=M] selectlike_Sup[where \=\ and V=\{x}\] sup_corr_eq_chain by simp (*proof- have "selectlike (\fst_proj M) \ {x} = \{selectlike A \ {x} | A. A\fst_proj M}" using selectlike_Sup by simp also have "... = \{selectlike B \ {x} | B. B\snd_proj M}" using sup_corr_eq_chain[OF IH] by simp also have "... = selectlike (\snd_proj M) \ {x}" using selectlike_Sup by simp finally show "selectlike (\fst_proj M) \ {x} = selectlike (\snd_proj M) \ {x}" . qed*) qed from this show "(\ \ game_sem I (Loop \) X) = (\ \ game_sem I (Loop \) (selectlike X \ {x}))" by (metis (mono_tags) game_sem.simps(6) lfp_def selectlike_self) (*by (simp add: selectlike_self game_sem_loop_back)*) qed then show "False" using a by blast qed qed lemma BVG_dual: "BVG(Dual \) \ BVG(\)" (*unfolding game_sem.simps using BVG_elem selectlike_co_selectlike co_selectlike selectlike_complement selectlike_antimon*) proof fix x assume a: "x\BVG(Dual \)" show "x\BVG \" proof- from a have "\I.\\.\X. \ \ game_sem I (Dual \) X \ \ \ game_sem I (Dual \) (selectlike X \ {x})" by simp then obtain I \ X where adef: "\ \ game_sem I (Dual \) X \ \ \ game_sem I (Dual \) (selectlike X \ {x})" by blast from adef have a1: "\ \ game_sem I \ (- X)" by simp from adef have a2: "\ \ game_sem I \ (- selectlike X \ {x})" by simp let ?Y = "- selectlike X \ {x}" have f1: "\ \ game_sem I \ ?Y" by (rule a2) have f2: "\ \ game_sem I \ (selectlike ?Y \ {x})" using a1 selectlike_co_selectlike by (metis (no_types, lifting) selectlike_shrinks monotone dual_order.trans subset_Compl_singleton) show "x\BVG(\)" using f1 f2 by auto qed qed end diff --git a/thys/Differential_Game_Logic/USubst.thy b/thys/Differential_Game_Logic/USubst.thy --- a/thys/Differential_Game_Logic/USubst.thy +++ b/thys/Differential_Game_Logic/USubst.thy @@ -1,1751 +1,1751 @@ theory "USubst" imports Complex_Main "Syntax" "Static_Semantics" "Coincidence" "Denotational_Semantics" begin section \Uniform Substitution\ text \uniform substitution representation as tuple of partial maps from identifiers to type-compatible replacements.\ type_synonym usubst = "(ident \ trm) \ (ident \ trm) \ (ident \ fml) \ (ident \ game)" abbreviation SConst:: "usubst \ (ident \ trm)" where "SConst \ (\(F0, _, _, _). F0)" abbreviation SFuncs:: "usubst \ (ident \ trm)" where "SFuncs \ (\(_, F, _, _). F)" abbreviation SPreds:: "usubst \ (ident \ fml)" where "SPreds \ (\(_, _, P, _). P)" abbreviation SGames:: "usubst \ (ident \ game)" where "SGames \ (\(_, _, _, G). G)" text \crude approximation of size which is enough for termination arguments\ definition usubstsize:: "usubst \ nat" where "usubstsize \ = (if (dom (SFuncs \) = {} \ dom (SPreds \) = {}) then 1 else 2)" text \dot is some fixed constant function symbol that is reserved for the purposes of the substitution\ definition dot:: "trm" where "dot = Const (dotid)" subsection \Strict Mechanism for Handling Substitution Partiality in Isabelle\ text \Optional terms that result from a substitution, either actually a term or just none to indicate that the substitution clashed\ type_synonym trmo = "trm option" abbreviation undeft:: trmo where "undeft \ None" abbreviation Aterm:: "trm \ trmo" where "Aterm \ Some" lemma undeft_None: "undeft=None" by simp lemma Aterm_Some: "Aterm \=Some \" by simp lemma undeft_equiv: "(\\undeft) = (\t. \=Aterm t)" by simp text \Plus on defined terms, strict undeft otherwise \ fun Pluso :: "trmo \ trmo \ trmo" where "Pluso (Aterm \) (Aterm \) = Aterm(Plus \ \)" | "Pluso undeft \ = undeft" | "Pluso \ undeft = undeft" text \Times on defined terms, strict undeft otherwise \ fun Timeso :: "trmo \ trmo \ trmo" where "Timeso (Aterm \) (Aterm \) = Aterm(Times \ \)" | "Timeso undeft \ = undeft" | "Timeso \ undeft = undeft" fun Differentialo :: "trmo \ trmo" where "Differentialo (Aterm \) = Aterm(Differential \)" | "Differentialo undeft = undeft" lemma Pluso_undef: "(Pluso \ \ = undeft) = (\=undeft \ \=undeft)" by (metis Pluso.elims option.distinct(1)) lemma Timeso_undef: "(Timeso \ \ = undeft) = (\=undeft \ \=undeft)" by (metis Timeso.elims option.distinct(1)) lemma Differentialo_undef: "(Differentialo \ = undeft) = (\=undeft)" by (metis Differentialo.elims option.distinct(1)) type_synonym fmlo = "fml option" abbreviation undeff:: fmlo where "undeff \ None" abbreviation Afml:: "fml \ fmlo" where "Afml \ Some" type_synonym gameo = "game option" abbreviation undefg:: gameo where "undefg \ None" abbreviation Agame:: "game \ gameo" where "Agame \ Some" lemma undeff_equiv: "(\\undeff) = (\f. \=Afml f)" by simp lemma undefg_equiv: "(\\undefg) = (\g. \=Agame g)" by simp text \Geq on defined terms, strict undeft otherwise \ fun Geqo :: "trmo \ trmo \ fmlo" where "Geqo (Aterm \) (Aterm \) = Afml(Geq \ \)" | "Geqo undeft \ = undeff" | "Geqo \ undeft = undeff" text \Not on defined formulas, strict undeft otherwise \ fun Noto :: "fmlo \ fmlo" where "Noto (Afml \) = Afml(Not \)" | "Noto undeff = undeff" text \And on defined formulas, strict undeft otherwise \ fun Ando :: "fmlo \ fmlo \ fmlo" where "Ando (Afml \) (Afml \) = Afml(And \ \)" | "Ando undeff \ = undeff" | "Ando \ undeff = undeff" text \Exists on defined formulas, strict undeft otherwise \ fun Existso :: "variable \ fmlo \ fmlo" where "Existso x (Afml \) = Afml(Exists x \)" | "Existso x undeff = undeff" text \Diamond on defined games/formulas, strict undeft otherwise \ fun Diamondo :: "gameo \ fmlo \ fmlo" where "Diamondo (Agame \) (Afml \) = Afml(Diamond \ \)" | "Diamondo undefg \ = undeff" | "Diamondo \ undeff = undeff" lemma Geqo_undef: "(Geqo \ \ = undeff) = (\=undeft \ \=undeft)" by (metis Geqo.elims option.distinct(1)) lemma Noto_undef: "(Noto \ = undeff) = (\=undeff)" by (metis Noto.elims option.distinct(1)) lemma Ando_undef: "(Ando \ \ = undeff) = (\=undeff \ \=undeff)" by (metis Ando.elims option.distinct(1)) lemma Existso_undef: "(Existso x \ = undeff) = (\=undeff)" by (metis Existso.elims option.distinct(1)) lemma Diamondo_undef: "(Diamondo \ \ = undeff) = (\=undefg \ \=undeff)" by (metis Diamondo.elims option.distinct(1)) text \Assign on defined terms, strict undefg otherwise \ fun Assigno :: "variable \ trmo \ gameo" where "Assigno x (Aterm \) = Agame(Assign x \)" | "Assigno x undeft = undefg" fun ODEo :: "ident \ trmo \ gameo" where "ODEo x (Aterm \) = Agame(ODE x \)" | "ODEo x undeft = undefg" text \Test on defined formulas, strict undefg otherwise \ fun Testo :: "fmlo \ gameo" where "Testo (Afml \) = Agame(Test \)" | "Testo undeff = undefg" text \Choice on defined games, strict undefg otherwise \ fun Choiceo :: "gameo \ gameo \ gameo" where "Choiceo (Agame \) (Agame \) = Agame(Choice \ \)" | "Choiceo \ undefg = undefg" | "Choiceo undefg \ = undefg" text \Compose on defined games, strict undefg otherwise \ fun Composeo :: "gameo \ gameo \ gameo" where "Composeo (Agame \) (Agame \) = Agame(Compose \ \)" | "Composeo \ undefg = undefg" | "Composeo undefg \ = undefg" text \Loop on defined games, strict undefg otherwise \ fun Loopo :: "gameo \ gameo" where "Loopo (Agame \) = Agame(Loop \)" | "Loopo undefg = undefg" text \Dual on defined games, strict undefg otherwise \ fun Dualo :: "gameo \ gameo" where "Dualo (Agame \) = Agame(Dual \)" | "Dualo undefg = undefg" lemma Assigno_undef: "(Assigno x \ = undefg) = (\=undeft)" by (metis Assigno.elims option.distinct(1)) lemma ODEo_undef: "(ODEo x \ = undefg) = (\=undeft)" by (metis ODEo.elims option.distinct(1)) lemma Testo_undef: "(Testo \ = undefg) = (\=undeff)" by (metis Testo.elims option.distinct(1)) lemma Choiceo_undef: "(Choiceo \ \ = undefg) = (\=undefg \ \=undefg)" by (metis Choiceo.elims option.distinct(1)) lemma Composeo_undef: "(Composeo \ \ = undefg) = (\=undefg \ \=undefg)" by (metis Composeo.elims option.distinct(1)) lemma Loopo_undef: "(Loopo \ = undefg) = (\=undefg)" by (metis Loopo.elims option.distinct(1)) lemma Dualo_undef: "(Dualo \ = undefg) = (\=undefg)" by (metis Dualo.elims option.distinct(1)) subsection \Recursive Application of One-Pass Uniform Substitution\ text \\dotsubstt \\ is the dot substitution \{. ~> \}\ substituting a term for the . function symbol\ definition dotsubstt:: "trm \ usubst" where "dotsubstt \ = ( (\f. (if f=dotid then (Some(\)) else None)), (\_. None), (\_. None), (\_. None) )" definition usappconst:: "usubst \ variable set \ ident \ (trmo)" where "usappconst \ U f \ (case SConst \ f of Some r \ if FVT(r)\U={} then Aterm(r) else undeft | None \ Aterm(Const f))" function usubstappt:: "usubst \ variable set \ (trm \ trmo)" where "usubstappt \ U (Var x) = Aterm (Var x)" | "usubstappt \ U (Number r) = Aterm (Number r)" | "usubstappt \ U (Const f) = usappconst \ U f" | "usubstappt \ U (Func f \) = (case usubstappt \ U \ of undeft \ undeft | Aterm \\ \ (case SFuncs \ f of Some r \ if FVT(r)\U={} then usubstappt(dotsubstt \\) {} r else undeft | None \ Aterm(Func f \\)))" | "usubstappt \ U (Plus \ \) = Pluso (usubstappt \ U \) (usubstappt \ U \)" | "usubstappt \ U (Times \ \) = Timeso (usubstappt \ U \) (usubstappt \ U \)" | "usubstappt \ U (Differential \) = Differentialo (usubstappt \ allvars \)" by pat_completeness auto termination by (relation "measures [\(\,U,\). usubstsize \ , \(\,U,\). size \]") (auto simp add: usubstsize_def dotsubstt_def) (* Expand let constructs automatically *) declare Let_def [simp] function usubstappf:: "usubst \ variable set \ (fml \ fmlo)" and usubstappp:: "usubst \ variable set \ (game \ variable set \ gameo)" where "usubstappf \ U (Pred p \) = (case usubstappt \ U \ of undeft \ undeff | Aterm \\ \ (case SPreds \ p of Some r \ if FVF(r)\U={} then usubstappf(dotsubstt \\) {} r else undeff | None \ Afml(Pred p \\)))" | "usubstappf \ U (Geq \ \) = Geqo (usubstappt \ U \) (usubstappt \ U \)" | "usubstappf \ U (Not \) = Noto (usubstappf \ U \)" | "usubstappf \ U (And \ \) = Ando (usubstappf \ U \) (usubstappf \ U \)" | "usubstappf \ U (Exists x \) = Existso x (usubstappf \ (U\{x}) \)" | "usubstappf \ U (Diamond \ \) = (let V\ = usubstappp \ U \ in Diamondo (snd V\) (usubstappf \ (fst V\) \))" | "usubstappp \ U (Game a) = (case SGames \ a of Some r \ (U\BVG(r),Agame r) | None \ (allvars,Agame(Game a)))" | "usubstappp \ U (Assign x \) = (U\{x}, Assigno x (usubstappt \ U \))" | "usubstappp \ U (Test \) = (U, Testo (usubstappf \ U \))" | "usubstappp \ U (Choice \ \) = (let V\ = usubstappp \ U \ in let W\ = usubstappp \ U \ in (fst V\\fst W\, Choiceo (snd V\) (snd W\)))" | "usubstappp \ U (Compose \ \) = (let V\ = usubstappp \ U \ in let W\ = usubstappp \ (fst V\) \ in (fst W\, Composeo (snd V\) (snd W\)))" | "usubstappp \ U (Loop \) = (let V = fst(usubstappp \ U \) in (V, Loopo (snd(usubstappp \ V \))))" | "usubstappp \ U (Dual \) = (let V\ = usubstappp \ U \ in (fst V\, Dualo (snd V\)))" | "usubstappp \ U (ODE x \) = (U\{RVar x,DVar x}, ODEo x (usubstappt \ (U\{RVar x,DVar x}) \))" by pat_completeness auto termination by (relation "measures [(\k. usubstsize (case k of Inl(\,U,\) \ \ | Inr(\,U,\) \ \)) , (\k. case k of Inl (\,U,\) \ size \ | Inr (\,U,\) \ size \)]") (auto simp add: usubstsize_def dotsubstt_def) text \Induction Principles for Uniform Substitutions\ lemmas usubstappt_induct = usubstappt.induct [case_names Var Number Const FuncMatch Plus Times Differential] lemmas usubstappfp_induct = usubstappf_usubstappp.induct [case_names Pred Geq Not And Exists Diamond Game Assign Test Choice Compose Loop Dual ODE] paragraph \Simple Observations for Automation\ text \More automation for Case\ lemma usappconst_simp [simp]: "SConst \ f = Some r \ FVT(r)\U={} \ usappconst \ U f = Aterm(r)" and "SConst \ f = None \ usappconst \ U f = Aterm(Const f)" and "SConst \ f = Some r \ FVT(r)\U\{} \ usappconst \ U f = undeft" unfolding usappconst_def by auto lemma usappconst_conv: "usappconst \ U f\undeft \ SConst \ f = None \ (\r. SConst \ f = Some r \ FVT(r)\U={})" (*by (smt option.case_eq_if option.collapse usappconst_def)*) proof- assume as: "usappconst \ U f\undeft" show "SConst \ f = None \ (\r. SConst \ f = Some r \ FVT(r)\U={})" proof (cases "SConst \ f") case None then show ?thesis by auto next case (Some a) then show ?thesis using as usappconst_def[where \=\ and U=U and f=f] option.distinct(1) by fastforce qed qed lemma usubstappt_const [simp]: "SConst \ f = Some r \ FVT(r)\U={} \ usubstappt \ U (Const f) = Aterm(r)" and "SConst \ f = None \ usubstappt \ U (Const f) = Aterm(Const f)" and "SConst \ f = Some r \ FVT(r)\U\{} \ usubstappt \ U (Const f) = undeft" by (auto simp add: usappconst_def) lemma usubstappt_const_conv: "usubstappt \ U (Const f)\undeft \ SConst \ f = None \ (\r. SConst \ f = Some r \ FVT(r)\U={})" using usappconst_conv by auto lemma usubstappt_func [simp]: "SFuncs \ f = Some r \ FVT(r)\U={} \ usubstappt \ U \ = Aterm \\ \ usubstappt \ U (Func f \) = usubstappt (dotsubstt \\) {} r" and "SFuncs \ f=None \ usubstappt \ U \ = Aterm \\ \ usubstappt \ U (Func f \) = Aterm(Func f \\)" and "usubstappt \ U \ = undeft \ usubstappt \ U (Func f \) = undeft" by auto lemma usubstappt_func2 [simp]: "SFuncs \ f = Some r \ FVT(r)\U\{} \ usubstappt \ U (Func f \) = undeft" by (cases "usubstappt \ U \") (auto) lemma usubstappt_func_conv: "usubstappt \ U (Func f \) \ undeft \ usubstappt \ U \ \ undeft \ (SFuncs \ f = None \ (\r. SFuncs \ f = Some r \ FVT(r)\U={}))" by (metis (lifting) option.simps(4) undeft_equiv usubstappt.simps(4) usubstappt_func2) (*by (cases "usubstappt \ U \") (auto) *) lemma usubstappt_plus_conv: "usubstappt \ U (Plus \ \) \ undeft \ usubstappt \ U \ \ undeft \ usubstappt \ U \ \ undeft" by (simp add: Pluso_undef) lemma usubstappt_times_conv: "usubstappt \ U (Times \ \) \ undeft \ usubstappt \ U \ \ undeft \ usubstappt \ U \ \ undeft" by (simp add: Timeso_undef) lemma usubstappt_differential_conv: "usubstappt \ U (Differential \) \ undeft \ usubstappt \ allvars \ \ undeft" by (simp add: Differentialo_undef) lemma usubstappf_pred [simp]: "SPreds \ p = Some r \ FVF(r)\U={} \ usubstappt \ U \ = Aterm \\ \ usubstappf \ U (Pred p \) = usubstappf (dotsubstt \\) {} r" and "SPreds \ p = None \ usubstappt \ U \ = Aterm \\ \ usubstappf \ U (Pred p \) = Afml(Pred p \\)" and "usubstappt \ U \ = undeft \ usubstappf \ U (Pred p \) = undeff" by auto lemma usubstappf_pred2 [simp]: "SPreds \ p = Some r \ FVF(r)\U\{} \ usubstappf \ U (Pred p \) = undeff" by (cases "usubstappt \ U \") (auto) lemma usubstappf_pred_conv: "usubstappf \ U (Pred p \) \ undeff \ usubstappt \ U \ \ undeft \ (SPreds \ p = None \ (\r. SPreds \ p = Some r \ FVF(r)\U={}))" by (metis (lifting) option.simps(4) undeff_equiv usubstappf.simps(1) usubstappf_pred2) lemma usubstappf_geq: "usubstappt \ U \ \ undeft \ usubstappt \ U \ \ undeft \ usubstappf \ U (Geq \ \) = Afml(Geq (the (usubstappt \ U \)) (the (usubstappt \ U \)))" by fastforce lemma usubstappf_geq_conv: "usubstappf \ U (Geq \ \) \ undeff \ usubstappt \ U \ \ undeft \ usubstappt \ U \ \ undeft" by (simp add: Geqo_undef) lemma usubstappf_geqr: "usubstappf \ U (Geq \ \) \ undeff \ usubstappf \ U (Geq \ \) = Afml(Geq (the (usubstappt \ U \)) (the (usubstappt \ U \)))" using usubstappf_geq usubstappf_geq_conv by blast lemma usubstappf_exists: "usubstappf \ U (Exists x \) \ undeff \ usubstappf \ U (Exists x \) = Afml(Exists x (the (usubstappf \ (U\{x}) \)))" using Existso_undef by auto lemma usubstappp_game [simp]: "SGames \ a = Some r \ usubstappp \ U (Game a) = (U\BVG(r),Agame(r))" and "SGames \ a = None \ usubstappp \ U (Game a) = (allvars,Agame(Game a))" by auto lemma usubstappp_choice [simp]: "usubstappp \ U (Choice \ \) = (fst(usubstappp \ U \)\fst(usubstappp \ U \), Choiceo (snd(usubstappp \ U \)) (snd(usubstappp \ U \)))" by auto lemma usubstappp_choice_conv : "snd(usubstappp \ U (Choice \ \)) \ undefg \ snd(usubstappp \ U \) \ undefg \ snd(usubstappp \ U \) \ undefg" by (simp add: Choiceo_undef) lemma usubstappp_compose [simp]: "usubstappp \ U (Compose \ \) = (fst(usubstappp \ (fst(usubstappp \ U \)) \), Composeo (snd(usubstappp \ U \)) (snd(usubstappp \ (fst(usubstappp \ U \)) \)))" by simp lemma usubstappp_loop: "usubstappp \ U (Loop \) = (fst(usubstappp \ U \), Loopo (snd(usubstappp \ (fst(usubstappp \ U \)) \)))" by auto lemma usubstappp_dual [simp]: "usubstappp \ U (Dual \) = (fst(usubstappp \ U \), Dualo (snd (usubstappp \ U \)))" by simp section \Soundness of Uniform Substitution\ subsection \USubst Application is a Function of Deterministic Result\ lemma usubstappt_det: "usubstappt \ U \ \ undeft \ usubstappt \ V \ \ undeft \ usubstappt \ U \ = usubstappt \ V \" proof (induction \) case (Var x) then show ?case by simp next case (Number x) then show ?case by simp next case (Const f) then show ?case (*by (smt option.case_eq_if usappconst_def usubstappt.simps(3))*) proof - (*sledgehammer*) have f1: "usubstappt \ U (Const f) = (case SConst \ f of None \ Aterm (Const f) | Some t \ if FVT t \ U = {} then Aterm t else undeft)" by (simp add: usappconst_def) have f2: "\z f za. if za = undeft then (case za of None \ z::trm option | Some x \ f x) = z else (case za of None \ z | Some x \ f x) = f (the za)" by force then have "SConst \ f \ undeft \ (if FVT (the (SConst \ f)) \ U = {} then Aterm (the (SConst \ f)) else undeft) = usappconst \ U f" by (simp add: usappconst_def) then have f3: "SConst \ f \ undeft \ FVT (the (SConst \ f)) \ U = {}" by (metis Const.prems(1) usubstappt.simps(3)) have "SConst \ f \ undeft \ (if FVT (the (SConst \ f)) \ V = {} then Aterm (the (SConst \ f)) else undeft) = usappconst \ V f" using f2 usappconst_def by presburger then have "SConst \ f \ undeft \ FVT (the (SConst \ f)) \ V = {}" by (metis (no_types) Const.prems(2) usubstappt.simps(3)) then have f4: "SConst \ f \ undeft \ usubstappt \ U (Const f) = usappconst \ V f" using f3 f2 f1 usappconst_def by presburger { assume "usubstappt \ U (Const f) \ usubstappt \ V (Const f)" then have "usubstappt \ U (Const f) \ (case SConst \ f of None \ Aterm (Const f) | Some t \ if FVT t \ V = {} then Aterm t else undeft)" by (simp add: usappconst_def) then have "SConst \ f \ undeft" using f2 f1 by (metis (no_types)) then have ?thesis using f4 by simp } then show ?thesis by blast qed next case (Func f \) then show ?case using usubstappt_func (*by (cases "SFuncs \ f") (auto simp add: usubstappt_func)*) (*by (smt option.case_eq_if usubstappt.simps(4))*) proof - (*sledgehammer*) have f1: "(case usubstappt \ U \ of None \ undeft | Some t \ (case SFuncs \ f of None \ Aterm (trm.Func f t) | Some ta \ if FVT ta \ U = {} then usubstappt (dotsubstt t) {} ta else undeft)) \ undeft" using Func(2) by auto have f2: "\z f za. if za = undeft then (case za of None \ z::trm option | Some x \ f x) = z else (case za of None \ z | Some x \ f x) = f (the za)" by force then have f3: "usubstappt \ U \ \ undeft" using f1 by meson have "(case usubstappt \ V \ of None \ undeft | Some t \ (case SFuncs \ f of None \ Aterm (trm.Func f t) | Some ta \ if FVT ta \ V = {} then usubstappt (dotsubstt t) {} ta else undeft)) \ undeft" using Func(3) by auto then have f4: "usubstappt \ V \ \ undeft" using f2 by meson then have f5: "usubstappt \ U (trm.Func f \) = (case SFuncs \ f of None \ Aterm (trm.Func f (the (usubstappt \ V \))) | Some t \ if FVT t \ U = {} then usubstappt (dotsubstt (the (usubstappt \ V \))) {} t else undeft)" using f3 f2 Func(1) usubstappt.simps(4) by presburger have "SFuncs \ f \ undeft \ (if FVT (the (SFuncs \ f)) \ U = {} then usubstappt (dotsubstt (the (usubstappt \ V \))) {} (the (SFuncs \ f)) else undeft) = usubstappt \ U (trm.Func f \)" using f4 f3 f2 Func(1) usubstappt.simps(4) by presburger then have f6: "SFuncs \ f \ undeft \ FVT (the (SFuncs \ f)) \ U = {}" by (metis Func(2)) have f7: "(case usubstappt \ V \ of None \ undeft | Some t \ (case SFuncs \ f of None \ Aterm (trm.Func f t) | Some ta \ if FVT ta \ V = {} then usubstappt (dotsubstt t) {} ta else undeft)) = (case SFuncs \ f of None \ Aterm (trm.Func f (the (usubstappt \ V \))) | Some t \ if FVT t \ V = {} then usubstappt (dotsubstt (the (usubstappt \ V \))) {} t else undeft)" using f4 f2 by presburger then have "SFuncs \ f \ undeft \ (if FVT (the (SFuncs \ f)) \ V = {} then usubstappt (dotsubstt (the (usubstappt \ V \))) {} (the (SFuncs \ f)) else undeft) = usubstappt \ V (trm.Func f \)" using f2 by simp then have f8: "SFuncs \ f \ undeft \ FVT (the (SFuncs \ f)) \ V = {}" by (metis (full_types) Func(3)) { assume "usubstappt \ U (trm.Func f \) \ usubstappt \ V (trm.Func f \)" moreover { assume "(case SFuncs \ f of None \ Aterm (trm.Func f (the (usubstappt \ V \))) | Some t \ if FVT t \ V = {} then usubstappt (dotsubstt (the (usubstappt \ V \))) {} t else undeft) \ Aterm (trm.Func f (the (usubstappt \ V \)))" then have "SFuncs \ f \ undeft" using f2 by meson } ultimately have "SFuncs \ f \ undeft" using f7 f5 by fastforce then have ?thesis using f8 f7 f6 f5 f2 by simp } then show ?thesis by blast qed next case (Plus \1 \2) then show ?case using Pluso_undef by auto next case (Times \1 \2) then show ?case using Timeso_undef by auto next case (Differential \) then show ?case using Differentialo_undef by auto qed lemma usubstappf_and_usubstappp_det: shows "usubstappf \ U \ \ undeff \ usubstappf \ V \ \ undeff \ usubstappf \ U \ = usubstappf \ V \" and "snd(usubstappp \ U \) \ undefg \ snd(usubstappp \ V \) \ undefg \ snd(usubstappp \ U \) = snd(usubstappp \ V \)" proof (induction \ and \ arbitrary: U V and U V) case (Pred p \) then show ?case using usubstappt_det usubstappf_pred (*by (metis usubstappf.simps(1)) *) (*by (smt option.case_eq_if usubstappf.simps(1)) *) proof - (*sledgehammer*) have f1: "(case usubstappt \ U \ of None \ undeff | Some t \ (case SPreds \ p of None \ Afml (Pred p t) | Some f \ if FVF f \ U = {} then usubstappf (dotsubstt t) {} f else undeff)) \ undeff" using Pred.prems(1) by auto have f2: "\z f za. if za = undeft then (case za of None \ z::fml option | Some x \ f x) = z else (case za of None \ z | Some x \ f x) = f (the za)" by (simp add: option.case_eq_if) then have f3: "(case usubstappt \ U \ of None \ undeff | Some t \ (case SPreds \ p of None \ Afml (Pred p t) | Some f \ if FVF f \ U = {} then usubstappf (dotsubstt t) {} f else undeff)) = (case SPreds \ p of None \ Afml (Pred p (the (usubstappt \ U \))) | Some f \ if FVF f \ U = {} then usubstappf (dotsubstt (the (usubstappt \ U \))) {} f else undeff)" using f1 by meson have f4: "(case usubstappt \ V \ of None \ undeff | Some t \ (case SPreds \ p of None \ Afml (Pred p t) | Some f \ if FVF f \ V = {} then usubstappf (dotsubstt t) {} f else undeff)) \ undeff" using Pred.prems(2) by auto then have f5: "usubstappt \ U \ = usubstappt \ V \" using f2 f1 by (meson usubstappt_det) then have f6: "usubstappf \ U (Pred p \) = (case SPreds \ p of None \ Afml (Pred p (the (usubstappt \ V \))) | Some f \ if FVF f \ U = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} f else undeff)" using f3 usubstappf.simps(1) by presburger have f7: "\z f za. if za = undeff then (case za of None \ z::fml option | Some x \ f x) = z else (case za of None \ z | Some x \ f x) = f (the za)" by (simp add: option.case_eq_if) have f8: "(case usubstappt \ V \ of None \ undeff | Some t \ (case SPreds \ p of None \ Afml (Pred p t) | Some f \ if FVF f \ V = {} then usubstappf (dotsubstt t) {} f else undeff)) = (case SPreds \ p of None \ Afml (Pred p (the (usubstappt \ V \))) | Some f \ if FVF f \ V = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} f else undeff)" using f4 f2 by meson then have f9: "SPreds \ p = undeff \ usubstappf \ U (Pred p \) = usubstappf \ V (Pred p \)" using f6 by fastforce { assume "usubstappf \ U (Pred p \) \ usubstappf \ V (Pred p \)" then have "usubstappf \ U (Pred p \) \ (case SPreds \ p of None \ Afml (Pred p (the (usubstappt \ V \))) | Some f \ if FVF f \ V = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} f else undeff)" using f8 by simp moreover { assume "usubstappf \ U (Pred p \) \ (if FVF (the (SPreds \ p)) \ V = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} (the (SPreds \ p)) else undeff)" moreover { assume "usubstappf \ U (Pred p \) \ usubstappf (dotsubstt (the (usubstappt \ V \))) {} (the (SPreds \ p))" moreover { assume "usubstappf \ U (Pred p \) \ (if FVF (the (SPreds \ p)) \ U = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} (the (SPreds \ p)) else undeff)" then have "(case SPreds \ p of None \ Afml (Pred p (the (usubstappt \ V \))) | Some f \ if FVF f \ U = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} f else undeff) \ (if FVF (the (SPreds \ p)) \ U = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} (the (SPreds \ p)) else undeff)" using f6 by force then have "SPreds \ p = undeff" using f7 by (metis (no_types, lifting) Pred.prems(1) calculation f5 option.collapse usubstappf_pred usubstappf_pred2 usubstappf_pred_conv) } moreover { assume "(if FVF (the (SPreds \ p)) \ U = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} (the (SPreds \ p)) else undeff) \ usubstappf (dotsubstt (the (usubstappt \ V \))) {} (the (SPreds \ p))" then have "(if FVF (the (SPreds \ p)) \ U = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} (the (SPreds \ p)) else undeff) \ (case SPreds \ p of None \ Afml (Pred p (the (usubstappt \ U \))) | Some f \ if FVF f \ U = {} then usubstappf (dotsubstt (the (usubstappt \ U \))) {} f else undeff)" using f3 Pred.prems(1) by auto then have "(if FVF (the (SPreds \ p)) \ U = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} (the (SPreds \ p)) else undeff) \ (case SPreds \ p of None \ Afml (Pred p (the (usubstappt \ V \))) | Some f \ if FVF f \ U = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} f else undeff)" using f5 using Pred.prems(1) \(if FVF (the (SPreds \ p)) \ U = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} (the (SPreds \ p)) else undeff) \ usubstappf (dotsubstt (the (usubstappt \ V \))) {} (the (SPreds \ p))\ f6 by auto then have "(case SPreds \ p of None \ Afml (Pred p (the (usubstappt \ V \))) | Some f \ if FVF f \ U = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} f else undeff) \ (if FVF (the (SPreds \ p)) \ U = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} (the (SPreds \ p)) else undeff)" by simp then have "SPreds \ p = undeff" using f7 proof - show ?thesis using \(case SPreds \ p of None \ Afml (Pred p (the (usubstappt \ V \))) | Some f \ if FVF f \ U = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} f else undeff) \ (if FVF (the (SPreds \ p)) \ U = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} (the (SPreds \ p)) else undeff)\ calculation(2) f6 by presburger qed} ultimately have "SPreds \ p = undeff" by fastforce } moreover { assume "(if FVF (the (SPreds \ p)) \ V = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} (the (SPreds \ p)) else undeff) \ usubstappf (dotsubstt (the (usubstappt \ V \))) {} (the (SPreds \ p))" then have "(case SPreds \ p of None \ Afml (Pred p (the (usubstappt \ V \))) | Some f \ if FVF f \ V = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} f else undeff) \ (if FVF (the (SPreds \ p)) \ V = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} (the (SPreds \ p)) else undeff)" using f8 Pred.prems(2) by auto then have "SPreds \ p = undeff" using f7 by (metis (no_types, lifting) Pred.prems(2) \(if FVF (the (SPreds \ p)) \ V = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} (the (SPreds \ p)) else undeff) \ usubstappf (dotsubstt (the (usubstappt \ V \))) {} (the (SPreds \ p))\ option.collapse usubstappf_pred2)} ultimately have "SPreds \ p = undeff" by fastforce } moreover { assume "(case SPreds \ p of None \ Afml (Pred p (the (usubstappt \ V \))) | Some f \ if FVF f \ V = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} f else undeff) \ (if FVF (the (SPreds \ p)) \ V = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} (the (SPreds \ p)) else undeff)" then have "SPreds \ p = undeff" using f7 by (metis (no_types, lifting) Pred.prems(1) Pred.prems(2) \usubstappf \ U (Pred p \) \ usubstappf \ V (Pred p \)\ calculation(2) option.collapse usubstappf_pred usubstappf_pred_conv)} ultimately have ?thesis using f9 by fastforce } then show ?thesis by blast qed next case (Geq \ \) then show ?case using usubstappt_det by (metis Geqo_undef usubstappf.simps(2)) next case (Not x) then show ?case by (metis Noto.simps(2) usubstappf.simps(3)) next case (And x1 x2) then show ?case by (metis Ando_undef usubstappf.simps(4)) next case (Exists x1 x2) then show ?case by (metis Existso_undef usubstappf.simps(5)) next case (Diamond x1 x2) then show ?case by (metis Diamondo_undef usubstappf.simps(6)) next case (Game a) then show ?case by (cases "SGames \ a") (auto) next case (Assign x \) then show ?case using usubstappt_det by (metis Assigno_undef snd_conv usubstappp.simps(2)) next case (ODE x \) then show ?case using usubstappt_det by (metis ODEo_undef snd_conv usubstappp.simps(8)) next case (Test \) then show ?case by (metis Testo_undef snd_conv usubstappp.simps(3)) next case (Choice \ \) then show ?case by (metis Choiceo_undef snd_conv usubstappp.simps(4)) next case (Compose \ \) then show ?case by (metis Composeo_undef snd_conv usubstappp.simps(5)) next case (Loop \) then show ?case by (metis Loopo_undef snd_conv usubstappp.simps(6)) next case (Dual \) then show ?case by (metis Dualo_undef snd_conv usubstappp.simps(7)) qed lemma usubstappf_det: "usubstappf \ U \ \ undeff \ usubstappf \ V \ \ undeff \ usubstappf \ U \ = usubstappf \ V \" using usubstappf_and_usubstappp_det by simp lemma usubstappp_det: "snd(usubstappp \ U \) \ undefg \ snd(usubstappp \ V \) \ undefg \ snd(usubstappp \ U \) = snd(usubstappp \ V \)" using usubstappf_and_usubstappp_det by simp subsection \Uniform Substitutions are Antimonotone in Taboos\ lemma usubst_taboos_mon: "fst(usubstappp \ U \) \ U" proof (induction \ arbitrary: U rule: game_induct) case (Game a) then show ?case by (cases "SGames \ a") (auto) next case (Assign x \) then show ?case by fastforce next case (ODE x \) then show ?case by fastforce next case (Test \) then show ?case by fastforce next case (Choice \ \) then show ?case by fastforce next case (Compose \ \) then show ?case by fastforce next case (Loop \) then show ?case by fastforce next case (Dual \) then show ?case by fastforce qed lemma fst_pair [simp]: "fst (a,b) = a" by simp lemma snd_pair [simp]: "snd (a,b) = b" by simp lemma usubstappt_antimon: "V\U \ usubstappt \ U \ \ undeft \ usubstappt \ U \ = usubstappt \ V \" proof (induction \) case (Var x) then show ?case by simp next case (Number x) then show ?case by simp next case (Const f) then show ?case (*by (smt disjoint_iff_not_equal option.case_eq_if set_rev_mp usappconst_def usubstappt.simps(3))*) proof - (*sledgehammer*) have f1: "usubstappt \ U (Const f) = (case SConst \ f of None \ Aterm (Const f) | Some t \ if FVT t \ U = {} then Aterm t else undeft)" by (simp add: usappconst_def) have f2: "\z f za. if za = undeft then (case za of None \ z::trm option | Some x \ f x) = z else (case za of None \ z | Some x \ f x) = f (the za)" by force then have "SConst \ f \ undeft \ (if FVT (the (SConst \ f)) \ U = {} then Aterm (the (SConst \ f)) else undeft) = usappconst \ U f" using usappconst_def by presburger then have f3: "SConst \ f \ undeft \ FVT (the (SConst \ f)) \ U = {}" by (metis (no_types) Const.prems(2) usubstappt.simps(3)) have f4: "\V Va. (V \ Va = {}) = (\v. (v::variable) \ V \ (\va. va \ Va \ v \ va))" by blast { assume "usubstappt \ U (Const f) \ usubstappt \ V (Const f)" then have "usubstappt \ U (Const f) \ (case SConst \ f of None \ Aterm (Const f) | Some t \ if FVT t \ V = {} then Aterm t else undeft)" by (simp add: usappconst_def) then have "SConst \ f \ undeft" using f2 f1 by metis then have "SConst \ f \ undeft \ FVT (the (SConst \ f)) \ V = {}" using f4 f3 by (meson Const.prems(1) subsetD) then have ?thesis using f3 f2 usappconst_def usubstappt.simps(3) by presburger } then show ?thesis by blast qed next case (Func f \) then show ?case using usubstappt_func (*by (smt disjoint_iff_not_equal option.case_eq_if subset_iff usubstappt.simps(4))*) proof - (*sledgehammer*) have f1: "(case usubstappt \ U \ of None \ undeft | Some t \ (case SFuncs \ f of None \ Aterm (trm.Func f t) | Some ta \ if FVT ta \ U = {} then usubstappt (dotsubstt t) {} ta else undeft)) \ undeft" using Func.prems(2) by fastforce have f2: "\z f za. if za = undeft then (case za of None \ z::trm option | Some x \ f x) = z else (case za of None \ z | Some x \ f x) = f (the za)" by fastforce then have f3: "usubstappt \ U \ \ undeft" using f1 by meson then have f4: "(case usubstappt \ U \ of None \ undeft | Some t \ (case SFuncs \ f of None \ Aterm (trm.Func f t) | Some ta \ if FVT ta \ U = {} then usubstappt (dotsubstt t) {} ta else undeft)) = (case SFuncs \ f of None \ Aterm (trm.Func f (the (usubstappt \ U \))) | Some t \ if FVT t \ U = {} then usubstappt (dotsubstt (the (usubstappt \ U \))) {} t else undeft)" using f2 by presburger have f5: "usubstappt \ U \ = undeft \ usubstappt \ U \ = usubstappt \ V \" using Func.IH Func.prems(1) by fastforce have "SFuncs \ f \ undeft \ (if FVT (the (SFuncs \ f)) \ U = {} then usubstappt (dotsubstt (the (usubstappt \ V \))) {} (the (SFuncs \ f)) else undeft) = (case SFuncs \ f of None \ Aterm (trm.Func f (the (usubstappt \ V \))) | Some t \ if FVT t \ U = {} then usubstappt (dotsubstt (the (usubstappt \ V \))) {} t else undeft)" using f2 by presburger then have "SFuncs \ f \ undeft \ (if FVT (the (SFuncs \ f)) \ U = {} then usubstappt (dotsubstt (the (usubstappt \ V \))) {} (the (SFuncs \ f)) else undeft) = usubstappt \ U (trm.Func f \)" using f5 f4 f3 usubstappt.simps(4) by presburger then have f6: "SFuncs \ f \ undeft \ FVT (the (SFuncs \ f)) \ U = {}" by (metis (no_types) Func.prems(2)) then have f7: "SFuncs \ f \ undeft \ V \ - FVT (the (SFuncs \ f))" using Func.prems(1) by blast have f8: "(case usubstappt \ V \ of None \ undeft | Some t \ (case SFuncs \ f of None \ Aterm (trm.Func f t) | Some ta \ if FVT ta \ V = {} then usubstappt (dotsubstt t) {} ta else undeft)) = (case SFuncs \ f of None \ Aterm (trm.Func f (the (usubstappt \ V \))) | Some t \ if FVT t \ V = {} then usubstappt (dotsubstt (the (usubstappt \ V \))) {} t else undeft)" using f5 f3 f2 by presburger have "SFuncs \ f \ undeft \ usubstappt (dotsubstt (the (usubstappt \ V \))) {} (the (SFuncs \ f)) = (case SFuncs \ f of None \ Aterm (trm.Func f (the (usubstappt \ V \))) | Some t \ if FVT t \ U = {} then usubstappt (dotsubstt (the (usubstappt \ V \))) {} t else undeft)" using f6 f2 by presburger then have f9: "SFuncs \ f \ undeft \ usubstappt (dotsubstt (the (usubstappt \ V \))) {} (the (SFuncs \ f)) = usubstappt \ U (trm.Func f \)" using f5 f4 f3 usubstappt.simps(4) by presburger { assume "usubstappt \ U (trm.Func f \) \ usubstappt \ V (trm.Func f \)" moreover { assume "(case SFuncs \ f of None \ Aterm (trm.Func f (the (usubstappt \ V \))) | Some t \ if FVT t \ V = {} then usubstappt (dotsubstt (the (usubstappt \ V \))) {} t else undeft) \ Aterm (trm.Func f (the (usubstappt \ V \)))" then have "SFuncs \ f \ undeft" using f2 by meson } ultimately have "SFuncs \ f \ undeft" using f5 f3 by fastforce then have ?thesis using f9 f8 f7 f2 by (simp add: disjoint_eq_subset_Compl inf.commute) } then show ?thesis by blast qed next case (Plus \1 \2) then show ?case using Pluso_undef by auto next case (Times \1 \2) then show ?case using Timeso_undef by auto next case (Differential \) then show ?case using Differentialo_undef by auto qed text \Uniform Substitutions of Games have monotone taboo output\ lemma usubstappp_fst_mon: "U\V \ fst(usubstappp \ U \) \ fst(usubstappp \ V \)" proof (induction \ arbitrary: U V rule: game_induct) case (Game a) then show ?case by (cases "SGames \ a") (auto) next case (Assign x \) then show ?case by auto next case (ODE x \) then show ?case by auto next case (Test \) then show ?case by auto next case (Choice \ \) then show ?case by (metis Un_mono fst_pair usubstappp_choice) next case (Compose \ \) then show ?case by (metis fst_pair usubstappp_compose) next case (Loop \) then show ?case by (metis fst_pair usubstappp_loop) next case (Dual \) then show ?case by (metis fst_pair usubstappp_dual) qed lemma usubstappf_and_usubstappp_antimon: shows "V\U \ usubstappf \ U \ \ undeff \ usubstappf \ U \ = usubstappf \ V \" and "V\U \ snd(usubstappp \ U \) \ undefg \ snd(usubstappp \ U \) = snd(usubstappp \ V \)" proof- have "V\U \ usubstappf \ U \ \ undeff \ usubstappf \ V \ \ undeff" and "V\U \ snd(usubstappp \ U \) \ undefg \ snd(usubstappp \ V \) \ undefg" proof (induction \ and \ arbitrary: U V and U V) case (Pred p \) then show ?case using usubstappt_antimon usubstappf_pred (*by (smt Un_mono disjoint_eq_subset_Compl empty_subsetI inf.commute option.case_eq_if sup.absorb_iff1 sup.absorb_iff2 usubstappf.simps(1)) *) proof - (*sledgehammer*) have f1: "\v. v \ V \ v \ U" using Pred.prems(1) by auto have f2: "\z f za. if za = undeff then (case za of None \ z::fml option | Some x \ f x) = z else (case za of None \ z | Some x \ f x) = f (the za)" by (simp add: option.case_eq_if) have f3: "(case usubstappt \ U \ of None \ undeff | Some t \ (case SPreds \ p of None \ Afml (Pred p t) | Some f \ if FVF f \ U = {} then usubstappf (dotsubstt t) {} f else undeff)) \ undeff" using Pred.prems(2) by auto have f4: "\z f za. if za = undeft then (case za of None \ z::fml option | Some x \ f x) = z else (case za of None \ z | Some x \ f x) = f (the za)" by (simp add: option.case_eq_if) then have f5: "usubstappt \ U \ \ undeft" using f3 by meson then have f6: "(case usubstappt \ U \ of None \ undeff | Some t \ (case SPreds \ p of None \ Afml (Pred p t) | Some f \ if FVF f \ U = {} then usubstappf (dotsubstt t) {} f else undeff)) = (case SPreds \ p of None \ Afml (Pred p (the (usubstappt \ U \))) | Some f \ if FVF f \ U = {} then usubstappf (dotsubstt (the (usubstappt \ U \))) {} f else undeff)" using f4 by presburger have f7: "usubstappt \ U \ = usubstappt \ V \" using f5 by (meson Pred.prems(1) usubstappt_antimon) then have f8: "SPreds \ p = undeff \ usubstappf \ U (Pred p \) = usubstappf \ V (Pred p \)" using f2 usubstappf.simps(1) by presburger obtain vv :: "variable set \ variable set \ variable" where "\x0 x1. (\v2. v2 \ x1 \ (\v3. v3 \ x0 \ v2 = v3)) = (vv x0 x1 \ x1 \ (\v3. v3 \ x0 \ vv x0 x1 = v3))" by moura then obtain vva :: "variable set \ variable set \ variable" where f9: "\V Va. (V \ Va \ {} \ (\v. v \ V \ (\va. va \ Va \ v \ va))) \ (V \ Va = {} \ vv Va V \ V \ vva Va V \ Va \ vv Va V = vva Va V)" - by moura + by auto then have f10: "(FVF (the (SPreds \ p)) \ V \ {} \ (\v. v \ FVF (the (SPreds \ p)) \ (\va. va \ V \ v \ va))) \ (FVF (the (SPreds \ p)) \ V = {} \ vv V (FVF (the (SPreds \ p))) \ FVF (the (SPreds \ p)) \ vva V (FVF (the (SPreds \ p))) \ V \ vv V (FVF (the (SPreds \ p))) = vva V (FVF (the (SPreds \ p))))" by presburger { assume "vv V (FVF (the (SPreds \ p))) \ FVF (the (SPreds \ p)) \ vva V (FVF (the (SPreds \ p))) \ V \ vv V (FVF (the (SPreds \ p))) \ vva V (FVF (the (SPreds \ p)))" moreover { assume "(if FVF (the (SPreds \ p)) \ V = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} (the (SPreds \ p)) else undeff) \ undeff" moreover { assume "(if FVF (the (SPreds \ p)) \ V = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} (the (SPreds \ p)) else undeff) \ usubstappf \ V (Pred p \)" then have "(case SPreds \ p of None \ Afml (Pred p (the (usubstappt \ V \))) | Some f \ if FVF f \ V = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} f else undeff) \ (if FVF (the (SPreds \ p)) \ V = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} (the (SPreds \ p)) else undeff)" using f7 f5 f4 by simp then have "SPreds \ p = undeff" using f2 by (metis (no_types, lifting) \(if FVF (the (SPreds \ p)) \ V = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} (the (SPreds \ p)) else undeff) \ usubstappf \ V (Pred p \)\ calculation f5 f7 option.collapse usubstappf_pred)} ultimately have "usubstappf \ V (Pred p \) = undeff \ SPreds \ p = undeff" by fastforce } moreover { assume "usubstappf (dotsubstt (the (usubstappt \ V \))) {} (the (SPreds \ p)) \ usubstappf \ U (Pred p \)" then have "usubstappf (dotsubstt (the (usubstappt \ V \))) {} (the (SPreds \ p)) \ (case SPreds \ p of None \ Afml (Pred p (the (usubstappt \ U \))) | Some f \ if FVF f \ U = {} then usubstappf (dotsubstt (the (usubstappt \ U \))) {} f else undeff)" using f6 by simp then have "usubstappf (dotsubstt (the (usubstappt \ V \))) {} (the (SPreds \ p)) \ (case SPreds \ p of None \ Afml (Pred p (the (usubstappt \ V \))) | Some f \ if FVF f \ U = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} f else undeff)" using f7 by (metis f7) moreover { assume "(case SPreds \ p of None \ Afml (Pred p (the (usubstappt \ V \))) | Some f \ if FVF f \ U = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} f else undeff) \ (if FVF (the (SPreds \ p)) \ U = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} (the (SPreds \ p)) else undeff)" then have "SPreds \ p = undeff" using f2 by (metis (no_types, lifting) Pred.prems(2) \usubstappf (dotsubstt (the (usubstappt \ V \))) {} (the (SPreds \ p)) \ usubstappf \ U (Pred p \)\ f5 f7 option.collapse usubstappf_pred usubstappf_pred2)} ultimately have "FVF (the (SPreds \ p)) \ U = {} \ SPreds \ p = undeff" by force } ultimately have "FVF (the (SPreds \ p)) \ U = {} \ usubstappf \ V (Pred p \) = undeff \ SPreds \ p = undeff" using f10 by (metis Pred.prems(2)) } moreover { assume "FVF (the (SPreds \ p)) \ U \ {}" then have "(if FVF (the (SPreds \ p)) \ U = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} (the (SPreds \ p)) else undeff) \ (case SPreds \ p of None \ Afml (Pred p (the (usubstappt \ U \))) | Some f \ if FVF f \ U = {} then usubstappf (dotsubstt (the (usubstappt \ U \))) {} f else undeff)" using f6 Pred.prems(2) by auto then have "(if FVF (the (SPreds \ p)) \ U = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} (the (SPreds \ p)) else undeff) \ (case SPreds \ p of None \ Afml (Pred p (the (usubstappt \ V \))) | Some f \ if FVF f \ U = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} f else undeff)" using f7 by (metis \FVF (the (SPreds \ p)) \ U \ {}\) then have "(case SPreds \ p of None \ Afml (Pred p (the (usubstappt \ V \))) | Some f \ if FVF f \ U = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} f else undeff) \ (if FVF (the (SPreds \ p)) \ U = {} then usubstappf (dotsubstt (the (usubstappt \ V \))) {} (the (SPreds \ p)) else undeff)" by simp then have "SPreds \ p = undeff" using f2 proof - show ?thesis by (metis (no_types) Pred.prems(2) \FVF (the (SPreds \ p)) \ U \ {}\ option.discI option.expand option.sel usubstappf_pred2) qed } ultimately have "usubstappf \ V (Pred p \) = undeff \ SPreds \ p = undeff" using f9 f1 by meson then show ?thesis using f8 by (metis (full_types) Pred.prems(2)) qed next case (Geq \ \) then show ?case using usubstappt_antimon using Geqo_undef by auto next case (Not x) then show ?case using Noto_undef by auto next case (And x1 x2) then show ?case using Ando_undef by auto next case (Exists x1 x2) then show ?case using Existso_undef by (metis (no_types, lifting) Un_mono subsetI usubstappf.simps(5)) next case (Diamond x1 x2) then show ?case using Diamondo_undef usubstappf.simps(6) usubstappp_fst_mon by metis next case (Game a) then show ?case by (cases "SGames \ a") (auto) next case (Assign x \) then show ?case using usubstappt_antimon by (metis Assigno_undef snd_conv usubstappp.simps(2)) next case (ODE x \) then show ?case using usubstappt_antimon ODEo_undef by (metis (no_types, opaque_lifting) Un_mono order_refl snd_conv usubstappp.simps(8)) next case (Test \) then show ?case by (metis Testo_undef snd_conv usubstappp.simps(3)) next case (Choice \ \) then show ?case using Choiceo_undef by auto next case (Compose \ \) then show ?case using usubstappp_compose[where \=\ and U=U and \=\ and \=\] usubstappp_compose[where \=\ and U=V and \=\ and \=\] Composeo_undef[where \=\snd (usubstappp \ U \)\ and \=\snd (usubstappp \ (fst (usubstappp \ U \)) \)\] Composeo_undef[where \=\snd (usubstappp \ V \)\ and \=\snd (usubstappp \ (fst (usubstappp \ V \)) \)\] snd_conv usubstappp_fst_mon by metis (*proof- from Compose have ca: "snd (usubstappp \ U (\ ;; \)) \ undefg" by simp have decU: "snd (usubstappp \ U (\ ;; \)) = Composeo (snd(usubstappp \ U \)) (snd(usubstappp \ (fst(usubstappp \ U \)) \))" using usubstappp_compose by simp have decV: "snd (usubstappp \ V (\ ;; \)) = Composeo (snd(usubstappp \ V \)) (snd(usubstappp \ (fst(usubstappp \ V \)) \))" using usubstappp_compose by simp from Compose have fact1: "snd(usubstappp \ V \) \ undefg" using Composeo_undef by auto from Compose have fact2: "snd(usubstappp \ (fst(usubstappp \ U \)) \) \ undefg" using Composeo_undef by auto have rel: "fst(usubstappp \ V \) \ fst(usubstappp \ U \)" using \V\U\ usubstappp_fst_mon by auto from Compose have fact3: "snd(usubstappp \ (fst(usubstappp \ V \)) \) \ undefg" using fact2 rel by auto then show ?thesis by (simp add: Composeo_undef fact1) qed*) next case (Loop \) then show ?case using Loopo_undef snd_conv usubstappp.simps(6) usubstappp_fst_mon by metis next case (Dual \) then show ?case by (metis Dualo_undef snd_conv usubstappp.simps(7)) qed from this show "V\U \ usubstappf \ U \ \ undeff \ usubstappf \ U \ = usubstappf \ V \" and "V\U \ snd(usubstappp \ U \) \ undefg \ snd(usubstappp \ U \) = snd(usubstappp \ V \)" using usubstappf_and_usubstappp_det by auto qed lemma usubstappf_antimon: "V\U \ usubstappf \ U \ \ undeff \ usubstappf \ U \ = usubstappf \ V \" using usubstappf_and_usubstappp_antimon by simp lemma usubstappp_antimon: "V\U \ snd(usubstappp \ U \) \ undefg \ snd(usubstappp \ U \) = snd(usubstappp \ V \)" using usubstappf_and_usubstappp_antimon by simp subsection \Taboo Lemmas\ lemma usubstappp_loop_conv: "snd (usubstappp \ U (Loop \)) \ undefg \ snd(usubstappp \ U \) \ undefg \ snd(usubstappp \ (fst(usubstappp \ U \)) \) \ undefg" (*using usubstappp_loop Loopo_undef*) proof assume a: "snd (usubstappp \ U (Loop \)) \ undefg" have fact: "fst(usubstappp \ U \) \ U" using usubst_taboos_mon by simp show "snd(usubstappp \ (fst(usubstappp \ U \)) \) \ undefg" using a usubstappp_loop Loopo_undef by simp then show "snd(usubstappp \ U \) \ undefg" using a usubstappp_loop Loopo_undef fact usubstappp_antimon by auto qed text \Lemma 13 of \<^url>\http://arxiv.org/abs/1902.07230\\ lemma usubst_taboos: "snd(usubstappp \ U \)\undefg \ fst(usubstappp \ U \) \ U \ BVG(the (snd(usubstappp \ U \)))" proof (induction \ arbitrary: U rule: game_induct) case (Game a) then show ?case by (cases "SGames \ a") (auto) next case (Assign x \) then show ?case using BVG_assign Assigno_undef by (metis (no_types, lifting) Assigno.elims BVG_assign_other fst_pair option.sel singletonI snd_pair subsetI union_or usubstappp.simps(2)) next case (ODE x \) then show ?case using BVG_ODE ODEo_undef by (metis (no_types, lifting) ODEo.elims Un_least fst_pair option.sel snd_conv sup.coboundedI2 usubst_taboos_mon usubstappp.simps(8)) next case (Test p) then show ?case using BVG_test Testo_undef usubst_taboos_mon by auto next case (Choice \ \) then show ?case (*using usubstappp.simps Product_Type.fst_conv Product_Type.snd_conv BVG_choice*) proof- from Choice have IHa: "fst(usubstappp \ U \) \ U \ BVG(the (snd(usubstappp \ U \)))" by (simp add: Choiceo_undef) from Choice have IHb: "fst(usubstappp \ U \) \ U \ BVG(the (snd(usubstappp \ U \)))" by (simp add: Choiceo_undef) have fact: "BVG(the (snd(usubstappp \ U \))) \ BVG(the (snd(usubstappp \ U \))) \ BVG(the (snd(usubstappp \ U (Choice \ \))))" using BVG_choice (*by (smt Choice.prems Choiceo.simps(1) Choiceo_undef option.collapse option.sel snd_pair usubstappp_choice)*) proof - have "Agame (the (snd (usubstappp \ U \)) \\ the (snd (usubstappp \ U \))) = Choiceo (snd (usubstappp \ U \)) (snd (usubstappp \ U \))" by (metis (no_types) Choice.prems Choiceo.simps(1) option.collapse usubstappp_choice_conv) then show ?thesis by (metis (no_types) BVG_choice Choice.prems Pair_inject option.collapse option.inject surjective_pairing usubstappp.simps(4)) qed from IHa and IHb have "fst(usubstappp \ U \) \ fst(usubstappp \ U \) \ U \ BVG(the (snd(usubstappp \ U \))) \ BVG(the (snd(usubstappp \ U \)))" by auto then have "fst(usubstappp \ U (Choice \ \)) \ U \ BVG(the (snd(usubstappp \ U \))) \ BVG(the (snd(usubstappp \ U \)))" using usubstappp.simps Let_def by auto then show "fst(usubstappp \ U (Choice \ \)) \ U \ BVG(the (snd(usubstappp \ U (Choice \ \))))" using usubstappp.simps fact by auto qed next case (Compose \ \) then show ?case proof- let ?V = "fst(usubstappp \ U \)" let ?W = "fst(usubstappp \ ?V \)" from Compose have IHa: "?V \ U \ BVG(the (snd(usubstappp \ U \)))" by (simp add: Composeo_undef) from Compose have IHb: "?W \ ?V \ BVG(the (snd(usubstappp \ ?V \)))" by (simp add: Composeo_undef) have fact: "BVG(the (snd(usubstappp \ U \))) \ BVG(the (snd(usubstappp \ ?V \))) \ BVG(the (snd(usubstappp \ U (Compose \ \))))" using usubstappp.simps BVG_compose (*by (smt Compose.prems Composeo.simps(1) Composeo_undef option.collapse option.sel snd_pair)*) proof - have f1: "\z. z = undefg \ Agame (the z) = z" using option.collapse by blast then have "Agame (the (snd (usubstappp \ U \)) ;; the (snd (usubstappp \ (fst (usubstappp \ U \)) \))) = snd (usubstappp \ U (\ ;; \))" using Compose.prems Composeo_undef by auto then show ?thesis using f1 by (metis (no_types) BVG_compose Compose.prems option.inject) qed have "?W \ U \ BVG(the (snd(usubstappp \ U \))) \ BVG(the (snd(usubstappp \ ?V \)))" using usubstappp.simps Let_def IHa IHb by auto then have "?W \ U \ BVG(the (snd(usubstappp \ U (Compose \ \))))" using fact by auto then show "fst(usubstappp \ U (Compose \ \)) \ U \ BVG(the (snd(usubstappp \ U (Compose \ \))))" using usubstappp.simps Let_def by simp qed next case (Loop \) then show ?case proof- let ?V = "fst(usubstappp \ U \)" let ?W = "fst(usubstappp \ ?V \)" from Loop have def\: "snd(usubstappp \ U \) \ undefg" using usubstappp_loop_conv by simp from Loop have IHdef: "?V \ U \ BVG(the (snd(usubstappp \ U \)))" using def\ usubstappp_loop[where \=\ and U=U and \=\] Loopo_undef[where \=\snd (usubstappp \ (fst (usubstappp \ U \)) \)\] by auto from Loop have IH: "?W \ ?V \ BVG(the (snd(usubstappp \ ?V \)))" by (simp add: Loopo_undef) then have Vfix: "?V \ BVG(the (snd(usubstappp \ ?V \)))" using usubstappp_det by (metis IHdef Loop.prems le_sup_iff usubstappp_loop_conv) then have "?V \ U \ BVG(the (snd(usubstappp \ U (Loop \))))" using usubstappp.simps Vfix IHdef BVG_loop usubst_taboos_mon usubstappp_loop_conv (*by (smt Loop.prems Loopo.simps(1) Un_mono option.collapse option.sel snd_pair sup.absorb_iff1)*) proof - (*sledgehammer*) have f1: "\z. z = undefg \ Agame (the z) = z" using option.collapse by blast have "snd (usubstappp \ U \) \ undefg \ snd (usubstappp \ (fst (usubstappp \ U \)) \) \ undefg" using Loop.prems usubstappp_loop_conv by blast then have "Agame (Loop (the (snd (usubstappp \ (fst (usubstappp \ U \)) \)))) = snd (usubstappp \ U (Loop \))" by force then show ?thesis using f1 by (metis (no_types) BVG_loop Loop.prems Vfix option.inject sup.absorb_iff1 sup.mono usubst_taboos_mon) qed then show "fst(usubstappp \ U (Loop \)) \ U \ BVG(the (snd(usubstappp \ U (Loop \))))" using usubstappp.simps Let_def by simp qed next case (Dual \) then show ?case (*using BVG_dual usubstappp.simps Let_def by auto*) proof- let ?V = "fst(usubstappp \ U \)" from Dual have IH: "?V \ U \ BVG(the (snd(usubstappp \ U \)))" by (simp add: Dualo_undef) have fact: "BVG(the (snd(usubstappp \ U \))) \ BVG(the (snd(usubstappp \ U (Dual \))))" using usubstappp.simps BVG_dual by (metis (no_types, lifting) Dual.prems Dualo.simps(1) Dualo.simps(2) option.collapse option.sel snd_pair) then have "?V \ U \ BVG(the (snd(usubstappp \ U (Dual \))))" using IH fact by auto then show "fst(usubstappp \ U (Dual \)) \ U \ BVG(the (snd(usubstappp \ U (Dual \))))" using usubstappp.simps Let_def by simp qed qed subsection \Substitution Adjoints\ text \Modified interpretation \repI I f d\ replaces the interpretation of constant function \f\ in the interpretation \I\ with \d\\ definition repc :: " interp \ ident \ real \ interp" where "repc I f d \ mkinterp((\c. if c = f then d else Consts I c), Funcs I, Preds I, Games I)" lemma repc_consts [simp]: "Consts (repc I f d) c = (if (c=f) then d else Consts I c)" unfolding repc_def by auto lemma repc_funcs [simp]: "Funcs (repc I f d) = Funcs I" unfolding repc_def by simp lemma repc_preds [simp]: "Preds (repc I f d) = Preds I" unfolding repc_def by simp lemma repc_games [simp]: "Games (repc I f d) = Games I" unfolding repc_def by (simp add: mon_mono) lemma adjoint_stays_mon: "mono (case SGames \ a of None \ Games I a | Some r \ \X. game_sem I r X)" using game_sem_mono game_sem.simps(1) by (metis disjE_realizer2 option.case_distrib) (*proof - have "\z p b i f. (mono (case z of None \ f | Some x \ game_sem i x) \ \ (case z of None \ b | Some x \ p x)) \ \ mono f" by (metis (no_types) disjE_realizer2 game_sem_mono) then show ?thesis by (metis (no_types) game_sem.simps(1) game_sem_mono option.case_distrib) qed*) text \adjoint interpretation \adjoint \ I \\ to \\\ of interpretation \I\ in state \\\\ definition adjoint:: "usubst \ (interp \ state \ interp)" where "adjoint \ I \ = mkinterp( (\f. (case SConst \ f of None \ Consts I f| Some r \ term_sem I r \)), (\f. (case SFuncs \ f of None \ Funcs I f | Some r \ \d. term_sem (repc I dotid d) r \)), (\p. (case SPreds \ p of None \ Preds I p | Some r \ \d. \\fml_sem (repc I dotid d) r)), (\a. (case SGames \ a of None \ Games I a | Some r \ \X. game_sem I r X)) )" paragraph \Simple Observations about Adjoints\ lemma adjoint_consts: "Consts (adjoint \ I \) f = term_sem I (case SConst \ f of Some r \ r | None \ Const f) \" unfolding adjoint_def by (cases "SConst \ f=None") (auto) lemma adjoint_funcs: "Funcs (adjoint \ I \) f = (case SFuncs \ f of None \ Funcs I f | Some r \ \d. term_sem (repc I dotid d) r \)" unfolding adjoint_def by auto lemma adjoint_funcs_match: "SFuncs \ f=Some r \ Funcs (adjoint \ I \) f = (\d. term_sem (repc I dotid d) r \)" using adjoint_funcs by auto lemma adjoint_funcs_skip: "SFuncs \ f=None \ Funcs (adjoint \ I \) f = Funcs I f" using adjoint_funcs by auto lemma adjoint_preds: "Preds (adjoint \ I \) p = (case SPreds \ p of None \ Preds I p | Some r \ \d. \\fml_sem (repc I dotid d) r)" unfolding adjoint_def by auto lemma adjoint_preds_skip: "SPreds \ p=None \ Preds (adjoint \ I \) p = Preds I p" using adjoint_preds by simp lemma adjoint_preds_match: "SPreds \ p=Some r \ Preds (adjoint \ I \) p = (\d. \\fml_sem (repc I dotid d) r)" using adjoint_preds by simp lemma adjoint_games [simp]: "Games (adjoint \ I \) a = (case SGames \ a of None \ Games I a | Some r \ \X. game_sem I r X)" unfolding adjoint_def using adjoint_stays_mon Games_mkinterp by simp lemma adjoint_dotsubstt: "adjoint (dotsubstt \) I \ = repc I dotid (term_sem I \ \)" (*unfolding adjoint_def dotsubstt_def adjoint_consts adjoint_funcs_skip adjoint_preds adjoint_games*) proof- let ?lhs = "adjoint (dotsubstt \) I \" let ?rhs = "repc I dotid (term_sem I \ \)" have feq: "Funcs ?lhs = Funcs ?rhs" using repc_funcs adjoint_funcs dotsubstt_def by auto moreover have peq: "Preds ?lhs = Preds ?rhs" using repc_preds adjoint_preds dotsubstt_def by auto moreover have geq: "Games ?lhs = Games ?rhs" using repc_games adjoint_games dotsubstt_def by auto moreover have ceq: "Consts ?lhs = Consts ?rhs" using repc_consts adjoint_consts dotsubstt_def by auto show ?thesis using mkinterp_eq[where I=\?lhs\ and J=\?rhs\] feq peq geq ceq by simp qed subsection \Uniform Substitution for Terms\ text \Lemma 15 of \<^url>\http://arxiv.org/abs/1902.07230\\ theorem usubst_term: "Uvariation \ \ U \ usubstappt \ U \\undeft \ term_sem I (the (usubstappt \ U \)) \ = term_sem (adjoint \ I \) \ \" proof- assume vaouter: "Uvariation \ \ U" assume defouter: "usubstappt \ U \\undeft" show "usubstappt \ U \\undeft \ term_sem I (the (usubstappt \ U \)) \ = term_sem (adjoint \ I \) \ \" for \ \ using vaouter proof (induction arbitrary: \ \ rule: usubstappt_induct) case (Var \ U x) then show ?case by simp next case (Number \ U r) then show ?case by simp next case (Const \ U f) then show ?case proof (cases "SConst \ f") case None then show ?thesis using adjoint_consts by (simp add: usappconst_def) next case (Some r) then have varcond: "FVT(r)\U={}" using Const usubstappt_const usubstappt_const_conv by (metis option.inject option.simps(3)) from Some have "term_sem I (the (usubstappt \ U (Const f))) \ = term_sem I r \" by (simp add: varcond) also have "... = term_sem I r \" using Const coincidence_term_cor[of \ \ U r] varcond by simp also have "... = Consts (adjoint \ I \) f" using Some adjoint_consts by simp also have "... = term_sem (adjoint \ I \) (Const f) \" by auto finally show "term_sem I (the (usubstappt \ U (Const f))) \ = term_sem (adjoint \ I \) (Const f) \" . qed next case (FuncMatch \ U f \) then have va: "Uvariation \ \ U" by simp then show ?case proof (cases "SFuncs \ f") case None from FuncMatch and None have IHsubterm: "term_sem I (the (usubstappt \ U \)) \ = term_sem (adjoint \ I \) \ \" using va by (simp add: FuncMatch.IH(1) usubstappt_func_conv) from None show ?thesis using usubstappt_func IHsubterm adjoint_funcs by (metis (no_types, lifting) FuncMatch.prems(1) option.case_eq_if option.sel term_sem.simps(4) usubstappt.simps(4)) next case (Some r) then have varcond: "FVT(r)\U={}" using FuncMatch usubstappt_func usubstappt_func2 usubstappt_func_conv by meson from FuncMatch have subdef: "usubstappt \ U \ \ undeft" using usubstappt_func_conv by auto from FuncMatch and Some have IHsubterm: "term_sem I (the (usubstappt \ U \)) \ = term_sem (adjoint \ I \) \ \" using va subdef by blast from FuncMatch and Some have IHsubsubst: "\\ \. Uvariation \ \ {} \ term_sem I (the (usubstappt (dotsubstt (the (usubstappt \ U \))) {} r)) \ = term_sem (adjoint (dotsubstt (the (usubstappt \ U \))) I \) r \" using subdef varcond by fastforce let ?d = "term_sem I (the (usubstappt \ U \)) \" have deq: "?d = term_sem (adjoint \ I \) \ \" by (rule IHsubterm) let ?dotIa = "adjoint (dotsubstt (the (usubstappt \ U \))) I \" from Some have "term_sem I (the (usubstappt \ U (Func f \))) \ = term_sem I (the (usubstappt (dotsubstt (the (usubstappt \ U \))) {} r)) \" using subdef varcond by force also have "... = term_sem ?dotIa r \" using IHsubsubst[where \=\ and \=\] Uvariation_empty by auto also have "... = term_sem (repc I dotid ?d) r \" using adjoint_dotsubstt[where \=\the (usubstappt \ U \)\ and I=I and \=\] by simp also have "... = term_sem (repc I dotid ?d) r \" using coincidence_term_cor[where \=\ and \'=\ and U=U and \=r and I=\repc I dotid ?d\] va varcond by simp (*I.^d\\\f(\)\ also have "... = term_sem ?dotIa r \" using coincidence_term_cor[of \ \ U r ?dotIa] uv varcond by simp*) also have "... = (Funcs (adjoint \ I \) f)(?d)" using adjoint_funcs_match Some by simp also have "... = (Funcs (adjoint \ I \) f)(term_sem (adjoint \ I \) \ \)" using deq by simp also have "... = term_sem (adjoint \ I \) (Func f \) \" by simp finally show "term_sem I (the (usubstappt \ U (Func f \))) \ = term_sem (adjoint \ I \) (Func f \) \" . qed next case (Plus \ U \ \) then show ?case using Pluso_undef by auto next case (Times \ U \ \) then show ?case using Timeso_undef by auto next case (Differential \ U \) from Differential have subdef: "usubstappt \ allvars \ \ undeft" using usubstappt_differential_conv by simp from Differential have IH: "\\. term_sem I (the (usubstappt \ allvars \)) \ = term_sem (adjoint \ I \) \ \" using subdef Uvariation_univ by blast (*by auto*) have "term_sem I (the (usubstappt \ U (Differential \))) \ = term_sem I (Differential (the (usubstappt \ allvars \))) \" using subdef by force also have "... = sum(\x. \(DVar x)*deriv(\X. term_sem I (the (usubstappt \ allvars \)) (repv \ (RVar x) X))(\(RVar x)))(allidents)" by simp also have "... = sum(\x. \(DVar x)*deriv(\X. term_sem (adjoint \ I \) \ (repv \ (RVar x) X))(\(RVar x)))(allidents)" using IH by auto also have "... = term_sem (adjoint \ I \) (Differential \) \" by simp finally show "term_sem I (the (usubstappt \ U (Differential \))) \ = term_sem (adjoint \ I \) (Differential \) \" . qed qed subsection \Uniform Substitution for Formulas and Games\ paragraph \Separately Prove Crucial Ingredient for the ODE Case of \usubst_fml_game\\ lemma same_ODE_same_sol: "(\\. Uvariation \ (F(0)) {RVar x,DVar x} \ term_sem I \ \ = term_sem J \ \) \ solves_ODE I F x \ = solves_ODE J F x \" using Uvariation_Vagree Vagree_def solves_ODE_def (*by (smt double_complement)*) proof- assume va: "\\. Uvariation \ (F(0)) {RVar x,DVar x} \ term_sem I \ \ = term_sem J \ \" then have va2: "\\. Uvariation \ (F(0)) {RVar x,DVar x} \ term_sem J \ \ = term_sem I \ \" by simp have one: "\I J \ \. (\\. Uvariation \ (F(0)) {RVar x,DVar x} \ term_sem I \ \ = term_sem J \ \) \ solves_ODE I F x \ \ solves_ODE J F x \" proof- fix I J \ \ assume vaflow: "\\. Uvariation \ (F(0)) {RVar x,DVar x} \ term_sem I \ \ = term_sem J \ \" assume sol: "solves_ODE I F x \" from vaflow and sol show "solves_ODE J F x \" unfolding solves_ODE_def using Uvariation_Vagree coincidence_term by (metis double_complement solves_Vagree sol) qed show "solves_ODE I F x \ = solves_ODE J F x \" using one[where \=\ and \=\, OF va] one[where \=\ and \=\, OF va2] by force qed lemma usubst_ode: assumes subdef: "usubstappt \ {RVar x,DVar x} \ \ undeft" shows "solves_ODE I F x (the (usubstappt \ {RVar x,DVar x} \)) = solves_ODE (adjoint \ I (F(0))) F x \" proof- have vaflow: "\F \ \. solves_ODE I F x \ \ Uvariation (F(\)) (F(0)) {RVar x,DVar x}" using solves_Vagree_trans by simp from subdef have IH: "\\. Uvariation \ (F(0)) {RVar x,DVar x} \ term_sem I (the (usubstappt \ {RVar x,DVar x} \)) \ = term_sem (adjoint \ I (F(0))) \ \" by (simp add: usubst_term) then show ?thesis using IH vaflow solves_ODE_def Uvariation_Vagree same_ODE_same_sol by blast qed lemma usubst_ode_ext: assumes uv: "Uvariation (F(0)) \ (U\{RVar x,DVar x})" assumes subdef: "usubstappt \ (U\{RVar x,DVar x}) \ \ undeft" shows "solves_ODE I F x (the (usubstappt \ (U\{RVar x,DVar x}) \)) = solves_ODE (adjoint \ I \) F x \" (*using usubst_ode usubstappt_det usubstappt_antimon Uvariation_Vagree Uvariation_mon *) proof- have vaflow1: "\F \ \. solves_ODE I F x (the (usubstappt \ (U\{RVar x,DVar x}) \)) \ Uvariation (F(\)) (F(0)) {RVar x,DVar x}" using solves_Vagree_trans by simp have vaflow2: "\F \ \. solves_ODE (adjoint \ I \) F x \ \ Uvariation (F(\)) (F(0)) {RVar x,DVar x}" using solves_Vagree_trans by simp from subdef have IH: "\\. Uvariation \ (F(0)) (U\{RVar x,DVar x}) \ term_sem I (the (usubstappt \ (U\{RVar x,DVar x}) \)) \ = term_sem (adjoint \ I (F(0))) \ \" using Uvariation_refl Uvariation_trans usubst_term by blast have l2r: "solves_ODE I F x (the (usubstappt \ (U\{RVar x,DVar x}) \)) \ solves_ODE (adjoint \ I \) F x \" using vaflow1 subdef same_ODE_same_sol Uvariation_trans usubst_term uv (*by (smt sup_commute sup_left_idem)*) proof - (*sledgehammer*) assume a1: "solves_ODE I F x (the (usubstappt \ (U \ {RVar x, DVar x}) \))" obtain rr :: "trm \ interp \ trm \ interp \ char \ (real \ variable \ real) \ variable \ real" where f2: "\x0 x1 x2 x3 x4 x5. (\v6. Uvariation v6 (x5 0) {RVar x4, DVar x4} \ term_sem x3 x2 v6 \ term_sem x1 x0 v6) = (Uvariation (rr x0 x1 x2 x3 x4 x5) (x5 0) {RVar x4, DVar x4} \ term_sem x3 x2 (rr x0 x1 x2 x3 x4 x5) \ term_sem x1 x0 (rr x0 x1 x2 x3 x4 x5))" by moura have f3: "Uvariation (F 0) \ (insert (RVar x) (U \ {DVar x}))" using uv by auto have f4: "{DVar x} \ {} \ {DVar x} = insert (DVar x) ({DVar x} \ {} \ {}) \ {RVar x} \ {DVar x} \ insert (RVar x) (U \ {DVar x}) = insert (RVar x) (U \ {DVar x})" by fastforce { assume "{DVar x} \ {} \ {DVar x} = insert (DVar x) ({DVar x} \ {} \ {}) \ Uvariation (rr (the (usubstappt \ (U \ {RVar x, DVar x}) \)) I \ (USubst.adjoint \ I \) x F) \ ({RVar x} \ {DVar x} \ insert (RVar x) (U \ {DVar x}))" then have "\ Uvariation (rr (the (usubstappt \ (U \ {RVar x, DVar x}) \)) I \ (USubst.adjoint \ I \) x F) (F 0) {RVar x, DVar x} \ term_sem (USubst.adjoint \ I \) \ (rr (the (usubstappt \ (U \ {RVar x, DVar x}) \)) I \ (USubst.adjoint \ I \) x F) = term_sem I (the (usubstappt \ (U \ {RVar x, DVar x}) \)) (rr (the (usubstappt \ (U \ {RVar x, DVar x}) \)) I \ (USubst.adjoint \ I \) x F)" using f4 subdef usubst_term by auto } then have "\ Uvariation (rr (the (usubstappt \ (U \ {RVar x, DVar x}) \)) I \ (USubst.adjoint \ I \) x F) (F 0) {RVar x, DVar x} \ term_sem (USubst.adjoint \ I \) \ (rr (the (usubstappt \ (U \ {RVar x, DVar x}) \)) I \ (USubst.adjoint \ I \) x F) = term_sem I (the (usubstappt \ (U \ {RVar x, DVar x}) \)) (rr (the (usubstappt \ (U \ {RVar x, DVar x}) \)) I \ (USubst.adjoint \ I \) x F)" using f3 by (metis (no_types) Uvariation_trans insert_is_Un) then show ?thesis using f2 a1 by (meson same_ODE_same_sol) qed have r2l: "solves_ODE (adjoint \ I \) F x \ \ solves_ODE I F x (the (usubstappt \ (U\{RVar x,DVar x}) \))" using vaflow2 subdef same_ODE_same_sol Uvariation_trans usubst_term uv (*by (smt sup_commute sup_left_idem)*) proof - (*sledgehammer*) assume a1: "solves_ODE (USubst.adjoint \ I \) F x \" obtain rr :: "trm \ interp \ trm \ interp \ char \ (real \ variable \ real) \ variable \ real" where "\x0 x1 x2 x3 x4 x5. (\v6. Uvariation v6 (x5 0) {RVar x4, DVar x4} \ term_sem x3 x2 v6 \ term_sem x1 x0 v6) = (Uvariation (rr x0 x1 x2 x3 x4 x5) (x5 0) {RVar x4, DVar x4} \ term_sem x3 x2 (rr x0 x1 x2 x3 x4 x5) \ term_sem x1 x0 (rr x0 x1 x2 x3 x4 x5))" by moura then have f2: "\f c i t ia ta. Uvariation (rr ta ia t i c f) (f 0) {RVar c, DVar c} \ term_sem i t (rr ta ia t i c f) \ term_sem ia ta (rr ta ia t i c f) \ solves_ODE i f c t = solves_ODE ia f c ta" by (meson same_ODE_same_sol) have f3: "Uvariation (F 0) \ ({RVar x, DVar x} \ U)" using uv by force have f4: "usubstappt \ ({RVar x, DVar x} \ U) \ \ undeft" using subdef by auto { assume "Uvariation (rr \ (USubst.adjoint \ I \) (the (usubstappt \ (U \ {RVar x, DVar x}) \)) I x F) \ ({RVar x, DVar x} \ ({RVar x, DVar x} \ U))" then have "\ Uvariation (rr \ (USubst.adjoint \ I \) (the (usubstappt \ (U \ {RVar x, DVar x}) \)) I x F) (F 0) {RVar x, DVar x} \ term_sem I (the (usubstappt \ (U \ {RVar x, DVar x}) \)) (rr \ (USubst.adjoint \ I \) (the (usubstappt \ (U \ {RVar x, DVar x}) \)) I x F) = term_sem (USubst.adjoint \ I \) \ (rr \ (USubst.adjoint \ I \) (the (usubstappt \ (U \ {RVar x, DVar x}) \)) I x F)" using f4 by (simp add: Un_commute usubst_term) } then show ?thesis using f3 f2 a1 by (meson Uvariation_trans) qed from l2r and r2l show ?thesis by auto qed lemma usubst_ode_ext2: assumes subdef: "usubstappt \ (U\{RVar x,DVar x}) \ \ undeft" assumes uv: "Uvariation (F(0)) \ (U\{RVar x,DVar x})" shows "solves_ODE I F x (the (usubstappt \ (U\{RVar x,DVar x}) \)) = solves_ODE (adjoint \ I \) F x \" using usubst_ode_ext subdef uv by blast paragraph \Separately Prove the Loop Case of \usubst_fml_game\\ lemma union_comm: "A\B=B\A" by auto definition loopfp\:: "game \ interp \ (state set \ state set)" where "loopfp\ \ I X = lfp(\Z. X \ game_sem I \ Z)" lemma usubst_game_loop: assumes uv: "Uvariation \ \ U" and IH\rec: "\\ \ X. Uvariation \ \ (fst(usubstappp \ U \)) \ snd (usubstappp \ (fst(usubstappp \ U \)) \)\undefg \ (\ \ game_sem I (the (snd (usubstappp \ (fst(usubstappp \ U \)) \))) X) = (\ \ game_sem (adjoint \ I \) \ X)" shows "snd (usubstappp \ U (Loop \))\undefg \ (\ \ game_sem I (the (snd (usubstappp \ U (Loop \)))) X) = (\ \ game_sem (adjoint \ I \) (Loop \) X)" proof- assume def: "snd (usubstappp \ U (Loop \))\undefg" have loopfix: "\\ I X. loopfp\ \ I X = game_sem I (Loop \) X" unfolding loopfp\_def using game_sem_loop by metis let ?\\loopoff = "the (snd (usubstappp \ U (Loop \)))" let ?\\ = "the (snd(usubstappp \ (fst(usubstappp \ U \)) \))" let ?\\loop = "Loop ?\\" have loopform: "?\\loopoff = ?\\loop" using usubstappp_loop by (metis (mono_tags, lifting) Loopo.simps(1) def option.exhaust_sel option.inject snd_conv usubstappp_loop_conv) let ?\ = "loopfp\ ?\\loop I" let ?\ = "loopfp\ \ (adjoint \ I \)" let ?V = "fst(usubstappp \ U \)" have fact1: "\V. snd(usubstappp \ V \)\undefg \ fst(usubstappp \ V \) \ BVG(the (snd(usubstappp \ V \)))" using usubst_taboos by simp have fact2: "\V W. snd(usubstappp \ V \)\undefg \ snd(usubstappp \ W \)\undefg \ (fst(usubstappp \ V \) \ BVG(the (snd(usubstappp \ W \))))" using fact1 usubst_taboos usubstappp_det by metis have VgeqBV: "?V \ BVG(?\\)" using usubst_taboos fact2 def usubstappp_loop_conv by simp have uvV: "Vagree \ \ (-?V)" using uv by (metis Uvariation_Vagree Uvariation_mon double_compl usubst_taboos_mon) have \eq: "?\(X) = game_sem I ?\\loop X" using loopfix by auto have \eq: "?\(X) = game_sem (adjoint \ I \) (Loop \) X" using loopfix by auto have \is\: "selectlike (?\(X)) \ (-?V)= selectlike (?\(X)) \ (-?V)" proof- let ?f = "\Z. X \ game_sem I ?\\ Z" let ?g = "\Y. X \ game_sem (adjoint \ I \) \ Y" let ?R = "\Z Y. selectlike Z \ (-?V) = selectlike Y \ (-?V)" have "?R (lfp ?f) (lfp ?g)" proof (induction rule: lfp_lockstep_induct[where f=\?f\ and g=\?g\ and R=\?R\]) case monof then show ?case using game_sem_loop_fixpoint_mono by simp next case monog then show ?case using game_sem_loop_fixpoint_mono by simp next case (step A B) then have IHfp: "selectlike A \ (-?V) = selectlike B \ (-?V)" by simp show "selectlike (X \ game_sem I ?\\ A) \ (-?V) = selectlike (X \ game_sem (adjoint \ I \) \ B) \ (-?V)" proof (rule selectlike_equal_cocond_corule (*[where \=\ and V=\?V\ and X=\X \ game_sem I ?\\ A\ and Y=\X \ game_sem (adjoint \ I \) \ B\]*)) fix \ assume muvar: "Uvariation \ \ ?V" have forw: "(\ \ X \ game_sem I ?\\ A) = (\ \ X \ game_sem I ?\\ (selectlike A \ (-BVG(?\\))))" using boundeffect by auto have "(\ \ X \ game_sem (adjoint \ I \) \ B) = (\ \ X \ game_sem I ?\\ B)" using IH\rec[OF muvar] using def usubstappp_loop_conv by auto also have "... = (\ \ X \ game_sem I ?\\ (selectlike B \ (-BVG(?\\))))" using boundeffect by simp finally have backw: "(\ \ X \ game_sem (adjoint \ I \) \ B) = (\ \ X \ game_sem I ?\\ (selectlike B \ (-BVG(?\\))))" . have samewin: "selectlike A \ (-BVG(?\\)) = selectlike B \ (-BVG(?\\))" using IHfp selectlike_antimon VgeqBV muvar Uvariation_trans selectlike_equal_cocond (*by (smt le_iff_sup)*) proof - have "Vagree \ \ (- fst (usubstappp \ U \))" by (metis Uvariation_Vagree double_complement muvar) then have "selectlike A \ (- fst (usubstappp \ U \)) = selectlike B \ (- fst (usubstappp \ U \))" using IHfp selectlike_Vagree by presburger then show ?thesis by (metis (no_types) Compl_subset_Compl_iff VgeqBV selectlike_compose sup.absorb_iff2) qed from forw and backw show "(\ \ X \ game_sem I ?\\ A) = (\ \ X \ game_sem (adjoint \ I \) \ B)" using samewin by auto qed next case (union M) then show ?case using selectlike_Sup[where \=\ and V=\-?V\] fst_proj_def snd_proj_def by (simp) (blast) qed then show ?thesis using \eq loopfix loopfp\_def by auto qed show ?thesis using \eq \eq \is\ similar_selectlike_mem[OF uvV] by (metis loopform) qed lemma usubst_fml_game: assumes vaouter: "Uvariation \ \ U" shows "usubstappf \ U \\undeff \ (\ \ fml_sem I (the (usubstappf \ U \))) = (\ \ fml_sem (adjoint \ I \) \)" and "snd (usubstappp \ U \)\undefg \ (\ \ game_sem I (the (snd (usubstappp \ U \))) X) = (\ \ game_sem (adjoint \ I \) \ X)" proof- show "usubstappf \ U \\undeff \ (\ \ fml_sem I (the (usubstappf \ U \))) = (\ \ fml_sem (adjoint \ I \) \)" and "snd (usubstappp \ U \)\undefg \ (\ \ game_sem I (the (snd (usubstappp \ U \))) X) = (\ \ game_sem (adjoint \ I \) \ X)" for \ \ \ using vaouter proof (induction \ and \ arbitrary: \ \ and \ \ X rule: usubstappfp_induct) case (Pred \ U p \) then have va: "Uvariation \ \ U" by simp then show ?case proof (cases "SPreds \ p") case None then show ?thesis using usubst_term[OF va] adjoint_preds_skip (*by (smt Pred.prems(1) fml_sem.simps(1) mem_Collect_eq option.case_eq_if option.sel usubstappf.simps(1))*) proof - (*sledgehammer*) have "\p V c t. usubstappf p V (Pred c t) = (if usubstappt p V t = undeft then undeff else case SPreds p c of None \ Afml (Pred c (the (usubstappt p V t))) | Some f \ if FVF f \ V = {} then usubstappf (dotsubstt (the (usubstappt p V t))) {} f else undeff)" by (simp add: option.case_eq_if) then have f1: "\p V c t. if usubstappt p V t = undeft then usubstappf p V (Pred c t) = undeff else usubstappf p V (Pred c t) = (case SPreds p c of None \ Afml (Pred c (the (usubstappt p V t))) | Some f \ if FVF f \ V = {} then usubstappf (dotsubstt (the (usubstappt p V t))) {} f else undeff)" by presburger then have "usubstappt \ U \ \ undeft" by (meson Pred.prems(1)) then have f2: "usubstappf \ U (Pred p \) = Afml (Pred p (the (usubstappt \ U \)))" using None by force have "usubstappt \ U \ \ undeft" using f1 by (meson Pred.prems(1)) then show ?thesis using f2 by (simp add: None \\\ \ I. usubstappt \ U \ \ undeft \ term_sem I (the (usubstappt \ U \)) \ = term_sem (USubst.adjoint \ I \) \ \\ adjoint_preds_skip) qed next case (Some r) then have varcond: "FVF(r)\U={}" using Pred usubstappf_pred usubstappf_pred2 usubstappf_pred_conv by meson from Pred have subdef: "usubstappt \ U \ \ undeft" using usubstappf_pred_conv by auto from Pred and Some have IHsubsubst: "\\ \. Uvariation \ \ {} \ (\ \ fml_sem I (the (usubstappf (dotsubstt (the (usubstappt \ U \))) {} r))) = (\ \ fml_sem (adjoint (dotsubstt (the (usubstappt \ U \))) I \) r)" using subdef varcond by fastforce let ?d = "term_sem I (the (usubstappt \ U \)) \" have deq: "?d = term_sem (adjoint \ I \) \ \" by (rule usubst_term[OF va subdef]) let ?dotIa = "adjoint (dotsubstt (the (usubstappt \ U \))) I \" from Some have "(\\fml_sem I (the (usubstappf \ U (Pred p \)))) = (\\fml_sem I (the (usubstappf (dotsubstt (the (usubstappt \ U \))) {} r)))" using subdef varcond by force also have "... = (\\fml_sem ?dotIa r)" using IHsubsubst[where \=\ and \=\] Uvariation_empty by auto also have "... = (\\fml_sem (repc I dotid ?d) r)" using adjoint_dotsubstt[where \=\the (usubstappt \ U \)\ and I=I and \=\] by simp also have "... = (\\fml_sem (repc I dotid ?d) r)" using coincidence_formula_cor[where \=\ and \'=\ and U=U and \=r and I=\repc I dotid ?d\] va varcond by simp (*I.^d\\\f(\)\ also have "... = term_sem ?dotIa r \" using coincidence_term_cor[of \ \ U r ?dotIa] uv varcond by simp*) also have "... = (Preds (adjoint \ I \) p)(?d)" using adjoint_preds_match Some by simp also have "... = (Preds (adjoint \ I \) p)(term_sem (adjoint \ I \) \ \)" using deq by simp also have "... = (\\fml_sem (adjoint \ I \) (Pred p \))" by simp finally show "(\\fml_sem I (the (usubstappf \ U (Pred p \)))) = (\\fml_sem (adjoint \ I \) (Pred p \))" . qed next case (Geq \ U \ \) (* then show ?case using usubst_term usubstappf_geq usubstappf_geq_conv by (smt fml_sem.simps(2) mem_Collect_eq option.sel)*) then have def1: "usubstappt \ U \ \ undeft" using usubstappf_geq_conv by simp moreover have def2: "usubstappt \ U \ \ undeft" using usubstappf_geq_conv Geq.prems(1) by blast show ?case using usubst_term[OF \Uvariation \ \ U\, OF def1] usubst_term[OF \Uvariation \ \ U\,OF def2] usubstappf_geqr[OF \usubstappf \ U (Geq \ \) \ undeff\] by force next case (Not \ U \) then show ?case using Noto_undef by auto next case (And \ U \ \) then show ?case using Ando_undef by auto next case (Exists \ U x \) then have IH: "\\ \. Uvariation \ \ (U\{x}) \ (\ \ fml_sem I (the (usubstappf \ (U\{x}) \))) = (\ \ fml_sem (adjoint \ I \) \)" by force from Exists have "Uvariation \ \ U" by simp (*from Exists have subdef: "usubstappt \ (U\{x}) \ \ undeft" by auto*) then have Uvar: "\d. Uvariation (repv \ x d) \ (U\{x})" using Uvariation_repv Uvariation_trans Uvariation_sym using Exists.prems(2) Uvariation_def by auto have "(\\fml_sem I (the (usubstappf \ U (Exists x \)))) = (\\fml_sem I (Exists x (the (usubstappf \ (U\{x}) \))))" using usubstappf_exists Exists.prems(1) by fastforce also have "... = (\d. (repv \ x d)\fml_sem I (the (usubstappf \ (U\{x}) \)))" by simp also have "... = (\d. (repv \ x d)\fml_sem (adjoint \ I \) \)" using IH Uvar by auto also have "... = (\\fml_sem (adjoint \ I \) (Exists x \))" by auto finally have "(\\fml_sem I (the (usubstappf \ U (Exists x \)))) = (\\fml_sem (adjoint \ I \) (Exists x \))" . from this show ?case by simp next case (Diamond \ U \ \) let ?V = "fst(usubstappp \ U \)" from Diamond have IH\: "\X. Uvariation \ \ U \ (\ \ game_sem I (the (snd (usubstappp \ U \))) X) = (\ \ game_sem (adjoint \ I \) \ X)" by fastforce from Diamond have IH\: "\\ \. Uvariation \ \ (fst(usubstappp \ U \)) \ (\ \ fml_sem I (the (usubstappf \ (fst(usubstappp \ U \)) \))) = (\ \ fml_sem (adjoint \ I \) \)" by (simp add: Diamondo_undef) from Diamond have uv: "Uvariation \ \ U" by simp have "(\ \ fml_sem I (the (usubstappf \ U (Diamond \ \)))) = (\ \ fml_sem I (let V\ = usubstappp \ U \ in Diamond (the (snd V\)) (the (usubstappf \ (fst V\) \))))" by (metis Diamond.prems(1) Diamondo.elims option.sel usubstappf.simps(6)) also have "... = (\ \ fml_sem I (Diamond (the (snd(usubstappp \ U \))) (the (usubstappf \ (fst(usubstappp \ U \)) \))))" by simp also have "... = (\ \ game_sem I (the (snd(usubstappp \ U \))) (fml_sem I (the (usubstappf \ (fst(usubstappp \ U \)) \))))" by simp also have "... = (\ \ game_sem I (the (snd(usubstappp \ U \))) (selectlike (fml_sem I (the (usubstappf \ (fst(usubstappp \ U \)) \))) \ (-BVG(the (snd(usubstappp \ U \))))))" using boundeffect by auto finally have forw: "(\ \ fml_sem I (the (usubstappf \ U (Diamond \ \)))) = (\ \ game_sem I (the (snd(usubstappp \ U \))) (selectlike (fml_sem I (the (usubstappf \ (fst(usubstappp \ U \)) \))) \ (-BVG(the (snd(usubstappp \ U \))))))" . have "(\ \ fml_sem (adjoint \ I \) (Diamond \ \)) = (\ \ game_sem (adjoint \ I \) \ (fml_sem (adjoint \ I \) \))" by simp also have "... = (\ \ game_sem I (the (snd(usubstappp \ U \))) (fml_sem (adjoint \ I \) \))" using IH\ uv by simp also have "... = (\ \ game_sem I (the (snd(usubstappp \ U \))) (selectlike (fml_sem (adjoint \ I \) \) \ (-BVG(the (snd(usubstappp \ U \))))))" using boundeffect by auto finally have backw: "(\ \ fml_sem (adjoint \ I \) (Diamond \ \)) = (\ \ game_sem I (the (snd(usubstappp \ U \))) (selectlike (fml_sem (adjoint \ I \) \) \ (-BVG(the (snd(usubstappp \ U \))))))" . have samewin: "selectlike (fml_sem I (the (usubstappf \ (fst(usubstappp \ U \)) \))) \ (-BVG(the (snd(usubstappp \ U \)))) = selectlike (fml_sem (adjoint \ I \) \) \ (-BVG(the (snd(usubstappp \ U \))))" proof (rule selectlike_equal_cocond_corule) fix \ assume muvar: "Uvariation \ \ (BVG(the (snd(usubstappp \ U \))))" have U\\: "Uvariation \ \ ?V" using muvar uv Uvariation_trans union_comm usubst_taboos Uvariation_mon (*by (smt Diamond.prems(1) Diamondo.simps(2) usubstappf.simps(6))*) proof- have U\\: "Uvariation \ \ (BVG(the (snd(usubstappp \ U \))))" by (rule muvar) have U\\: "Uvariation \ \ U" by (rule uv) have "Uvariation \ \ (U \ BVG(the (snd(usubstappp \ U \))))" using Uvariation_trans[OF U\\ U\\] union_comm by (rule HOL.back_subst) thus ?thesis using usubst_taboos Uvariation_mon by (metis (mono_tags, lifting) Diamond.prems(1) Diamondo_undef Uvariation_mon usubst_taboos usubstappf.simps(6)) qed have "(\ \ fml_sem (adjoint \ I \) \) = (\ \ fml_sem I (the (usubstappf \ (fst(usubstappp \ U \)) \)))" using muvar Uvariation_trans uv IH\ boundeffect Uvariation_mon usubst_taboos U\\ by auto then show "(\ \ fml_sem I (the (usubstappf \ (fst (usubstappp \ U \)) \))) = (\ \ fml_sem (adjoint \ I \) \)" by simp qed from forw and backw show "(\ \ fml_sem I (the (usubstappf \ U (Diamond \ \)))) = (\ \ fml_sem (adjoint \ I \) (Diamond \ \))" using samewin by auto next (* games *) case (Game \ U a) then show ?case using adjoint_games usubstappp_game by (cases "SGames \ a") auto next case (Assign \ U x \) then show ?case using usubst_term Assigno_undef (*by (smt Assigno.elims game_sem.simps(2) mem_Collect_eq option.sel snd_pair usubstappp.simps(2))*) proof - (*sledgehammer*) have f1: "usubstappt \ U \ \ undeft" using Assign.prems(1) Assigno_undef by auto { assume "repv \ x (term_sem (USubst.adjoint \ I \) \ \) \ X" then have "repv \ x (term_sem I (the (usubstappt \ U \)) \) \ X" using f1 Assign.prems(2) usubst_term by auto then have "repv \ x (term_sem (USubst.adjoint \ I \) \ \) \ X \ (\ \ game_sem I (the (snd (usubstappp \ U (x := \)))) X) = (\ \ game_sem (USubst.adjoint \ I \) (x := \) X)" using f1 by force } moreover { assume "repv \ x (term_sem (USubst.adjoint \ I \) \ \) \ X" then have "repv \ x (term_sem I (the (usubstappt \ U \)) \) \ X \ repv \ x (term_sem (USubst.adjoint \ I \) \ \) \ X" using f1 Assign.prems(2) usubst_term by presburger then have ?thesis using f1 by force } ultimately show ?thesis by fastforce qed next case (Test \ U \) then show ?case using Testo_undef by auto next case (Choice \ U \ \) from Choice have IH\: "\X. Uvariation \ \ U \ (\ \ game_sem I (the (snd (usubstappp \ U \))) X) = (\ \ game_sem (adjoint \ I \) \ X)" by (simp add: Choiceo_undef) from Choice have IH\: "\X. Uvariation \ \ U \ (\ \ game_sem I (the (snd (usubstappp \ U \))) X) = (\ \ game_sem (adjoint \ I \) \ X)" by (simp add: Choiceo_undef) from Choice have uv: "Uvariation \ \ U" by simp show ?case using IH\ IH\ uv (*by (smt Choice.prems(1) Choiceo.elims game_sem.simps(4) option.sel snd_pair union_or usubstappp_choice) *) proof - have f1: "Agame (the (snd (usubstappp \ U \))) = snd (usubstappp \ U \)" by (meson Choice(3) option.collapse usubstappp_choice_conv) have "Agame (the (snd (usubstappp \ U \))) = snd (usubstappp \ U \)" by (meson Choice(3) option.collapse usubstappp_choice_conv) then have "snd (usubstappp \ U (\ \\ \)) = Agame (the (snd (usubstappp \ U \)) \\ the (snd (usubstappp \ U \)))" using f1 by (metis Choiceo.simps(1) snd_conv usubstappp.simps(4)) then have f2: "game_sem I (the (snd (usubstappp \ U \))) X \ game_sem I (the (snd (usubstappp \ U \))) X = game_sem I (the (snd (usubstappp \ U (\ \\ \)))) X" by simp moreover { assume "\ \ game_sem I (the (snd (usubstappp \ U \))) X" have "(\ \ game_sem I (the (snd (usubstappp \ U (\ \\ \)))) X) = (\ \ game_sem (adjoint \ I \) (\ \\ \) X) \ \ \ game_sem I (the (snd (usubstappp \ U \))) X \ \ \ game_sem (adjoint \ I \) (\ \\ \) X" using f2 Choice(4) IH\ IH\ by auto then have "(\ \ game_sem I (the (snd (usubstappp \ U (\ \\ \)))) X) \ (\ \ game_sem (adjoint \ I \) (\ \\ \) X)" using f2 Choice(4) IH\ by auto } ultimately show ?thesis using Choice(4) IH\ by auto qed next case (Compose \ U \ \) let ?V = "fst(usubstappp \ U \)" from Compose have IH\: "\X. Uvariation \ \ U \ (\ \ game_sem I (the (snd (usubstappp \ U \))) X) = (\ \ game_sem (adjoint \ I \) \ X)" by (simp add: Composeo_undef) from Compose have IH\: "\\ \ X. Uvariation \ \ ?V \ (\ \ game_sem I (the (snd (usubstappp \ ?V \))) X) = (\ \ game_sem (adjoint \ I \) \ X)" by (simp add: Composeo_undef) from Compose have uv: "Uvariation \ \ U" by simp have "(\ \ game_sem I (the (snd (usubstappp \ U (Compose \ \)))) X) = (\ \ game_sem I (Compose (the (snd(usubstappp \ U \))) (the (snd(usubstappp \ ?V \)))) X)" by (metis (no_types, lifting) Compose.prems(1) Composeo.elims option.sel snd_pair usubstappp.simps(5)) also have "... = (\ \ game_sem I (the (snd(usubstappp \ U \))) (game_sem I (the (snd(usubstappp \ ?V \))) X))" by simp also have "... = (\ \ game_sem I (the (snd(usubstappp \ U \))) (selectlike (game_sem I (the (snd(usubstappp \ ?V \))) X) \ (-BVG(the(snd(usubstappp \ U \))))))" using boundeffect by auto finally have forw: "(\ \ game_sem I (the (snd (usubstappp \ U (Compose \ \)))) X) = (\ \ game_sem I (the (snd(usubstappp \ U \))) (selectlike (game_sem I (the (snd(usubstappp \ ?V \))) X) \ (-BVG(the(snd(usubstappp \ U \))))))" . have "(\ \ game_sem (adjoint \ I \) (Compose \ \) X) = (\ \ game_sem (adjoint \ I \) \ ((game_sem (adjoint \ I \) \) X))" by simp also have "... = (\ \ game_sem I (the (snd(usubstappp \ U \))) ((game_sem (adjoint \ I \) \) X))" using IH\[OF uv] by auto also have "... = (\ \ game_sem I (the (snd(usubstappp \ U \))) (selectlike ((game_sem (adjoint \ I \) \) X) \ (-BVG(the(snd(usubstappp \ U \))))))" using boundeffect by auto finally have backw: "(\ \ game_sem (adjoint \ I \) (Compose \ \) X) = (\ \ game_sem I (the (snd(usubstappp \ U \))) (selectlike ((game_sem (adjoint \ I \) \) X) \ (-BVG(the(snd(usubstappp \ U \))))))" . have samewin: "selectlike (game_sem I (the (snd(usubstappp \ ?V \))) X) \ (-BVG(the(snd(usubstappp \ U \)))) = selectlike ((game_sem (adjoint \ I \) \) X) \ (-BVG(the(snd(usubstappp \ U \))))" proof (rule selectlike_equal_cocond_corule) fix \ assume muvar: "Uvariation \ \ (BVG(the(snd(usubstappp \ U \))))" have U\\: "Uvariation \ \ ?V" using muvar uv Uvariation_trans union_comm usubst_taboos Uvariation_mon (*by (smt Compose.prems(1) Composeo_undef snd_pair usubstappp.simps(5))*) proof - have "Uvariation \ \ (BVG (the (snd (usubstappp \ U \))) \ U)" by (meson Uvariation_trans muvar uv) then show ?thesis using Uvariation_mon union_comm usubst_taboos by (metis (no_types, lifting) Compose.prems(1) Composeo_undef Pair_inject prod.collapse usubstappp_compose) (*by (metis (no_types) Uvariation_mon union_comm usubst_taboos)*) qed have "(\ \ game_sem I (the(snd(usubstappp \ ?V \))) X) = (\ \ game_sem (adjoint \ I \) \ X)" using muvar Uvariation_trans uv IH\ boundeffect Uvariation_mon usubst_taboos U\\ by auto then show "(\ \ game_sem I (the(snd(usubstappp \ ?V \))) X) = (\ \ game_sem (adjoint \ I \) \ X)" by simp qed from forw and backw show "(\ \ game_sem I (the(snd (usubstappp \ U (Compose \ \)))) X) = (\ \ game_sem (adjoint \ I \) (Compose \ \) X)" using samewin by auto next case (Loop \ U \) let ?V = "fst(usubstappp \ U \)" from Loop have selfdef: "snd (usubstappp \ U (Loop \)) \ undefg" by auto from Loop have IH\rec: "\\ \ X. Uvariation \ \ ?V \ (\ \ game_sem I (the (snd (usubstappp \ ?V \))) X) = (\ \ game_sem (adjoint \ I \) \ X)" by fastforce from Loop have uv: "Uvariation \ \ U" by simp show "(\ \ game_sem I (the (snd (usubstappp \ U (Loop \)))) X) = (\ \ game_sem (adjoint \ I \) (Loop \) X)" using usubst_game_loop IH\rec Loop.prems(2) selfdef by blast (*by (rule usubst_game_loop[OF uv (*IH\*) IH\rec])*) next case (Dual \ U \) from Dual have IH\: "\X. Uvariation \ \ U \ (\ \ game_sem I (the (snd (usubstappp \ U \))) X) = (\ \ game_sem (adjoint \ I \) \ X)" by force from Dual have uv: "Uvariation \ \ U" by simp from Dual have def: "snd (usubstappp \ U (\^d)) \ undefg" by simp (*show ?case using IH\[OF uv] by (smt Compl_iff Dual.prems(1) Dualo.elims game_sem.simps(7) option.sel snd_pair usubstappp.simps(7))*) have "(\ \ -game_sem I (the (snd (usubstappp \ U \))) (-X)) = (\ \ -game_sem (adjoint \ I \) \ (-X))" using IH\[OF uv] by simp then have "(\ \ game_sem I ((the (snd (usubstappp \ U \)))^d) X) = (\ \ game_sem (adjoint \ I \) (\^d) X)" using game_sem.simps(7) by auto then show ?case using usubstappp_dual Dualo_undef proof - have "\\ V \. snd (usubstappp \ U (\^d)) = Dualo (snd (usubstappp \ U \))" by simp then have "snd (usubstappp \ U \) \ undefg" using Dualo_undef def by presburger then show ?thesis using \(\ \ game_sem I ((the (snd (usubstappp \ U \)))^d) X) = (\ \ game_sem (USubst.adjoint \ I \) (\^d) X)\ by force qed next case (ODE \ U x \) then have va: "Uvariation \ \ U" by simp from ODE have subdef: "usubstappt \ (U\{RVar x,DVar x}) \ \ undeft" by (simp add: ODEo_undef) from ODE have IH: "term_sem I (the (usubstappt \ (U\{RVar x,DVar x}) \)) \ = term_sem (adjoint \ I \) \ \" using va by (metis ODEo_undef fst_pair snd_conv usubst_taboos_mon usubst_term usubstappp.simps(8) usubstappt_antimon) have "(\ \ game_sem I (the (snd (usubstappp \ U (ODE x \)))) X) = (\ \ game_sem I (the (ODEo x (usubstappt \ (U\{RVar x,DVar x}) \))) X)" by simp also have "... = (\ \ game_sem I (ODE x (the (usubstappt \ (U\{RVar x,DVar x}) \))) X)" using subdef by force also have "... = (\F T. Vagree \ (F(0)) (-{DVar x}) \ F(T) \ X \ solves_ODE I F x (the (usubstappt \ (U\{RVar x,DVar x}) \)))" by simp also have "... = (\F T. Uvariation \ (F(0)) {DVar x} \ F(T) \ X \ solves_ODE I F x (the (usubstappt \ (U\{RVar x,DVar x}) \)))" using Uvariation_Vagree by (metis double_compl) also have "... = (\F T. Uvariation \ (F(0)) {DVar x} \ F(T) \ X \ solves_ODE (adjoint \ I \) F x \)" using usubst_ode_ext2[OF subdef] va solves_Vagree_trans Uvariation_trans Uvariation_sym_rel Uvariation_mon by (meson subset_insertI) also have "... = (\F T. Vagree \ (F(0)) (-{DVar x}) \ F(T) \ X \ solves_ODE (adjoint \ I \) F x \)" using Uvariation_Vagree by (metis double_compl) also have "... = (\ \ game_sem (adjoint \ I \) (ODE x \) X)" using solves_ODE_def by simp finally show "(\ \ game_sem I (the (snd (usubstappp \ U (ODE x \)))) X) = (\ \ game_sem (adjoint \ I \) (ODE x \) X)" . qed qed text \Lemma 16 of \<^url>\http://arxiv.org/abs/1902.07230\\ theorem usubst_fml: "Uvariation \ \ U \ usubstappf \ U \ \ undeff \ (\ \ fml_sem I (the (usubstappf \ U \))) = (\ \ fml_sem (adjoint \ I \) \)" using usubst_fml_game by simp text \Lemma 17 of \<^url>\http://arxiv.org/abs/1902.07230\\ theorem usubst_game: "Uvariation \ \ U \ snd (usubstappp \ U \) \ undefg \ (\ \ game_sem I (the (snd (usubstappp \ U \))) X) = (\ \ game_sem (adjoint \ I \) \ X)" using usubst_fml_game by simp subsection \Soundness of Uniform Substitution of Formulas\ abbreviation usubsta:: "usubst \ fml \ fmlo" where "usubsta \ \ \ usubstappf \ {} \" text \Theorem 18 of \<^url>\http://arxiv.org/abs/1902.07230\\ theorem usubst_sound: "usubsta \ \ \ undeff \ valid \ \ valid (the (usubsta \ \))" proof- assume def: "usubsta \ \ \ undeff" assume prem: "valid \" from prem have premc: "\I \. \ \ fml_sem I \" using valid_def by auto show "valid (the (usubsta \ \))" unfolding valid_def proof (clarify) fix I \ have "(\ \ fml_sem I (the (usubsta \ \))) = (\ \ fml_sem (adjoint \ I \) \)" using usubst_fml by (simp add: usubst_fml def) also have "... = True" using premc by simp finally have "(\ \ fml_sem I (the (usubsta \ \))) = True" . from this show "\ \ fml_sem I (the (usubstappf \ {} \))" by simp qed qed subsection \Soundness of Uniform Substitution of Rules\ text \Uniform Substitution applied to a rule or inference\ definition usubstr:: "usubst \ inference \ inference option" where "usubstr \ R \ if (usubstappf \ allvars (snd R) \ undeff \ (\\\set (fst R). usubstappf \ allvars \ \ undeff)) then Some(map(the o (usubstappf \ allvars))(fst R), the (usubstappf \ allvars (snd R))) else None" text \Simple observations about applying uniform substitutions to a rule\ lemma usubstr_conv: "usubstr \ R \ None \ usubstappf \ allvars (snd R) \ undeff \ (\\\set (fst R). usubstappf \ allvars \ \ undeff)" by (metis usubstr_def) lemma usubstr_union_undef: "(usubstr \ ((append A B), C) \ None) = (usubstr \ (A, C) \ None \ usubstr \ (B, C) \ None)" using usubstr_def by auto lemma usubstr_union_undef2: "(usubstr \ ((append A B), C) \ None) \ (usubstr \ (A, C) \ None \ usubstr \ (B, C) \ None)" using usubstr_union_undef by blast lemma usubstr_cons_undef: "(usubstr \ ((Cons A B), C) \ None) = (usubstr \ ([A], C) \ None \ usubstr \ (B, C) \ None)" using usubstr_def by auto lemma usubstr_cons_undef2: "(usubstr \ ((Cons A B), C) \ None) \ (usubstr \ ([A], C) \ None \ usubstr \ (B, C) \ None)" using usubstr_cons_undef by blast lemma usubstr_cons: "(usubstr \ ((Cons A B), C) \ None) \ the (usubstr \ ((Cons A B), C)) = (Cons (the (usubstappf \ allvars A)) (fst (the (usubstr \ (B, C)))), snd (the (usubstr \ ([A], C))))" using usubstr_union_undef map_cons usubstr_def proof- assume def: "(usubstr \ ((Cons A B), C) \ None)" let ?R = "((Cons A B), C)" have "the (usubstr \ ?R) = (map(the o (usubstappf \ allvars))(fst ?R) , the (usubstappf \ allvars (snd ?R)))" using def usubstr_def by (metis option.sel) also have "... = (Cons (the (usubstappf \ allvars A)) (map(the o (usubstappf \ allvars))(B)) , the (usubstappf \ allvars (snd ?R)))" using map_cons by auto also have "... = (Cons (the (usubstappf \ allvars A)) (fst (the (usubstr \ (B, C)))) , the (usubstappf \ allvars (snd ?R)))" using usubstr_cons_undef2[OF def] usubstr_def by (metis (no_types, lifting) fst_conv option.sel) also have "... = (Cons (the (usubstappf \ allvars A)) (fst (the (usubstr \ (B, C)))) , snd (the (usubstr \ ([A], C))))" using def usubstr_def by auto ultimately show "the (usubstr \ ((Cons A B), C)) = (Cons (the (usubstappf \ allvars A)) (fst (the (usubstr \ (B, C)))) , snd (the (usubstr \ ([A], C))))" by simp qed lemma usubstr_union: "(usubstr \ ((append A B), C) \ None) \ the (usubstr \ ((append A B), C)) = (append (fst (the (usubstr \ (A, C)))) (fst (the (usubstr \ (B, C)))), snd (the (usubstr \ (A, C))))" using usubstr_union_undef2 (*by (smt fst_pair map_append option.sel snd_pair usubstr_def)*) proof- assume def: "(usubstr \ ((append A B), C) \ None)" let ?R = "((append A B), C)" have "the (usubstr \ ?R) = (map(the o (usubstappf \ allvars))(fst ?R) , the (usubstappf \ allvars (snd ?R)))" using def usubstr_def by (metis option.sel) also have "... = (map(the o (usubstappf \ allvars))(fst ?R) , snd (the (usubstr \ (A, C))))" using usubstr_union_undef2[OF def] usubstr_def by (metis option.sel sndI) also have "... = (append (map(the o (usubstappf \ allvars))(A)) (map(the o (usubstappf \ allvars))(B)) , snd (the (usubstr \ (A, C))))" using usubstr_union_undef2[OF def] map_append by simp also have "... = (append (fst (the (usubstr \ (A, C)))) (fst (the (usubstr \ (B, C)))), snd (the (usubstr \ (A, C))))" using usubstr_union_undef2[OF def] usubstr_def by (metis (no_types, lifting) fst_conv option.sel) ultimately show "the (usubstr \ ((append A B), C)) = (append (fst (the (usubstr \ (A, C)))) (fst (the (usubstr \ (B, C)))), snd (the (usubstr \ (A, C))))" by simp qed lemma usubstr_length: "usubstr \ R \ None \ length (fst (the (usubstr \ R))) = length (fst R)" by (metis fst_pair length_map option.sel usubstr_def) lemma usubstr_nth: "usubstr \ R \ None \ 0\k \ k nth (fst (the (usubstr \ R))) k = the (usubstappf \ allvars (nth (fst R) k))" (*unfolding usubstr_def using usubstr_length by (smt comp_apply fst_pair nth_map option.sel)*) proof- assume a1: "usubstr \ R \ None" assume a2: "0\k" assume a3: "k R))) k = the (usubstappf \ allvars (nth (fst R) k))" using a1 a2 a3 proof (induction R arbitrary: k) case (Pair A C) then show ?case proof (induction A arbitrary: k) case Nil then show ?case by simp next case (Cons D E) then have IH: "\k. usubstr \ (E, C) \ None \ 0 \ k \ k < length E \ nth (fst (the (usubstr \ (E, C)))) k = the (usubstappf \ allvars (nth E k))" by simp then show ?case proof (cases k) case 0 then show ?thesis using Cons usubstr_cons by simp next case (Suc n) then have smaller: "nn" by simp have def: "usubstr \ (E, C) \ None" using usubstr_cons_undef2[OF Cons.prems(1)] by blast have "nth (fst (the (usubstr \ (E, C)))) n = the (usubstappf \ allvars (nth (fst (E,C)) n))" using IH[OF def, OF nati, OF smaller] by simp then show ?thesis using Cons usubstr_cons by (simp add: Suc) qed qed qed qed text \Theorem 19 of \<^url>\http://arxiv.org/abs/1902.07230\\ theorem usubst_rule_sound: "usubstr \ R \ None \ locally_sound R \ locally_sound (the (usubstr \ R))" proof- assume def: "usubstr \ R \ None" assume prem: "locally_sound R" let ?\D = "usubstr \ R" fix \ from usubst_fml have substeq: "\I \ \. usubstappf \ allvars \ \ undeff \ (\ \ fml_sem I (the (usubstappf \ allvars \))) = (\ \ fml_sem (adjoint \ I \) \)" using Uvariation_univ by blast then have substval: "\I. usubstappf \ allvars \ \ undeff \ valid_in I (the (usubstappf \ allvars \)) = valid_in (adjoint \ I \) \" unfolding valid_in_def by auto show "locally_sound (the (usubstr \ R))" unfolding locally_sound_def proof (clarify) fix I assume "\k\0. k < length (fst (the (usubstr \ R))) \ valid_in I (nth (fst (the (usubstr \ R))) k)" then have "\k\0. k < length (fst R) \ valid_in (adjoint \ I \) (nth (fst R) k)" using substval usubstr_nth usubstr_length substeq valid_in_def by (metis def nth_mem usubstr_def) then have "valid_in (adjoint \ I \) (snd R)" using prem unfolding locally_sound_def by simp from this show "valid_in I (snd (the (usubstr \ R)))" using usubst_fml substeq usubstr_def valid_in_def by (metis def option.sel snd_conv) qed qed end diff --git a/thys/FSM_Tests/AdaptiveStateCounting/Test_Suite.thy b/thys/FSM_Tests/AdaptiveStateCounting/Test_Suite.thy --- a/thys/FSM_Tests/AdaptiveStateCounting/Test_Suite.thy +++ b/thys/FSM_Tests/AdaptiveStateCounting/Test_Suite.thy @@ -1,3014 +1,3014 @@ section \Test Suites\ text \This theory introduces a predicate @{text "implies_completeness"} and proves that any test suite satisfying this predicate is sufficient to check the reduction conformance relation between two (possibly nondeterministic FSMs)\ theory Test_Suite imports Helper_Algorithms Adaptive_Test_Case Traversal_Set begin subsection \Preliminary Definitions\ type_synonym ('a,'b,'c) preamble = "('a,'b,'c) fsm" type_synonym ('a,'b,'c) traversal_path = "('a \ 'b \ 'c \ 'a) list" type_synonym ('a,'b,'c) separator = "('a,'b,'c) fsm" text \A test suite contains of 1) a set of d-reachable states with their associated preambles 2) a map from d-reachable states to their associated m-traversal paths 3) a map from d-reachable states and associated m-traversal paths to the set of states to r-distinguish the targets of those paths from 4) a map from pairs of r-distinguishable states to a separator\ datatype ('a,'b,'c,'d) test_suite = Test_Suite "('a \ ('a,'b,'c) preamble) set" "'a \ ('a,'b,'c) traversal_path set" "('a \ ('a,'b,'c) traversal_path) \ 'a set" "('a \ 'a) \ (('d,'b,'c) separator \ 'd \ 'd) set" subsection \A Sufficiency Criterion for Reduction Testing\ (* to simplify the soundness proof, this formalisation also requires tps to contain nothing but the m-traversal paths; this could be lifted by requiring that for every additional path the suite retains additional paths such that all "non-deadlock" (w.r.t. the set of all (tps q) paths) states are output complete for the inputs applied *) fun implies_completeness_for_repetition_sets :: "('a,'b,'c,'d) test_suite \ ('a,'b,'c) fsm \ nat \ ('a set \ 'a set) list \ bool" where "implies_completeness_for_repetition_sets (Test_Suite prs tps rd_targets separators) M m repetition_sets = ( (initial M,initial_preamble M) \ prs \ (\ q P . (q,P) \ prs \ (is_preamble P M q) \ (tps q) \ {}) \ (\ q1 q2 A d1 d2 . (A,d1,d2) \ separators (q1,q2) \ (A,d2,d1) \ separators (q2,q1) \ is_separator M q1 q2 A d1 d2) \ (\ q . q \ states M \ (\d \ set repetition_sets. q \ fst d)) \ (\ d . d \ set repetition_sets \ ((fst d \ states M) \ (snd d = fst d \ fst ` prs) \ (\ q1 q2 . q1 \ fst d \ q2 \ fst d \ q1 \ q2 \ separators (q1,q2) \ {}))) \ (\ q . q \ image fst prs \ tps q \ {p1 . \ p2 d . (p1@p2,d) \ m_traversal_paths_with_witness M q repetition_sets m} \ fst ` (m_traversal_paths_with_witness M q repetition_sets m) \ tps q) \ (\ q p d . q \ image fst prs \ (p,d) \ m_traversal_paths_with_witness M q repetition_sets m \ ( (\ p1 p2 p3 . p=p1@p2@p3 \ p2 \ [] \ target q p1 \ fst d \ target q (p1@p2) \ fst d \ target q p1 \ target q (p1@p2) \ (p1 \ tps q \ (p1@p2) \ tps q \ target q p1 \ rd_targets (q,(p1@p2)) \ target q (p1@p2) \ rd_targets (q,p1))) \ (\ p1 p2 q' . p=p1@p2 \ q' \ image fst prs \ target q p1 \ fst d \ q' \ fst d \ target q p1 \ q' \ (p1 \ tps q \ [] \ tps q' \ target q p1 \ rd_targets (q',[]) \ q' \ rd_targets (q,p1)))) \ (\ q1 q2 . q1 \ q2 \ q1 \ snd d \ q2 \ snd d \ ([] \ tps q1 \ [] \ tps q2 \ q1 \ rd_targets (q2,[]) \ q2 \ rd_targets (q1,[])))) )" definition implies_completeness :: "('a,'b,'c,'d) test_suite \ ('a,'b,'c) fsm \ nat \ bool" where "implies_completeness T M m = (\ repetition_sets . implies_completeness_for_repetition_sets T M m repetition_sets)" lemma implies_completeness_for_repetition_sets_simps : assumes "implies_completeness_for_repetition_sets (Test_Suite prs tps rd_targets separators) M m repetition_sets" shows "(initial M,initial_preamble M) \ prs" and "\ q P . (q,P) \ prs \ (is_preamble P M q) \ (tps q) \ {}" and "\ q1 q2 A d1 d2 . (A,d1,d2) \ separators (q1,q2) \ (A,d2,d1) \ separators (q2,q1) \ is_separator M q1 q2 A d1 d2" and "\ q . q \ states M \ (\d \ set repetition_sets. q \ fst d)" and "\ d . d \ set repetition_sets \ (fst d \ states M) \ (snd d = fst d \ fst ` prs)" and "\ d q1 q2 . d \ set repetition_sets \ q1 \ fst d \ q2 \ fst d \ q1 \ q2 \ separators (q1,q2) \ {}" and "\ q . q \ image fst prs \ tps q \ {p1 . \ p2 d . (p1@p2,d) \ m_traversal_paths_with_witness M q repetition_sets m} \ fst ` (m_traversal_paths_with_witness M q repetition_sets m) \ tps q" and "\ q p d p1 p2 p3 . q \ image fst prs \ (p,d) \ m_traversal_paths_with_witness M q repetition_sets m \ p=p1@p2@p3 \ p2 \ [] \ target q p1 \ fst d \ target q (p1@p2) \ fst d \ target q p1 \ target q (p1@p2) \ (p1 \ tps q \ (p1@p2) \ tps q \ target q p1 \ rd_targets (q,(p1@p2)) \ target q (p1@p2) \ rd_targets (q,p1))" and "\ q p d p1 p2 q' . q \ image fst prs \ (p,d) \ m_traversal_paths_with_witness M q repetition_sets m \ p=p1@p2 \ q' \ image fst prs \ target q p1 \ fst d \ q' \ fst d \ target q p1 \ q' \ (p1 \ tps q \ [] \ tps q' \ target q p1 \ rd_targets (q',[]) \ q' \ rd_targets (q,p1))" and "\ q p d q1 q2 . q \ image fst prs \ (p,d) \ m_traversal_paths_with_witness M q repetition_sets m \ q1 \ q2 \ q1 \ snd d \ q2 \ snd d \ ([] \ tps q1 \ [] \ tps q2 \ q1 \ rd_targets (q2,[]) \ q2 \ rd_targets (q1,[]))" proof- show "(initial M,initial_preamble M) \ prs" and "\ q . q \ image fst prs \ tps q \ {p1 . \ p2 d . (p1@p2,d) \ m_traversal_paths_with_witness M q repetition_sets m} \ fst ` (m_traversal_paths_with_witness M q repetition_sets m) \ tps q" using assms unfolding implies_completeness_for_repetition_sets.simps by blast+ show "\ q1 q2 A d1 d2 . (A,d1,d2) \ separators (q1,q2) \ (A,d2,d1) \ separators (q2,q1) \ is_separator M q1 q2 A d1 d2" and "\ q P . (q,P) \ prs \ (is_preamble P M q) \ (tps q) \ {}" and "\ q . q \ states M \ (\d \ set repetition_sets. q \ fst d)" and "\ d . d \ set repetition_sets \ (fst d \ states M) \ (snd d = fst d \ fst ` prs)" and "\ d q1 q2 . d \ set repetition_sets \ q1 \ fst d \ q2 \ fst d \ q1 \ q2 \ separators (q1,q2) \ {}" using assms unfolding implies_completeness_for_repetition_sets.simps by force+ show "\ q p d p1 p2 p3 . q \ image fst prs \ (p,d) \ m_traversal_paths_with_witness M q repetition_sets m \ p=p1@p2@p3 \ p2 \ [] \ target q p1 \ fst d \ target q (p1@p2) \ fst d \ target q p1 \ target q (p1@p2) \ (p1 \ tps q \ (p1@p2) \ tps q \ target q p1 \ rd_targets (q,(p1@p2)) \ target q (p1@p2) \ rd_targets (q,p1))" using assms unfolding implies_completeness_for_repetition_sets.simps by (metis (no_types, lifting)) show "\ q p d p1 p2 q' . q \ image fst prs \ (p,d) \ m_traversal_paths_with_witness M q repetition_sets m \ p=p1@p2 \ q' \ image fst prs \ target q p1 \ fst d \ q' \ fst d \ target q p1 \ q' \ (p1 \ tps q \ [] \ tps q' \ target q p1 \ rd_targets (q',[]) \ q' \ rd_targets (q,p1))" using assms unfolding implies_completeness_for_repetition_sets.simps by (metis (no_types, lifting)) show "\ q p d q1 q2 . q \ image fst prs \ (p,d) \ m_traversal_paths_with_witness M q repetition_sets m \ q1 \ q2 \ q1 \ snd d \ q2 \ snd d \ ([] \ tps q1 \ [] \ tps q2 \ q1 \ rd_targets (q2,[]) \ q2 \ rd_targets (q1,[]))" using assms unfolding implies_completeness_for_repetition_sets.simps by (metis (no_types, lifting)) qed subsection \A Pass Relation for Test Suites and Reduction Testing\ fun passes_test_suite :: "('a,'b,'c) fsm \ ('a,'b,'c,'d) test_suite \ ('e,'b,'c) fsm \ bool" where "passes_test_suite M (Test_Suite prs tps rd_targets separators) M' = ( \ \Reduction on preambles: as the preambles contain all responses of M to their chosen inputs, M' must not exhibit any other response\ (\ q P io x y y' . (q,P) \ prs \ io@[(x,y)] \ L P \ io@[(x,y')] \ L M' \ io@[(x,y')] \ L P) \ \Reduction on traversal-paths applied after preambles (i.e., completed paths in preambles) - note that tps q is not necessarily prefix-complete\ \ (\ q P pP ioT pT x y y' . (q,P) \ prs \ path P (initial P) pP \ target (initial P) pP = q \ pT \ tps q \ ioT@[(x,y)] \ set (prefixes (p_io pT)) \ (p_io pP)@ioT@[(x,y')] \ L M' \ (\ pT' . pT' \ tps q \ ioT@[(x,y')] \ set (prefixes (p_io pT')))) \ \Passing separators: if M' contains an IO-sequence that in the test suite leads through a preamble and an m-traversal path and the target of the latter is to be r-distinguished from some other state, then M' passes the corresponding ATC\ \ (\ q P pP pT . (q,P) \ prs \ path P (initial P) pP \ target (initial P) pP = q \ pT \ tps q \ (p_io pP)@(p_io pT) \ L M' \ (\ q' A d1 d2 qT . q' \ rd_targets (q,pT) \ (A,d1,d2) \ separators (target q pT, q') \ qT \ io_targets M' ((p_io pP)@(p_io pT)) (initial M') \ pass_separator_ATC M' A qT d2)) )" subsection \Soundness of Sufficient Test Suites\ lemma passes_test_suite_soundness_helper_1 : assumes "is_preamble P M q" and "observable M" and "io@[(x,y)] \ L P" and "io@[(x,y')] \ L M" shows "io@[(x,y')] \ L P" proof - have "is_submachine P M" and *: "\ q' x t t' . q'\reachable_states P \ x\FSM.inputs M \ t\FSM.transitions P \ t_source t = q' \ t_input t = x \ t'\FSM.transitions M \ t_source t' = q' \ t_input t' = x \ t' \ FSM.transitions P" using assms(1) unfolding is_preamble_def by blast+ have "initial P = initial M" unfolding submachine_simps[OF \is_submachine P M\] by simp obtain p where "path M (initial M) p" and "p_io p = io @ [(x,y')]" using assms(4) unfolding submachine_simps[OF \is_submachine P M\] by auto obtain p' t where "p = p'@[t]" using \p_io p = io @ [(x,y')]\ by (induction p rule: rev_induct; auto) have "path M (initial M) p'" and "t \ transitions M" and "t_source t = target (initial M) p'" using \path M (initial M) p\ path_append_transition_elim unfolding \p = p'@[t]\ by force+ have "p_io p' = io" and "t_input t = x" and "t_output t = y'" using \p_io p = io @ [(x,y')]\ unfolding \p = p'@[t]\ by force+ have "p_io p' \ LS P (FSM.initial M)" using assms(3) unfolding \p_io p' = io\ \initial P = initial M\ by (meson language_prefix) have "FSM.initial M \ reachable_states P" unfolding submachine_simps(1)[OF \is_submachine P M\, symmetric] using reachable_states_initial by blast obtain pp where "path P (initial P) pp" and "p_io pp = io @ [(x,y)]" using assms(3) by auto then obtain pp' t' where "pp = pp'@[t']" proof - assume a1: "\pp' t'. pp = pp' @ [t'] \ thesis" have "pp \ []" using \p_io pp = io @ [(x, y)]\ by auto then show ?thesis using a1 by (metis (no_types) rev_exhaust) qed have "path P (initial P) pp'" and "t' \ transitions P" and "t_source t' = target (initial P) pp'" using \path P (initial P) pp\ path_append_transition_elim unfolding \pp = pp'@[t']\ by force+ have "p_io pp' = io" and "t_input t' = x" using \p_io pp = io @ [(x,y)]\ unfolding \pp = pp'@[t']\ by force+ have "path M (initial M) pp'" using \path P (initial P) pp'\ submachine_path_initial[OF \is_submachine P M\] by blast have "pp' = p'" using observable_path_unique[OF assms(2) \path M (initial M) pp'\ \path M (initial M) p'\ ] unfolding \p_io pp' = io\ \p_io p' = io\ by blast then have "t_source t' = target (initial M) p'" using \t_source t' = target (initial P) pp'\ unfolding \initial P = initial M\ by blast have "path P (FSM.initial M) p'" using observable_preamble_paths[OF assms(1,2) \path M (initial M) p'\ \p_io p' \ LS P (FSM.initial M)\ \FSM.initial M \ reachable_states P\] by assumption then have "target (initial M) p' \ reachable_states P" using reachable_states_intro unfolding \initial P = initial M\[symmetric] by blast moreover have "x \ inputs M" using \t \ transitions M\ \t_input t = x\ fsm_transition_input by blast have "t \ transitions P" using *[OF \target (initial M) p' \ reachable_states P\ \x \ inputs M\ \t' \ transitions P\ \t_source t' = target (initial M) p'\ \t_input t' = x\ \t \ transitions M\ \t_source t = target (FSM.initial M) p'\ \t_input t = x\] by assumption then have "path P (initial P) (p'@[t])" using \path P (initial P) pp'\ \t_source t = target (initial M) p'\ unfolding \pp' = p'\ \initial P = initial M\ using path_append_transition by simp then show ?thesis unfolding \p = p'@[t]\[symmetric] LS.simps using \p_io p = io@[(x,y')]\ by force qed lemma passes_test_suite_soundness : assumes "implies_completeness (Test_Suite prs tps rd_targets separators) M m" and "observable M" and "observable M'" and "inputs M' = inputs M" and "completely_specified M" and "L M' \ L M" shows "passes_test_suite M (Test_Suite prs tps rd_targets separators) M'" proof - obtain repetition_sets where repetition_sets_def: "implies_completeness_for_repetition_sets (Test_Suite prs tps rd_targets separators) M m repetition_sets" using assms(1) unfolding implies_completeness_def by blast have t1: "(initial M, initial_preamble M) \ prs" using implies_completeness_for_repetition_sets_simps(1)[OF repetition_sets_def] by assumption have t2: "\ q P. (q, P) \ prs \ is_preamble P M q \ tps q \ {}" using implies_completeness_for_repetition_sets_simps(2)[OF repetition_sets_def] by assumption have t3: "\ q1 q2 A d1 d2. (A, d1, d2) \ separators (q1, q2) \ (A, d2, d1) \ separators (q2, q1) \ is_separator M q1 q2 A d1 d2" using implies_completeness_for_repetition_sets_simps(3)[OF repetition_sets_def] by assumption have t5: "\q. q \ FSM.states M \ (\d\set repetition_sets. q \ fst d)" using implies_completeness_for_repetition_sets_simps(4)[OF repetition_sets_def] by assumption have t6: "\ q. q \ fst ` prs \ tps q \ {p1 . \ p2 d . (p1@p2,d) \ m_traversal_paths_with_witness M q repetition_sets m} \ fst ` (m_traversal_paths_with_witness M q repetition_sets m) \ tps q" using implies_completeness_for_repetition_sets_simps(7)[OF repetition_sets_def] by assumption have t7: "\ d. d \ set repetition_sets \ fst d \ FSM.states M" and t8: "\ d. d \ set repetition_sets \ snd d \ fst d" and t9: "\ d q1 q2. d \ set repetition_sets \ q1 \ fst d \ q2 \ fst d \ q1 \ q2 \ separators (q1, q2) \ {}" using implies_completeness_for_repetition_sets_simps(5,6)[OF repetition_sets_def] by blast+ have t10: "\ q p d p1 p2 p3. q \ fst ` prs \ (p, d) \ m_traversal_paths_with_witness M q repetition_sets m \ p = p1 @ p2 @ p3 \ p2 \ [] \ target q p1 \ fst d \ target q (p1 @ p2) \ fst d \ target q p1 \ target q (p1 @ p2) \ p1 \ tps q \ p1 @ p2 \ tps q \ target q p1 \ rd_targets (q, p1 @ p2) \ target q (p1 @ p2) \ rd_targets (q, p1)" using implies_completeness_for_repetition_sets_simps(8)[OF repetition_sets_def] by assumption have t11: "\ q p d p1 p2 q'. q \ fst ` prs \ (p, d) \ m_traversal_paths_with_witness M q repetition_sets m \ p = p1 @ p2 \ q' \ fst ` prs \ target q p1 \ fst d \ q' \ fst d \ target q p1 \ q' \ p1 \ tps q \ [] \ tps q' \ target q p1 \ rd_targets (q', []) \ q' \ rd_targets (q, p1)" using implies_completeness_for_repetition_sets_simps(9)[OF repetition_sets_def] by assumption have "\ q P io x y y' . (q,P) \ prs \ io@[(x,y)] \ L P \ io@[(x,y')] \ L M' \ io@[(x,y')] \ L P" proof - fix q P io x y y' assume "(q,P) \ prs" and "io@[(x,y)] \ L P" and "io@[(x,y')] \ L M'" have "is_preamble P M q" using \(q,P) \ prs\ \\ q P. (q, P) \ prs \ is_preamble P M q \ tps q \ {}\ by blast have "io@[(x,y')] \ L M" using \io@[(x,y')] \ L M'\ assms(6) by blast show "io@[(x,y')] \ L P" using passes_test_suite_soundness_helper_1[OF \is_preamble P M q\ assms(2) \io@[(x,y)] \ L P\ \io@[(x,y')] \ L M\] by assumption qed then have p1: "(\ q P io x y y' . (q,P) \ prs \ io@[(x,y)] \ L P \ io@[(x,y')] \ L M' \ io@[(x,y')] \ L P)" by blast have "\ q P pP ioT pT x x' y y' . (q,P) \ prs \ path P (initial P) pP \ target (initial P) pP = q \ pT \ tps q \ ioT @ [(x, y)] \ set (prefixes (p_io pT)) \ (p_io pP)@ioT@[(x',y')] \ L M' \ (\ pT' . pT' \ tps q \ ioT @ [(x', y')] \ set (prefixes (p_io pT')))" proof - fix q P pP ioT pT x x' y y' assume "(q,P) \ prs" and "path P (initial P) pP" and "target (initial P) pP = q" and "pT \ tps q" and "ioT @ [(x, y)] \ set (prefixes (p_io pT))" and "(p_io pP)@ioT@[(x',y')] \ L M'" have "is_preamble P M q" using \(q,P) \ prs\ \\ q P. (q, P) \ prs \ is_preamble P M q \ tps q \ {}\ by blast then have "q \ states M" unfolding is_preamble_def by (metis \path P (FSM.initial P) pP\ \target (FSM.initial P) pP = q\ path_target_is_state submachine_path) have "initial P = initial M" using \is_preamble P M q\ unfolding is_preamble_def by auto have "path M (initial M) pP" using \is_preamble P M q\ unfolding is_preamble_def using submachine_path_initial using \path P (FSM.initial P) pP\ by blast have "(p_io pP)@ioT@[(x',y')] \ L M" using \(p_io pP)@ioT@[(x',y')] \ L M'\ assms(6) by blast then obtain pM' where "path M (initial M) pM'" and "p_io pM' = (p_io pP)@ioT@[(x',y')]" by auto let ?pP = "take (length pP) pM'" let ?pT = "take (length ioT) (drop (length pP) pM')" let ?t = "last pM'" have "pM' = ?pP @ ?pT @ [?t]" proof - have "length pM' = (length pP) + (length ioT) + 1" using \p_io pM' = (p_io pP)@ioT@[(x',y')]\ unfolding length_map[of "(\ t . (t_input t, t_output t))", of pM', symmetric] length_map[of "(\ t . (t_input t, t_output t))", of pP, symmetric] by auto then show ?thesis by (metis (no_types, lifting) add_diff_cancel_right' antisym_conv antisym_conv2 append_butlast_last_id append_eq_append_conv2 butlast_conv_take drop_Nil drop_eq_Nil le_add1 less_add_one take_add) qed have "p_io ?pP = p_io pP" using \p_io pM' = (p_io pP)@ioT@[(x',y')]\ by (metis (no_types, lifting) append_eq_conv_conj length_map take_map) have "p_io ?pT = ioT" using \p_io pM' = (p_io pP)@ioT@[(x',y')]\ using \pM' = ?pP @ ?pT @ [?t]\ by (metis (no_types, lifting) append_eq_conv_conj length_map map_append take_map) have "p_io [?t] = [(x',y')]" using \p_io pM' = (p_io pP)@ioT@[(x',y')]\ using \pM' = ?pP @ ?pT @ [?t]\ by (metis (no_types, lifting) append_is_Nil_conv last_appendR last_map last_snoc list.simps(8) list.simps(9)) have "path M (initial M) ?pP" using \path M (initial M) pM'\ \pM' = ?pP @ ?pT @ [?t]\ by (meson path_prefix_take) have "?pP = pP" using observable_path_unique[OF \observable M\ \path M (initial M) ?pP\ \path M (initial M) pP\ \p_io ?pP = p_io pP\] by assumption then have "path M q (?pT@[?t])" by (metis \FSM.initial P = FSM.initial M\ \pM' = take (length pP) pM' @ take (length ioT) (drop (length pP) pM') @ [last pM']\ \path M (FSM.initial M) pM'\ \target (FSM.initial P) pP = q\ path_suffix) then have "path M q ?pT" and "?t \ transitions M" and "t_source ?t = target q ?pT" by auto have "inputs M \ {}" using language_io(1)[OF \(p_io pP)@ioT@[(x',y')] \ L M\, of x' y'] by auto have "q \ fst ` prs" using \(q,P) \ prs\ using image_iff by fastforce obtain ioT' where "p_io pT = (ioT @ [(x, y)]) @ ioT'" using \ioT @ [(x, y)] \ set (prefixes (p_io pT))\ - unfolding prefixes_set - by moura + unfolding prefixes_set mem_Collect_eq by metis then have "length pT > length ioT" using length_map[of "(\ t . (t_input t, t_output t))" pT] by auto obtain pT' d' where "(pT @ pT', d') \ m_traversal_paths_with_witness M q repetition_sets m" using t6[OF \q \ fst ` prs\] \pT \ tps q\ by blast let ?p = "pT @ pT'" have "path M q ?p" and "find (\d. Suc (m - card (snd d)) \ length (filter (\t. t_target t \ fst d) ?p)) repetition_sets = Some d'" and "\ p' p''. ?p = p' @ p'' \ p'' \ [] \ find (\d. Suc (m - card (snd d)) \ length (filter (\t. t_target t \ fst d) p')) repetition_sets = None" using \(pT @ pT', d') \ m_traversal_paths_with_witness M q repetition_sets m\ m_traversal_paths_with_witness_set[OF t5 t8 \q \ states M\, of m] by blast+ let ?pIO = "take (length ioT) pT" have "?pIO = take (length ioT) ?p" using \length pT > length ioT\ by auto then have "?p = ?pIO@(drop (length ioT) ?p)" by auto have "(drop (length ioT) ?p) \ []" using \length pT > length ioT\ by auto have "p_io ?pIO = ioT" proof - have "p_io ?pIO = take (length ioT) (p_io pT)" by (simp add: take_map) moreover have "take (length ioT) (p_io pT) = ioT" using \p_io pT = (ioT @ [(x, y)]) @ ioT'\ by auto ultimately show ?thesis by simp qed then have "p_io ?pIO = p_io ?pT" using \p_io ?pT = ioT\ by simp have "path M q ?pIO" using \path M q ?p\ unfolding \?pIO = take (length ioT) ?p\ using path_prefix_take by metis have "?pT = ?pIO" using observable_path_unique[OF \observable M\ \path M q ?pIO\ \path M q ?pT\ \p_io ?pIO = p_io ?pT\] by simp show "(\ pT' . pT' \ tps q \ ioT @ [(x', y')] \ set (prefixes (p_io pT')))" proof (cases "find (\d. Suc (m - card (snd d)) \ length (filter (\t. t_target t \ fst d) (?pT@[?t]))) repetition_sets = None") case True obtain pT' d' where "(?pT @ [?t] @ pT', d') \ m_traversal_paths_with_witness M q repetition_sets m" using m_traversal_path_extension_exist[OF \completely_specified M\ \q \ states M\ \inputs M \ {}\ t5 t8 \path M q (?pT@[?t])\ True] by auto then have "?pT @ [?t] @ pT' \ tps q" using t6[OF \q \ fst ` prs\] by force moreover have "ioT @ [(x', y')] \ set (prefixes (p_io (?pT @ [?t] @ pT')))" using \p_io ?pIO = ioT\ \p_io [?t] = [(x',y')]\ unfolding \?pT = ?pIO\ prefixes_set by force ultimately show ?thesis by blast next case False note \path M q (?pT @ [?t])\ moreover obtain d' where "find (\d. Suc (m - card (snd d)) \ length (filter (\t. t_target t \ fst d) (?pT@[?t]))) repetition_sets = Some d'" using False by blast moreover have "\ p' p''. (?pT @ [?t]) = p' @ p'' \ p'' \ [] \ find (\d. Suc (m - card (snd d)) \ length (filter (\t. t_target t \ fst d) p')) repetition_sets = None" proof - have "\ p' p''. (?pT @ [?t]) = p' @ p'' \ p'' \ [] \ find (\d. Suc (m - card (snd d)) \ length (filter (\t. t_target t \ fst d) p')) repetition_sets = None" proof - fix p' p'' assume "(?pT @ [?t]) = p' @ p''" and "p'' \ []" then obtain pIO' where "?pIO = p' @ pIO'" unfolding \?pT = ?pIO\ by (metis butlast_append butlast_snoc) then have "?p = p'@pIO'@(drop (length ioT) ?p)" using \?p = ?pIO@((drop (length ioT) ?p))\ by (metis append.assoc) have "pIO' @ drop (length ioT) (pT @ pT') \ []" using \(drop (length ioT) ?p) \ []\ by auto show "find (\d. Suc (m - card (snd d)) \ length (filter (\t. t_target t \ fst d) p')) repetition_sets = None" using \\ p' p''. ?p = p' @ p'' \ p'' \ [] \ find (\d. Suc (m - card (snd d)) \ length (filter (\t. t_target t \ fst d) p')) repetition_sets = None\ [of p' "pIO'@(drop (length ioT) ?p)", OF \?p = p'@pIO'@(drop (length ioT) ?p)\ \pIO' @ drop (length ioT) (pT @ pT') \ []\] by assumption qed then show ?thesis by blast qed ultimately have "((?pT @ [?t]),d') \ m_traversal_paths_with_witness M q repetition_sets m" using m_traversal_paths_with_witness_set[OF t5 t8 \q \ states M\, of m] by auto then have "(?pT @ [?t]) \ tps q" using t6[OF \q \ fst ` prs\] by force moreover have "ioT @ [(x', y')] \ set (prefixes (p_io (?pT @ [?t])))" using \p_io ?pT = ioT\ \p_io [?t] = [(x',y')]\ unfolding prefixes_set by force ultimately show ?thesis by blast qed qed then have p2: "(\ q P pP ioT pT x y y' . (q,P) \ prs \ path P (initial P) pP \ target (initial P) pP = q \ pT \ tps q \ ioT @ [(x, y)] \ set (prefixes (p_io pT)) \ (p_io pP)@ioT@[(x,y')] \ L M' \ (\ pT' . pT' \ tps q \ ioT @ [(x, y')] \ set (prefixes (p_io pT'))))" by blast have "\ q P pP pT q' A d1 d2 qT . (q,P) \ prs \ path P (initial P) pP \ target (initial P) pP = q \ pT \ tps q \ q' \ rd_targets (q,pT) \ (A,d1,d2) \ separators (target q pT, q') \ qT \ io_targets M' ((p_io pP)@(p_io pT)) (initial M') \ pass_separator_ATC M' A qT d2" proof - fix q P pP pT q' A d1 d2 qT assume "(q,P) \ prs" and "path P (initial P) pP" and "target (initial P) pP = q" and "pT \ tps q" and "q' \ rd_targets (q,pT)" and "(A,d1,d2) \ separators (target q pT, q')" and "qT \ io_targets M' ((p_io pP)@(p_io pT)) (initial M')" have "q \ fst ` prs" using \(q,P) \ prs\ by force have "is_preamble P M q" using \(q,P) \ prs\ \\ q P. (q, P) \ prs \ is_preamble P M q \ tps q \ {}\ by blast then have "q \ states M" unfolding is_preamble_def by (metis \path P (FSM.initial P) pP\ \target (FSM.initial P) pP = q\ path_target_is_state submachine_path) have "initial P = initial M" using \is_preamble P M q\ unfolding is_preamble_def by auto have "path M (initial M) pP" using \is_preamble P M q\ unfolding is_preamble_def using submachine_path_initial using \path P (FSM.initial P) pP\ by blast have "is_separator M (target q pT) q' A d1 d2" using t3[OF \(A,d1,d2) \ separators (target q pT, q')\] by blast have "qT \ states M'" using \qT \ io_targets M' ((p_io pP)@(p_io pT)) (initial M')\ io_targets_states by (metis (no_types, lifting) subsetD) obtain pT' d' where "(pT @ pT', d') \ m_traversal_paths_with_witness M q repetition_sets m" using t6[OF \q \ fst ` prs\] \pT \ tps q\ by blast then have "path M q pT" using m_traversal_paths_with_witness_set[OF t5 t8 \q \ states M\, of m] by auto then have "target q pT \ FSM.states M" using path_target_is_state by metis have "q' \ FSM.states M" using is_separator_separated_state_is_state[OF \is_separator M (target q pT) q' A d1 d2\] by simp have "\ pass_separator_ATC M' A qT d2 \ \ LS M' qT \ LS M (target q pT)" using pass_separator_ATC_fail_no_reduction[OF \observable M'\ \observable M\ \qT \ states M'\ \target q pT \ FSM.states M\ \q' \ FSM.states M\ \is_separator M (target q pT) q' A d1 d2\ \inputs M' = inputs M\] by assumption moreover have "LS M' qT \ LS M (target q pT)" proof - have "(target q pT) = target (initial M) (pP@pT)" using \target (initial P) pP = q\ unfolding \initial P = initial M\ by auto have "path M (initial M) (pP@pT)" using \path M (initial M) pP\ \target (initial P) pP = q\ \path M q pT\ unfolding \initial P = initial M\ by auto then have "(target q pT) \ io_targets M (p_io pP @ p_io pT) (FSM.initial M)" unfolding io_targets.simps \(target q pT) = target (initial M) (pP@pT)\ using map_append by blast show ?thesis using observable_language_target[OF \observable M\ \(target q pT) \ io_targets M (p_io pP @ p_io pT) (FSM.initial M)\ \qT \ io_targets M' ((p_io pP)@(p_io pT)) (initial M')\ \L M' \ L M\] by assumption qed ultimately show "pass_separator_ATC M' A qT d2" by blast qed then have p3: "(\ q P pP pT . (q,P) \ prs \ path P (initial P) pP \ target (initial P) pP = q \ pT \ tps q \ (p_io pP)@(p_io pT) \ L M' \ (\ q' A d1 d2 qT . q' \ rd_targets (q,pT) \ (A,d1,d2) \ separators (target q pT, q') \ qT \ io_targets M' ((p_io pP)@(p_io pT)) (initial M') \ pass_separator_ATC M' A qT d2))" by blast show ?thesis using p1 p2 p3 unfolding passes_test_suite.simps by blast qed subsection \Exhaustiveness of Sufficient Test Suites\ text \This subsection shows that test suites satisfying the sufficiency criterion are exhaustive. That is, for a System Under Test with at most m states that contains an error (i.e.: is not a reduction) a test suite sufficient for m will not pass.\ subsubsection \R Functions\ (* collect all proper suffixes of p that target q' if applied to q *) definition R :: "('a,'b,'c) fsm \ 'a \ 'a \ ('a \ 'b \ 'c \ 'a) list \ ('a \ 'b \ 'c \ 'a) list \ ('a \ 'b \ 'c \ 'a) list set" where "R M q q' pP p = {pP @ p' | p' . p' \ [] \ target q p' = q' \ (\ p'' . p = p'@p'')}" (* add one completed path of some Preamble of q' to R if a preamble exists *) definition RP :: "('a,'b,'c) fsm \ 'a \ 'a \ ('a \ 'b \ 'c \ 'a) list \ ('a \ 'b \ 'c \ 'a) list \ ('a \ ('a,'b,'c) preamble) set \ ('d,'b,'c) fsm \ ('a \ 'b \ 'c \ 'a) list set" where "RP M q q' pP p PS M' = (if \ P' . (q',P') \ PS then insert (SOME pP' . \ P' . (q',P') \ PS \ path P' (initial P') pP' \ target (initial P') pP' = q' \ p_io pP' \ L M') (R M q q' pP p) else (R M q q' pP p))" lemma RP_from_R : assumes "\ q P . (q,P) \ PS \ is_preamble P M q" and "\ q P io x y y' . (q,P) \ PS \ io@[(x,y)] \ L P \ io@[(x,y')] \ L M' \ io@[(x,y')] \ L P" and "completely_specified M'" and "inputs M' = inputs M" shows "(RP M q q' pP p PS M' = R M q q' pP p) \ (\ P' pP' . (q',P') \ PS \ path P' (initial P') pP' \ target (initial P') pP' = q' \ path M (initial M) pP' \ target (initial M) pP' = q' \ p_io pP' \ L M' \ RP M q q' pP p PS M' = insert pP' (R M q q' pP p))" proof (rule ccontr) assume "\ (RP M q q' pP p PS M' = R M q q' pP p \ (\ P' pP' . (q',P') \ PS \ path P' (initial P') pP' \ target (initial P') pP' = q' \ path M (initial M) pP' \ target (initial M) pP' = q' \ p_io pP' \ L M' \ RP M q q' pP p PS M' = insert pP' (R M q q' pP p)))" then have "(RP M q q' pP p PS M' \ R M q q' pP p)" and "\ (\ P' pP' . (q',P') \ PS \ path P' (initial P') pP' \ target (initial P') pP' = q' \ path M (initial M) pP' \ target (initial M) pP' = q' \ p_io pP' \ L M' \ RP M q q' pP p PS M' = insert pP' (R M q q' pP p))" by blast+ let ?p = "SOME pP' . \ P' . (q',P') \ PS \ path P' (initial P') pP' \ target (initial P') pP' = q' \ p_io pP' \ L M'" have "\ P' . (q',P') \ PS" using \(RP M q q' pP p PS M' \ R M q q' pP p)\ unfolding RP_def by auto then obtain P' where "(q',P') \ PS" by auto then have "is_preamble P' M q'" using assms by blast obtain pP' where "path P' (initial P') pP'" and "target (initial P') pP' = q'" and "p_io pP' \ L M'" using preamble_pass_path[OF \is_preamble P' M q'\ assms(2)[OF \(q',P') \ PS\] assms(3,4)] by force then have "\ pP' . \ P' . (q',P') \ PS \ path P' (initial P') pP' \ target (initial P') pP' = q' \ p_io pP' \ L M'" using \(q',P') \ PS\ by blast have "\ P' . (q',P') \ PS \ path P' (initial P') ?p \ target (initial P') ?p = q' \ p_io ?p \ L M'" using someI_ex[OF \\ pP' . \ P' . (q',P') \ PS \ path P' (initial P') pP' \ target (initial P') pP' = q' \ p_io pP' \ L M'\] by blast then obtain P'' where "(q',P'') \ PS" and "path P'' (initial P'') ?p" and "target (initial P'') ?p = q'" and "p_io ?p \ L M'" by auto then have "is_preamble P'' M q'" using assms by blast have "initial P'' = initial M" using \is_preamble P'' M q'\ unfolding is_preamble_def by auto have "path M (initial M) ?p" using \is_preamble P'' M q'\ unfolding is_preamble_def using submachine_path_initial using \path P'' (FSM.initial P'') ?p\ by blast have "target (initial M) ?p = q'" using \target (initial P'') ?p = q'\ unfolding \initial P'' = initial M\ by assumption have "RP M q q' pP p PS M' = insert ?p (R M q q' pP p)" using \\ P' . (q',P') \ PS\ unfolding RP_def by auto then have "(\ P' pP' . (q',P') \ PS \ path P' (initial P') pP' \ target (initial P') pP' = q' \ path M (initial M) pP' \ target (initial M) pP' = q' \ p_io pP' \ L M' \ RP M q q' pP p PS M' = insert pP' (R M q q' pP p))" using \(q',P'') \ PS\ \path P'' (initial P'') ?p\ \target (initial P'') ?p = q'\ \path M (initial M) ?p\ \target (initial M) ?p = q'\ \p_io ?p \ L M'\ by blast then show "False" using \\ (\ P' pP' . (q',P') \ PS \ path P' (initial P') pP' \ target (initial P') pP' = q' \ path M (initial M) pP' \ target (initial M) pP' = q' \ p_io pP' \ L M' \ RP M q q' pP p PS M' = insert pP' (R M q q' pP p))\ by blast qed lemma RP_from_R_inserted : assumes "\ q P . (q,P) \ PS \ is_preamble P M q" and "\ q P io x y y' . (q,P) \ PS \ io@[(x,y)] \ L P \ io@[(x,y')] \ L M' \ io@[(x,y')] \ L P" and "completely_specified M'" and "inputs M' = inputs M" and "pP' \ RP M q q' pP p PS M'" and "pP' \ R M q q' pP p" obtains P' where "(q',P') \ PS" "path P' (initial P') pP'" "target (initial P') pP' = q'" "path M (initial M) pP'" "target (initial M) pP' = q'" "p_io pP' \ L M'" "RP M q q' pP p PS M' = insert pP' (R M q q' pP p)" proof - have "(RP M q q' pP p PS M' \ R M q q' pP p)" using assms(5,6) by blast then have "(\P' pP'. (q', P') \ PS \ path P' (FSM.initial P') pP' \ target (FSM.initial P') pP' = q' \ path M (FSM.initial M) pP' \ target (FSM.initial M) pP' = q' \ p_io pP' \ L M' \ RP M q q' pP p PS M' = insert pP' (R M q q' pP p))" using RP_from_R[OF assms(1-4), of PS _ _ q q' pP p] by force then obtain P' pP'' where "(q', P') \ PS" "path P' (FSM.initial P') pP''" "target (FSM.initial P') pP'' = q'" "path M (FSM.initial M) pP''" "target (FSM.initial M) pP'' = q'" "p_io pP'' \ L M'" "RP M q q' pP p PS M' = insert pP'' (R M q q' pP p)" by blast moreover have "pP'' = pP'" using \RP M q q' pP p PS M' = insert pP'' (R M q q' pP p)\ assms(5,6) by simp ultimately show ?thesis using that[of P'] unfolding \pP'' = pP'\ by blast qed lemma finite_R : assumes "path M q p" shows "finite (R M q q' pP p)" proof - have "\ p' . p' \ (R M q q' pP p) \ p' \ set (prefixes (pP@p))" proof - fix p' assume "p' \ (R M q q' pP p)" then obtain p'' where "p' = pP @ p''" unfolding R_def by blast then obtain p''' where "p = p'' @ p'''" using \p' \ (R M q q' pP p)\ unfolding R_def by blast show "p' \ set (prefixes (pP@p))" unfolding prefixes_set \p' = pP @ p''\ \p = p'' @ p'''\ by auto qed then have "(R M q q' pP p) \ set (prefixes (pP@p))" by blast then show ?thesis using rev_finite_subset by auto qed lemma finite_RP : assumes "path M q p" and "\ q P . (q,P) \ PS \ is_preamble P M q" and "\ q P io x y y' . (q,P) \ PS \ io@[(x,y)] \ L P \ io@[(x,y')] \ L M' \ io@[(x,y')] \ L P" and "completely_specified M'" and "inputs M' = inputs M" shows "finite (RP M q q' pP p PS M')" using finite_R[OF assms(1), of q' pP ] RP_from_R[OF assms(2,3,4,5), of PS _ _ q q' pP p] by force lemma R_component_ob : assumes "pR' \ R M q q' pP p" obtains pR where "pR' = pP@pR" using assms unfolding R_def by blast lemma R_component : assumes "(pP@pR) \ R M q q' pP p" shows "pR = take (length pR) p" and "length pR \ length p" and "t_target (p ! (length pR - 1)) = q'" and "pR \ []" proof - let ?R = "R M q q' p" have "pR \ []" and "target q pR = q'" and "\ p'' . p = pR@p''" using \pP@pR \ R M q q' pP p\ unfolding R_def by blast+ then obtain pR' where "p = pR@pR'" by blast then show "pR = take (length pR) p" and "length pR \ length p" by auto show "t_target (p ! (length pR - 1)) = q'" using \pR \ []\ \target q pR = q'\ unfolding target.simps visited_states.simps by (metis (no_types, lifting) Suc_diff_1 \pR = take (length pR) p\ append_butlast_last_id last.simps last_map length_butlast lessI list.map_disc_iff not_gr_zero nth_append_length nth_take take_eq_Nil) show "pR \ []" using \pR \ []\ by assumption qed lemma R_component_observable : assumes "pP@pR \ R M (target (initial M) pP) q' pP p" and "observable M" and "path M (initial M) pP" and "path M (target (initial M) pP) p" shows "io_targets M (p_io pP @ p_io pR) (initial M) = {target (target (initial M) pP) (take (length pR) p)}" proof - have "pR = take (length pR) p" and "length pR \ length p" and "t_target (p ! (length pR - 1)) = q'" using R_component[OF assms(1)] by blast+ let ?q = "(target (initial M) pP)" have "path M ?q (take (length pR) p)" using assms(4) by (simp add: path_prefix_take) have "p_io (take (length pR) p) = p_io pR" using \pR = take (length pR) p\ by auto have *:"path M (initial M) (pP @ (take (length pR) p))" using \path M (initial M) pP\ \path M ?q (take (length pR) p)\ by auto have **:"p_io (pP @ (take (length pR) p)) = (p_io pP @ p_io pR)" using \p_io (take (length pR) p) = p_io pR\ by auto have "target (initial M) (pP @ (take (length pR) p)) = target ?q (take (length pR) p)" by auto then have "target ?q (take (length pR) p) \ io_targets M (p_io pP @ p_io pR) (initial M)" unfolding io_targets.simps using * ** by (metis (mono_tags, lifting) mem_Collect_eq) show "io_targets M (p_io pP @ p_io pR) (initial M) = {target ?q (take (length pR) p)}" using observable_io_targets[OF \observable M\ language_state_containment[OF * **]] by (metis (no_types) \target (target (FSM.initial M) pP) (take (length pR) p) \ io_targets M (p_io pP @ p_io pR) (FSM.initial M)\ singleton_iff) qed lemma R_count : assumes "minimal_sequence_to_failure_extending_preamble_path M M' PS pP io" and "observable M" and "observable M'" and "\ q P. (q, P) \ PS \ is_preamble P M q" and "path M (target (initial M) pP) p" and "butlast io = p_io p @ ioX" shows "card (\ (image (\ pR . io_targets M' (p_io pR) (initial M')) (R M (target (initial M) pP) q' pP p))) = card (R M (target (initial M) pP) q' pP p)" (is "card ?Tgts = card ?R") and "\ pR . pR \ (R M (target (initial M) pP) q' pP p) \ \ q . io_targets M' (p_io pR) (initial M') = {q}" and "\ pR1 pR2 . pR1 \ (R M (target (initial M) pP) q' pP p) \ pR2 \ (R M (target (initial M) pP) q' pP p) \ pR1 \ pR2 \ io_targets M' (p_io pR1) (initial M') \ io_targets M' (p_io pR2) (initial M') = {}" proof - have "sequence_to_failure_extending_preamble_path M M' PS pP io" and "\ p' io' . sequence_to_failure_extending_preamble_path M M' PS p' io' \ length io \ length io'" using \minimal_sequence_to_failure_extending_preamble_path M M' PS pP io\ unfolding minimal_sequence_to_failure_extending_preamble_path_def by blast+ obtain q P where "(q,P) \ PS" and "path P (initial P) pP" and "target (initial P) pP = q" and "((p_io pP) @ butlast io) \ L M" and "((p_io pP) @ io) \ L M" and "((p_io pP) @ io) \ L M'" using \sequence_to_failure_extending_preamble_path M M' PS pP io\ unfolding sequence_to_failure_extending_preamble_path_def by blast have "is_preamble P M q" using \(q,P) \ PS\ \\ q P. (q, P) \ PS \ is_preamble P M q\ by blast then have "q \ states M" unfolding is_preamble_def by (metis \path P (FSM.initial P) pP\ \target (FSM.initial P) pP = q\ path_target_is_state submachine_path) have "initial P = initial M" using \is_preamble P M q\ unfolding is_preamble_def by auto have "path M (initial M) pP" using \is_preamble P M q\ unfolding is_preamble_def using submachine_path_initial using \path P (FSM.initial P) pP\ by blast have "target (initial M) pP = q" using \target (initial P) pP = q\ unfolding \initial P = initial M\ by assumption then have "path M q p" using \path M (target (initial M) pP) p\ by auto have "io \ []" using \((p_io pP) @ butlast io) \ L M\ \((p_io pP) @ io) \ L M\ by auto (* get the full path along (butlast io) in M, of which p is a (possibly proper) prefix *) obtain pX where "path M (target (initial M) pP) (p@pX)" and "p_io (p@pX) = butlast io" proof - have "p_io pP @ p_io p @ ioX \ L M" using \((p_io pP) @ butlast io) \ L M\ unfolding \butlast io = p_io p @ ioX\ by assumption obtain p1 p23 where "path M (FSM.initial M) p1" and "path M (target (FSM.initial M) p1) p23" and "p_io p1 = p_io pP" and "p_io p23 = p_io p @ ioX" using language_state_split[OF \p_io pP @ p_io p @ ioX \ L M\] by blast have "p1 = pP" using observable_path_unique[OF \observable M\ \path M (FSM.initial M) p1\ \path M (FSM.initial M) pP\ \p_io p1 = p_io pP\] by assumption then have "path M (target (FSM.initial M) pP) p23" using \path M (target (FSM.initial M) p1) p23\ by auto then have "p_io p @ ioX \ LS M (target (initial M) pP)" using \p_io p23 = p_io p @ ioX\ language_state_containment by auto obtain p2 p3 where "path M (target (FSM.initial M) pP) p2" and "path M (target (target (FSM.initial M) pP) p2) p3" and "p_io p2 = p_io p" and "p_io p3 = ioX" using language_state_split[OF \p_io p @ ioX \ LS M (target (initial M) pP)\] by blast have "p2 = p" using observable_path_unique[OF \observable M\ \path M (target (FSM.initial M) pP) p2\ \path M (target (FSM.initial M) pP) p\ \p_io p2 = p_io p\] by assumption then have "path M (target (FSM.initial M) pP) (p@p3)" using \path M (target (FSM.initial M) pP) p\ \path M (target (target (FSM.initial M) pP) p2) p3\ by auto moreover have "p_io (p@p3) = butlast io" unfolding \butlast io = p_io p @ ioX\ using \p_io p3 = ioX\ by auto ultimately show ?thesis using that[of p3] by simp qed (* finiteness properties *) have "finite ?R" using finite_R[OF \path M (target (initial M) pP) p\] by assumption moreover have "\ pR . pR \ ?R \ finite (io_targets M' (p_io pR) (initial M'))" using io_targets_finite by metis ultimately have "finite ?Tgts" by blast (* path properties *) obtain pP' p' where "path M' (FSM.initial M') pP'" and "path M' (target (FSM.initial M') pP') p'" and "p_io pP' = p_io pP" and "p_io p' = io" using language_state_split[OF \((p_io pP) @ io) \ L M'\] by blast have "length p \ length (butlast io)" using \butlast io = p_io p @ ioX\ by auto moreover have "length (butlast io) < length io" using \io \ []\ by auto ultimately have "length p < length p'" unfolding \p_io p' = io\ length_map[of "(\ t . (t_input t, t_output t))", symmetric] by simp let ?q = "(target (FSM.initial M') pP')" have "\ pR . pP@pR \ ?R \ path M' ?q (take (length pR) p') \ p_io (take (length pR) p') = p_io pR" proof - fix pR assume "pP@pR \ ?R" then have "pR = take (length pR) p \ length pR \ length p" using R_component(1,2) by metis then have "p_io pR = take (length pR) (butlast io)" unfolding \butlast io = p_io p @ ioX\ by (metis (no_types, lifting) length_map take_le take_map) moreover have "p_io (take (length pR) p') = take (length pR) io" by (metis (full_types) \p_io p' = io\ take_map) moreover have "take (length pR) (butlast io) = take (length pR) io" by (meson \length (butlast io) < length io\ \length p \ length (butlast io)\ \pR = take (length pR) p \ length pR \ length p\ dual_order.strict_trans2 take_butlast) ultimately have "p_io (take (length pR) p') = p_io pR" by simp moreover have "path M' ?q (take (length pR) p')" using \path M' (target (FSM.initial M') pP') p'\ by (simp add: path_prefix_take) ultimately show "path M' ?q (take (length pR) p') \ p_io (take (length pR) p') = p_io pR" by blast qed (* every element in R reaches exactly one state in M' *) have singleton_prop': "\ pR . pP@pR \ ?R \ io_targets M' (p_io (pP@pR)) (initial M') = {target ?q (take (length pR) p')}" proof - fix pR assume "pP@pR \ ?R" then have "path M' ?q (take (length pR) p')" and "p_io (take (length pR) p') = p_io pR" using \\ pR . pP@pR \ ?R \ path M' ?q (take (length pR) p') \ p_io (take (length pR) p') = p_io pR\ by blast+ have *:"path M' (initial M') (pP' @ (take (length pR) p'))" using \path M' (initial M') pP'\ \path M' ?q (take (length pR) p')\ by auto have **:"p_io (pP' @ (take (length pR) p')) = (p_io (pP@pR))" using \p_io pP' = p_io pP\ \p_io (take (length pR) p') = p_io pR\ by auto have "target (initial M') (pP' @ (take (length pR) p')) = target ?q (take (length pR) p')" by auto then have "target ?q (take (length pR) p') \ io_targets M' (p_io (pP@pR)) (initial M')" unfolding io_targets.simps using * ** by (metis (mono_tags, lifting) mem_Collect_eq) show "io_targets M' (p_io (pP@pR)) (initial M') = {target ?q (take (length pR) p')}" using observable_io_targets[OF \observable M'\ language_state_containment[OF * **]] by (metis (no_types) \target (target (FSM.initial M') pP') (take (length pR) p') \ io_targets M' (p_io (pP@pR)) (FSM.initial M')\ singleton_iff) qed have singleton_prop: "\ pR . pR \ ?R \ io_targets M' (p_io pR) (initial M') = {target ?q (take (length pR - length pP) p')}" proof - fix pR assume "pR \ ?R" then obtain pR' where "pR = pP@pR'" using R_component_ob[of _ M "(target (FSM.initial M) pP)" q' pP p] by blast have **: "(length (pP @ pR') - length pP) = length pR'" by auto show "io_targets M' (p_io pR) (initial M') = {target ?q (take (length pR - length pP) p')}" using singleton_prop'[of pR'] \pR \ ?R\ unfolding \pR = pP@pR'\ ** by blast qed then show "\ pR . pR \ ?R \ \ q . io_targets M' (p_io pR) (initial M') = {q}" by blast (* distinct elements in R reach distinct states in M' *) have pairwise_dist_prop': "\ pR1 pR2 . pP@pR1 \ ?R \ pP@pR2 \ ?R \ pR1 \ pR2 \ io_targets M' (p_io (pP@pR1)) (initial M') \ io_targets M' (p_io (pP@pR2)) (initial M') = {}" proof - have diff_prop: "\ pR1 pR2 . pP@pR1 \ ?R \ pP@pR2 \ ?R \ length pR1 < length pR2 \ io_targets M' (p_io (pP@pR1)) (initial M') \ io_targets M' (p_io (pP@pR2)) (initial M') = {}" proof - fix pR1 pR2 assume "pP@pR1 \ ?R" and "pP@pR2 \ ?R" and "length pR1 < length pR2" let ?i = "length pR1 - 1" let ?j = "length pR2 - 1" have "pR1 = take (length pR1) p" and \length pR1 \ length p\ and "t_target (p ! ?i) = q'" using R_component[OF \pP@pR1 \ ?R\] by simp+ have "length pR1 \ 0" using \pP@pR1 \ ?R\ unfolding R_def by simp then have "?i < ?j" using \length pR1 < length pR2\ by (simp add: less_diff_conv) have "pR2 = take (length pR2) p" and \length pR2 \ length p\ and "t_target (p ! ?j) = q'" using R_component[OF \pP@pR2 \ ?R\] by simp+ then have "?j < length (butlast io)" using \length p \ length (butlast io)\ \length pR1 < length pR2\ by linarith have "?q \ io_targets M' (p_io pP) (FSM.initial M')" unfolding \p_io pP' = p_io pP\[symmetric] io_targets.simps using \path M' (initial M') pP'\ by auto have "t_target (p ! ?i) = t_target (p ! ?j)" using \t_target (p ! ?i) = q'\ \t_target (p ! ?j) = q'\ by simp moreover have "(p @ pX) ! ?i = p ! ?i" by (meson \length pR1 < length pR2\ \length pR2 \ length p\ less_imp_diff_less less_le_trans nth_append) moreover have "(p @ pX) ! ?j = p ! ?j" by (metis (no_types) \length pR1 < length pR2\ \pR2 = take (length pR2) p\ diff_less less_imp_diff_less less_nat_zero_code less_numeral_extra(1) not_le_imp_less not_less_iff_gr_or_eq nth_append take_all) ultimately have "t_target (p' ! ?i) \ t_target (p' ! ?j)" using minimal_sequence_to_failure_extending_preamble_no_repetitions_along_path[OF assms(1,2) \path M (target (initial M) pP) (p@pX)\ \p_io (p @ pX) = butlast io\ \?q \ io_targets M' (p_io pP) (FSM.initial M')\ \path M' (target (FSM.initial M') pP') p'\ \p_io p' = io\ \?i < ?j\ \?j < length (butlast io)\ assms(4)] by auto have t1: "io_targets M' (p_io (pP@pR1)) (initial M') = {t_target (p' ! ?i)}" proof - have "(p' ! ?i) = last (take (length pR1) p')" using \length pR1 \ length p\ \length p < length p'\ by (metis Suc_diff_1 \length pR1 \ 0\ dual_order.strict_trans2 length_0_conv length_greater_0_conv less_imp_diff_less take_last_index) then have *: "target (target (FSM.initial M') pP') (take (length pR1) p') = t_target (p' ! ?i)" unfolding target.simps visited_states.simps by (metis (no_types, lifting) \length p < length p'\ \length pR1 \ 0\ gr_implies_not_zero last.simps last_map length_0_conv map_is_Nil_conv take_eq_Nil) have **: "(length (pP @ pR1) - length pP) = length pR1" by auto show ?thesis using singleton_prop[OF \pP@pR1 \ ?R\] unfolding * ** by assumption qed have t2: "io_targets M' (p_io (pP@pR2)) (initial M') = {t_target (p' ! ?j)}" proof - have "(p' ! ?j) = last (take (length pR2) p')" using \length pR2 \ length p\ \length p < length p'\ by (metis Suc_diff_1 \length pR1 - 1 < length pR2 - 1\ le_less_trans less_imp_diff_less linorder_neqE_nat not_less_zero take_last_index zero_less_diff) then have *: "target (target (FSM.initial M') pP') (take (length pR2) p') = t_target (p' ! ?j)" unfolding target.simps visited_states.simps by (metis (no_types, lifting) Nil_is_map_conv \length p < length p'\ \length pR1 < length pR2\ last.simps last_map list.size(3) not_less_zero take_eq_Nil) have **: "(length (pP @ pR2) - length pP) = length pR2" by auto show ?thesis using singleton_prop'[OF \pP@pR2 \ ?R\] unfolding * ** by assumption qed show "io_targets M' (p_io (pP@pR1)) (initial M') \ io_targets M' (p_io (pP@pR2)) (initial M') = {}" using \t_target (p' ! ?i) \ t_target (p' ! ?j)\ unfolding t1 t2 by simp qed fix pR1 pR2 assume "pP@pR1 \ ?R" and "pP@pR2 \ ?R" and "pR1 \ pR2" then have "length pR1 \ length pR2" unfolding R_def by auto then consider (a) "length pR1 < length pR2" | (b) "length pR2 < length pR1" using nat_neq_iff by blast then show "io_targets M' (p_io (pP@pR1)) (initial M') \ io_targets M' (p_io (pP@pR2)) (initial M') = {}" proof cases case a show ?thesis using diff_prop[OF \pP@pR1 \ ?R\ \pP@pR2 \ ?R\ a] by blast next case b show ?thesis using diff_prop[OF \pP@pR2 \ ?R\ \pP@pR1 \ ?R\ b] by blast qed qed then show pairwise_dist_prop: "\ pR1 pR2 . pR1 \ ?R \ pR2 \ ?R \ pR1 \ pR2 \ io_targets M' (p_io pR1) (initial M') \ io_targets M' (p_io pR2) (initial M') = {}" using R_component_ob by (metis (no_types, lifting)) (* combining results *) let ?f = "(\ pR . io_targets M' (p_io pR) (initial M'))" have p1: "(\S1 S2. S1 \ ?R \ S2 \ ?R \ S1 = S2 \ ?f S1 \ ?f S2 = {})" using pairwise_dist_prop by blast have p2: "(\S. S \ R M (target (FSM.initial M) pP) q' pP p \ io_targets M' (p_io S) (FSM.initial M') \ {})" using singleton_prop by blast have c1: "card (R M (target (FSM.initial M) pP) q' pP p) = card ((\S. io_targets M' (p_io S) (FSM.initial M')) ` R M (target (FSM.initial M) pP) q' pP p)" using card_union_of_distinct[of ?R, OF p1 \finite ?R\ p2] by force have p3: "(\S. S \ (\S. io_targets M' (p_io S) (FSM.initial M')) ` R M (target (FSM.initial M) pP) q' pP p \ \t. S = {t})" using singleton_prop by blast have c2:"card ((\S. io_targets M' (p_io S) (FSM.initial M')) ` R M (target (FSM.initial M) pP) q' pP p) = card (\S\R M (target (FSM.initial M) pP) q' pP p. io_targets M' (p_io S) (FSM.initial M'))" using card_union_of_singletons[of "((\S. io_targets M' (p_io S) (FSM.initial M')) ` R M (target (FSM.initial M) pP) q' pP p)", OF p3] by force show "card ?Tgts = card ?R" unfolding c1 c2 by blast qed lemma R_update : "R M q q' pP (p@[t]) = (if (target q (p@[t]) = q') then insert (pP@p@[t]) (R M q q' pP p) else (R M q q' pP p))" (is "?R1 = ?R2") proof (cases "(target q (p@[t]) = q')") case True then have *: "?R2 = insert (pP@p@[t]) (R M q q' pP p)" by auto have "\ p' . p' \ R M q q' pP (p@[t]) \ p' \ insert (pP@p@[t]) (R M q q' pP p)" proof - fix p' assume "p' \ R M q q' pP (p@[t])" obtain p'' where "p' = pP @ p''" using R_component_ob[OF \p' \ R M q q' pP (p@[t])\] by blast obtain p''' where "p'' \ []" and "target q p'' = q'" and "p @ [t] = p'' @ p'''" using \p' \ R M q q' pP (p@[t])\ unfolding R_def \p' = pP @ p''\ by auto show "p' \ insert (pP@p@[t]) (R M q q' pP p)" proof (cases p''' rule: rev_cases) case Nil then have "p' = pP@(p@[t])" using \p' = pP @ p''\ \p @ [t] = p'' @ p'''\ by auto then show ?thesis by blast next case (snoc p'''' t') then have "p = p'' @ p''''" using \p @ [t] = p'' @ p'''\ by auto then show ?thesis unfolding R_def using \p'' \ []\ \target q p'' = q'\ by (simp add: \p' = pP @ p''\) qed qed moreover have "\ p' . p' \ insert (pP@p@[t]) (R M q q' pP p) \ p' \ R M q q' pP (p@[t])" proof - fix p' assume "p' \ insert (pP@p@[t]) (R M q q' pP p)" then consider (a) "p' = (pP@p@[t])" | (b) "p' \ (R M q q' pP p)" by blast then show "p' \ R M q q' pP (p@[t])" proof cases case a then show ?thesis using True unfolding R_def by simp next case b then show ?thesis unfolding R_def using append.assoc by blast qed qed ultimately show ?thesis unfolding * by blast next case False then have *: "?R2 = (R M q q' pP p)" by auto have "\ p' . p' \ R M q q' pP (p@[t]) \ p' \ (R M q q' pP p)" proof - fix p' assume "p' \ R M q q' pP (p@[t])" obtain p'' where "p' = pP @ p''" using R_component_ob[OF \p' \ R M q q' pP (p@[t])\] by blast obtain p''' where "p'' \ []" and "target q p'' = q'" and "p @ [t] = p'' @ p'''" using \p' \ R M q q' pP (p@[t])\ unfolding R_def \p' = pP @ p''\ by blast show "p' \ (R M q q' pP p)" proof (cases p''' rule: rev_cases) case Nil then have "p' = pP@(p@[t])" using \p' = pP @ p''\ \p @ [t] = p'' @ p'''\ by auto then show ?thesis using False \p @ [t] = p'' @ p'''\ \target q p'' = q'\ local.Nil by auto next case (snoc p'''' t') then have "p = p'' @ p''''" using \p @ [t] = p'' @ p'''\ by auto then show ?thesis unfolding R_def using \p'' \ []\ \target q p'' = q'\ by (simp add: \p' = pP @ p''\) qed qed moreover have "\ p' . p' \ (R M q q' pP p) \ p' \ R M q q' pP (p@[t])" proof - fix p' assume "p' \ (R M q q' pP p)" then show "p' \ R M q q' pP (p@[t])" unfolding R_def using append.assoc by blast qed ultimately show ?thesis unfolding * by blast qed lemma R_union_card_is_suffix_length : assumes "path M (initial M) pP" and "path M (target (initial M) pP) p" shows "(\ q \ states M . card (R M (target (initial M) pP) q pP p)) = length p" using assms(2) proof (induction p rule: rev_induct) case Nil have "\ q' . R M (target (initial M) pP) q' pP [] = {}" unfolding R_def by auto then show ?case by simp next case (snoc t p) then have "path M (target (initial M) pP) p" by auto let ?q = "(target (initial M) pP)" let ?q' = "target ?q (p @ [t])" have "\ q . q \ ?q' \ R M ?q q pP (p@[t]) = R M ?q q pP p" using R_update[of M ?q _ pP p t] by force then have *: "(\ q \ states M - {?q'} . card (R M (target (initial M) pP) q pP (p@[t]))) = (\ q \ states M - {?q'} . card (R M (target (initial M) pP) q pP p))" by force have "R M ?q ?q' pP (p@[t]) = insert (pP@p@[t]) (R M ?q ?q' pP p)" using R_update[of M ?q ?q' pP p t] by force moreover have "(pP@p@[t]) \ (R M ?q ?q' pP p)" unfolding R_def by simp ultimately have **: "card (R M (target (initial M) pP) ?q' pP (p@[t])) = Suc (card (R M (target (initial M) pP) ?q' pP p))" using finite_R[OF \path M (target (initial M) pP) (p@[t])\] finite_R[OF \path M (target (initial M) pP) p\] by simp have "?q' \ states M" using path_target_is_state[OF \path M (target (FSM.initial M) pP) (p @ [t])\] by assumption then have ***: "(\ q \ states M . card (R M (target (initial M) pP) q pP (p@[t]))) = (\ q \ states M - {?q'} . card (R M (target (initial M) pP) q pP (p@[t]))) + (card (R M (target (initial M) pP) ?q' pP (p@[t])))" and ****: "(\ q \ states M . card (R M (target (initial M) pP) q pP p)) = (\ q \ states M - {?q'} . card (R M (target (initial M) pP) q pP p)) + (card (R M (target (initial M) pP) ?q' pP p))" by (metis (no_types, lifting) Diff_insert_absorb add.commute finite_Diff fsm_states_finite mk_disjoint_insert sum.insert)+ have "(\ q \ states M . card (R M (target (initial M) pP) q pP (p@[t]))) = Suc (\ q \ states M . card (R M (target (initial M) pP) q pP p))" unfolding **** *** ** * by simp then show ?case unfolding snoc.IH[OF \path M (target (initial M) pP) p\] by auto qed lemma RP_count : assumes "minimal_sequence_to_failure_extending_preamble_path M M' PS pP io" and "observable M" and "observable M'" and "\ q P. (q, P) \ PS \ is_preamble P M q" and "path M (target (initial M) pP) p" and "butlast io = p_io p @ ioX" and "\ q P io x y y' . (q,P) \ PS \ io@[(x,y)] \ L P \ io@[(x,y')] \ L M' \ io@[(x,y')] \ L P" and "completely_specified M'" and "inputs M' = inputs M" shows "card (\ (image (\ pR . io_targets M' (p_io pR) (initial M')) (RP M (target (initial M) pP) q' pP p PS M'))) = card (RP M (target (initial M) pP) q' pP p PS M')" (is "card ?Tgts = card ?RP") and "\ pR . pR \ (RP M (target (initial M) pP) q' pP p PS M') \ \ q . io_targets M' (p_io pR) (initial M') = {q}" and "\ pR1 pR2 . pR1 \ (RP M (target (initial M) pP) q' pP p PS M') \ pR2 \ (RP M (target (initial M) pP) q' pP p PS M') \ pR1 \ pR2 \ io_targets M' (p_io pR1) (initial M') \ io_targets M' (p_io pR2) (initial M') = {}" proof - let ?P1 = "card (\ (image (\ pR . io_targets M' (p_io pR) (initial M')) (RP M (target (initial M) pP) q' pP p PS M'))) = card (RP M (target (initial M) pP) q' pP p PS M')" let ?P2 = "\ pR . pR \ (RP M (target (initial M) pP) q' pP p PS M') \ (\ q . io_targets M' (p_io pR) (initial M') = {q})" let ?P3 = "\ pR1 pR2 . pR1 \ (RP M (target (initial M) pP) q' pP p PS M') \ pR2 \ (RP M (target (initial M) pP) q' pP p PS M') \ pR1 \ pR2 \ io_targets M' (p_io pR1) (initial M') \ io_targets M' (p_io pR2) (initial M') = {}" let ?combined_goals = "?P1 \ ?P2 \ ?P3" let ?q = "(target (initial M) pP)" let ?R = "R M ?q q' pP p" consider (a) "(?RP = ?R)" | (b) "(\ P' pP' . (q',P') \ PS \ path P' (initial P') pP' \ target (initial P') pP' = q' \ path M (initial M) pP' \ target (initial M) pP' = q' \ p_io pP' \ L M' \ ?RP = insert pP' ?R)" using RP_from_R[OF assms(4,7,8,9), of PS _ _ ?q q' pP p] by force then have ?combined_goals proof cases case a show ?thesis unfolding a using R_count[OF assms(1-6)] by blast next case b (* properties from R_count *) have "sequence_to_failure_extending_preamble_path M M' PS pP io" and "\ p' io' . sequence_to_failure_extending_preamble_path M M' PS p' io' \ length io \ length io'" using \minimal_sequence_to_failure_extending_preamble_path M M' PS pP io\ unfolding minimal_sequence_to_failure_extending_preamble_path_def by blast+ obtain q P where "(q,P) \ PS" and "path P (initial P) pP" and "target (initial P) pP = q" and "((p_io pP) @ butlast io) \ L M" and "((p_io pP) @ io) \ L M" and "((p_io pP) @ io) \ L M'" using \sequence_to_failure_extending_preamble_path M M' PS pP io\ unfolding sequence_to_failure_extending_preamble_path_def by blast have "is_preamble P M q" using \(q,P) \ PS\ \\ q P. (q, P) \ PS \ is_preamble P M q\ by blast then have "q \ states M" unfolding is_preamble_def by (metis \path P (FSM.initial P) pP\ \target (FSM.initial P) pP = q\ path_target_is_state submachine_path) have "initial P = initial M" using \is_preamble P M q\ unfolding is_preamble_def by auto have "path M (initial M) pP" using \is_preamble P M q\ unfolding is_preamble_def using submachine_path_initial using \path P (FSM.initial P) pP\ by blast have "target (initial M) pP = q" using \target (initial P) pP = q\ unfolding \initial P = initial M\ by assumption then have "path M q p" using \path M (target (initial M) pP) p\ by auto have "io \ []" using \((p_io pP) @ butlast io) \ L M\ \((p_io pP) @ io) \ L M\ by auto (* finiteness properties *) have "finite ?RP" using finite_RP[OF \path M (target (initial M) pP) p\ assms(4,7,8,9)] by force moreover have "\ pR . pR \ ?RP \ finite (io_targets M' (p_io pR) (initial M'))" using io_targets_finite by metis ultimately have "finite ?Tgts" by blast (* path properties *) obtain pP' p' where "path M' (FSM.initial M') pP'" and "path M' (target (FSM.initial M') pP') p'" and "p_io pP' = p_io pP" and "p_io p' = io" using language_state_split[OF \((p_io pP) @ io) \ L M'\] by blast have "length p \ length (butlast io)" using \butlast io = p_io p @ ioX\ by auto moreover have "length (butlast io) < length io" using \io \ []\ by auto ultimately have "length p < length p'" unfolding \p_io p' = io\ length_map[of "(\ t . (t_input t, t_output t))", symmetric] by simp let ?q = "(target (FSM.initial M') pP')" have "\ pR . pP@pR \ ?R \ path M' ?q (take (length pR) p') \ p_io (take (length pR) p') = p_io pR" proof - fix pR assume "pP@pR \ ?R" then have "pR = take (length pR) p \ length pR \ length p" using R_component(1,2) by metis then have "p_io pR = take (length pR) (butlast io)" by (metis (no_types, lifting) assms(6) length_map take_le take_map) moreover have "p_io (take (length pR) p') = take (length pR) io" by (metis (full_types) \p_io p' = io\ take_map) moreover have "take (length pR) (butlast io) = take (length pR) io" using \length p \ length (butlast io)\ \pR = take (length pR) p \ length pR \ length p\ butlast_take_le dual_order.trans by blast ultimately have "p_io (take (length pR) p') = p_io pR" by simp moreover have "path M' ?q (take (length pR) p')" using \path M' (target (FSM.initial M') pP') p'\ by (simp add: path_prefix_take) ultimately show "path M' ?q (take (length pR) p') \ p_io (take (length pR) p') = p_io pR" by blast qed (* every element in R reaches exactly one state in M' *) have singleton_prop'_R: "\ pR . pP@pR \ ?R \ io_targets M' (p_io (pP@pR)) (initial M') = {target ?q (take (length pR) p')}" proof - fix pR assume "pP@pR \ ?R" then have "path M' ?q (take (length pR) p')" and "p_io (take (length pR) p') = p_io pR" using \\ pR . pP@pR \ ?R \ path M' ?q (take (length pR) p') \ p_io (take (length pR) p') = p_io pR\ by blast+ have *:"path M' (initial M') (pP' @ (take (length pR) p'))" using \path M' (initial M') pP'\ \path M' ?q (take (length pR) p')\ by auto have **:"p_io (pP' @ (take (length pR) p')) = (p_io (pP@pR))" using \p_io pP' = p_io pP\ \p_io (take (length pR) p') = p_io pR\ by auto have "target (initial M') (pP' @ (take (length pR) p')) = target ?q (take (length pR) p')" by auto then have "target ?q (take (length pR) p') \ io_targets M' (p_io (pP@pR)) (initial M')" unfolding io_targets.simps using * ** by (metis (mono_tags, lifting) mem_Collect_eq) show "io_targets M' (p_io (pP@pR)) (initial M') = {target ?q (take (length pR) p')}" using observable_io_targets[OF \observable M'\ language_state_containment[OF * **]] by (metis (no_types) \target (target (FSM.initial M') pP') (take (length pR) p') \ io_targets M' (p_io (pP@pR)) (FSM.initial M')\ singleton_iff) qed have singleton_prop_R: "\ pR . pR \ ?R \ io_targets M' (p_io pR) (initial M') = {target ?q (take (length pR - length pP) p')}" proof - fix pR assume "pR \ ?R" then obtain pR' where "pR = pP@pR'" using R_component_ob[of _ M "(target (FSM.initial M) pP)" q' pP p] by blast have **: "(length (pP @ pR') - length pP) = length pR'" by auto show "io_targets M' (p_io pR) (initial M') = {target ?q (take (length pR - length pP) p')}" using singleton_prop'_R[of pR'] \pR \ ?R\ unfolding \pR = pP@pR'\ ** by blast qed (* RP props *) from b obtain P' pP'' where "(q',P') \ PS" and "path P' (initial P') pP''" and "target (initial P') pP'' = q'" and "path M (initial M) pP''" and "target (initial M) pP'' = q'" and "p_io pP'' \ L M'" and "?RP = insert pP'' ?R" by blast have "initial P' = initial M" using assms(4)[OF \(q',P') \ PS\] unfolding is_preamble_def by auto (* revisiting singleton_prop *) have "\ pR . pR \ ?RP \ pR \ ?R \ pR = pP''" using \?RP = insert pP'' ?R\ by blast then have rp_cases[consumes 1, case_names in_R inserted]: "\ pR P . (pR \ ?RP) \ (pR \ ?R \ P) \ (pR = pP'' \ P) \ P" by force have singleton_prop_RP: "\ pR . pR \ ?RP \ \ q . io_targets M' (p_io pR) (initial M') = {q}" proof - fix pR assume "pR \ ?RP" then show "\ q . io_targets M' (p_io pR) (initial M') = {q}" proof (cases rule: rp_cases) case in_R then show ?thesis using singleton_prop_R by blast next case inserted show ?thesis using observable_io_targets[OF \observable M'\ \p_io pP'' \ L M'\] unfolding inserted by meson qed qed then have ?P2 by blast (* distinctness in RP *) have pairwise_dist_prop_RP: "\ pR1 pR2 . pR1 \ ?RP \ pR2 \ ?RP \ pR1 \ pR2 \ io_targets M' (p_io pR1) (initial M') \ io_targets M' (p_io pR2) (initial M') = {}" proof - have pairwise_dist_prop_R: "\ pR1 pR2 . pR1 \ ?R \ pR2 \ ?R \ pR1 \ pR2 \ io_targets M' (p_io pR1) (initial M') \ io_targets M' (p_io pR2) (initial M') = {}" using R_count(3)[OF assms(1-6)] by force have pairwise_dist_prop_PS: "\ pR1 . pR1 \ ?RP \ pR1 \ pP'' \ io_targets M' (p_io pR1) (initial M') \ io_targets M' (p_io pP'') (initial M') = {}" proof - fix pR1 assume "pR1 \ ?RP" and "pR1 \ pP''" then have "pR1 \ ?R" using \\ pR . pR \ ?RP \ pR \ ?R \ pR = pP''\ by blast obtain pR' where "pR1 = pP@pR'" using R_component_ob[OF \pR1 \ ?R\] by blast then have "pP@pR' \ ?R" using \pR1 \ ?R\ by blast have "pR' = take (length pR') p" and "length pR' \ length p" and "t_target (p ! (length pR' - 1)) = q'" and "pR' \ []" using R_component[OF \pP@pR' \ ?R\] by blast+ let ?i = "(length pR') - 1" have "?i < length p" using \length pR' \ length p\ \pR' \ []\ using diff_less dual_order.strict_trans1 less_numeral_extra(1) by blast then have "?i < length (butlast io)" using \length p \ length (butlast io)\ less_le_trans by blast have "io_targets M' (p_io pR1) (initial M') = {t_target (p' ! ?i)}" proof - have "(p' ! ?i) = last (take (length pR') p')" using \length pR' \ length p\ \length p < length p'\ by (metis Suc_diff_1 \pR' \ []\ dual_order.strict_trans2 length_greater_0_conv less_imp_diff_less take_last_index) then have *: "target ?q (take (length pR') p') = t_target (p' ! ?i)" unfolding target.simps visited_states.simps by (metis (no_types, lifting) \length p < length p'\ \pR' \ []\ gr_implies_not_zero last.simps last_map length_0_conv map_is_Nil_conv take_eq_Nil) moreover have "io_targets M' (p_io pR1) (initial M') = {target ?q (take (length pR') p')}" using singleton_prop'_R \pR1 \ ?R\ unfolding \pR1 = pP@pR'\ by auto ultimately show ?thesis by auto qed have "t_target (p' ! (length pR' - 1)) \ io_targets M' (p_io pP'') (FSM.initial M')" proof - (* get the full path along (butlast io) in M, of which p is a (possibly proper) prefix *) obtain pX where "path M (target (initial M) pP) (p@pX)" and "p_io (p@pX) = butlast io" proof - have "p_io pP @ p_io p @ ioX \ L M" using \((p_io pP) @ butlast io) \ L M\ unfolding \butlast io = p_io p @ ioX\ by assumption obtain p1 p23 where "path M (FSM.initial M) p1" and "path M (target (FSM.initial M) p1) p23" and "p_io p1 = p_io pP" and "p_io p23 = p_io p @ ioX" using language_state_split[OF \p_io pP @ p_io p @ ioX \ L M\] by blast have "p1 = pP" using observable_path_unique[OF \observable M\ \path M (FSM.initial M) p1\ \path M (FSM.initial M) pP\ \p_io p1 = p_io pP\] by assumption then have "path M (target (FSM.initial M) pP) p23" using \path M (target (FSM.initial M) p1) p23\ by auto then have "p_io p @ ioX \ LS M (target (initial M) pP)" using \p_io p23 = p_io p @ ioX\ language_state_containment by auto obtain p2 p3 where "path M (target (FSM.initial M) pP) p2" and "path M (target (target (FSM.initial M) pP) p2) p3" and "p_io p2 = p_io p" and "p_io p3 = ioX" using language_state_split[OF \p_io p @ ioX \ LS M (target (initial M) pP)\] by blast have "p2 = p" using observable_path_unique[OF \observable M\ \path M (target (FSM.initial M) pP) p2\ \path M (target (FSM.initial M) pP) p\ \p_io p2 = p_io p\] by assumption then have "path M (target (FSM.initial M) pP) (p@p3)" using \path M (target (FSM.initial M) pP) p\ \path M (target (target (FSM.initial M) pP) p2) p3\ by auto moreover have "p_io (p@p3) = butlast io" unfolding \butlast io = p_io p @ ioX\ using \p_io p3 = ioX\ by auto ultimately show ?thesis using that[of p3] by simp qed (* get index properties *) have "target (FSM.initial M') pP' \ io_targets M' (p_io pP) (FSM.initial M')" using \p_io pP' = p_io pP\ \path M' (FSM.initial M') pP'\ observable_path_io_target by auto have "(t_target (p ! (length pR' - 1)), P') \ PS" using \(q',P') \ PS\ unfolding \t_target (p ! (length pR' - 1)) = q'\ by assumption then have "(t_target ((p @ pX) ! ?i), P') \ PS" by (metis \length pR' - 1 < length p\ nth_append) have "target (FSM.initial P') pP'' = t_target (p ! (length pR' - 1))" unfolding \target (initial M) pP'' = q'\ \t_target (p ! (length pR' - 1)) = q'\ \initial P' = initial M\ by simp then have "target (FSM.initial P') pP'' = t_target ((p @ pX) ! ?i)" by (metis \length pR' - 1 < length p\ nth_append) show ?thesis using minimal_sequence_to_failure_extending_preamble_no_repetitions_with_other_preambles [OF assms(1,2) \path M (target (initial M) pP) (p@pX)\ \p_io (p@pX) = butlast io\ \target (FSM.initial M') pP' \ io_targets M' (p_io pP) (FSM.initial M')\ \path M' (target (FSM.initial M') pP') p'\ \p_io p' = io\ assms(4) \?i < length (butlast io)\ \(t_target ((p @ pX) ! ?i), P') \ PS\ \path P' (initial P') pP''\ \target (FSM.initial P') pP'' = t_target ((p @ pX) ! ?i)\] by blast qed then show "io_targets M' (p_io pR1) (initial M') \ io_targets M' (p_io pP'') (initial M') = {}" unfolding \io_targets M' (p_io pR1) (initial M') = {t_target (p' ! ?i)}\ by blast qed fix pR1 pR2 assume "pR1 \ ?RP" and "pR2 \ ?RP" and "pR1 \ pR2" then consider (a) "pR1 \ ?R \ pR2 \ ?R" | (b) "pR1 = pP''" | (c) "pR2 = pP''" using \\ pR . pR \ ?RP \ pR \ ?R \ pR = pP''\ \pR1 \ pR2\ by blast then show "io_targets M' (p_io pR1) (initial M') \ io_targets M' (p_io pR2) (initial M') = {}" proof cases case a then show ?thesis using pairwise_dist_prop_R[of pR1 pR2, OF _ _ \pR1 \ pR2\] by blast next case b then show ?thesis using pairwise_dist_prop_PS[OF \pR2 \ ?RP\] \pR1 \ pR2\ by blast next case c then show ?thesis using pairwise_dist_prop_PS[OF \pR1 \ ?RP\] \pR1 \ pR2\ by blast qed qed then have ?P3 by blast let ?f = "(\ pR . io_targets M' (p_io pR) (initial M'))" have p1: "(\S1 S2. S1 \ ?RP \ S2 \ ?RP \ S1 = S2 \ ?f S1 \ ?f S2 = {})" using pairwise_dist_prop_RP by blast have p2: "(\S. S \ ?RP \ io_targets M' (p_io S) (FSM.initial M') \ {})" using singleton_prop_RP by blast have c1: "card ?RP = card ((\S. io_targets M' (p_io S) (FSM.initial M')) ` ?RP)" using card_union_of_distinct[of ?RP, OF p1 \finite ?RP\ p2] by force have p3: "(\S. S \ (\S. io_targets M' (p_io S) (FSM.initial M')) ` ?RP \ \t. S = {t})" using singleton_prop_RP by blast have c2:"card ((\S. io_targets M' (p_io S) (FSM.initial M')) ` ?RP) = card (\S\?RP. io_targets M' (p_io S) (FSM.initial M'))" using card_union_of_singletons[of "((\S. io_targets M' (p_io S) (FSM.initial M')) ` ?RP)", OF p3] by force have ?P1 unfolding c1 c2 by blast show ?combined_goals using \?P1\ \?P2\ \?P3\ by blast qed then show "card (\ (image (\ pR . io_targets M' (p_io pR) (initial M')) (RP M (target (initial M) pP) q' pP p PS M'))) = card (RP M (target (initial M) pP) q' pP p PS M')" and "\ pR . pR \ (RP M (target (initial M) pP) q' pP p PS M') \ \ q . io_targets M' (p_io pR) (initial M') = {q}" and "\ pR1 pR2 . pR1 \ (RP M (target (initial M) pP) q' pP p PS M') \ pR2 \ (RP M (target (initial M) pP) q' pP p PS M') \ pR1 \ pR2 \ io_targets M' (p_io pR1) (initial M') \ io_targets M' (p_io pR2) (initial M') = {}" by blast+ qed lemma RP_target: assumes "pR \ (RP M q q' pP p PS M')" assumes "\ q P . (q,P) \ PS \ is_preamble P M q" and "\ q P io x y y' . (q,P) \ PS \ io@[(x,y)] \ L P \ io@[(x,y')] \ L M' \ io@[(x,y')] \ L P" and "completely_specified M'" and "inputs M' = inputs M" shows "target (initial M) pR = q'" proof - show "target (initial M) pR = q'" proof (cases "pR \ R M q q' pP p") case True then show ?thesis unfolding R_def by force next case False then have "RP M q q' pP p PS M' \ R M q q' pP p" using assms(1) by blast then have "(\P' pP'. (q', P') \ PS \ path P' (FSM.initial P') pP' \ target (FSM.initial P') pP' = q' \ path M (FSM.initial M) pP' \ target (FSM.initial M) pP' = q' \ p_io pP' \ L M' \ RP M q q' pP p PS M' = insert pP' (R M q q' pP p))" using RP_from_R[OF assms(2-5), of PS _ _ q q' pP p] by force then obtain pP' where "target (FSM.initial M) pP' = q'" and "RP M q q' pP p PS M' = insert pP' (R M q q' pP p)" by blast have "pR = pP'" using \RP M q q' pP p PS M' = insert pP' (R M q q' pP p)\ \pR \ (RP M q q' pP p PS M')\ False by blast show ?thesis using \target (FSM.initial M) pP' = q'\ unfolding \pR = pP'\ by assumption qed qed subsubsection \Proof of Exhaustiveness\ lemma passes_test_suite_exhaustiveness_helper_1 : assumes "completely_specified M'" and "inputs M' = inputs M" and "observable M" and "observable M'" and "(q,P) \ PS" and "path P (initial P) pP" and "target (initial P) pP = q" and "p_io pP @ p_io p \ L M'" and "(p, d) \ m_traversal_paths_with_witness M q repetition_sets m" and "implies_completeness_for_repetition_sets (Test_Suite PS tps rd_targets separators) M m repetition_sets" and "passes_test_suite M (Test_Suite PS tps rd_targets separators) M'" and "q' \ q''" and "q' \ fst d" and "q'' \ fst d" and "pR1 \ (RP M q q' pP p PS M')" and "pR2 \ (RP M q q'' pP p PS M')" shows "io_targets M' (p_io pR1) (initial M') \ io_targets M' (p_io pR2) (initial M') = {}" proof - let ?RP1 = "(RP M q q' pP p PS M')" let ?RP2 = "(RP M q q'' pP p PS M')" let ?R1 = "(R M q q' pP p)" let ?R2 = "(R M q q'' pP p)" (* sufficiency properties *) have t1: "(initial M, initial_preamble M) \ PS" using \implies_completeness_for_repetition_sets (Test_Suite PS tps rd_targets separators) M m repetition_sets\ unfolding implies_completeness_for_repetition_sets.simps by blast have t2: "\ q P. (q, P) \ PS \ is_preamble P M q" using \implies_completeness_for_repetition_sets (Test_Suite PS tps rd_targets separators) M m repetition_sets\ unfolding implies_completeness_for_repetition_sets.simps by force have t3: "\ q1 q2 A d1 d2. (A, d1, d2) \ separators (q1, q2) \ (A, d2, d1) \ separators (q2, q1) \ is_separator M q1 q2 A d1 d2" using \implies_completeness_for_repetition_sets (Test_Suite PS tps rd_targets separators) M m repetition_sets\ unfolding implies_completeness_for_repetition_sets.simps by force have t5: "\q. q \ FSM.states M \ (\d\set repetition_sets. q \ fst d)" using \implies_completeness_for_repetition_sets (Test_Suite PS tps rd_targets separators) M m repetition_sets\ unfolding implies_completeness_for_repetition_sets.simps by force have t6: "\ q. q \ fst ` PS \ tps q \ {p1 . \ p2 d . (p1@p2,d) \ m_traversal_paths_with_witness M q repetition_sets m} \ fst ` (m_traversal_paths_with_witness M q repetition_sets m) \ tps q" using \implies_completeness_for_repetition_sets (Test_Suite PS tps rd_targets separators) M m repetition_sets\ unfolding implies_completeness_for_repetition_sets.simps by auto have "\ d. d \ set repetition_sets \ fst d \ FSM.states M \ snd d = fst d \ fst ` PS \ (\q1 q2. q1 \ fst d \ q2 \ fst d \ q1 \ q2 \ separators (q1, q2) \ {})" using \implies_completeness_for_repetition_sets (Test_Suite PS tps rd_targets separators) M m repetition_sets\ unfolding implies_completeness_for_repetition_sets.simps by force then have t7: "\ d. d \ set repetition_sets \ fst d \ FSM.states M" and t8: "\ d. d \ set repetition_sets \ snd d \ fst d" and t8': "\ d. d \ set repetition_sets \ snd d = fst d \ fst ` PS" and t9: "\ d q1 q2. d \ set repetition_sets \ q1 \ fst d \ q2 \ fst d \ q1 \ q2 \ separators (q1, q2) \ {}" by blast+ have t10: "\ q p d p1 p2 p3. q \ fst ` PS \ (p, d) \ m_traversal_paths_with_witness M q repetition_sets m \ p = p1 @ p2 @ p3 \ p2 \ [] \ target q p1 \ fst d \ target q (p1 @ p2) \ fst d \ target q p1 \ target q (p1 @ p2) \ p1 \ tps q \ p1 @ p2 \ tps q \ target q p1 \ rd_targets (q, p1 @ p2) \ target q (p1 @ p2) \ rd_targets (q, p1)" using \implies_completeness_for_repetition_sets (Test_Suite PS tps rd_targets separators) M m repetition_sets\ unfolding implies_completeness_for_repetition_sets.simps by (metis (no_types, lifting)) have t11: "\ q p d p1 p2 q'. q \ fst ` PS \ (p, d) \ m_traversal_paths_with_witness M q repetition_sets m \ p = p1 @ p2 \ q' \ fst ` PS \ target q p1 \ fst d \ q' \ fst d \ target q p1 \ q' \ p1 \ tps q \ [] \ tps q' \ target q p1 \ rd_targets (q', []) \ q' \ rd_targets (q, p1)" using \implies_completeness_for_repetition_sets (Test_Suite PS tps rd_targets separators) M m repetition_sets\ unfolding implies_completeness_for_repetition_sets.simps by (metis (no_types, lifting)) have t12: "\ q p d q1 q2. q \ fst ` PS \ (p, d) \ m_traversal_paths_with_witness M q repetition_sets m \ q1 \ q2 \ q1 \ snd d \ q2 \ snd d \ [] \ tps q1 \ [] \ tps q2 \ q1 \ rd_targets (q2, []) \ q2 \ rd_targets (q1, [])" using \implies_completeness_for_repetition_sets (Test_Suite PS tps rd_targets separators) M m repetition_sets\ unfolding implies_completeness_for_repetition_sets.simps by (metis (no_types, lifting)) (* pass properties *) have pass1: "\ q P io x y y' . (q,P) \ PS \ io@[(x,y)] \ L P \ io@[(x,y')] \ L M' \ io@[(x,y')] \ L P" using \passes_test_suite M (Test_Suite PS tps rd_targets separators) M'\ unfolding passes_test_suite.simps by meson have pass2: "\ q P pP ioT pT x y y' . (q,P) \ PS \ path P (initial P) pP \ target (initial P) pP = q \ pT \ tps q \ ioT@[(x,y)] \ set (prefixes (p_io pT)) \ (p_io pP)@ioT@[(x,y')] \ L M' \ (\ pT' . pT' \ tps q \ ioT@[(x,y')] \ set (prefixes (p_io pT')))" using \passes_test_suite M (Test_Suite PS tps rd_targets separators) M'\ unfolding passes_test_suite.simps by blast have pass3: "\ q P pP pT q' A d1 d2 qT . (q,P) \ PS \ path P (initial P) pP \ target (initial P) pP = q \ pT \ tps q \ (p_io pP)@(p_io pT) \ L M' \ q' \ rd_targets (q,pT) \ (A,d1,d2) \ separators (target q pT, q') \ qT \ io_targets M' ((p_io pP)@(p_io pT)) (initial M') \ pass_separator_ATC M' A qT d2" using \passes_test_suite M (Test_Suite PS tps rd_targets separators) M'\ unfolding passes_test_suite.simps by blast (* additional props *) have "is_preamble P M q" using \(q,P) \ PS\ \\ q P. (q, P) \ PS \ is_preamble P M q\ by blast then have "q \ states M" unfolding is_preamble_def by (metis \path P (FSM.initial P) pP\ \target (FSM.initial P) pP = q\ path_target_is_state submachine_path) have "initial P = initial M" using \is_preamble P M q\ unfolding is_preamble_def by auto have "path M (initial M) pP" using \is_preamble P M q\ submachine_path_initial \path P (FSM.initial P) pP\ unfolding is_preamble_def by blast moreover have "target (initial M) pP = q" using \target (initial P) pP = q\ unfolding \initial P = initial M\ by assumption ultimately have "q \ states M" using path_target_is_state by metis have "q \ fst ` PS" using \(q,P) \ PS\ by force have "d \ set repetition_sets" using \(p, d) \ m_traversal_paths_with_witness M q repetition_sets m\ using m_traversal_paths_with_witness_set[OF t5 t8 \q \ states M\, of m] using find_set by force have "q' \ states M" by (meson \d \ set repetition_sets\ assms(13) subset_iff t7) have "q'' \ states M" by (meson \d \ set repetition_sets\ assms(14) subset_iff t7) have "target (initial M) pR1 = q'" using RP_target[OF \pR1 \ ?RP1\ t2 pass1 \completely_specified M'\ \inputs M' = inputs M\] by force then have "target (initial M) pR1 \ fst d" using \q' \ fst d\ by blast have "target (initial M) pR2 = q''" using RP_target[OF \pR2 \ ?RP2\ t2 pass1 \completely_specified M'\ \inputs M' = inputs M\] by force then have "target (initial M) pR2 \ fst d" using \q'' \ fst d\ by blast have "pR1 \ pR2" using \target (initial M) pR1 = q'\ \target (initial M) pR2 = q''\ \q' \ q''\ by auto obtain A t1 t2 where "(A,t1,t2) \ separators (q',q'')" using t9[OF \d \ set repetition_sets\ \q' \ fst d\ \q'' \ fst d\ \q' \ q''\] by auto have "(A,t2,t1) \ separators (q'',q')" and "is_separator M q' q'' A t1 t2" using t3[OF \(A,t1,t2) \ separators (q',q'')\] by simp+ then have "is_separator M q'' q' A t2 t1" using is_separator_sym by force show "io_targets M' (p_io pR1) (initial M') \ io_targets M' (p_io pR2) (initial M') = {}" proof (rule ccontr) assume "io_targets M' (p_io pR1) (FSM.initial M') \ io_targets M' (p_io pR2) (FSM.initial M') \ {}" then obtain qT where "qT \ io_targets M' (p_io pR1) (FSM.initial M')" and "qT \ io_targets M' (p_io pR2) (FSM.initial M')" by blast then have "qT \ states M'" using path_target_is_state unfolding io_targets.simps by force consider (a) "pR1 \ ?R1 \ pR2 \ ?R2" | (b) "pR1 \ ?R1 \ pR2 \ ?R2" | (c) "pR1 \ ?R1 \ pR2 \ ?R2" | (d) "pR1 \ ?R1 \ pR2 \ ?R2" by blast then show "False" proof cases case a then have "pR1 \ ?R1" and "pR2 \ ?R2" by auto obtain pR1' where "pR1 = pP@pR1'" using R_component_ob[OF \pR1 \ ?R1\] by blast obtain pR2' where "pR2 = pP@pR2'" using R_component_ob[OF \pR2 \ ?R2\] by blast have "pR1' = take (length pR1') p" and "length pR1' \ length p" and "t_target (p ! (length pR1' - 1)) = q'" and "pR1' \ []" using R_component[of pP pR1' M q q' p] \pR1 \ ?R1\ unfolding \pR1 = pP@pR1'\ by blast+ have "pR2' = take (length pR2') p" and "length pR2' \ length p" and "t_target (p ! (length pR2' - 1)) = q''" and "pR2' \ []" using R_component[of pP pR2' M q q'' p] \pR2 \ ?R2\ unfolding \pR2 = pP@pR2'\ by blast+ have "target q pR1' = q'" using \target (initial M) pR1 = q'\ \pR1' \ []\ unfolding target.simps visited_states.simps \pR1 = pP@pR1'\ by simp then have "target q pR1' \ fst d" using \q' \ fst d\ by blast have "target q pR2' = q''" using \target (initial M) pR2 = q''\ \pR2' \ []\ unfolding target.simps visited_states.simps \pR2 = pP@pR2'\ by simp then have "target q pR2' \ fst d" using \q'' \ fst d\ by blast have "pR1' \ pR2'" using \pR1 \ pR2\ unfolding \pR1 = pP@pR1'\ \pR2 = pP@pR2'\ by simp then have "length pR1' \ length pR2'" using \pR1' = take (length pR1') p\ \pR2' = take (length pR2') p\ by auto then consider (a1) "length pR1' < length pR2'" | (a2) "length pR2' < length pR1'" using nat_neq_iff by blast then have "pR1' \ tps q \ pR2' \ tps q \ q' \ rd_targets (q, pR2') \ q'' \ rd_targets (q, pR1')" proof cases case a1 then have "pR2' = pR1' @ (drop (length pR1') pR2')" using \pR1' = take (length pR1') p\ \pR2' = take (length pR2') p\ by (metis append_take_drop_id less_imp_le_nat take_le) then have "p = pR1' @ (drop (length pR1') pR2') @ (drop (length pR2') p)" using \pR2' = take (length pR2') p\ by (metis append.assoc append_take_drop_id) have "(drop (length pR1') pR2') \ []" using a1 \pR2' = take (length pR2') p\ by auto have "target q (pR1' @ drop (length pR1') pR2') \ fst d" using \pR2' = pR1' @ (drop (length pR1') pR2')\[symmetric] \target q pR2' \ fst d\ by auto show ?thesis using t10[OF \q \ fst ` PS\ \(p, d) \ m_traversal_paths_with_witness M q repetition_sets m\ \p = pR1' @ (drop (length pR1') pR2') @ (drop (length pR2') p)\ \(drop (length pR1') pR2') \ []\ \target q pR1' \ fst d\ \target q (pR1' @ drop (length pR1') pR2') \ fst d\] unfolding \pR2' = pR1' @ (drop (length pR1') pR2')\[symmetric] \target q pR1' = q'\ \target q pR2' = q''\ using \q' \ q''\ by blast next case a2 then have "pR1' = pR2' @ (drop (length pR2') pR1')" using \pR1' = take (length pR1') p\ \pR2' = take (length pR2') p\ by (metis append_take_drop_id less_imp_le_nat take_le) then have "p = pR2' @ (drop (length pR2') pR1') @ (drop (length pR1') p)" using \pR1' = take (length pR1') p\ by (metis append.assoc append_take_drop_id) have "(drop (length pR2') pR1') \ []" using a2 \pR1' = take (length pR1') p\ by auto have "target q (pR2' @ drop (length pR2') pR1') \ fst d" using \pR1' = pR2' @ (drop (length pR2') pR1')\[symmetric] \target q pR1' \ fst d\ by auto show ?thesis using t10[OF \q \ fst ` PS\ \(p, d) \ m_traversal_paths_with_witness M q repetition_sets m\ \p = pR2' @ (drop (length pR2') pR1') @ (drop (length pR1') p)\ \(drop (length pR2') pR1') \ []\ \target q pR2' \ fst d\ \target q (pR2' @ drop (length pR2') pR1') \ fst d\] unfolding \pR1' = pR2' @ (drop (length pR2') pR1')\[symmetric] \target q pR1' = q'\ \target q pR2' = q''\ using \q' \ q''\ by blast qed then have "pR1' \ tps q" and "pR2' \ tps q" and "q' \ rd_targets (q, pR2')" and "q'' \ rd_targets (q, pR1')" by simp+ have "p_io pP @ p_io pR1' \ L M'" using language_prefix_append[OF \p_io pP @ p_io p \ L M'\, of "length pR1'"] using \pR1' = take (length pR1') p\ by simp have "pass_separator_ATC M' A qT t2" using pass3[OF \(q, P) \ PS\ \path P (initial P) pP\ \target (initial P) pP = q\ \pR1' \ tps q\ \p_io pP @ p_io pR1' \ L M'\ \q'' \ rd_targets (q, pR1')\, of A t1 t2] \(A, t1, t2) \ separators (q', q'')\ \qT \ io_targets M' (p_io pR1) (FSM.initial M')\ unfolding \target q pR1' = q'\ \pR1 = pP @ pR1'\ by auto have "p_io pP @ p_io pR2' \ L M'" using language_prefix_append[OF \p_io pP @ p_io p \ L M'\, of "length pR2'"] using \pR2' = take (length pR2') p\ by simp have "pass_separator_ATC M' A qT t1" using pass3[OF \(q, P) \ PS\ \path P (initial P) pP\ \target (initial P) pP = q\ \pR2' \ tps q\ \p_io pP @ p_io pR2' \ L M'\ \q' \ rd_targets (q, pR2')\, of A t2 t1] \(A, t2, t1) \ separators (q'', q')\ \qT \ io_targets M' (p_io pR2) (FSM.initial M')\ unfolding \target q pR2' = q''\ \pR2 = pP @ pR2'\ by auto (* the state qT reached by both paths cannot satisfy the ATC that r-distinguishes their respective targets for both targets at the same time *) have "qT \ qT" using pass_separator_ATC_reduction_distinction[OF \observable M\ \observable M'\ \inputs M' = inputs M\ \pass_separator_ATC M' A qT t2\ \pass_separator_ATC M' A qT t1\ \q' \ states M\ \q'' \ states M\ \q' \ q''\ \qT \ states M'\ \qT \ states M'\ \is_separator M q' q'' A t1 t2\ \completely_specified M'\] by assumption then show False by simp next case b then have "pR1 \ ?R1" and "pR2 \ ?R2" using \pR1 \ ?RP1\ by auto obtain pR1' where "pR1 = pP@pR1'" using R_component_ob[OF \pR1 \ ?R1\] by blast have "pR1' = take (length pR1') p" and "length pR1' \ length p" and "t_target (p ! (length pR1' - 1)) = q'" and "pR1' \ []" using R_component[of pP pR1' M q q' p] \pR1 \ ?R1\ unfolding \pR1 = pP@pR1'\ by blast+ have "target q pR1' = q'" using \target (initial M) pR1 = q'\ \pR1' \ []\ unfolding target.simps visited_states.simps \pR1 = pP@pR1'\ by simp then have "target q pR1' \ fst d" and "target q pR1' \ q''" using \q' \ fst d\ \q' \ q''\ by blast+ obtain P' where "(q'', P') \ PS" "path P' (FSM.initial P') pR2" "target (FSM.initial P') pR2 = q''" "path M (FSM.initial M) pR2" "target (FSM.initial M) pR2 = q''" "p_io pR2 \ L M'" "RP M q q'' pP p PS M' = insert pR2 (R M q q'' pP p)" using RP_from_R_inserted[OF t2 pass1 \completely_specified M'\ \inputs M' = inputs M\ \pR2 \ ?RP2\ \pR2 \ ?R2\, of "\ q P io x y y' . q" "\ q P io x y y' . y"] by blast have "q'' \ fst ` PS" using \(q'',P') \ PS\ by force have "p = pR1' @ (drop (length pR1') p)" using \pR1' = take (length pR1') p\ by (metis append_take_drop_id) have "pR1' \ tps q" and "[] \ tps q''" and "target q pR1' \ rd_targets (q'', [])" and "q'' \ rd_targets (q, pR1')" using t11[OF \q \ fst ` PS\ \(p, d) \ m_traversal_paths_with_witness M q repetition_sets m\ \p = pR1' @ (drop (length pR1') p)\ \q'' \ fst ` PS\ \target q pR1' \ fst d\ \q'' \ fst d\ \target q pR1' \ q''\] by simp+ have "p_io pP @ p_io pR1' \ L M'" using language_prefix_append[OF \p_io pP @ p_io p \ L M'\, of "length pR1'"] using \pR1' = take (length pR1') p\ by simp have "pass_separator_ATC M' A qT t2" using pass3[OF \(q, P) \ PS\ \path P (initial P) pP\ \target (initial P) pP = q\ \pR1' \ tps q\ \p_io pP @ p_io pR1' \ L M'\ \q'' \ rd_targets (q, pR1')\, of A t1 t2] \(A, t1, t2) \ separators (q', q'')\ \qT \ io_targets M' (p_io pR1) (FSM.initial M')\ unfolding \target q pR1' = q'\ \pR1 = pP @ pR1'\ by auto have "pass_separator_ATC M' A qT t1" using pass3[OF \(q'', P') \ PS\ \path P' (FSM.initial P') pR2\ \target (FSM.initial P') pR2 = q''\ \[] \ tps q''\ _ \target q pR1' \ rd_targets (q'', [])\, of A t2 t1 qT] \(A, t2, t1) \ separators (q'', q')\ \qT \ io_targets M' (p_io pR2) (FSM.initial M')\ \p_io pR2 \ L M'\ unfolding \target q pR1' = q'\ by auto have "qT \ qT" using pass_separator_ATC_reduction_distinction[OF \observable M\ \observable M'\ \inputs M' = inputs M\ \pass_separator_ATC M' A qT t2\ \pass_separator_ATC M' A qT t1\ \q' \ states M\ \q'' \ states M\ \q' \ q''\ \qT \ states M'\ \qT \ states M'\ \is_separator M q' q'' A t1 t2\ \completely_specified M'\] by assumption then show False by simp next case c then have "pR2 \ ?R2" and "pR1 \ ?R1" using \pR2 \ ?RP2\ by auto obtain pR2' where "pR2 = pP@pR2'" using R_component_ob[OF \pR2 \ ?R2\] by blast have "pR2' = take (length pR2') p" and "length pR2' \ length p" and "t_target (p ! (length pR2' - 1)) = q''" and "pR2' \ []" using R_component[of pP pR2' M q q'' p] \pR2 \ ?R2\ unfolding \pR2 = pP@pR2'\ by blast+ have "target q pR2' = q''" using \target (initial M) pR2 = q''\ \pR2' \ []\ unfolding target.simps visited_states.simps \pR2 = pP@pR2'\ by simp then have "target q pR2' \ fst d" and "target q pR2' \ q'" using \q'' \ fst d\ \q' \ q''\ by blast+ obtain P' where "(q', P') \ PS" "path P' (FSM.initial P') pR1" "target (FSM.initial P') pR1 = q'" "path M (FSM.initial M) pR1" "target (FSM.initial M) pR1 = q'" "p_io pR1 \ L M'" "RP M q q' pP p PS M' = insert pR1 (R M q q' pP p)" using RP_from_R_inserted[OF t2 pass1 \completely_specified M'\ \inputs M' = inputs M\ \pR1 \ ?RP1\ \pR1 \ ?R1\, of "\ q P io x y y' . q" "\ q P io x y y' . y"] by blast have "q' \ fst ` PS" using \(q',P') \ PS\ by force have "p = pR2' @ (drop (length pR2') p)" using \pR2' = take (length pR2') p\ by (metis append_take_drop_id) have "pR2' \ tps q" and "[] \ tps q'" and "target q pR2' \ rd_targets (q', [])" and "q' \ rd_targets (q, pR2')" using t11[OF \q \ fst ` PS\ \(p, d) \ m_traversal_paths_with_witness M q repetition_sets m\ \p = pR2' @ (drop (length pR2') p)\ \q' \ fst ` PS\ \target q pR2' \ fst d\ \q' \ fst d\ \target q pR2' \ q'\] by simp+ have "p_io pP @ p_io pR2' \ L M'" using language_prefix_append[OF \p_io pP @ p_io p \ L M'\, of "length pR2'"] using \pR2' = take (length pR2') p\ by simp have "pass_separator_ATC M' A qT t1" using pass3[OF \(q, P) \ PS\ \path P (initial P) pP\ \target (initial P) pP = q\ \pR2' \ tps q\ \p_io pP @ p_io pR2' \ L M'\ \q' \ rd_targets (q, pR2')\, of A t2 t1] \(A, t2, t1) \ separators (q'', q')\ \qT \ io_targets M' (p_io pR2) (FSM.initial M')\ unfolding \target q pR2' = q''\ \pR2 = pP @ pR2'\ by auto have "pass_separator_ATC M' A qT t2" using pass3[OF \(q', P') \ PS\ \path P' (FSM.initial P') pR1\ \target (FSM.initial P') pR1 = q'\ \[] \ tps q'\ _ \target q pR2' \ rd_targets (q', [])\, of A t1 t2 qT] \(A, t1, t2) \ separators (q', q'')\ \qT \ io_targets M' (p_io pR1) (FSM.initial M')\ \p_io pR1 \ L M'\ unfolding \target q pR2' = q''\ by auto have "qT \ qT" using pass_separator_ATC_reduction_distinction[OF \observable M\ \observable M'\ \inputs M' = inputs M\ \pass_separator_ATC M' A qT t1\ \pass_separator_ATC M' A qT t2\ \q'' \ states M\ \q' \ states M\ _ \qT \ states M'\ \qT \ states M'\ \is_separator M q'' q' A t2 t1\ \completely_specified M'\] \q' \ q''\ by simp then show False by simp next case d then have "pR1 \ ?R1" and "pR2 \ ?R2" by auto obtain P' where "(q', P') \ PS" "path P' (FSM.initial P') pR1" "target (FSM.initial P') pR1 = q'" "path M (FSM.initial M) pR1" "target (FSM.initial M) pR1 = q'" "p_io pR1 \ L M'" "RP M q q' pP p PS M' = insert pR1 (R M q q' pP p)" using RP_from_R_inserted[OF t2 pass1 \completely_specified M'\ \inputs M' = inputs M\ \pR1 \ ?RP1\ \pR1 \ ?R1\, of "\ q P io x y y' . q" "\ q P io x y y' . y"] by blast have "q' \ snd d" by (metis IntI \(q', P') \ PS\ \d \ set repetition_sets\ assms(13) fst_eqD image_eqI t8') obtain P'' where "(q'', P'') \ PS" "path P'' (FSM.initial P'') pR2" "target (FSM.initial P'') pR2 = q''" "path M (FSM.initial M) pR2" "target (FSM.initial M) pR2 = q''" "p_io pR2 \ L M'" "RP M q q'' pP p PS M' = insert pR2 (R M q q'' pP p)" using RP_from_R_inserted[OF t2 pass1 \completely_specified M'\ \inputs M' = inputs M\ \pR2 \ ?RP2\ \pR2 \ ?R2\, of "\ q P io x y y' . q" "\ q P io x y y' . y"] by blast have "q'' \ snd d" by (metis IntI \(q'', P'') \ PS\ \d \ set repetition_sets\ assms(14) fst_eqD image_eqI t8') have "[] \ tps q'" and "[] \ tps q''" and "q' \ rd_targets (q'', [])" and "q'' \ rd_targets (q', [])" using t12[OF \q \ fst ` PS\ \(p, d) \ m_traversal_paths_with_witness M q repetition_sets m\ \q' \ q''\ \q' \ snd d\ \q'' \ snd d\] by simp+ have "pass_separator_ATC M' A qT t1" using pass3[OF \(q'', P'') \ PS\ \path P'' (initial P'') pR2\ \target (initial P'') pR2 = q''\ \[] \ tps q''\ _ \q' \ rd_targets (q'', [])\, of A t2 t1 qT] \p_io pR2 \ L M'\ \(A, t2, t1) \ separators (q'', q')\ \qT \ io_targets M' (p_io pR2) (FSM.initial M')\ by auto have "pass_separator_ATC M' A qT t2" using pass3[OF \(q', P') \ PS\ \path P' (initial P') pR1\ \target (initial P') pR1 = q'\ \[] \ tps q'\ _ \q'' \ rd_targets (q', [])\, of A t1 t2 qT] \p_io pR1 \ L M'\ \(A, t1, t2) \ separators (q', q'')\ \qT \ io_targets M' (p_io pR1) (FSM.initial M')\ by auto have "qT \ qT" using pass_separator_ATC_reduction_distinction[OF \observable M\ \observable M'\ \inputs M' = inputs M\ \pass_separator_ATC M' A qT t1\ \pass_separator_ATC M' A qT t2\ \q'' \ states M\ \q' \ states M\ _ \qT \ states M'\ \qT \ states M'\ \is_separator M q'' q' A t2 t1\ \completely_specified M'\] \q' \ q''\ by simp then show False by simp qed qed qed lemma passes_test_suite_exhaustiveness : assumes "passes_test_suite M (Test_Suite prs tps rd_targets separators) M'" and "implies_completeness (Test_Suite prs tps rd_targets separators) M m" and "observable M" and "observable M'" and "inputs M' = inputs M" and "inputs M \ {}" and "completely_specified M" and "completely_specified M'" and "size M' \ m" shows "L M' \ L M" proof (rule ccontr) assume "\ L M' \ L M" (* sufficiency properties *) obtain repetition_sets where repetition_sets_def: "implies_completeness_for_repetition_sets (Test_Suite prs tps rd_targets separators) M m repetition_sets" using assms(2) unfolding implies_completeness_def by blast have t1: "(initial M, initial_preamble M) \ prs" using implies_completeness_for_repetition_sets_simps(1)[OF repetition_sets_def] by assumption have t2: "\ q P. (q, P) \ prs \ is_preamble P M q" using implies_completeness_for_repetition_sets_simps(2)[OF repetition_sets_def] by blast have t3: "\ q1 q2 A d1 d2. (A, d1, d2) \ separators (q1, q2) \ (A, d2, d1) \ separators (q2, q1) \ is_separator M q1 q2 A d1 d2" using implies_completeness_for_repetition_sets_simps(3)[OF repetition_sets_def] by assumption have t5: "\q. q \ FSM.states M \ (\d\set repetition_sets. q \ fst d)" using implies_completeness_for_repetition_sets_simps(4)[OF repetition_sets_def] by assumption have t6: "\ q. q \ fst ` prs \ tps q \ {p1 . \ p2 d . (p1@p2,d) \ m_traversal_paths_with_witness M q repetition_sets m} \ fst ` (m_traversal_paths_with_witness M q repetition_sets m) \ tps q" using implies_completeness_for_repetition_sets_simps(7)[OF repetition_sets_def] by assumption have t7: "\ d. d \ set repetition_sets \ fst d \ FSM.states M" and t8: "\ d. d \ set repetition_sets \ snd d \ fst d" and t8': "\ d. d \ set repetition_sets \ snd d = fst d \ fst ` prs" and t9: "\ d q1 q2. d \ set repetition_sets \ q1 \ fst d \ q2 \ fst d \ q1 \ q2 \ separators (q1, q2) \ {}" using implies_completeness_for_repetition_sets_simps(5,6)[OF repetition_sets_def] by blast+ have t10: "\ q p d p1 p2 p3. q \ fst ` prs \ (p, d) \ m_traversal_paths_with_witness M q repetition_sets m \ p = p1 @ p2 @ p3 \ p2 \ [] \ target q p1 \ fst d \ target q (p1 @ p2) \ fst d \ target q p1 \ target q (p1 @ p2) \ p1 \ tps q \ p1 @ p2 \ tps q \ target q p1 \ rd_targets (q, p1 @ p2) \ target q (p1 @ p2) \ rd_targets (q, p1)" using implies_completeness_for_repetition_sets_simps(8)[OF repetition_sets_def] by assumption have t11: "\ q p d p1 p2 q'. q \ fst ` prs \ (p, d) \ m_traversal_paths_with_witness M q repetition_sets m \ p = p1 @ p2 \ q' \ fst ` prs \ target q p1 \ fst d \ q' \ fst d \ target q p1 \ q' \ p1 \ tps q \ [] \ tps q' \ target q p1 \ rd_targets (q', []) \ q' \ rd_targets (q, p1)" using implies_completeness_for_repetition_sets_simps(9)[OF repetition_sets_def] by assumption have t12: "\ q p d q1 q2. q \ fst ` prs \ (p, d) \ m_traversal_paths_with_witness M q repetition_sets m \ q1 \ q2 \ q1 \ snd d \ q2 \ snd d \ [] \ tps q1 \ [] \ tps q2 \ q1 \ rd_targets (q2, []) \ q2 \ rd_targets (q1, [])" using implies_completeness_for_repetition_sets_simps(10)[OF repetition_sets_def] by assumption (* pass properties *) have pass1: "\ q P io x y y' . (q,P) \ prs \ io@[(x,y)] \ L P \ io@[(x,y')] \ L M' \ io@[(x,y')] \ L P" using \passes_test_suite M (Test_Suite prs tps rd_targets separators) M'\ unfolding passes_test_suite.simps by meson have pass2: "\ q P pP ioT pT x y y' . (q,P) \ prs \ path P (initial P) pP \ target (initial P) pP = q \ pT \ tps q \ ioT@[(x,y)] \ set (prefixes (p_io pT)) \ (p_io pP)@ioT@[(x,y')] \ L M' \ (\ pT' . pT' \ tps q \ ioT@[(x,y')] \ set (prefixes (p_io pT')))" using \passes_test_suite M (Test_Suite prs tps rd_targets separators) M'\ unfolding passes_test_suite.simps by blast have pass3: "\ q P pP pT q' A d1 d2 qT . (q,P) \ prs \ path P (initial P) pP \ target (initial P) pP = q \ pT \ tps q \ (p_io pP)@(p_io pT) \ L M' \ q' \ rd_targets (q,pT) \ (A,d1,d2) \ separators (target q pT, q') \ qT \ io_targets M' ((p_io pP)@(p_io pT)) (initial M') \ pass_separator_ATC M' A qT d2" using \passes_test_suite M (Test_Suite prs tps rd_targets separators) M'\ unfolding passes_test_suite.simps by blast (* seq to failure *) obtain pP io where "minimal_sequence_to_failure_extending_preamble_path M M' prs pP io" using minimal_sequence_to_failure_extending_preamble_ex[OF t1 \\ L M' \ L M\] by blast then have "sequence_to_failure_extending_preamble_path M M' prs pP io" "\ io'. sequence_to_failure_extending_preamble_path M M' prs pP io' \ length io \ length io'" unfolding minimal_sequence_to_failure_extending_preamble_path_def by blast+ obtain q P where "q \ states M" and "(q,P) \ prs" and "path P (initial P) pP" and "target (initial P) pP = q" and "((p_io pP) @ butlast io) \ L M" and "((p_io pP) @ io) \ L M" and "((p_io pP) @ io) \ L M'" using \sequence_to_failure_extending_preamble_path M M' prs pP io\ unfolding sequence_to_failure_extending_preamble_path_def by blast let ?xF = "fst (last io)" let ?yF = "snd (last io)" let ?xyF = "(?xF,?yF)" let ?ioF = "butlast io" have "io \ []" using \((p_io pP) @ io) \ L M\ \((p_io pP) @ butlast io) \ L M\ by auto then have "io = ?ioF@[?xyF]" by auto have "?xF \ inputs M'" using language_io(1)[OF \((p_io pP) @ io) \ L M'\, of ?xF ?yF] \io \ []\ by auto then have "?xF \ inputs M" using \inputs M' = inputs M\ by simp have "q \ fst ` prs" using \(q,P) \ prs\ by force have "is_preamble P M q" using \(q,P) \ prs\ t2 by blast then have "q \ states M" unfolding is_preamble_def by (metis \path P (FSM.initial P) pP\ \target (FSM.initial P) pP = q\ path_target_is_state submachine_path) have "initial P = initial M" using \is_preamble P M q\ unfolding is_preamble_def by auto have "path M (initial M) pP" using \is_preamble P M q\ unfolding is_preamble_def using submachine_path_initial using \path P (FSM.initial P) pP\ by blast have "target (initial M) pP = q" using \target (initial P) pP = q\ unfolding \initial P = initial M\ by assumption (* io must be a proper extension of some m-traversal path, as all m-traversal paths pass *) obtain pM dM ioEx where "(pM,dM) \ m_traversal_paths_with_witness M q repetition_sets m" and "io = (p_io pM)@ioEx" and "ioEx \ []" proof - obtain pF where "path M q pF" and "p_io pF = ?ioF" using observable_path_suffix[OF \((p_io pP) @ ?ioF) \ L M\ \path M (initial M) pP\ \observable M\ ] unfolding \target (initial M) pP = q\ by blast obtain tM where "tM \ transitions M" and "t_source tM = target q pF" and "t_input tM = ?xF" using \?xF \ inputs M\ path_target_is_state[OF \path M q pF\] \completely_specified M\ unfolding completely_specified.simps by blast then have "path M q (pF@[tM])" using \path M q pF\ path_append_transition by simp show ?thesis proof (cases "find (\d. Suc (m - card (snd d)) \ length (filter (\t. t_target t \ fst d) (pF@[tM]))) repetition_sets") case None (* if no repetition_sets witness exists for (pF@[tM]), then it is a proper prefix of some m-traversal path and hence also in L M, providing a contradiction*) obtain pF' d' where "((pF@[tM]) @ pF', d') \ m_traversal_paths_with_witness M q repetition_sets m" using m_traversal_path_extension_exist[OF \completely_specified M\ \q \ states M\ \inputs M \ {}\ t5 t8 \path M q (pF@[tM])\ None] by blast then have "(pF@[tM]) @ pF' \ tps q" using t6[OF \q \ fst ` prs\] by force have "(p_io pF) @ [(?xF,t_output tM)] \ set (prefixes (p_io ((pF@[tM])@pF')))" using \t_input tM = ?xF\ unfolding prefixes_set by auto have "p_io pP @ p_io pF @ [?xyF] \ L M'" using \((p_io pP) @ io) \ L M'\ unfolding \p_io pF = ?ioF\ \io = ?ioF@[?xyF]\[symmetric] by assumption obtain pT' where "pT' \ tps q" and "p_io pF @ [(fst (last io), snd (last io))] \ set (prefixes (p_io pT'))" using pass2[OF \(q,P) \ prs\ \path P (initial P) pP\ \target (initial P) pP = q\ \(pF@[tM]) @ pF' \ tps q\ \(p_io pF) @ [(?xF,t_output tM)] \ set (prefixes (p_io ((pF@[tM])@pF')))\ \p_io pP @ p_io pF @ [?xyF] \ L M'\] by blast have "path M q pT'" proof - obtain pT'' d'' where "(pT'@pT'', d'') \ m_traversal_paths_with_witness M q repetition_sets m" using \pT' \ tps q\ t6[OF \q \ fst ` prs\] by blast then have "path M q (pT'@pT'')" using m_traversal_paths_with_witness_set[OF t5 t8 \q \ states M\] by force then show ?thesis by auto qed then have "path M (initial M) (pP@pT')" using \path M (initial M) pP\ \target (initial M) pP = q\ by auto then have "(p_io (pP@pT')) \ L M" unfolding LS.simps by blast then have "(p_io pP)@(p_io pT') \ L M" by auto have "io \ set (prefixes (p_io pT'))" using \p_io pF @ [(fst (last io), snd (last io))] \ set (prefixes (p_io pT'))\ unfolding \p_io pF = ?ioF\ \io = ?ioF@[?xyF]\[symmetric] by assumption then obtain io' where "p_io pT' = io @ io'" - unfolding prefixes_set by moura + unfolding prefixes_set mem_Collect_eq by metis have " p_io pP @ io \ L M" using \(p_io pP)@(p_io pT') \ L M\ unfolding \p_io pT' = io @ io'\ unfolding append.assoc[symmetric] using language_prefix[of "p_io pP @ io" io', of M "initial M"] by blast then show ?thesis using \(p_io pP) @ io \ L M\ by simp next case (Some d) (* get the shortest prefix of pF that still has a RepSet witness *) let ?ps = "{ p1 . \ p2 . (pF@[tM]) = p1 @ p2 \ find (\d. Suc (m - card (snd d)) \ length (filter (\t. t_target t \ fst d) p1)) repetition_sets \ None}" have "finite ?ps" proof - have "?ps \ set (prefixes (pF@[tM]))" unfolding prefixes_set by force moreover have "finite (set (prefixes (pF@[tM])))" by simp ultimately show ?thesis by (simp add: finite_subset) qed moreover have "?ps \ {}" proof - have "pF @ [tM] = (pF @ [tM]) @ [] \ find (\d. Suc (m - card (snd d)) \ length (filter (\t. t_target t \ fst d) (pF @ [tM]))) repetition_sets \ None" using Some by auto then have "(pF@[tM]) \ ?ps" by blast then show ?thesis by blast qed ultimately obtain pMin where "pMin \ ?ps" and "\ p' . p' \ ?ps \ length pMin \ length p'" by (meson leI min_length_elem) obtain pMin' dMin where "(pF@[tM]) = pMin @ pMin'" and "find (\d. Suc (m - card (snd d)) \ length (filter (\t. t_target t \ fst d) pMin)) repetition_sets = Some dMin" using \pMin \ ?ps\ by blast then have "path M q pMin" using \path M q (pF@[tM])\ by auto moreover have "(\p' p''. pMin = p' @ p'' \ p'' \ [] \ find (\d. Suc (m - card (snd d)) \ length (filter (\t. t_target t \ fst d) p')) repetition_sets = None)" proof - have "\ p' p''. pMin = p' @ p'' \ p'' \ [] \ find (\d. Suc (m - card (snd d)) \ length (filter (\t. t_target t \ fst d) p')) repetition_sets = None" proof - fix p' p'' assume "pMin = p' @ p''" and "p'' \ []" show "find (\d. Suc (m - card (snd d)) \ length (filter (\t. t_target t \ fst d) p')) repetition_sets = None" proof (rule ccontr) assume "find (\d. Suc (m - card (snd d)) \ length (filter (\t. t_target t \ fst d) p')) repetition_sets \ None" then have "p' \ ?ps" using \(pF@[tM]) = pMin @ pMin'\ unfolding \pMin = p' @ p''\ append.assoc by blast have "length p' < length pMin" using \pMin = p' @ p''\ \p'' \ []\ by auto then show "False" using \\ p' . p' \ ?ps \ length pMin \ length p'\[OF \p' \ ?ps\] by simp qed qed then show ?thesis by blast qed ultimately have "(pMin,dMin) \ m_traversal_paths_with_witness M q repetition_sets m" using \find (\d. Suc (m - card (snd d)) \ length (filter (\t. t_target t \ fst d) pMin)) repetition_sets = Some dMin\ m_traversal_paths_with_witness_set[OF t5 t8 \q \ states M\, of m] by blast then have "pMin \ tps q" using t6[OF \q \ fst ` prs\] by force show ?thesis proof (cases "pMin = (pF@[tM])") (* if the m-traversal path is not shorter than io, then again the failure is observed *) case True then have "?ioF @ [(?xF, t_output tM)] \ set (prefixes (p_io pMin))" using \p_io pF = ?ioF\ \t_input tM = ?xF\ unfolding prefixes_set by force obtain pMinF where "pMinF \ tps q" and "io \ set (prefixes (p_io pMinF))" using pass2[OF \(q,P) \ prs\ \path P (initial P) pP\ \target (initial P) pP = q\ \pMin \ tps q\ \?ioF @ [(?xF, t_output tM)] \ set (prefixes (p_io pMin))\, of ?yF] using \p_io pP @ io \ L M'\ unfolding \io = ?ioF@[?xyF]\[symmetric] by blast have "path M q pMinF" proof - obtain pT'' d'' where "(pMinF@pT'', d'') \ m_traversal_paths_with_witness M q repetition_sets m" using \pMinF \ tps q\ t6[OF \q \ fst ` prs\] by blast then have "path M q (pMinF@pT'')" using m_traversal_paths_with_witness_set[OF t5 t8 \q \ states M\] by force then show ?thesis by auto qed then have "path M (initial M) (pP@pMinF)" using \path M (initial M) pP\ \target (initial M) pP = q\ by auto then have "(p_io (pP@pMinF)) \ L M" unfolding LS.simps by blast then have "(p_io pP)@(p_io pMinF) \ L M" by auto obtain io' where "p_io pMinF = io @ io'" - using \io \ set (prefixes (p_io pMinF))\ unfolding prefixes_set by moura + using \io \ set (prefixes (p_io pMinF))\ + unfolding prefixes_set mem_Collect_eq by metis have " p_io pP @ io \ L M" using \(p_io pP)@(p_io pMinF) \ L M\ unfolding \p_io pMinF = io @ io'\ unfolding append.assoc[symmetric] using language_prefix[of "p_io pP @ io" io', of M "initial M"] by blast then show ?thesis using \(p_io pP) @ io \ L M\ by simp next case False then obtain pMin'' where "pF = pMin @ pMin''" using \(pF@[tM]) = pMin @ pMin'\ by (metis butlast_append butlast_snoc) then have "io = (p_io pMin) @ (p_io pMin'') @ [?xyF]" using \io = ?ioF @ [?xyF]\ \p_io pF = ?ioF\ by (metis (no_types, lifting) append_assoc map_append) then show ?thesis using that[OF \(pMin,dMin) \ m_traversal_paths_with_witness M q repetition_sets m\, of "(p_io pMin'') @ [?xyF]"] by auto qed qed qed (* collect length properties on pM to ultimately show that M' must have at least m+1 states *) have "p_io pP @ p_io pM \ L M'" using \((p_io pP) @ io) \ L M'\ unfolding \io = (p_io pM)@ioEx\ append.assoc[symmetric] using language_prefix[of "p_io pP @ p_io pM" ioEx M' "initial M'"] by blast have no_shared_targets_for_distinct_states : "\ q' q'' pR1 pR2. q' \ q'' \ q' \ fst dM \ q'' \ fst dM \ pR1 \ RP M q q' pP pM prs M' \ pR2 \ RP M q q'' pP pM prs M' \ io_targets M' (p_io pR1) (initial M') \ io_targets M' (p_io pR2) (initial M') = {}" using passes_test_suite_exhaustiveness_helper_1[OF \completely_specified M'\ \inputs M' = inputs M\ \observable M\ \observable M'\ \(q,P) \ prs\ \path P (initial P) pP\ \target (FSM.initial P) pP = q\ \p_io pP @ p_io pM \ L M'\ \(pM, dM) \ m_traversal_paths_with_witness M q repetition_sets m\ repetition_sets_def \passes_test_suite M (Test_Suite prs tps rd_targets separators) M'\] by blast have "path M q pM" and "find (\d. Suc (m - card (snd d)) \ length (filter (\t. t_target t \ fst d) pM)) repetition_sets = Some dM" using \(pM,dM) \ m_traversal_paths_with_witness M q repetition_sets m\ using m_traversal_paths_with_witness_set[OF t5 t8 \q \ states M\, of m] by force+ then have "path M (target (FSM.initial M) pP) pM" unfolding \(target (FSM.initial M) pP) = q\ by simp have "dM \ set repetition_sets" using find_set[OF \find (\d. Suc (m - card (snd d)) \ length (filter (\t. t_target t \ fst d) pM)) repetition_sets = Some dM\] by assumption have "Suc (m - card (snd dM)) \ length (filter (\t. t_target t \ fst dM) pM)" using find_condition[OF \find (\d. Suc (m - card (snd d)) \ length (filter (\t. t_target t \ fst d) pM)) repetition_sets = Some dM\] by assumption obtain ioX where "butlast io = (p_io pM)@ioX" using \io = (p_io pM)@ioEx\ by (simp add: \ioEx \ []\ butlast_append) have RP_card : "\ q' . card (\pR\RP M (target (FSM.initial M) pP) q' pP pM prs M'. io_targets M' (p_io pR) (FSM.initial M')) = card (RP M (target (FSM.initial M) pP) q' pP pM prs M')" and RP_targets: "\ q' pR . pR \ RP M (target (FSM.initial M) pP) q' pP pM prs M' \ \q. io_targets M' (p_io pR) (FSM.initial M') = {q}" and no_shared_targets_for_identical_states: "\ q' pR1 pR2 . pR1 \ RP M (target (FSM.initial M) pP) q' pP pM prs M' \ pR2 \ RP M (target (FSM.initial M) pP) q' pP pM prs M' \ pR1 \ pR2 \ io_targets M' (p_io pR1) (FSM.initial M') \ io_targets M' (p_io pR2) (FSM.initial M') = {}" using RP_count[OF \minimal_sequence_to_failure_extending_preamble_path M M' prs pP io\ \observable M\ \observable M'\ t2 \path M (target (FSM.initial M) pP) pM\ \butlast io = (p_io pM)@ioX\ pass1 \completely_specified M'\ \inputs M' = inputs M\, of "\ q P io x y y' . q" "\ q P io x y y' . y"] by blast+ have snd_dM_prop: "\ q' . q' \ snd dM \ (\ pR \ (RP M q q' pP pM prs M') . io_targets M' (p_io pR) (initial M')) \ (\ pR \ (R M q q' pP pM) . io_targets M' (p_io pR) (initial M'))" proof - fix q' assume "q' \ snd dM" let ?RP = "(RP M q q' pP pM prs M')" let ?R = "(R M q q' pP pM)" let ?P = "\ pP' . \P'. (q', P') \ prs \ path P' (FSM.initial P') pP' \ target (FSM.initial P') pP' = q' \ p_io pP' \ L M'" obtain PQ where "(q',PQ) \ prs" using \q' \ snd dM\ t8'[OF \dM \ set repetition_sets\] by auto then have "is_preamble PQ M q'" and "\P'. (q', P') \ prs" using t2 by blast+ obtain pq where "path PQ (initial PQ) pq" and "target (initial PQ) pq = q'" and "p_io pq \ L M'" using preamble_pass_path[OF \is_preamble PQ M q'\ pass1[OF \(q',PQ) \ prs\] \completely_specified M'\ \inputs M' = inputs M\] by force then have "\ pP' . ?P pP'" using \(q',PQ) \ prs\ by blast define pPQ where pPQ_def: "pPQ = (SOME pP'. ?P pP')" have "?P pPQ" unfolding pPQ_def using someI_ex[OF \\ pP' . ?P pP'\] by assumption then obtain PQ' where "(q',PQ') \ prs" and "path PQ' (initial PQ') pPQ" and "target (initial PQ') pPQ = q'" and "p_io pPQ \ L M'" by blast have "?RP = insert pPQ (R M q q' pP pM)" unfolding RP_def pPQ_def using \\P'. (q', P') \ prs\ by auto obtain pPQ' where "path M' (initial M') pPQ'" and "p_io pPQ' = p_io pPQ" using \p_io pPQ \ L M'\ by auto then have "io_targets M' (p_io pPQ) (initial M') = {target (initial M') pPQ'}" using \observable M'\ by (metis (mono_tags, lifting) observable_path_io_target) moreover have "target (initial M') pPQ' \ (\ (image (\ pR . io_targets M' (p_io pR) (initial M')) ?R))" proof assume "target (initial M') pPQ' \ (\ (image (\ pR . io_targets M' (p_io pR) (initial M')) ?R))" then obtain pR where "pR \ ?R" and "target (initial M') pPQ' \ io_targets M' (p_io pR) (initial M')" by blast obtain pR' where "pR = pP@pR'" using R_component_ob[OF \pR \ ?R\] by blast then have "pP@pR' \ ?R" using \pR \ ?R\ by simp have "pR' = take (length pR') pM" and "length pR' \ length pM" and "t_target (pM ! (length pR' - 1)) = q'" and "pR' \ []" using R_component[OF \pP@pR' \ ?R\] by auto (* get the full path along (butlast io) in M, of which p is a (possibly proper) prefix *) obtain pX where "path M (target (initial M) pP) (pM@pX)" and "p_io (pM@pX) = butlast io" proof - have "p_io pP @ p_io pM @ ioX \ L M" using \((p_io pP) @ butlast io) \ L M\ unfolding \butlast io = p_io pM @ ioX\ by assumption obtain p1 p23 where "path M (FSM.initial M) p1" and "path M (target (FSM.initial M) p1) p23" and "p_io p1 = p_io pP" and "p_io p23 = p_io pM @ ioX" using language_state_split[OF \p_io pP @ p_io pM @ ioX \ L M\] by blast have "p1 = pP" using observable_path_unique[OF \observable M\ \path M (FSM.initial M) p1\ \path M (FSM.initial M) pP\ \p_io p1 = p_io pP\] by assumption then have "path M (target (FSM.initial M) pP) p23" using \path M (target (FSM.initial M) p1) p23\ by auto then have "p_io pM @ ioX \ LS M (target (initial M) pP)" using \p_io p23 = p_io pM @ ioX\ language_state_containment by auto obtain p2 p3 where "path M (target (FSM.initial M) pP) p2" and "path M (target (target (FSM.initial M) pP) p2) p3" and "p_io p2 = p_io pM" and "p_io p3 = ioX" using language_state_split[OF \p_io pM @ ioX \ LS M (target (initial M) pP)\] by blast have "p2 = pM" using observable_path_unique[OF \observable M\ \path M (target (FSM.initial M) pP) p2\ \path M (target (FSM.initial M) pP) pM\ \p_io p2 = p_io pM\] by assumption then have "path M (target (FSM.initial M) pP) (pM@p3)" using \path M (target (FSM.initial M) pP) pM\ \path M (target (target (FSM.initial M) pP) p2) p3\ by auto moreover have "p_io (pM@p3) = butlast io" unfolding \butlast io = p_io pM @ ioX\ using \p_io p3 = ioX\ by auto ultimately show ?thesis using that[of p3] by simp qed obtain pP' pIO where "path M' (FSM.initial M') pP'" and "path M' (target (FSM.initial M') pP') pIO" and "p_io pP' = p_io pP" and "p_io pIO = io" using language_state_split[OF \((p_io pP) @ io) \ L M'\] by blast have "target (initial M') pP' \ io_targets M' (p_io pP) (FSM.initial M')" using \path M' (FSM.initial M') pP'\ unfolding \p_io pP' = p_io pP\[symmetric] by auto let ?i = "length pR' - 1" have "?i < length pR'" using \pR' \ []\ by auto have "?i < length (butlast io)" using \pR' = take (length pR') pM\ \pR' \ []\ unfolding \p_io (pM@pX) = butlast io\[symmetric] using leI by fastforce have "t_target ((pM @ pX) ! (length pR' - 1)) = q'" by (metis \length pR' - 1 < length pR'\ \length pR' \ length pM\ \t_target (pM ! (length pR' - 1)) = q'\ dual_order.strict_trans1 nth_append) then have "(t_target ((pM @ pX) ! (length pR' - 1)), PQ') \ prs" using \(q',PQ') \ prs\ by simp have "target (FSM.initial PQ') pPQ = t_target ((pM @ pX) ! (length pR' - 1))" using \t_target ((pM @ pX) ! (length pR' - 1)) = q'\ \target (FSM.initial PQ') pPQ = q'\ by blast have "t_target (pIO ! ?i) \ io_targets M' (p_io pPQ) (FSM.initial M')" using minimal_sequence_to_failure_extending_preamble_no_repetitions_with_other_preambles [OF \minimal_sequence_to_failure_extending_preamble_path M M' prs pP io\ \observable M\ \ path M (target (initial M) pP) (pM@pX)\ \p_io (pM@pX) = butlast io\ \target (initial M') pP' \ io_targets M' (p_io pP) (FSM.initial M')\ \path M' (target (FSM.initial M') pP') pIO\ \p_io pIO = io\ t2 \?i < length (butlast io)\ \(t_target ((pM @ pX) ! (length pR' - 1)), PQ') \ prs\ \path PQ' (initial PQ') pPQ\ \target (FSM.initial PQ') pPQ = t_target ((pM @ pX) ! (length pR' - 1))\] by blast moreover have "io_targets M' (p_io pPQ) (FSM.initial M') = {target (initial M') pPQ'}" using \path M' (initial M') pPQ'\ using \io_targets M' (p_io pPQ) (FSM.initial M') = {target (FSM.initial M') pPQ'}\ \p_io pPQ' = p_io pPQ\ by auto moreover have "io_targets M' (p_io (pP@pR')) (FSM.initial M') = {t_target (pIO ! ?i)}" proof - have "(take (length pR') pIO) \ []" using \p_io pIO = io\ \?i < length pR'\ using \io = p_io pM @ ioEx\ \pR' = take (length pR') pM\ by auto moreover have "pIO ! ?i = last (take (length pR') pIO)" using \p_io pIO = io\ \?i < length pR'\ by (metis (no_types, lifting) \io = p_io pM @ ioEx\ \length pR' \ length pM\ \pR' = take (length pR') pM\ butlast.simps(1) last_conv_nth length_butlast length_map neq_iff nth_take take_le take_map) ultimately have "t_target (pIO ! ?i) = target (target (FSM.initial M') pP') (take (length pR') pIO)" unfolding target.simps visited_states.simps by (simp add: last_map) then have "t_target (pIO ! ?i) = target (initial M') (pP' @ (take (length pR') pIO))" by auto have "path M' (target (FSM.initial M') pP') (take (length pR') pIO)" using \path M' (target (FSM.initial M') pP') pIO\ by (simp add: path_prefix_take) then have "path M' (initial M') (pP' @ (take (length pR') pIO))" using \path M' (FSM.initial M') pP'\ by auto moreover have "p_io (pP' @ (take (length pR') pIO)) = (p_io (pP@pR'))" proof - have "p_io (take (length pR') pIO) = p_io pR'" using \p_io pIO = io\ \pR' = take (length pR') pM\ \p_io (pM@pX) = butlast io\ \length pR' \ length pM\ by (metis (no_types, lifting) \io = p_io pM @ ioEx\ length_map take_le take_map) then show ?thesis using \p_io pP' = p_io pP\ by auto qed ultimately have "io_targets M' (p_io (pP@pR')) (FSM.initial M') = {target (initial M') (pP' @ (take (length pR') pIO))}" by (metis (mono_tags, lifting) assms(4) observable_path_io_target) then show ?thesis unfolding \t_target (pIO ! ?i) = target (initial M') (pP' @ (take (length pR') pIO))\ by assumption qed ultimately have "target (initial M') pPQ' \ io_targets M' (p_io pR) (initial M')" unfolding \pR = pP@pR'\ by auto then show "False" using \target (initial M') pPQ' \ io_targets M' (p_io pR) (initial M')\ by blast qed ultimately have "io_targets M' (p_io pPQ) (initial M') \ (\ (image (\ pR . io_targets M' (p_io pR) (initial M')) ?R)) = {}" by force then show "(\ pR \ (RP M q q' pP pM prs M') . io_targets M' (p_io pR) (initial M')) \ (\ pR \ (R M q q' pP pM) . io_targets M' (p_io pR) (initial M'))" unfolding \?RP = insert pPQ (R M q q' pP pM)\ using \io_targets M' (p_io pPQ) (FSM.initial M') = {target (FSM.initial M') pPQ'}\ by force qed (* create a function that separates the additional entries for the preambles in the RP sets of states in (snd dM) *) then obtain f where f_def: "\ q' . q' \ snd dM \ (RP M q q' pP pM prs M') = insert (f q') (R M q q' pP pM) \ (f q') \ (R M q q' pP pM)" proof - define f where f_def : "f = (\ q' . SOME p . (RP M q q' pP pM prs M') = insert p (R M q q' pP pM) \ p \ (R M q q' pP pM))" have "\ q' . q' \ snd dM \ (RP M q q' pP pM prs M') = insert (f q') (R M q q' pP pM) \ (RP M q q' pP pM prs M') \ (R M q q' pP pM)" proof - fix q' assume "q' \ snd dM" have "(\pR\RP M q q' pP pM prs M'. io_targets M' (p_io pR) (FSM.initial M')) \ (\pR\R M q q' pP pM. io_targets M' (p_io pR) (FSM.initial M'))" using snd_dM_prop[OF \q' \ snd dM\] by assumption then have "(RP M q q' pP pM prs M') \ (R M q q' pP pM)" by blast then obtain x where "(RP M q q' pP pM prs M') = insert x (R M q q' pP pM)" using RP_from_R[OF t2 pass1 \completely_specified M'\ \inputs M' = inputs M\, of prs "\ q P io x y y' . q" "\ q P io x y y' . y" q q' pP pM] by force then have "x \ (R M q q' pP pM)" using \(RP M q q' pP pM prs M') \ (R M q q' pP pM)\ by auto then have "\ p . (RP M q q' pP pM prs M') = insert p (R M q q' pP pM) \ p \ (R M q q' pP pM)" using \(RP M q q' pP pM prs M') = insert x (R M q q' pP pM)\ by blast show "(RP M q q' pP pM prs M') = insert (f q') (R M q q' pP pM) \ (RP M q q' pP pM prs M') \ (R M q q' pP pM)" using someI_ex[OF \\ p . (RP M q q' pP pM prs M') = insert p (R M q q' pP pM) \ p \ (R M q q' pP pM)\] unfolding f_def by auto qed then show ?thesis using that by force qed (* combine counting results on RP (from R) to show that the path along pM collects at least m+1 distinct io-targets (i.e. visites at least m+1 distinct states) in M' *) have "(\ q' \ fst dM . card (\ pR \ (RP M q q' pP pM prs M') . io_targets M' (p_io pR) (initial M'))) \ Suc m" proof - have "\ nds . finite nds \ nds \ fst dM \ (\ q' \ nds . card (RP M q q' pP pM prs M')) \ length (filter (\t. t_target t \ nds) pM) + card (nds \ snd dM)" proof - fix nds assume "finite nds" and "nds \ fst dM" then show "(\ q' \ nds . card (RP M q q' pP pM prs M')) \ length (filter (\t. t_target t \ nds) pM) + card (nds \ snd dM)" proof induction case empty then show ?case by auto next case (insert q' nds) then have leq1: "length (filter (\t. t_target t \ nds) pM) + card (nds \ snd dM) \ (\q'\nds. card (RP M q q' pP pM prs M'))" by blast have p4: "(card (R M q q' pP pM)) = length (filter (\t. t_target t = q') pM)" using \path M q pM\ proof (induction pM rule: rev_induct) case Nil then show ?case unfolding R_def by auto next case (snoc t pM) then have "path M q pM" and "card (R M q q' pP pM) = length (filter (\t. t_target t = q') pM)" by auto show ?case proof (cases "target q (pM @ [t]) = q'") case True then have "(R M q q' pP (pM @ [t])) = insert (pP @ pM @ [t]) (R M q q' pP pM)" unfolding R_update[of M q q' pP pM t] by simp moreover have "(pP @ pM @ [t]) \ (R M q q' pP pM)" unfolding R_def by auto ultimately have "card (R M q q' pP (pM @ [t])) = Suc (card (R M q q' pP pM))" using finite_R[OF \path M q pM\, of q' pP] by simp then show ?thesis using True unfolding \card (R M q q' pP pM) = length (filter (\t. t_target t = q') pM)\ by auto next case False then have "card (R M q q' pP (pM @ [t])) = card (R M q q' pP pM)" unfolding R_update[of M q q' pP pM t] by simp then show ?thesis using False unfolding \card (R M q q' pP pM) = length (filter (\t. t_target t = q') pM)\ by auto qed qed show ?case proof (cases "q' \ snd dM") case True then have p0: "(RP M q q' pP pM prs M') = insert (f q') (R M q q' pP pM)" and "(f q') \ (R M q q' pP pM)" using f_def by blast+ then have "card (RP M q q' pP pM prs M') = Suc (card (R M q q' pP pM))" by (simp add: \path M q pM\ finite_R) then have p1: "(\q' \ (insert q' nds). card (RP M q q' pP pM prs M')) = (\q'\nds. card (RP M q q' pP pM prs M')) + Suc (card (R M q q' pP pM))" by (simp add: insert.hyps(1) insert.hyps(2)) have p2: "length (filter (\t. t_target t \ insert q' nds) pM) = length (filter (\t. t_target t \ nds) pM) + length (filter (\t. t_target t = q') pM)" using \q' \ nds\ by (induction pM; auto) have p3: "card ((insert q' nds) \ snd dM) = Suc (card (nds \ snd dM))" using True \finite nds\ \q' \ nds\ by simp show ?thesis using leq1 unfolding p1 p2 p3 p4 by simp next case False have "card (RP M q q' pP pM prs M') \ (card (R M q q' pP pM))" proof (cases "(RP M q q' pP pM prs M') = (R M q q' pP pM)") case True then show ?thesis using finite_R[OF \path M q pM\, of q' pP] by auto next case False then obtain pX where "(RP M q q' pP pM prs M') = insert pX (R M q q' pP pM)" using RP_from_R[OF t2 pass1 \completely_specified M'\ \inputs M' = inputs M\, of prs "\ q P io x y y' . q" "\ q P io x y y' . y" q q' pP pM] by force then show ?thesis using finite_R[OF \path M q pM\, of q' pP] by (simp add: card_insert_le) qed then have p1: "(\q' \ (insert q' nds). card (RP M q q' pP pM prs M')) \ ((\q'\nds. card (RP M q q' pP pM prs M')) + (card (R M q q' pP pM)))" by (simp add: insert.hyps(1) insert.hyps(2)) have p2: "length (filter (\t. t_target t \ insert q' nds) pM) = length (filter (\t. t_target t \ nds) pM) + length (filter (\t. t_target t = q') pM)" using \q' \ nds\ by (induction pM; auto) have p3: "card ((insert q' nds) \ snd dM) = (card (nds \ snd dM))" using False \finite nds\ \q' \ nds\ by simp have "length (filter (\t. t_target t \ nds) pM) + length (filter (\t. t_target t = q') pM) + card (nds \ snd dM) \ (\q'\nds. card (RP M q q' pP pM prs M')) + length (filter (\t. t_target t = q') pM)" using leq1 add_le_cancel_right by auto then show ?thesis using p1 unfolding p2 p3 p4 by simp qed qed qed moreover have "finite (fst dM)" using t7[OF \dM \ set repetition_sets\] fsm_states_finite[of M] using rev_finite_subset by auto ultimately have "(\ q' \ fst dM . card (RP M q q' pP pM prs M')) \ length (filter (\t. t_target t \ fst dM) pM) + card (fst dM \ snd dM)" by blast have "(fst dM \ snd dM) = (snd dM)" using t8[OF \dM \ set repetition_sets\] by blast have "(\ q' \ fst dM . card (RP M q q' pP pM prs M')) \ length (filter (\t. t_target t \ fst dM) pM) + card (snd dM)" using \(\ q' \ fst dM . card (RP M q q' pP pM prs M')) \ length (filter (\t. t_target t \ fst dM) pM) + card (fst dM \ snd dM)\ unfolding \(fst dM \ snd dM) = (snd dM)\ by assumption moreover have "(\q'\fst dM. card (\pR\RP M q q' pP pM prs M'. io_targets M' (p_io pR) (FSM.initial M'))) = (\ q' \ fst dM . card (RP M q q' pP pM prs M'))" using RP_card \FSM.initial P = FSM.initial M\ \target (FSM.initial P) pP = q\ by auto ultimately have "(\q'\fst dM. card (\pR\RP M q q' pP pM prs M'. io_targets M' (p_io pR) (FSM.initial M'))) \ length (filter (\t. t_target t \ fst dM) pM) + card (snd dM)" by linarith moreover have "Suc m \ length (filter (\t. t_target t \ fst dM) pM) + card (snd dM)" using \Suc (m - card (snd dM)) \ length (filter (\t. t_target t \ fst dM) pM)\ by linarith ultimately show ?thesis by linarith qed moreover have "(\ q' \ fst dM . card (\ pR \ (RP M q q' pP pM prs M') . io_targets M' (p_io pR) (initial M'))) \ card (states M')" proof - have "finite (fst dM)" by (meson \dM \ set repetition_sets\ fsm_states_finite rev_finite_subset t7) have "(\x1. finite (RP M q x1 pP pM prs M'))" using finite_RP[OF \path M q pM\ t2 pass1 \completely_specified M'\ \inputs M' = inputs M\] by force have "(\y1. finite (io_targets M' (p_io y1) (FSM.initial M')))" by (meson io_targets_finite) have "(\y1. io_targets M' (p_io y1) (FSM.initial M') \ states M')" by (meson io_targets_states) show ?thesis using distinct_union_union_card [ of "fst dM" "\ q' . (RP M q q' pP pM prs M')" "\ pR . io_targets M' (p_io pR) (initial M')" , OF \finite (fst dM)\ no_shared_targets_for_distinct_states no_shared_targets_for_identical_states \(\x1. finite (RP M q x1 pP pM prs M'))\ io_targets_finite io_targets_states fsm_states_finite[of M']] unfolding \(target (FSM.initial M) pP) = q\ by force qed (* create a contradiction due to the assumption that M' has at most m states *) moreover have "card (states M') \ m" using \size M' \ m\ by auto ultimately show False by linarith qed subsection \Completeness of Sufficient Test Suites\ text \This subsection combines the soundness and exhaustiveness properties of sufficient test suites to show completeness: for any System Under Test with at most m states a test suite sufficient for m passes if and only if the System Under Test is a reduction of the specification.\ lemma passes_test_suite_completeness : assumes "implies_completeness T M m" and "observable M" and "observable M'" and "inputs M' = inputs M" and "inputs M \ {}" and "completely_specified M" and "completely_specified M'" and "size M' \ m" shows "(L M' \ L M) \ passes_test_suite M T M'" using passes_test_suite_exhaustiveness[OF _ _ assms(2-8)] passes_test_suite_soundness[OF _ assms(2,3,4,6)] assms(1) test_suite.exhaust[of T] by metis subsection \Additional Test Suite Properties\ (* Whether a (sufficient) test suite can be represented as a finite set of IO sequences *) (* (tps q) must already be finite for a sufficient test suite due to being a subset of the maximal m-traversal paths *) fun is_finite_test_suite :: "('a,'b,'c,'d) test_suite \ bool" where "is_finite_test_suite (Test_Suite prs tps rd_targets separators) = ((finite prs) \ (\ q p . q \ fst ` prs \ finite (rd_targets (q,p))) \ (\ q q' . finite (separators (q,q'))))" end \ No newline at end of file diff --git a/thys/FSM_Tests/AdaptiveStateCounting/Test_Suite_IO.thy b/thys/FSM_Tests/AdaptiveStateCounting/Test_Suite_IO.thy --- a/thys/FSM_Tests/AdaptiveStateCounting/Test_Suite_IO.thy +++ b/thys/FSM_Tests/AdaptiveStateCounting/Test_Suite_IO.thy @@ -1,1816 +1,1810 @@ section \Representing Test Suites as Sets of Input-Output Sequences\ text \This theory describes the representation of test suites as sets of input-output sequences and defines a pass relation for this representation.\ theory Test_Suite_IO imports Test_Suite Maximal_Path_Trie begin fun test_suite_to_io :: "('a,'b,'c) fsm \ ('a,'b,'c,'d) test_suite \ ('b \ 'c) list set" where "test_suite_to_io M (Test_Suite prs tps rd_targets atcs) = (\ (q,P) \ prs . L P) \ (\{(\ io' . p_io p @ io') ` (set (prefixes (p_io pt))) | p pt . \ q P . (q,P) \ prs \ path P (initial P) p \ target (initial P) p = q \ pt \ tps q}) \ (\{(\ io_atc . p_io p @ p_io pt @ io_atc) ` (atc_to_io_set (from_FSM M (target q pt)) A) | p pt q A . \ P q' t1 t2 . (q,P) \ prs \ path P (initial P) p \ target (initial P) p = q \ pt \ tps q \ q' \ rd_targets (q,pt) \ (A,t1,t2) \ atcs (target q pt,q') })" lemma test_suite_to_io_language : assumes "implies_completeness T M m" shows "(test_suite_to_io M T) \ L M" proof fix io assume "io \ test_suite_to_io M T" obtain prs tps rd_targets atcs where "T = Test_Suite prs tps rd_targets atcs" by (meson test_suite.exhaust) then obtain repetition_sets where repetition_sets_def: "implies_completeness_for_repetition_sets (Test_Suite prs tps rd_targets atcs) M m repetition_sets" using assms(1) unfolding implies_completeness_def by blast then have "implies_completeness (Test_Suite prs tps rd_targets atcs) M m" unfolding implies_completeness_def by blast have t2: "\ q P. (q, P) \ prs \ is_preamble P M q" using implies_completeness_for_repetition_sets_simps(2)[OF repetition_sets_def] by blast have t5: "\q. q \ FSM.states M \ (\d\set repetition_sets. q \ fst d)" using implies_completeness_for_repetition_sets_simps(4)[OF repetition_sets_def] by assumption have t6: "\ q. q \ fst ` prs \ tps q \ {p1 . \ p2 d . (p1@p2,d) \ m_traversal_paths_with_witness M q repetition_sets m} \ fst ` (m_traversal_paths_with_witness M q repetition_sets m) \ tps q" using implies_completeness_for_repetition_sets_simps(7)[OF repetition_sets_def] by assumption have t8: "\ d. d \ set repetition_sets \ snd d \ fst d" using implies_completeness_for_repetition_sets_simps(5,6)[OF repetition_sets_def] by blast from \io \ test_suite_to_io M T\ consider (a) "io \ (\ (q,P) \ prs . L P)" | (b) "io \ (\{(\ io' . p_io p @ io') ` (set (prefixes (p_io pt))) | p pt . \ q P . (q,P) \ prs \ path P (initial P) p \ target (initial P) p = q \ pt \ tps q})" | (c) "io \ (\{(\ io_atc . p_io p @ p_io pt @ io_atc) ` (atc_to_io_set (from_FSM M (target q pt)) A) | p pt q A . \ P q' t1 t2 . (q,P) \ prs \ path P (initial P) p \ target (initial P) p = q \ pt \ tps q \ q' \ rd_targets (q,pt) \ (A,t1,t2) \ atcs (target q pt,q') })" unfolding \T = Test_Suite prs tps rd_targets atcs\ test_suite_to_io.simps by blast then show "io \ L M" proof cases case a then obtain q P where "(q, P) \ prs" and "io \ L P" by blast have "is_submachine P M" using t2[OF \(q, P) \ prs\] unfolding is_preamble_def by blast show "io \ L M" using submachine_language[OF \is_submachine P M\] \io \ L P\ by blast next case b then obtain p pt q P where "io \ (\ io' . p_io p @ io') ` (set (prefixes (p_io pt)))" and "(q,P) \ prs" and "path P (initial P) p" and "target (initial P) p = q" and "pt \ tps q" by blast then obtain io' where "io = p_io p @ io'" and "io' \ (set (prefixes (p_io pt)))" by blast then obtain io'' where "p_io pt = io' @ io''" and "io = p_io p @ io'" - unfolding prefixes_set - proof - - assume a1: "io' \ {xs'. \xs''. xs' @ xs'' = p_io pt}" - assume "\io''. \p_io pt = io' @ io''; io = p_io p @ io'\ \ thesis" - then show ?thesis - using a1 \io = p_io p @ io'\ by moura - qed + unfolding prefixes_set using \io' \ set (prefixes (p_io pt))\ prefixes_set_ob by blast have "q \ fst ` prs" using \(q,P) \ prs\ by force have "is_submachine P M" using t2[OF \(q, P) \ prs\] unfolding is_preamble_def by blast then have "initial P = initial M" by auto have "path M (initial M) p" using submachine_path[OF \is_submachine P M\ \path P (initial P) p\] unfolding \initial P = initial M\ by assumption have "target (initial M) p = q" using \target (initial P) p = q\ unfolding \initial P = initial M\ by assumption obtain p2 d where "(pt @ p2, d) \ m_traversal_paths_with_witness M q repetition_sets m" using t6[OF \q \ fst ` prs\] \pt \ tps q\ by blast then have "path M q (pt @ p2)" using m_traversal_paths_with_witness_set[OF t5 t8 path_target_is_state[OF \path M (initial M) p\], of m] unfolding \target (initial M) p = q\ by blast then have "path M (initial M) (p@pt)" using \path M (initial M) p\ \target (initial M) p = q\ by auto then have "p_io p @ p_io pt \ L M" by (metis (mono_tags, lifting) language_intro map_append) then show "io \ L M" unfolding \io = p_io p @ io'\ \p_io pt = io' @ io''\ append.assoc[symmetric] using language_prefix[of "p_io p @ io'" io'' M "initial M"] by blast next case c then obtain p pt q A P q' t1 t2 where "io \ (\ io_atc . p_io p @ p_io pt @ io_atc) ` (atc_to_io_set (from_FSM M (target q pt)) A)" and "(q,P) \ prs" and "path P (initial P) p" and "target (initial P) p = q" and "pt \ tps q" and "q' \ rd_targets (q,pt)" and "(A,t1,t2) \ atcs (target q pt,q')" by blast obtain ioA where "io = p_io p @ p_io pt @ ioA" and "ioA \ (atc_to_io_set (from_FSM M (target q pt)) A)" using \io \ (\ io_atc . p_io p @ p_io pt @ io_atc) ` (atc_to_io_set (from_FSM M (target q pt)) A)\ by blast then have "ioA \ L (from_FSM M (target q pt))" unfolding atc_to_io_set.simps by blast have "q \ fst ` prs" using \(q,P) \ prs\ by force have "is_submachine P M" using t2[OF \(q, P) \ prs\] unfolding is_preamble_def by blast then have "initial P = initial M" by auto have "path M (initial M) p" using submachine_path[OF \is_submachine P M\ \path P (initial P) p\] unfolding \initial P = initial M\ by assumption have "target (initial M) p = q" using \target (initial P) p = q\ unfolding \initial P = initial M\ by assumption obtain p2 d where "(pt @ p2, d) \ m_traversal_paths_with_witness M q repetition_sets m" using t6[OF \q \ fst ` prs\] \pt \ tps q\ by blast then have "path M q (pt @ p2)" using m_traversal_paths_with_witness_set[OF t5 t8 path_target_is_state[OF \path M (initial M) p\], of m] unfolding \target (initial M) p = q\ by blast then have "path M (initial M) (p@pt)" using \path M (initial M) p\ \target (initial M) p = q\ by auto moreover have "(target q pt) = target (initial M) (p@pt)" using \target (initial M) p = q\ by auto ultimately have "(target q pt) \ states M" using path_target_is_state by metis have "ioA \ LS M (target q pt)" using from_FSM_language[OF \(target q pt) \ states M\] \ioA \ L (from_FSM M (target q pt))\ by blast then obtain pA where "path M (target q pt) pA" and "p_io pA = ioA" by auto then have "path M (initial M) (p @ pt @ pA)" using \path M (initial M) (p@pt)\ unfolding \(target q pt) = target (initial M) (p@pt)\ by auto then have "p_io p @ p_io pt @ ioA \ L M" unfolding \p_io pA = ioA\[symmetric] using language_intro by fastforce then show "io \ L M" unfolding \io = p_io p @ p_io pt @ ioA\ by assumption qed qed lemma minimal_io_seq_to_failure : assumes "\ (L M' \ L M)" and "inputs M' = inputs M" and "completely_specified M" obtains io x y y' where "io@[(x,y)] \ L M" and "io@[(x,y')] \ L M" and "io@[(x,y')] \ L M'" proof - obtain ioF where "ioF \ L M'" and "ioF \ L M" using assms(1) by blast let ?prefs = "{ioF' \ set (prefixes ioF) . ioF' \ L M' \ ioF' \ L M}" have "finite ?prefs" using prefixes_finite by auto moreover have "?prefs \ {}" unfolding prefixes_set using \ioF \ L M'\ \ioF \ L M\ by auto ultimately obtain ioF' where "ioF' \ ?prefs" and "\ ioF'' . ioF'' \ ?prefs \ length ioF' \ length ioF''" by (meson leI min_length_elem) then have "ioF' \ L M'" and "ioF' \ L M" by auto then have "ioF' \ []" by auto then have "ioF' = (butlast ioF')@[last ioF']" and "length (butlast ioF') < length ioF'" by auto then have "butlast ioF' \ ?prefs" using \\ ioF'' . ioF'' \ ?prefs \ length ioF' \ length ioF''\ leD by blast moreover have "butlast ioF' \ L M'" using \ioF' \ L M'\ language_prefix[of "butlast ioF'" "[last ioF']" M' "initial M'"] unfolding \ioF' = (butlast ioF')@[last ioF']\[symmetric] by blast moreover have "butlast ioF' \ set (prefixes ioF)" using \ioF' = (butlast ioF')@[last ioF']\ \ioF' \ ?prefs\ prefixes_set proof - have "\ps. (butlast ioF' @ [last ioF']) @ ps = ioF" using \ioF' = butlast ioF' @ [last ioF']\ \ioF' \ {ioF' \ set (prefixes ioF). ioF' \ L M' \ ioF' \ L M}\ unfolding prefixes_set by auto then show ?thesis using prefixes_set by fastforce qed ultimately have "butlast ioF' \ L M" by blast have *: "(butlast ioF')@[(fst (last ioF'), snd (last ioF'))] \ L M'" using \ioF' \ L M'\ \ioF' = (butlast ioF')@[last ioF']\ by auto have **: "(butlast ioF')@[(fst (last ioF'), snd (last ioF'))] \ L M" using \ioF' \ L M\ \ioF' = (butlast ioF')@[last ioF']\ by auto obtain p where "path M (initial M) p" and "p_io p = butlast ioF'" using \butlast ioF' \ L M\ by auto moreover obtain t where "t \ transitions M" and "t_source t = target (initial M) p" and "t_input t = fst (last ioF')" proof - have "fst (last ioF') \ inputs M'" using language_io(1)[OF *, of "fst (last ioF')" "snd (last ioF')"] by simp then have "fst (last ioF') \ inputs M" using assms(2) by auto then show ?thesis using that \completely_specified M\ path_target_is_state[OF \path M (initial M) p\] unfolding completely_specified.simps by blast qed ultimately have ***: "(butlast ioF')@[(fst (last ioF'), t_output t)] \ L M" proof - have "p_io (p @ [t]) \ L M" by (metis (no_types) \path M (FSM.initial M) p\ \t \ FSM.transitions M\ \t_source t = target (FSM.initial M) p\ language_intro path_append single_transition_path) then show ?thesis by (simp add: \p_io p = butlast ioF'\ \t_input t = fst (last ioF')\) qed show ?thesis using that[OF *** ** *] by assumption qed lemma observable_minimal_path_to_failure : assumes "\ (L M' \ L M)" and "observable M" and "observable M'" and "inputs M' = inputs M" and "completely_specified M" and "completely_specified M'" obtains p p' t t' where "path M (initial M) (p@[t])" and "path M' (initial M') (p'@[t'])" and "p_io p' = p_io p" and "t_input t' = t_input t" and "\(\ t'' . t'' \ transitions M \ t_source t'' = target (initial M) p \ t_input t'' = t_input t \ t_output t'' = t_output t')" proof - obtain io x y y' where "io@[(x,y)] \ L M" and "io@[(x,y')] \ L M" and "io@[(x,y')] \ L M'" using minimal_io_seq_to_failure[OF assms(1,4,5)] by blast obtain p t where "path M (initial M) (p@[t])" and "p_io p = io" and "t_input t = x" and "t_output t = y" using language_append_path_ob[OF \io@[(x,y)] \ L M\] by blast moreover obtain p' t' where "path M' (initial M') (p'@[t'])" and "p_io p' = io" and "t_input t' = x" and "t_output t' = y'" using language_append_path_ob[OF \io@[(x,y')] \ L M'\] by blast moreover have "\(\ t'' . t'' \ transitions M \ t_source t'' = target (initial M) p \ t_input t'' = t_input t \ t_output t'' = t_output t')" proof assume "\t''. t'' \ FSM.transitions M \ t_source t'' = target (FSM.initial M) p \ t_input t'' = t_input t \ t_output t'' = t_output t'" then obtain t'' where "t'' \ FSM.transitions M" and "t_source t'' = target (FSM.initial M) p" and "t_input t'' = x" and "t_output t'' = y'" unfolding \t_input t = x\ \t_output t' = y'\ by blast then have "path M (initial M) (p@[t''])" using \path M (initial M) (p@[t])\ by (meson path_append_elim path_append_transition) moreover have "p_io (p@[t'']) = io@[(x,y')]" using \p_io p = io\ \t_input t'' = x\ \t_output t'' = y'\ by auto ultimately have "io@[(x,y')] \ L M" using language_state_containment by (metis (mono_tags, lifting)) then show "False" using \io@[(x,y')] \ L M\ by blast qed ultimately show ?thesis using that[of p t p' t'] by force qed lemma test_suite_to_io_pass : assumes "implies_completeness T M m" and "observable M" and "observable M'" and "inputs M' = inputs M" and "inputs M \ {}" and "completely_specified M" and "completely_specified M'" shows "pass_io_set M' (test_suite_to_io M T) = passes_test_suite M T M'" proof - obtain prs tps rd_targets atcs where "T = Test_Suite prs tps rd_targets atcs" by (meson test_suite.exhaust) then obtain repetition_sets where repetition_sets_def: "implies_completeness_for_repetition_sets (Test_Suite prs tps rd_targets atcs) M m repetition_sets" using assms(1) unfolding implies_completeness_def by blast then have "implies_completeness (Test_Suite prs tps rd_targets atcs) M m" unfolding implies_completeness_def by blast then have test_suite_language_prop: "test_suite_to_io M (Test_Suite prs tps rd_targets atcs) \ L M" using test_suite_to_io_language by blast have t1: "(initial M, initial_preamble M) \ prs" using implies_completeness_for_repetition_sets_simps(1)[OF repetition_sets_def] by assumption have t2: "\ q P. (q, P) \ prs \ is_preamble P M q" using implies_completeness_for_repetition_sets_simps(2)[OF repetition_sets_def] by blast have t3: "\ q1 q2 A d1 d2. (A, d1, d2) \ atcs (q1, q2) \ (A, d2, d1) \ atcs (q2, q1) \ is_separator M q1 q2 A d1 d2" using implies_completeness_for_repetition_sets_simps(3)[OF repetition_sets_def] by assumption have t5: "\q. q \ FSM.states M \ (\d\set repetition_sets. q \ fst d)" using implies_completeness_for_repetition_sets_simps(4)[OF repetition_sets_def] by assumption have t6: "\ q. q \ fst ` prs \ tps q \ {p1 . \ p2 d . (p1@p2,d) \ m_traversal_paths_with_witness M q repetition_sets m} \ fst ` (m_traversal_paths_with_witness M q repetition_sets m) \ tps q" using implies_completeness_for_repetition_sets_simps(7)[OF repetition_sets_def] by assumption have t7: "\ d. d \ set repetition_sets \ fst d \ FSM.states M" and t8: "\ d. d \ set repetition_sets \ snd d \ fst d" and t8': "\ d. d \ set repetition_sets \ snd d = fst d \ fst ` prs" and t9: "\ d q1 q2. d \ set repetition_sets \ q1 \ fst d \ q2 \ fst d \ q1 \ q2 \ atcs (q1, q2) \ {}" using implies_completeness_for_repetition_sets_simps(5,6)[OF repetition_sets_def] by blast+ have "pass_io_set M' (test_suite_to_io M (Test_Suite prs tps rd_targets atcs)) \ passes_test_suite M (Test_Suite prs tps rd_targets atcs) M'" proof - assume "pass_io_set M' (test_suite_to_io M (Test_Suite prs tps rd_targets atcs))" then have pass_io_prop: "\ io x y y' . io @ [(x, y)] \ test_suite_to_io M (Test_Suite prs tps rd_targets atcs) \ io @ [(x, y')] \ L M' \ io @ [(x, y')] \ test_suite_to_io M (Test_Suite prs tps rd_targets atcs)" unfolding pass_io_set_def by blast show "passes_test_suite M (Test_Suite prs tps rd_targets atcs) M'" proof (rule ccontr) assume "\ passes_test_suite M (Test_Suite prs tps rd_targets atcs) M'" then consider (a) "\ (\q P io x y y'. (q, P) \ prs \ io @ [(x, y)] \ L P \ io @ [(x, y')] \ L M' \ io @ [(x, y')] \ L P)" | (b) "\ ((\q P pP ioT pT x y y'. (q, P) \ prs \ path P (FSM.initial P) pP \ target (FSM.initial P) pP = q \ pT \ tps q \ ioT @ [(x, y)] \ set (prefixes (p_io pT)) \ p_io pP @ ioT @ [(x, y')] \ L M' \ (\pT'. pT' \ tps q \ ioT @ [(x, y')] \ set (prefixes (p_io pT')))))" | (c) "\ ((\q P pP pT. (q, P) \ prs \ path P (FSM.initial P) pP \ target (FSM.initial P) pP = q \ pT \ tps q \ p_io pP @ p_io pT \ L M' \ (\q' A d1 d2 qT. q' \ rd_targets (q, pT) \ (A, d1, d2) \ atcs (target q pT, q') \ qT \ io_targets M' (p_io pP @ p_io pT) (FSM.initial M') \ pass_separator_ATC M' A qT d2)))" unfolding passes_test_suite.simps by blast then show False proof cases case a then obtain q P io x y y' where "(q, P) \ prs" and "io @ [(x, y)] \ L P" and "io @ [(x, y')] \ L M'" and "io @ [(x, y')] \ L P" by blast have "is_preamble P M q" using t2[OF \(q, P) \ prs\] by assumption have "io @ [(x, y)] \ test_suite_to_io M (Test_Suite prs tps rd_targets atcs)" unfolding test_suite_to_io.simps using \(q, P) \ prs\ \io @ [(x, y)] \ L P\ by fastforce have "io @ [(x, y')] \ test_suite_to_io M (Test_Suite prs tps rd_targets atcs)" using pass_io_prop[OF \io @ [(x, y)] \ test_suite_to_io M (Test_Suite prs tps rd_targets atcs)\ \io @ [(x, y')] \ L M'\] by assumption then have "io @ [(x, y')] \ L M" using test_suite_language_prop by blast have "io @ [(x, y')] \ L P" using passes_test_suite_soundness_helper_1[OF \is_preamble P M q\ \observable M\ \io @ [(x, y)] \ L P\ \io @ [(x, y')] \ L M\] by assumption then show "False" using \io @ [(x, y')] \ L P\ by blast next case b then obtain q P pP ioT pT x y y' where "(q, P) \ prs" and "path P (FSM.initial P) pP" and "target (FSM.initial P) pP = q" and "pT \ tps q" and "ioT @ [(x, y)] \ set (prefixes (p_io pT))" and "p_io pP @ ioT @ [(x, y')] \ L M'" and "\ (\pT'. pT' \ tps q \ ioT @ [(x, y')] \ set (prefixes (p_io pT')))" by blast have "\q P. (q, P) \ prs \ path P (FSM.initial P) pP \ target (FSM.initial P) pP = q \ pT \ tps q" using \(q, P) \ prs\ \path P (FSM.initial P) pP\ \target (FSM.initial P) pP = q\ \pT \ tps q\ by blast moreover have "p_io pP @ ioT @ [(x, y)] \ (@) (p_io pP) ` set (prefixes (p_io pT))" using \ioT @ [(x, y)] \ set (prefixes (p_io pT))\ by auto ultimately have "p_io pP @ ioT @ [(x, y)] \ (\{(@) (p_io p) ` set (prefixes (p_io pt)) |p pt. \q P. (q, P) \ prs \ path P (FSM.initial P) p \ target (FSM.initial P) p = q \ pt \ tps q})" by blast then have "p_io pP @ ioT @ [(x, y)] \ test_suite_to_io M (Test_Suite prs tps rd_targets atcs)" unfolding test_suite_to_io.simps by blast then have *: "(p_io pP @ ioT) @ [(x, y)] \ test_suite_to_io M (Test_Suite prs tps rd_targets atcs)" by auto have **: "(p_io pP @ ioT) @ [(x, y')] \ L M'" using \p_io pP @ ioT @ [(x, y')] \ L M'\ by auto have "(p_io pP @ ioT) @ [(x, y')] \ test_suite_to_io M (Test_Suite prs tps rd_targets atcs)" using pass_io_prop[OF * ** ] by assumption then have "(p_io pP @ ioT) @ [(x, y')] \ L M" using test_suite_language_prop by blast have "q \ states M" using is_preamble_is_state[OF t2[OF \(q, P) \ prs\]] by assumption have "q \ fst ` prs" using \(q, P) \ prs\ by force obtain pT' d' where "(pT @ pT', d') \ m_traversal_paths_with_witness M q repetition_sets m" using t6[OF \q \ fst ` prs\] \pT \ tps q\ by blast then have "path M q (pT @ pT')" and "find (\d. Suc (m - card (snd d)) \ length (filter (\t. t_target t \ fst d) (pT @ pT'))) repetition_sets = Some d'" and "\ p' p''. (pT @ pT') = p' @ p'' \ p'' \ [] \ find (\d. Suc (m - card (snd d)) \ length (filter (\t. t_target t \ fst d) p')) repetition_sets = None" using m_traversal_paths_with_witness_set[OF t5 t8 \q \ states M\, of m] by blast+ obtain ioT' where "p_io pT = ioT @ [(x,y)] @ ioT'" using prefixes_set_ob[OF \ioT @ [(x, y)] \ set (prefixes (p_io pT))\] unfolding prefixes_set append.assoc[symmetric] by blast let ?pt = "take (length (ioT @ [(x,y)])) pT" let ?p = "butlast ?pt" let ?t = "last ?pt" have "length ?pt > 0" using \p_io pT = ioT @ [(x,y)] @ ioT'\ unfolding length_map[of "(\ t . (t_input t, t_output t))", symmetric] by auto then have "?pt = ?p @ [?t]" by auto moreover have "path M q ?pt" using \path M q (pT @ pT')\ by (meson path_prefix path_prefix_take) ultimately have "path M q (?p@[?t])" by simp have "p_io ?p = ioT" proof - have "p_io ?pt = take (length (ioT @ [(x,y)])) (p_io pT)" by (simp add: take_map) then have "p_io ?pt = ioT @ [(x,y)]" using \p_io pT = ioT @ [(x,y)] @ ioT'\ by auto then show ?thesis by (simp add: map_butlast) qed have "path M q ?p" using path_append_transition_elim[OF \path M q (?p@[?t])\] by blast have "is_submachine P M" using t2[OF \(q, P) \ prs\] unfolding is_preamble_def by blast then have "initial P = initial M" by auto have "path M (initial M) pP" using submachine_path[OF \is_submachine P M\ \path P (initial P) pP\] unfolding \initial P = initial M\ by assumption moreover have "target (initial M) pP = q" using \target (initial P) pP = q\ unfolding \initial P = initial M\ by assumption ultimately have "path M (initial M) (pP@?p)" using \path M q ?p\ by auto have "find (\d. Suc (m - card (snd d)) \ length (filter (\t. t_target t \ fst d) ?p)) repetition_sets = None" proof - have *: "(pT @ pT') = ?p @ ([?t] @ (drop (length (ioT @ [(x,y)])) pT) @ pT')" using \?pt = ?p @ [?t]\ by (metis append_assoc append_take_drop_id) have **: "([?t] @ (drop (length (ioT @ [(x,y)])) pT) @ pT') \ []" by simp show ?thesis using \\ p' p''. (pT @ pT') = p' @ p'' \ p'' \ [] \ find (\d. Suc (m - card (snd d)) \ length (filter (\t. t_target t \ fst d) p')) repetition_sets = None\[OF * **] by assumption qed (* obtain paths from the transition ending in y' to get a transition from ?p *) obtain p' t' where "path M (FSM.initial M) (p' @ [t'])" and "p_io p' = p_io pP @ ioT" and "t_input t' = x" and "t_output t' = y'" using language_append_path_ob[OF \(p_io pP @ ioT) @ [(x, y')] \ L M\] by blast then have "path M (FSM.initial M) p'" and "t_source t' = target (initial M) p'" and "t' \ transitions M" by auto have "p' = pP @ ?p" using observable_path_unique[OF \observable M\ \path M (FSM.initial M) p'\ \path M (initial M) (pP@?p)\] \p_io ?p = ioT\ unfolding \p_io p' = p_io pP @ ioT\ by simp then have "t_source t' = target q ?p" unfolding \t_source t' = target (initial M) p'\ using \target (initial M) pP = q\ by auto obtain pTt' dt' where "(?p @ [t'] @ pTt', dt') \ m_traversal_paths_with_witness M q repetition_sets m" using m_traversal_path_extension_exist_for_transition[OF \completely_specified M\ \q \ states M\ \FSM.inputs M \ {}\ t5 t8 \path M q ?p\ \find (\d. Suc (m - card (snd d)) \ length (filter (\t. t_target t \ fst d) ?p)) repetition_sets = None\ \t' \ transitions M\ \t_source t' = target q ?p\] by blast have "?p @ [t'] @ pTt' \ tps q" using t6[OF \q \ fst ` prs\ ] \(?p @ [t'] @ pTt', dt') \ m_traversal_paths_with_witness M q repetition_sets m\ by force moreover have "ioT @ [(x, y')] \ set (prefixes (p_io (?p @ [t'] @ pTt')))" proof - have "p_io (?p @ [t'] @ pTt') = ioT @ [(x,y')] @ p_io pTt'" using \p_io ?p = ioT\ using \t_input t' = x\ \t_output t' = y'\ by auto then show ?thesis unfolding prefixes_set by force qed ultimately show "False" using \\ (\pT'. pT' \ tps q \ ioT @ [(x, y')] \ set (prefixes (p_io pT')))\ by blast next case c then obtain q P pP pT q' A d1 d2 qT where "(q, P) \ prs" and "path P (FSM.initial P) pP" and "target (FSM.initial P) pP = q" and "pT \ tps q" and "p_io pP @ p_io pT \ L M'" and "q' \ rd_targets (q, pT)" and "(A, d1, d2) \ atcs (target q pT, q')" and "qT \ io_targets M' (p_io pP @ p_io pT) (FSM.initial M')" and "\pass_separator_ATC M' A qT d2" by blast have "is_submachine P M" using t2[OF \(q, P) \ prs\] unfolding is_preamble_def by blast then have "initial P = initial M" by auto have "path M (initial M) pP" using submachine_path[OF \is_submachine P M\ \path P (initial P) pP\] unfolding \initial P = initial M\ by assumption have "target (initial M) pP = q" using \target (initial P) pP = q\ unfolding \initial P = initial M\ by assumption have "q \ states M" using is_preamble_is_state[OF t2[OF \(q, P) \ prs\]] by assumption have "q \ fst ` prs" using \(q, P) \ prs\ by force obtain pT' d' where "(pT @ pT', d') \ m_traversal_paths_with_witness M q repetition_sets m" using t6[OF \q \ fst ` prs\] \pT \ tps q\ by blast then have "path M q (pT @ pT')" and "find (\d. Suc (m - card (snd d)) \ length (filter (\t. t_target t \ fst d) (pT @ pT'))) repetition_sets = Some d'" and "\ p' p''. (pT @ pT') = p' @ p'' \ p'' \ [] \ find (\d. Suc (m - card (snd d)) \ length (filter (\t. t_target t \ fst d) p')) repetition_sets = None" using m_traversal_paths_with_witness_set[OF t5 t8 \q \ states M\, of m] by blast+ then have "path M q pT" by auto have "qT \ states M'" using \qT \ io_targets M' (p_io pP @ p_io pT) (FSM.initial M')\ io_targets_states subset_iff by fastforce have "is_separator M (target q pT) q' A d1 d2" using t3[OF \(A, d1, d2) \ atcs (target q pT, q')\] by blast have "\ pass_io_set (FSM.from_FSM M' qT) (atc_to_io_set (FSM.from_FSM M (target q pT)) A)" using \\pass_separator_ATC M' A qT d2\ pass_separator_pass_io_set_iff[OF \is_separator M (target q pT) q' A d1 d2\ \observable M\ \observable M'\ path_target_is_state[OF \path M q pT\] \qT \ states M'\ \inputs M' = inputs M\ \completely_specified M\] by simp have "pass_io_set (FSM.from_FSM M' qT) (atc_to_io_set (FSM.from_FSM M (target q pT)) A)" proof - have "\ io x y y' . io @ [(x, y)] \ atc_to_io_set (FSM.from_FSM M (target q pT)) A \ io @ [(x, y')] \ L (FSM.from_FSM M' qT) \ io @ [(x, y')] \ atc_to_io_set (FSM.from_FSM M (target q pT)) A" proof - fix io x y y' assume "io @ [(x, y)] \ atc_to_io_set (FSM.from_FSM M (target q pT)) A" and "io @ [(x, y')] \ L (FSM.from_FSM M' qT)" (* subsets of test_suite_to_io *) define tmp where tmp_def : "tmp = (\ {(\io_atc. p_io p @ p_io pt @ io_atc) ` atc_to_io_set (FSM.from_FSM M (target q pt)) A |p pt q A. \P q' t1 t2. (q, P) \ prs \ path P (FSM.initial P) p \ target (FSM.initial P) p = q \ pt \ tps q \ q' \ rd_targets (q, pt) \ (A, t1, t2) \ atcs (target q pt, q')})" define tmp2 where tmp2_def : "tmp2 = \ {(@) (p_io p) ` set (prefixes (p_io pt)) |p pt. \q P. (q, P) \ prs \ path P (FSM.initial P) p \ target (FSM.initial P) p = q \ pt \ tps q}" have "\P q' t1 t2. (q, P) \ prs \ path P (FSM.initial P) pP \ target (FSM.initial P) pP = q \ pT \ tps q \ q' \ rd_targets (q, pT) \ (A, t1, t2) \ atcs (target q pT, q')" using \(q, P) \ prs\ \path P (FSM.initial P) pP\ \target (FSM.initial P) pP = q\ \pT \ tps q\ \q' \ rd_targets (q, pT)\ \(A, d1, d2) \ atcs (target q pT, q')\ by blast then have "(\io_atc. p_io pP @ p_io pT @ io_atc) ` atc_to_io_set (FSM.from_FSM M (target q pT)) A \ tmp" unfolding tmp_def by blast then have "(\io_atc. p_io pP @ p_io pT @ io_atc) ` atc_to_io_set (FSM.from_FSM M (target q pT)) A \ test_suite_to_io M (Test_Suite prs tps rd_targets atcs)" unfolding test_suite_to_io.simps tmp_def[symmetric] tmp2_def[symmetric] by blast moreover have "(p_io pP @ p_io pT @ (io @ [(x, y)])) \ (\io_atc. p_io pP @ p_io pT @ io_atc) ` atc_to_io_set (FSM.from_FSM M (target q pT)) A" using \io @ [(x, y)] \ atc_to_io_set (FSM.from_FSM M (target q pT)) A\ by auto ultimately have "(p_io pP @ p_io pT @ (io @ [(x, y)])) \ test_suite_to_io M (Test_Suite prs tps rd_targets atcs)" by blast then have *: "(p_io pP @ p_io pT @ io) @ [(x, y)] \ test_suite_to_io M (Test_Suite prs tps rd_targets atcs)" by simp have "io @ [(x, y')] \ LS M' qT" using \io @ [(x, y')] \ L (FSM.from_FSM M' qT)\ \qT \ states M'\ by (metis from_FSM_language) have **: "(p_io pP @ p_io pT @ io) @ [(x, y')] \ L M'" using io_targets_language_append[OF \qT \ io_targets M' (p_io pP @ p_io pT) (FSM.initial M')\ \io @ [(x, y')] \ LS M' qT\] by simp have "(p_io pP @ p_io pT) @ (io @ [(x, y')]) \ test_suite_to_io M (Test_Suite prs tps rd_targets atcs)" using pass_io_prop[OF * ** ] by simp then have "(p_io pP @ p_io pT) @ (io @ [(x, y')]) \ L M" using test_suite_language_prop by blast moreover have "target q pT \ io_targets M (p_io pP @ p_io pT) (initial M)" proof - have "target (initial M) (pP@pT) = target q pT" unfolding \target (initial M) pP = q\[symmetric] by auto moreover have "path M (initial M) (pP@pT)" using \path M (initial M) pP\ \path M q pT\ unfolding \target (initial M) pP = q\[symmetric] by auto moreover have "p_io (pP@pT) = (p_io pP @ p_io pT)" by auto ultimately show ?thesis unfolding io_targets.simps by (metis (mono_tags, lifting) mem_Collect_eq) qed ultimately have "io @ [(x, y')] \ LS M (target q pT)" using observable_io_targets_language[OF _ \observable M\, of "(p_io pP @ p_io pT)" "(io @ [(x, y')])" "initial M" "target q pT"] by blast then have "io @ [(x, y')] \ L (FSM.from_FSM M (target q pT))" unfolding from_FSM_language[OF path_target_is_state[OF \path M q pT\]] by assumption moreover have "io @ [(x, y')] \ L A" by (metis Int_iff \io @ [(x, y')] \ LS M (target q pT)\ \io @ [(x, y)] \ atc_to_io_set (FSM.from_FSM M (target q pT)) A\ \is_separator M (target q pT) q' A d1 d2\ atc_to_io_set.simps is_separator_simps(9)) ultimately show "io @ [(x, y')] \ atc_to_io_set (FSM.from_FSM M (target q pT)) A" unfolding atc_to_io_set.simps by blast qed then show ?thesis unfolding pass_io_set_def by blast qed then show "False" using pass_separator_from_pass_io_set[OF \is_separator M (target q pT) q' A d1 d2\ _ \observable M\ \observable M'\ path_target_is_state[OF \path M q pT\] \qT \ states M'\ \inputs M' = inputs M\ \completely_specified M\] \\pass_separator_ATC M' A qT d2\ by simp qed qed qed moreover have "passes_test_suite M (Test_Suite prs tps rd_targets atcs) M' \ pass_io_set M' (test_suite_to_io M (Test_Suite prs tps rd_targets atcs))" proof - assume "passes_test_suite M (Test_Suite prs tps rd_targets atcs) M'" (* pass properties *) have pass1: "\ q P io x y y' . (q,P) \ prs \ io@[(x,y)] \ L P \ io@[(x,y')] \ L M' \ io@[(x,y')] \ L P" using \passes_test_suite M (Test_Suite prs tps rd_targets atcs) M'\ unfolding passes_test_suite.simps by meson have pass2: "\ q P pP ioT pT x y y' . (q,P) \ prs \ path P (initial P) pP \ target (initial P) pP = q \ pT \ tps q \ ioT@[(x,y)] \ set (prefixes (p_io pT)) \ (p_io pP)@ioT@[(x,y')] \ L M' \ (\ pT' . pT' \ tps q \ ioT@[(x,y')] \ set (prefixes (p_io pT')))" using \passes_test_suite M (Test_Suite prs tps rd_targets atcs) M'\ unfolding passes_test_suite.simps by blast have pass3: "\ q P pP pT q' A d1 d2 qT . (q,P) \ prs \ path P (initial P) pP \ target (initial P) pP = q \ pT \ tps q \ (p_io pP)@(p_io pT) \ L M' \ q' \ rd_targets (q,pT) \ (A,d1,d2) \ atcs (target q pT, q') \ qT \ io_targets M' ((p_io pP)@(p_io pT)) (initial M') \ pass_separator_ATC M' A qT d2" using \passes_test_suite M (Test_Suite prs tps rd_targets atcs) M'\ unfolding passes_test_suite.simps by blast show "pass_io_set M' (test_suite_to_io M (Test_Suite prs tps rd_targets atcs))" proof (rule ccontr) assume "\ pass_io_set M' (test_suite_to_io M (Test_Suite prs tps rd_targets atcs))" then obtain io x y y' where "io @ [(x, y)] \ test_suite_to_io M (Test_Suite prs tps rd_targets atcs)" and "io @ [(x, y')] \ L M'" and "io @ [(x, y')] \ test_suite_to_io M (Test_Suite prs tps rd_targets atcs)" unfolding pass_io_set_def by blast have preamble_prop: "\ q P . (q, P) \ prs \ io @ [(x, y)] \ L P \ False" proof - fix q P assume "(q, P)\prs" and "io @ [(x, y)] \ L P" have "io @ [(x, y')] \ L P" using pass1[OF \(q, P)\prs\ \io @ [(x, y)] \ L P\ \io @ [(x, y')] \ L M'\ ] by assumption then have "io @ [(x, y')] \ test_suite_to_io M (Test_Suite prs tps rd_targets atcs)" unfolding test_suite_to_io.simps using \(q, P)\prs\ by blast then show "False" using \io @ [(x, y')] \ test_suite_to_io M (Test_Suite prs tps rd_targets atcs)\ by simp qed have traversal_path_prop : "\ pP pt q P . io @ [(x, y)] \ (@) (p_io pP) ` set (prefixes (p_io pt)) \ (q, P) \ prs \ path P (FSM.initial P) pP \ target (FSM.initial P) pP = q \ pt \ tps q \ False" proof - fix pP pt q P assume "io @ [(x, y)] \ (@) (p_io pP) ` set (prefixes (p_io pt))" and "(q, P) \ prs" and "path P (FSM.initial P) pP" and "target (FSM.initial P) pP = q" and "pt \ tps q" obtain io' io'' where "io @ [(x, y)] = (p_io pP) @ io'" and "io'@io'' = p_io pt" using \io @ [(x, y)] \ (@) (p_io pP) ` set (prefixes (p_io pt))\ unfolding prefixes_set by blast have "is_submachine P M" using t2[OF \(q, P) \ prs\] unfolding is_preamble_def by blast then have "initial P = initial M" by auto have "path M (initial M) pP" using submachine_path[OF \is_submachine P M\ \path P (initial P) pP\] unfolding \initial P = initial M\ by assumption have "target (initial M) pP = q" using \target (initial P) pP = q\ unfolding \initial P = initial M\ by assumption have "q \ states M" using is_preamble_is_state[OF t2[OF \(q, P) \ prs\]] by assumption have "q \ fst ` prs" using \(q, P) \ prs\ by force show "False" proof (cases io' rule: rev_cases) case Nil then have "p_io pP = io @ [(x, y)]" using \io @ [(x, y)] = (p_io pP) @ io'\ by auto show ?thesis using preamble_prop[OF \(q,P) \ prs\ language_state_containment[OF \path P (FSM.initial P) pP\ \p_io pP = io @ [(x, y)]\]] by assumption next case (snoc ioI xy) then have "xy = (x,y)" and "io = (p_io pP) @ ioI" using \io @ [(x, y)] = (p_io pP) @ io'\ by auto then have "p_io pP @ ioI @ [(x, y')] \ L M'" using \io @ [(x, y')] \ L M'\ by simp have "ioI @ [(x, y)] \ set (prefixes (p_io pt))" unfolding prefixes_set using \io' @ io'' = p_io pt\ \xy = (x, y)\ snoc by auto obtain pT' where "pT' \ tps q" and "ioI @ [(x, y')] \ set (prefixes (p_io pT'))" using pass2[OF \(q, P) \ prs\ \path P (FSM.initial P) pP\ \target (FSM.initial P) pP = q\ \pt \ tps q\ \ioI @ [(x, y)] \ set (prefixes (p_io pt))\ \p_io pP @ ioI @ [(x, y')] \ L M'\] by blast have "io @ [(x, y')] \ (@) (p_io pP) ` set (prefixes (p_io pT'))" using \ioI @ [(x, y')] \ set (prefixes (p_io pT'))\ unfolding \io = (p_io pP) @ ioI\ by simp have "io @ [(x, y')] \ (\ {(@) (p_io p) ` set (prefixes (p_io pt)) |p pt. \q P. (q, P) \ prs \ path P (FSM.initial P) p \ target (FSM.initial P) p = q \ pt \ tps q})" using \(q, P) \ prs\ \path P (FSM.initial P) pP\ \target (FSM.initial P) pP = q\ \pT' \ tps q\ \io @ [(x, y')] \ (@) (p_io pP) ` set (prefixes (p_io pT'))\ by blast then have "io @ [(x, y')] \ test_suite_to_io M (Test_Suite prs tps rd_targets atcs)" unfolding test_suite_to_io.simps by blast then show "False" using \io @ [(x, y')] \ test_suite_to_io M (Test_Suite prs tps rd_targets atcs)\ by blast qed qed consider (a) "io @ [(x, y)] \ (\(q, P)\prs. L P)" | (b) "io @ [(x, y)] \ (\ {(@) (p_io p) ` set (prefixes (p_io pt)) |p pt. \q P. (q, P) \ prs \ path P (FSM.initial P) p \ target (FSM.initial P) p = q \ pt \ tps q})" | (c) "io @ [(x, y)] \ (\ {(\io_atc. p_io p @ p_io pt @ io_atc) ` atc_to_io_set (FSM.from_FSM M (target q pt)) A |p pt q A. \P q' t1 t2. (q, P) \ prs \ path P (FSM.initial P) p \ target (FSM.initial P) p = q \ pt \ tps q \ q' \ rd_targets (q, pt) \ (A, t1, t2) \ atcs (target q pt, q')})" using \io @ [(x, y)] \ test_suite_to_io M (Test_Suite prs tps rd_targets atcs)\ unfolding test_suite_to_io.simps by blast then show "False" proof cases case a then show ?thesis using preamble_prop by blast next case b then show ?thesis using traversal_path_prop by blast next case c then obtain pP pt q A P q' t1 t2 where "io @ [(x, y)] \ (\io_atc. p_io pP @ p_io pt @ io_atc) ` atc_to_io_set (FSM.from_FSM M (target q pt)) A" and "(q, P) \ prs" and "path P (FSM.initial P) pP" and "target (FSM.initial P) pP = q" and "pt \ tps q" and "q' \ rd_targets (q, pt)" and "(A, t1, t2) \ atcs (target q pt, q')" by blast obtain ioA where "io @ [(x, y)] = p_io pP @ p_io pt @ ioA" using \io @ [(x, y)] \ (\io_atc. p_io pP @ p_io pt @ io_atc) ` atc_to_io_set (FSM.from_FSM M (target q pt)) A\ unfolding prefixes_set by blast show "False" proof (cases ioA rule: rev_cases) case Nil then have "io @ [(x, y)] = p_io pP @ p_io pt" using \io @ [(x, y)] = p_io pP @ p_io pt @ ioA\ by simp then have "io @ [(x, y)] \ (@) (p_io pP) ` set (prefixes (p_io pt))" unfolding prefixes_set by blast show ?thesis using traversal_path_prop[OF \io @ [(x, y)] \ (@) (p_io pP) ` set (prefixes (p_io pt))\ \(q, P) \ prs\ \path P (FSM.initial P) pP\ \target (FSM.initial P) pP = q\ \pt \ tps q\] by assumption next case (snoc ioAI xy) then have "xy = (x,y)" and "io = p_io pP @ p_io pt @ ioAI" using \io @ [(x, y)] = p_io pP @ p_io pt @ ioA\ by simp+ then have "p_io pP @ p_io pt @ ioAI @ [(x,y)] \ (\io_atc. p_io pP @ p_io pt @ io_atc) ` atc_to_io_set (FSM.from_FSM M (target q pt)) A" using \io @ [(x, y)] \ (\io_atc. p_io pP @ p_io pt @ io_atc) ` atc_to_io_set (FSM.from_FSM M (target q pt)) A\ by auto then have "ioAI @ [(x,y)] \ atc_to_io_set (FSM.from_FSM M (target q pt)) A" by auto have "p_io pP @ p_io pt \ L M'" using \io @ [(x,y')] \ L M'\ language_prefix[of "p_io pP @ p_io pt" "ioAI @ [(x, y')]" M' "initial M'"] unfolding \io = p_io pP @ p_io pt @ ioAI\ by simp then obtain pt' where "path M' (initial M') pt'" and "p_io pt' = p_io pP @ p_io pt" by auto then have "target (initial M') pt' \ io_targets M' (p_io pP @ p_io pt) (FSM.initial M')" by fastforce have "pass_separator_ATC M' A (target (FSM.initial M') pt') t2" using pass3[OF \(q, P) \ prs\ \path P (FSM.initial P) pP\ \target (FSM.initial P) pP = q\ \pt \ tps q\ \p_io pP @ p_io pt \ L M'\ \q' \ rd_targets (q, pt)\ \(A, t1, t2) \ atcs (target q pt, q')\ \target (initial M') pt' \ io_targets M' (p_io pP @ p_io pt) (FSM.initial M')\] by assumption have "is_separator M (target q pt) q' A t1 t2" using t3[OF \(A, t1, t2) \ atcs (target q pt, q')\] by blast have "is_submachine P M" using t2[OF \(q, P) \ prs\] unfolding is_preamble_def by blast then have "initial P = initial M" by auto have "path M (initial M) pP" using submachine_path[OF \is_submachine P M\ \path P (initial P) pP\] unfolding \initial P = initial M\ by assumption have "target (initial M) pP = q" using \target (initial P) pP = q\ unfolding \initial P = initial M\ by assumption have "q \ states M" using is_preamble_is_state[OF t2[OF \(q, P) \ prs\]] by assumption have "q \ fst ` prs" using \(q, P) \ prs\ by force obtain pT' d' where "(pt @ pT', d') \ m_traversal_paths_with_witness M q repetition_sets m" using t6[OF \q \ fst ` prs\] \pt \ tps q\ by blast then have "path M q (pt @ pT')" and "find (\d. Suc (m - card (snd d)) \ length (filter (\t. t_target t \ fst d) (pt @ pT'))) repetition_sets = Some d'" and "\ p' p''. (pt @ pT') = p' @ p'' \ p'' \ [] \ find (\d. Suc (m - card (snd d)) \ length (filter (\t. t_target t \ fst d) p')) repetition_sets = None" using m_traversal_paths_with_witness_set[OF t5 t8 \q \ states M\, of m] by blast+ then have "path M q pt" by auto have "target (initial M') pt' \ states M'" using \target (initial M') pt' \ io_targets M' (p_io pP @ p_io pt) (FSM.initial M')\ io_targets_states using subset_iff by fastforce have "pass_io_set (FSM.from_FSM M' (target (FSM.initial M') pt')) (atc_to_io_set (FSM.from_FSM M (target q pt)) A)" using pass_io_set_from_pass_separator[OF \is_separator M (target q pt) q' A t1 t2\ \pass_separator_ATC M' A (target (FSM.initial M') pt') t2\ \observable M\ \observable M'\ path_target_is_state[OF \path M q pt\] \target (FSM.initial M') pt' \ FSM.states M'\ \inputs M' = inputs M\] by assumption moreover note \ioAI @ [(x,y)] \ atc_to_io_set (FSM.from_FSM M (target q pt)) A\ moreover have "ioAI @ [(x, y')] \ L (FSM.from_FSM M' (target (FSM.initial M') pt'))" using \io @ [(x,y')] \ L M'\ unfolding \io = p_io pP @ p_io pt @ ioAI\ by (metis (no_types, lifting) \target (FSM.initial M') pt' \ FSM.states M'\ \target (FSM.initial M') pt' \ io_targets M' (p_io pP @ p_io pt) (FSM.initial M')\ append_assoc assms(3) from_FSM_language observable_io_targets_language) ultimately have "ioAI @ [(x,y')] \ atc_to_io_set (FSM.from_FSM M (target q pt)) A" unfolding pass_io_set_def by blast (* subsets of test_suite_to_io *) define tmp where tmp_def : "tmp = (\ {(\io_atc. p_io p @ p_io pt @ io_atc) ` atc_to_io_set (FSM.from_FSM M (target q pt)) A |p pt q A. \P q' t1 t2. (q, P) \ prs \ path P (FSM.initial P) p \ target (FSM.initial P) p = q \ pt \ tps q \ q' \ rd_targets (q, pt) \ (A, t1, t2) \ atcs (target q pt, q')})" define tmp2 where tmp2_def : "tmp2 = \ {(@) (p_io p) ` set (prefixes (p_io pt)) |p pt. \q P. (q, P) \ prs \ path P (FSM.initial P) p \ target (FSM.initial P) p = q \ pt \ tps q}" have "\P q' t1 t2. (q, P) \ prs \ path P (FSM.initial P) pP \ target (FSM.initial P) pP = q \ pt \ tps q \ q' \ rd_targets (q, pt) \ (A, t1, t2) \ atcs (target q pt, q')" using \(q, P) \ prs\ \path P (FSM.initial P) pP\ \target (FSM.initial P) pP = q\ \pt \ tps q\ \q' \ rd_targets (q, pt)\ \(A, t1, t2) \ atcs (target q pt, q')\ by blast then have "(\io_atc. p_io pP @ p_io pt @ io_atc) ` atc_to_io_set (FSM.from_FSM M (target q pt)) A \ tmp" unfolding tmp_def by blast then have "(\io_atc. p_io pP @ p_io pt @ io_atc) ` atc_to_io_set (FSM.from_FSM M (target q pt)) A \ test_suite_to_io M (Test_Suite prs tps rd_targets atcs)" unfolding test_suite_to_io.simps tmp_def[symmetric] tmp2_def[symmetric] by blast moreover have "(p_io pP @ p_io pt @ (ioAI @ [(x, y')])) \ (\io_atc. p_io pP @ p_io pt @ io_atc) ` atc_to_io_set (FSM.from_FSM M (target q pt)) A" using \ioAI @ [(x, y')] \ atc_to_io_set (FSM.from_FSM M (target q pt)) A\ by auto ultimately have "(p_io pP @ p_io pt @ (ioAI @ [(x, y')])) \ test_suite_to_io M (Test_Suite prs tps rd_targets atcs)" by blast then have "io @ [(x, y')] \ test_suite_to_io M (Test_Suite prs tps rd_targets atcs)" unfolding \io = p_io pP @ p_io pt @ ioAI\ by auto then show "False" using \io @ [(x, y')] \ test_suite_to_io M (Test_Suite prs tps rd_targets atcs)\ by blast qed qed qed qed ultimately show ?thesis unfolding \T = Test_Suite prs tps rd_targets atcs\ by blast qed lemma test_suite_to_io_finite : assumes "implies_completeness T M m" and "is_finite_test_suite T" shows "finite (test_suite_to_io M T)" proof - obtain prs tps rd_targets atcs where "T = Test_Suite prs tps rd_targets atcs" by (meson test_suite.exhaust) then obtain repetition_sets where repetition_sets_def: "implies_completeness_for_repetition_sets (Test_Suite prs tps rd_targets atcs) M m repetition_sets" using assms(1) unfolding implies_completeness_def by blast then have "implies_completeness (Test_Suite prs tps rd_targets atcs) M m" unfolding implies_completeness_def by blast then have test_suite_language_prop: "test_suite_to_io M (Test_Suite prs tps rd_targets atcs) \ L M" using test_suite_to_io_language by blast have f1: "(finite prs)" and f2: "\ q p . q \ fst ` prs \ finite (rd_targets (q,p))" and f3: "\ q q' . finite (atcs (q,q'))" using assms(2) unfolding \T = Test_Suite prs tps rd_targets atcs\ is_finite_test_suite.simps by blast+ have t1: "(initial M, initial_preamble M) \ prs" using implies_completeness_for_repetition_sets_simps(1)[OF repetition_sets_def] by assumption have t2: "\ q P. (q, P) \ prs \ is_preamble P M q" using implies_completeness_for_repetition_sets_simps(2)[OF repetition_sets_def] by blast have t3: "\ q1 q2 A d1 d2. (A, d1, d2) \ atcs (q1, q2) \ (A, d2, d1) \ atcs (q2, q1) \ is_separator M q1 q2 A d1 d2" using implies_completeness_for_repetition_sets_simps(3)[OF repetition_sets_def] by assumption have t5: "\q. q \ FSM.states M \ (\d\set repetition_sets. q \ fst d)" using implies_completeness_for_repetition_sets_simps(4)[OF repetition_sets_def] by assumption have t6: "\ q. q \ fst ` prs \ tps q \ {p1 . \ p2 d . (p1@p2,d) \ m_traversal_paths_with_witness M q repetition_sets m} \ fst ` (m_traversal_paths_with_witness M q repetition_sets m) \ tps q" using implies_completeness_for_repetition_sets_simps(7)[OF repetition_sets_def] by assumption have t7: "\ d. d \ set repetition_sets \ fst d \ FSM.states M" and t8: "\ d. d \ set repetition_sets \ snd d \ fst d" and t8': "\ d. d \ set repetition_sets \ snd d = fst d \ fst ` prs" and t9: "\ d q1 q2. d \ set repetition_sets \ q1 \ fst d \ q2 \ fst d \ q1 \ q2 \ atcs (q1, q2) \ {}" using implies_completeness_for_repetition_sets_simps(5,6)[OF repetition_sets_def] by blast+ have f4: "\ q . q \ fst ` prs \ finite (tps q)" proof - fix q assume "q \ fst ` prs" then have "tps q \ {p1 . \ p2 d . (p1@p2,d) \ m_traversal_paths_with_witness M q repetition_sets m}" using t6 by blast moreover have "{p1 . \ p2 d . (p1@p2,d) \ m_traversal_paths_with_witness M q repetition_sets m} \ (\ p2 \ fst ` m_traversal_paths_with_witness M q repetition_sets m . set (prefixes p2))" proof fix p1 assume "p1 \ {p1 . \ p2 d . (p1@p2,d) \ m_traversal_paths_with_witness M q repetition_sets m}" then obtain p2 d where "(p1@p2,d) \ m_traversal_paths_with_witness M q repetition_sets m" by blast then have "p1@p2 \ fst ` m_traversal_paths_with_witness M q repetition_sets m" by force moreover have "p1 \ set (prefixes (p1@p2))" unfolding prefixes_set by blast ultimately show "p1 \ (\ p2 \ fst ` m_traversal_paths_with_witness M q repetition_sets m . set (prefixes p2))" by blast qed ultimately have "tps q \ (\ p2 \ fst ` m_traversal_paths_with_witness M q repetition_sets m . set (prefixes p2))" by simp moreover have "finite (\ p2 \ fst ` m_traversal_paths_with_witness M q repetition_sets m . set (prefixes p2))" proof - have "finite (fst ` m_traversal_paths_with_witness M q repetition_sets m)" using m_traversal_paths_with_witness_finite[of M q repetition_sets m] by auto moreover have "\ p2 . finite (set (prefixes p2))" by auto ultimately show ?thesis by blast qed ultimately show "finite (tps q)" using finite_subset by blast qed then have f4' : "\ q P . (q,P) \ prs \ finite (tps q)" by force define T1 where T1_def : "T1 = (\(q, P)\prs. L P)" define T2 where T2_def : "T2 = \{(@) (p_io p) ` set (prefixes (p_io pt)) |p pt. \q P. (q, P) \ prs \ path P (FSM.initial P) p \ target (FSM.initial P) p = q \ pt \ tps q}" define T3 where T3_def : "T3 = \ {(\io_atc. p_io p @ p_io pt @ io_atc) ` atc_to_io_set (FSM.from_FSM M (target q pt)) A |p pt q A. \P q' t1 t2. (q, P) \ prs \ path P (FSM.initial P) p \ target (FSM.initial P) p = q \ pt \ tps q \ q' \ rd_targets (q, pt) \ (A, t1, t2) \ atcs (target q pt, q')}" have "test_suite_to_io M T = T1 \ T2 \ T3" unfolding \T = Test_Suite prs tps rd_targets atcs\ test_suite_to_io.simps T1_def T2_def T3_def by simp moreover have "finite T1" proof - have "\ q P . (q, P)\prs \ finite (L P)" proof - fix q P assume "(q, P)\prs" have "acyclic P" using t2[OF \(q, P)\prs\] unfolding is_preamble_def by blast then show "finite (L P)" using acyclic_alt_def by blast qed then show ?thesis using f1 unfolding T1_def by auto qed moreover have "finite T2" proof - have *: "T2 = (\ (p,pt) \ {(p,pt) | p pt. \q P. (q, P) \ prs \ path P (FSM.initial P) p \ target (FSM.initial P) p = q \ pt \ tps q} . ((@) (p_io p) ` set (prefixes (p_io pt))))" unfolding T2_def by auto have "\ p pt . finite ((@) (p_io p) ` set (prefixes (p_io pt)))" by auto moreover have "finite {(p,pt) | p pt. \q P. (q, P) \ prs \ path P (FSM.initial P) p \ target (FSM.initial P) p = q \ pt \ tps q}" proof - have "{(p,pt) | p pt. \q P. (q, P) \ prs \ path P (FSM.initial P) p \ target (FSM.initial P) p = q \ pt \ tps q} \ (\ (q,P) \ prs . {p . path P (initial P) p} \ (tps q))" by auto moreover have "finite (\ (q,P) \ prs . {p . path P (initial P) p} \ (tps q))" proof - note \finite prs\ moreover have "\ q P . (q,P) \ prs \ finite ({p . path P (initial P) p} \ (tps q))" proof - fix q P assume "(q,P) \ prs" have "acyclic P" using t2[OF \(q, P)\prs\] unfolding is_preamble_def by blast then have "finite {p . path P (initial P) p}" using acyclic_paths_finite[of P "initial P"] unfolding acyclic.simps by (metis (no_types, lifting) Collect_cong) then show "finite ({p . path P (initial P) p} \ (tps q))" using f4'[OF \(q,P) \ prs\] by simp qed ultimately show ?thesis by force qed ultimately show ?thesis by (meson rev_finite_subset) qed ultimately show ?thesis unfolding * by auto qed moreover have "finite T3" proof - have scheme: "\ f P . (\ {f a b c d | a b c d . P a b c d}) = (\ (a,b,c,d) \ {(a,b,c,d) | a b c d . P a b c d} . f a b c d)" by blast have *: "T3 = (\ (p,pt,q,A) \ {(p, pt, q, A) | p pt q A . \P q' t1 t2. (q, P) \ prs \ path P (FSM.initial P) p \ target (FSM.initial P) p = q \ pt \ tps q \ q' \ rd_targets (q, pt) \ (A, t1, t2) \ atcs (target q pt, q')} . (\io_atc. p_io p @ p_io pt @ io_atc) ` atc_to_io_set (FSM.from_FSM M (target q pt)) A)" unfolding T3_def scheme by blast have "{(p, pt, q, A) | p pt q A . \P q' t1 t2. (q, P) \ prs \ path P (FSM.initial P) p \ target (FSM.initial P) p = q \ pt \ tps q \ q' \ rd_targets (q, pt) \ (A, t1, t2) \ atcs (target q pt, q')} \ (\ (q,P) \ prs . \ pt \ tps q . \ q' \ rd_targets (q, pt) . (\ (A, t1, t2) \ atcs (target q pt, q') . {p . path P (initial P) p} \ {pt} \ {q} \ {A}))" by blast moreover have "finite (\ (q,P) \ prs . \ pt \ tps q . \ q' \ rd_targets (q, pt) . (\ (A, t1, t2) \ atcs (target q pt, q') . {p . path P (initial P) p} \ {pt} \ {q} \ {A}))" proof - note \finite prs\ moreover have "\ q P . (q,P) \ prs \ finite (\ pt \ tps q . \ q' \ rd_targets (q, pt) . (\ (A, t1, t2) \ atcs (target q pt, q') . {p . path P (initial P) p} \ {pt} \ {q} \ {A}))" proof - fix q P assume "(q,P) \ prs" then have "q \ fst ` prs" by force have "finite (tps q)" using f4'[OF \(q,P) \ prs\] by assumption moreover have "\ pt . pt \ tps q \ finite (\ q' \ rd_targets (q, pt) . (\ (A, t1, t2) \ atcs (target q pt, q') . {p . path P (initial P) p} \ {pt} \ {q} \ {A}))" proof - fix pt assume "pt \ tps q" have "finite (rd_targets (q,pt))" using f2[OF \q \ fst ` prs\] by blast moreover have "\ q' . q' \ rd_targets (q, pt) \ finite (\ (A, t1, t2) \ atcs (target q pt, q') . {p . path P (initial P) p} \ {pt} \ {q} \ {A})" proof - fix q' assume "q' \ rd_targets (q, pt)" have "finite (atcs (target q pt, q'))" using f3 by blast moreover have "finite {p . path P (initial P) p}" proof - have "acyclic P" using t2[OF \(q, P)\prs\] unfolding is_preamble_def by blast then show ?thesis using acyclic_paths_finite[of P "initial P"] unfolding acyclic.simps by (metis (no_types, lifting) Collect_cong) qed ultimately show "finite (\ (A, t1, t2) \ atcs (target q pt, q') . {p . path P (initial P) p} \ {pt} \ {q} \ {A})" by force qed ultimately show "finite (\ q' \ rd_targets (q, pt) . (\ (A, t1, t2) \ atcs (target q pt, q') . {p . path P (initial P) p} \ {pt} \ {q} \ {A}))" by force qed ultimately show "finite (\ pt \ tps q . \ q' \ rd_targets (q, pt) . (\ (A, t1, t2) \ atcs (target q pt, q') . {p . path P (initial P) p} \ {pt} \ {q} \ {A}))" by force qed ultimately show ?thesis by force qed ultimately have "finite {(p, pt, q, A) | p pt q A . \P q' t1 t2. (q, P) \ prs \ path P (FSM.initial P) p \ target (FSM.initial P) p = q \ pt \ tps q \ q' \ rd_targets (q, pt) \ (A, t1, t2) \ atcs (target q pt, q')}" by (meson rev_finite_subset) moreover have "\ p pt q A . (p,pt,q,A) \ {(p, pt, q, A) | p pt q A . \P q' t1 t2. (q, P) \ prs \ path P (FSM.initial P) p \ target (FSM.initial P) p = q \ pt \ tps q \ q' \ rd_targets (q, pt) \ (A, t1, t2) \ atcs (target q pt, q')} \ finite ((\io_atc. p_io p @ p_io pt @ io_atc) ` atc_to_io_set (FSM.from_FSM M (target q pt)) A)" proof - fix p pt q A assume "(p,pt,q,A) \ {(p, pt, q, A) | p pt q A . \P q' t1 t2. (q, P) \ prs \ path P (FSM.initial P) p \ target (FSM.initial P) p = q \ pt \ tps q \ q' \ rd_targets (q, pt) \ (A, t1, t2) \ atcs (target q pt, q')}" then obtain P q' t1 t2 where "(q, P) \ prs" and "path P (FSM.initial P) p" and "target (FSM.initial P) p = q" and "pt \ tps q" and "q' \ rd_targets (q, pt)" and "(A, t1, t2) \ atcs (target q pt, q')" by blast have "is_separator M (target q pt) q' A t1 t2" using t3[OF \(A, t1, t2) \ atcs (target q pt, q')\] by blast then have "acyclic A" using is_separator_simps(2) by simp then have "finite (L A)" unfolding acyclic_alt_def by assumption then have "finite (atc_to_io_set (FSM.from_FSM M (target q pt)) A)" unfolding atc_to_io_set.simps by blast then show "finite ((\io_atc. p_io p @ p_io pt @ io_atc) ` atc_to_io_set (FSM.from_FSM M (target q pt)) A)" by blast qed ultimately show ?thesis unfolding * by force qed ultimately show ?thesis by simp qed subsection \Calculating the Sets of Sequences\ abbreviation "L_acyclic M \ LS_acyclic M (initial M)" (* collect the prefix-closed IO-sequence representation of a test suite for some reference FSM *) fun test_suite_to_io' :: "('a,'b,'c) fsm \ ('a,'b,'c,'d) test_suite \ ('b \ 'c) list set" where "test_suite_to_io' M (Test_Suite prs tps rd_targets atcs) = (\ (q,P) \ prs . L_acyclic P \ (\ ioP \ remove_proper_prefixes (L_acyclic P) . \ pt \ tps q . ((\ io' . ioP @ io') ` (set (prefixes (p_io pt)))) \ (\ q' \ rd_targets (q,pt) . \ (A,t1,t2) \ atcs (target q pt,q') . (\ io_atc . ioP @ p_io pt @ io_atc) ` (acyclic_language_intersection (from_FSM M (target q pt)) A))))" lemma test_suite_to_io_code : assumes "implies_completeness T M m" and "is_finite_test_suite T" and "observable M" shows "test_suite_to_io M T = test_suite_to_io' M T" proof - obtain prs tps rd_targets atcs where "T = Test_Suite prs tps rd_targets atcs" by (meson test_suite.exhaust) then obtain repetition_sets where repetition_sets_def: "implies_completeness_for_repetition_sets (Test_Suite prs tps rd_targets atcs) M m repetition_sets" using assms(1) unfolding implies_completeness_def by blast have t2: "\ q P. (q, P) \ prs \ is_preamble P M q" using implies_completeness_for_repetition_sets_simps(2)[OF repetition_sets_def] by blast have t3: "\ q1 q2 A d1 d2. (A, d1, d2) \ atcs (q1, q2) \ (A, d2, d1) \ atcs (q2, q1) \ is_separator M q1 q2 A d1 d2" using implies_completeness_for_repetition_sets_simps(3)[OF repetition_sets_def] by assumption have test_suite_to_io'_alt_def: "test_suite_to_io' M T = (\ (q,P) \ prs . L_acyclic P) \ (\ (q,P) \ prs . \ ioP \ remove_proper_prefixes (L_acyclic P) . \ pt \ tps q . ((\ io' . ioP @ io') ` (set (prefixes (p_io pt))))) \ (\ (q,P) \ prs . \ ioP \ remove_proper_prefixes (L_acyclic P) . \ pt \ tps q . \ q' \ rd_targets (q,pt) . \ (A,t1,t2) \ atcs (target q pt,q') . (\ io_atc . ioP @ p_io pt @ io_atc) ` (acyclic_language_intersection (from_FSM M (target q pt)) A))" unfolding test_suite_to_io'.simps \T = Test_Suite prs tps rd_targets atcs\ by fast have test_suite_to_io_alt_def: "test_suite_to_io M T = (\ (q,P) \ prs . L P) \ (\{(\ io' . p_io p @ io') ` (set (prefixes (p_io pt))) | p pt . \ q P . (q,P) \ prs \ path P (initial P) p \ target (initial P) p = q \ pt \ tps q}) \ (\{(\ io_atc . p_io p @ p_io pt @ io_atc) ` (atc_to_io_set (from_FSM M (target q pt)) A) | p pt q A . \ P q' t1 t2 . (q,P) \ prs \ path P (initial P) p \ target (initial P) p = q \ pt \ tps q \ q' \ rd_targets (q,pt) \ (A,t1,t2) \ atcs (target q pt,q') })" unfolding \T = Test_Suite prs tps rd_targets atcs\ test_suite_to_io.simps by force have preamble_language_prop: "\ q P . (q,P) \ prs \ L_acyclic P = L P" proof - fix q P assume "(q,P) \ prs" have "acyclic P" using t2[OF \(q, P)\prs\] unfolding is_preamble_def by blast then show "L_acyclic P = L P" using LS_from_LS_acyclic by blast qed have preamble_path_prop: "\ q P ioP . (q,P) \ prs \ ioP \ remove_proper_prefixes (L_acyclic P) \ (\ p . path P (initial P) p \ target (initial P) p = q \ p_io p = ioP)" proof - fix q P ioP assume "(q,P) \ prs" have "is_preamble P M q" using t2[OF \(q, P)\prs\] by assumption have "ioP \ remove_proper_prefixes (L_acyclic P) \ (\ p . path P (initial P) p \ target (initial P) p = q \ p_io p = ioP)" proof - assume "ioP \ remove_proper_prefixes (L_acyclic P)" then have "ioP \ L P" and "\x'. x' \ [] \ ioP @ x' \ L P" unfolding preamble_language_prop[OF \(q,P) \ prs\] remove_proper_prefixes_def by blast+ show "(\ p . path P (initial P) p \ target (initial P) p = q \ p_io p = ioP)" using preamble_maximal_io_paths_rev[OF \is_preamble P M q\ \observable M\ \ioP \ L P\ \\x'. x' \ [] \ ioP @ x' \ L P\] by blast qed moreover have "(\ p . path P (initial P) p \ target (initial P) p = q \ p_io p = ioP) \ ioP \ remove_proper_prefixes (L_acyclic P)" proof - assume "(\ p . path P (initial P) p \ target (initial P) p = q \ p_io p = ioP)" then obtain p where "path P (initial P) p" and "target (initial P) p = q" and "p_io p = ioP" by blast then have "\io'. io' \ [] \ p_io p @ io' \ L P" using preamble_maximal_io_paths[OF \is_preamble P M q\ \observable M\] by blast then show "ioP \ remove_proper_prefixes (L_acyclic P)" using language_state_containment[OF \path P (initial P) p\ \p_io p = ioP\] unfolding preamble_language_prop[OF \(q,P) \ prs\] remove_proper_prefixes_def \p_io p = ioP\ by blast qed ultimately show "ioP \ remove_proper_prefixes (L_acyclic P) \ (\ p . path P (initial P) p \ target (initial P) p = q \ p_io p = ioP)" by blast qed have eq1: "(\ (q,P) \ prs . L_acyclic P) = (\ (q,P) \ prs . L P)" using preamble_language_prop by blast have eq2: "(\ (q,P) \ prs . \ ioP \ remove_proper_prefixes (L_acyclic P) . \ pt \ tps q . ((\ io' . ioP @ io') ` (set (prefixes (p_io pt))))) = (\{(\ io' . p_io p @ io') ` (set (prefixes (p_io pt))) | p pt . \ q P . (q,P) \ prs \ path P (initial P) p \ target (initial P) p = q \ pt \ tps q})" proof show "(\(q, P)\prs. \ioP\remove_proper_prefixes (L_acyclic P). \pt\tps q. (@) ioP ` set (prefixes (p_io pt))) \ \ {(@) (p_io p) ` set (prefixes (p_io pt)) |p pt. \q P. (q, P) \ prs \ path P (FSM.initial P) p \ target (FSM.initial P) p = q \ pt \ tps q}" proof fix io assume "io \ (\ (q,P) \ prs . (\ ioP \ remove_proper_prefixes (L_acyclic P) . \ pt \ tps q . ((\ io' . ioP @ io') ` (set (prefixes (p_io pt))))))" then obtain q P where "(q,P) \ prs" and "io \ (\ ioP \ remove_proper_prefixes (L_acyclic P) . \ pt \ tps q . ((\ io' . ioP @ io') ` (set (prefixes (p_io pt)))))" by blast then obtain ioP where "ioP \ remove_proper_prefixes (L_acyclic P)" and "io \ (\ pt \ tps q . ((\ io' . ioP @ io') ` (set (prefixes (p_io pt)))))" by blast obtain p where "path P (initial P) p" and "target (initial P) p = q" and "ioP = p_io p" using preamble_path_prop[OF \(q,P) \ prs\, of ioP] \ioP \ remove_proper_prefixes (L_acyclic P)\ by auto obtain pt where "pt \ tps q" and "io \ ((\ io' . p_io p @ io') ` (set (prefixes (p_io pt))))" using \io \ (\ pt \ tps q . ((\ io' . ioP @ io') ` (set (prefixes (p_io pt)))))\ unfolding \ioP = p_io p\ by blast show "io \ (\{(\ io' . p_io p @ io') ` (set (prefixes (p_io pt))) | p pt . \ q P . (q,P) \ prs \ path P (initial P) p \ target (initial P) p = q \ pt \ tps q})" using \io \ ((\ io' . p_io p @ io') ` (set (prefixes (p_io pt))))\ \(q,P) \ prs\ \path P (initial P) p\ \target (initial P) p = q\ \pt \ tps q\ by blast qed show "\ {(@) (p_io p) ` set (prefixes (p_io pt)) |p pt. \q P. (q, P) \ prs \ path P (FSM.initial P) p \ target (FSM.initial P) p = q \ pt \ tps q} \ (\(q, P)\prs. \ioP\remove_proper_prefixes (L_acyclic P). \pt\tps q. (@) ioP ` set (prefixes (p_io pt)))" proof fix io assume "io \ (\{(\ io' . p_io p @ io') ` (set (prefixes (p_io pt))) | p pt . \ q P . (q,P) \ prs \ path P (initial P) p \ target (initial P) p = q \ pt \ tps q})" then obtain p pt q P where "io \ (\ io' . p_io p @ io') ` (set (prefixes (p_io pt)))" and "(q,P) \ prs" and "path P (initial P) p" and "target (initial P) p = q" and "pt \ tps q" by blast then obtain ioP where "ioP \ remove_proper_prefixes (L_acyclic P)" and "p_io p = ioP" using preamble_path_prop[OF \(q,P) \ prs\, of "p_io p"] by blast show "io \ (\ (q,P) \ prs . (\ ioP \ remove_proper_prefixes (L_acyclic P) . \ pt \ tps q . ((\ io' . ioP @ io') ` (set (prefixes (p_io pt))))))" using \(q,P) \ prs\ \ioP \ remove_proper_prefixes (L_acyclic P)\ \pt \ tps q\ \io \ (\ io' . p_io p @ io') ` (set (prefixes (p_io pt)))\ unfolding \p_io p = ioP\ by blast qed qed have eq3: "(\ (q,P) \ prs . \ ioP \ remove_proper_prefixes (L_acyclic P) . \ pt \ tps q . \ q' \ rd_targets (q,pt) . \ (A,t1,t2) \ atcs (target q pt,q') . (\ io_atc . ioP @ p_io pt @ io_atc) ` (acyclic_language_intersection (from_FSM M (target q pt)) A)) = (\{(\ io_atc . p_io p @ p_io pt @ io_atc) ` (atc_to_io_set (from_FSM M (target q pt)) A) | p pt q A . \ P q' t1 t2 . (q,P) \ prs \ path P (initial P) p \ target (initial P) p = q \ pt \ tps q \ q' \ rd_targets (q,pt) \ (A,t1,t2) \ atcs (target q pt,q') })" proof show "(\ (q,P) \ prs . \ ioP \ remove_proper_prefixes (L_acyclic P) . \ pt \ tps q . \ q' \ rd_targets (q,pt) . \ (A,t1,t2) \ atcs (target q pt,q') . (\ io_atc . ioP @ p_io pt @ io_atc) ` (acyclic_language_intersection (from_FSM M (target q pt)) A)) \ (\{(\ io_atc . p_io p @ p_io pt @ io_atc) ` (atc_to_io_set (from_FSM M (target q pt)) A) | p pt q A . \ P q' t1 t2 . (q,P) \ prs \ path P (initial P) p \ target (initial P) p = q \ pt \ tps q \ q' \ rd_targets (q,pt) \ (A,t1,t2) \ atcs (target q pt,q') })" proof fix io assume "io \ (\ (q,P) \ prs . \ ioP \ remove_proper_prefixes (L_acyclic P) . \ pt \ tps q . \ q' \ rd_targets (q,pt) . \ (A,t1,t2) \ atcs (target q pt,q') . (\ io_atc . ioP @ p_io pt @ io_atc) ` (acyclic_language_intersection (from_FSM M (target q pt)) A))" then obtain q P ioP pt q' A t1 t2 where "(q,P) \ prs" and "ioP \ remove_proper_prefixes (L_acyclic P)" and "pt \ tps q" and "q' \ rd_targets (q,pt)" and "(A,t1,t2) \ atcs (target q pt,q')" and "io \ (\ io_atc . ioP @ p_io pt @ io_atc) ` (acyclic_language_intersection (from_FSM M (target q pt)) A)" by blast obtain p where "path P (initial P) p" and "target (initial P) p = q" and "ioP = p_io p" using preamble_path_prop[OF \(q,P) \ prs\, of ioP] \ioP \ remove_proper_prefixes (L_acyclic P)\ by auto have "acyclic A" using t3[OF \(A,t1,t2) \ atcs (target q pt,q')\] is_separator_simps(2) by metis have "(acyclic_language_intersection (from_FSM M (target q pt)) A) = (atc_to_io_set (from_FSM M (target q pt)) A)" unfolding acyclic_language_intersection_completeness[OF \acyclic A\] atc_to_io_set.simps by simp have "io \ (\ io_atc . p_io p @ p_io pt @ io_atc) ` (atc_to_io_set (from_FSM M (target q pt)) A)" using \io \ (\ io_atc . ioP @ p_io pt @ io_atc) ` (acyclic_language_intersection (from_FSM M (target q pt)) A)\ unfolding \(acyclic_language_intersection (from_FSM M (target q pt)) A) = (atc_to_io_set (from_FSM M (target q pt)) A)\ \ioP = p_io p\ by simp then show "io \ (\{(\ io_atc . p_io p @ p_io pt @ io_atc) ` (atc_to_io_set (from_FSM M (target q pt)) A) | p pt q A . \ P q' t1 t2 . (q,P) \ prs \ path P (initial P) p \ target (initial P) p = q \ pt \ tps q \ q' \ rd_targets (q,pt) \ (A,t1,t2) \ atcs (target q pt,q') })" using \(q,P) \ prs\ \path P (initial P) p\ \target (initial P) p = q\ \pt \ tps q\ \q' \ rd_targets (q,pt)\ \(A,t1,t2) \ atcs (target q pt,q')\ by blast qed show "(\{(\ io_atc . p_io p @ p_io pt @ io_atc) ` (atc_to_io_set (from_FSM M (target q pt)) A) | p pt q A . \ P q' t1 t2 . (q,P) \ prs \ path P (initial P) p \ target (initial P) p = q \ pt \ tps q \ q' \ rd_targets (q,pt) \ (A,t1,t2) \ atcs (target q pt,q') }) \ (\ (q,P) \ prs . \ ioP \ remove_proper_prefixes (L_acyclic P) . \ pt \ tps q . \ q' \ rd_targets (q,pt) . \ (A,t1,t2) \ atcs (target q pt,q') . (\ io_atc . ioP @ p_io pt @ io_atc) ` (acyclic_language_intersection (from_FSM M (target q pt)) A))" proof fix io assume "io \ (\{(\ io_atc . p_io p @ p_io pt @ io_atc) ` (atc_to_io_set (from_FSM M (target q pt)) A) | p pt q A . \ P q' t1 t2 . (q,P) \ prs \ path P (initial P) p \ target (initial P) p = q \ pt \ tps q \ q' \ rd_targets (q,pt) \ (A,t1,t2) \ atcs (target q pt,q') })" then obtain p pt q A P q' t1 t2 where "io \ (\ io_atc . p_io p @ p_io pt @ io_atc) ` (atc_to_io_set (from_FSM M (target q pt)) A)" and "(q,P) \ prs" and "path P (initial P) p" and "target (initial P) p = q" and "pt \ tps q" and "q' \ rd_targets (q,pt)" and "(A,t1,t2) \ atcs (target q pt,q')" by blast then obtain ioP where "ioP \ remove_proper_prefixes (L_acyclic P)" and "p_io p = ioP" using preamble_path_prop[OF \(q,P) \ prs\, of "p_io p"] by blast have "acyclic A" using t3[OF \(A,t1,t2) \ atcs (target q pt,q')\] is_separator_simps(2) by metis have *: "(atc_to_io_set (from_FSM M (target q pt)) A) = (acyclic_language_intersection (from_FSM M (target q pt)) A)" unfolding acyclic_language_intersection_completeness[OF \acyclic A\] atc_to_io_set.simps by simp have "io \ (\ io_atc . ioP @ p_io pt @ io_atc) ` (acyclic_language_intersection (from_FSM M (target q pt)) A)" using \io \ (\ io_atc . p_io p @ p_io pt @ io_atc) ` (atc_to_io_set (from_FSM M (target q pt)) A)\ unfolding * \p_io p = ioP\ by simp then show "io \ (\ (q,P) \ prs . \ ioP \ remove_proper_prefixes (L_acyclic P) . \ pt \ tps q . \ q' \ rd_targets (q,pt) . \ (A,t1,t2) \ atcs (target q pt,q') . (\ io_atc . ioP @ p_io pt @ io_atc) ` (acyclic_language_intersection (from_FSM M (target q pt)) A))" using \(q,P) \ prs\ \ioP \ remove_proper_prefixes (L_acyclic P)\ \pt \ tps q\ \q' \ rd_targets (q,pt)\ \(A,t1,t2) \ atcs (target q pt,q')\ by force qed qed show ?thesis unfolding test_suite_to_io'_alt_def test_suite_to_io_alt_def eq1 eq2 eq3 by simp qed subsection \Using Maximal Sequences Only\ fun test_suite_to_io_maximal :: "('a::linorder,'b::linorder,'c) fsm \ ('a,'b,'c,'d::linorder) test_suite \ ('b \ 'c) list set" where "test_suite_to_io_maximal M (Test_Suite prs tps rd_targets atcs) = remove_proper_prefixes (\ (q,P) \ prs . L_acyclic P \ (\ ioP \ remove_proper_prefixes (L_acyclic P) . \ pt \ tps q . Set.insert (ioP @ p_io pt) (\ q' \ rd_targets (q,pt) . \ (A,t1,t2) \ atcs (target q pt,q') . (\ io_atc . ioP @ p_io pt @ io_atc) ` (remove_proper_prefixes (acyclic_language_intersection (from_FSM M (target q pt)) A)))))" lemma test_suite_to_io_maximal_code : assumes "implies_completeness T M m" and "is_finite_test_suite T" and "observable M" shows "{io' \ (test_suite_to_io M T) . \ (\ io'' . io'' \ [] \ io'@io'' \ (test_suite_to_io M T))} = test_suite_to_io_maximal M T" proof - obtain prs tps rd_targets atcs where "T = Test_Suite prs tps rd_targets atcs" by (meson test_suite.exhaust) have t_def: "test_suite_to_io M T = test_suite_to_io' M T" using test_suite_to_io_code[OF assms] by assumption have t1_def : "test_suite_to_io' M T = (\ (q,P) \ prs . L_acyclic P \ (\ ioP \ remove_proper_prefixes (L_acyclic P) . \ pt \ tps q . ((\ io' . ioP @ io') ` (set (prefixes (p_io pt)))) \ (\ q' \ rd_targets (q,pt) . \ (A,t1,t2) \ atcs (target q pt,q') . (\ io_atc . ioP @ p_io pt @ io_atc) ` (acyclic_language_intersection (from_FSM M (target q pt)) A))))" unfolding \T = Test_Suite prs tps rd_targets atcs\ by simp define tmax where tmax_def: "tmax = (\ (q,P) \ prs . L_acyclic P \ (\ ioP \ remove_proper_prefixes (L_acyclic P) . \ pt \ tps q . Set.insert (ioP @ p_io pt) (\ q' \ rd_targets (q,pt) . \ (A,t1,t2) \ atcs (target q pt,q') . (\ io_atc . ioP @ p_io pt @ io_atc) ` (remove_proper_prefixes (acyclic_language_intersection (from_FSM M (target q pt)) A)))))" have t2_def : "test_suite_to_io_maximal M T = remove_proper_prefixes tmax" unfolding \T = Test_Suite prs tps rd_targets atcs\ tmax_def by simp have tmax_sub: "tmax \ (test_suite_to_io M T)" unfolding tmax_def t_def t1_def proof fix io assume "io \ (\ (q,P) \ prs . L_acyclic P \ (\ ioP \ remove_proper_prefixes (L_acyclic P) . \ pt \ tps q . Set.insert (ioP @ p_io pt) (\ q' \ rd_targets (q,pt) . \ (A,t1,t2) \ atcs (target q pt,q') . (\ io_atc . ioP @ p_io pt @ io_atc) ` (remove_proper_prefixes (acyclic_language_intersection (from_FSM M (target q pt)) A)))))" then obtain q P where "(q,P) \ prs" and "io \ L_acyclic P \ (\ ioP \ remove_proper_prefixes (L_acyclic P) . \ pt \ tps q . Set.insert (ioP @ p_io pt) (\ q' \ rd_targets (q,pt) . \ (A,t1,t2) \ atcs (target q pt,q') . (\ io_atc . ioP @ p_io pt @ io_atc) ` (remove_proper_prefixes (acyclic_language_intersection (from_FSM M (target q pt)) A))))" by force then consider (a) "io \ L_acyclic P" | (b) "io \ (\ ioP \ remove_proper_prefixes (L_acyclic P) . \ pt \ tps q . Set.insert (ioP @ p_io pt) (\ q' \ rd_targets (q,pt) . \ (A,t1,t2) \ atcs (target q pt,q') . (\ io_atc . ioP @ p_io pt @ io_atc) ` (remove_proper_prefixes (acyclic_language_intersection (from_FSM M (target q pt)) A))))" by blast then show "io \ (\ (q,P) \ prs . L_acyclic P \ (\ ioP \ remove_proper_prefixes (L_acyclic P) . \ pt \ tps q . ((\ io' . ioP @ io') ` (set (prefixes (p_io pt)))) \ (\ q' \ rd_targets (q,pt) . \ (A,t1,t2) \ atcs (target q pt,q') . (\ io_atc . ioP @ p_io pt @ io_atc) ` (acyclic_language_intersection (from_FSM M (target q pt)) A))))" proof cases case a then show ?thesis using \(q,P) \ prs\ by blast next case b then obtain ioP pt where "ioP \ remove_proper_prefixes (L_acyclic P)" and "pt \ tps q" and "io \ Set.insert (ioP @ p_io pt) (\ q' \ rd_targets (q,pt) . \ (A,t1,t2) \ atcs (target q pt,q') . (\ io_atc . ioP @ p_io pt @ io_atc) ` (remove_proper_prefixes (acyclic_language_intersection (from_FSM M (target q pt)) A)))" by blast then consider (b1) "io = (ioP @ p_io pt)" | (b2) "io \ (\ q' \ rd_targets (q,pt) . \ (A,t1,t2) \ atcs (target q pt,q') . (\ io_atc . ioP @ p_io pt @ io_atc) ` (remove_proper_prefixes (acyclic_language_intersection (from_FSM M (target q pt)) A)))" by blast then show ?thesis proof cases case b1 then have "io \ ((\ io' . ioP @ io') ` (set (prefixes (p_io pt))))" unfolding prefixes_set by blast then show ?thesis using \(q,P) \ prs\ \ioP \ remove_proper_prefixes (L_acyclic P)\ \pt \ tps q\ by blast next case b2 then obtain q' A t1 t2 where "q' \ rd_targets (q,pt)" and "(A,t1,t2) \ atcs (target q pt,q')" and "io \ (\ io_atc . ioP @ p_io pt @ io_atc) ` (remove_proper_prefixes (acyclic_language_intersection (from_FSM M (target q pt)) A))" by blast then have "io \ (\ io_atc . ioP @ p_io pt @ io_atc) ` (acyclic_language_intersection (from_FSM M (target q pt)) A)" unfolding remove_proper_prefixes_def by blast then show ?thesis using \(q,P) \ prs\ \ioP \ remove_proper_prefixes (L_acyclic P)\ \pt \ tps q\ \q' \ rd_targets (q,pt)\ \(A,t1,t2) \ atcs (target q pt,q')\ by force qed qed qed have tmax_max: "\ io . io \ test_suite_to_io M T \ io \ tmax \ \ io'' . io'' \ [] \ io@io'' \ (test_suite_to_io M T)" proof - fix io assume "io \ test_suite_to_io M T" and "io \ tmax" then have "io \ (\ (q,P) \ prs . L_acyclic P \ (\ ioP \ remove_proper_prefixes (L_acyclic P) . \ pt \ tps q . ((\ io' . ioP @ io') ` (set (prefixes (p_io pt)))) \ (\ q' \ rd_targets (q,pt) . \ (A,t1,t2) \ atcs (target q pt,q') . (\ io_atc . ioP @ p_io pt @ io_atc) ` (acyclic_language_intersection (from_FSM M (target q pt)) A))))" unfolding t_def t1_def by blast then obtain q P where "(q,P) \ prs" and "io \ (L_acyclic P \ (\ ioP \ remove_proper_prefixes (L_acyclic P) . \ pt \ tps q . ((\ io' . ioP @ io') ` (set (prefixes (p_io pt)))) \ (\ q' \ rd_targets (q,pt) . \ (A,t1,t2) \ atcs (target q pt,q') . (\ io_atc . ioP @ p_io pt @ io_atc) ` (acyclic_language_intersection (from_FSM M (target q pt)) A))))" by force then consider (a) "io \ L_acyclic P" | (b) "io \ (\ ioP \ remove_proper_prefixes (L_acyclic P) . \ pt \ tps q . ((\ io' . ioP @ io') ` (set (prefixes (p_io pt)))) \ (\ q' \ rd_targets (q,pt) . \ (A,t1,t2) \ atcs (target q pt,q') . (\ io_atc . ioP @ p_io pt @ io_atc) ` (acyclic_language_intersection (from_FSM M (target q pt)) A)))" by blast then show "\ io'' . io'' \ [] \ io@io'' \ (test_suite_to_io M T)" proof cases case a then have "io \ tmax" using \(q,P) \ prs\ unfolding tmax_def by blast then show ?thesis using \io \ tmax\ by simp next case b then obtain ioP pt where "ioP \ remove_proper_prefixes (L_acyclic P)" and "pt \ tps q" and "io \ ((\ io' . ioP @ io') ` (set (prefixes (p_io pt)))) \ (\ q' \ rd_targets (q,pt) . \ (A,t1,t2) \ atcs (target q pt,q') . (\ io_atc . ioP @ p_io pt @ io_atc) ` (acyclic_language_intersection (from_FSM M (target q pt)) A))" by blast then consider (b1) "io \ (\ io' . ioP @ io') ` (set (prefixes (p_io pt)))" | (b2) "io \ (\ q' \ rd_targets (q,pt) . \ (A,t1,t2) \ atcs (target q pt,q') . (\ io_atc . ioP @ p_io pt @ io_atc) ` (acyclic_language_intersection (from_FSM M (target q pt)) A))" by blast then show ?thesis proof cases case b1 then obtain pt1 pt2 where "io = ioP @ pt1" and "p_io pt = pt1@pt2" by (metis (no_types, lifting) b1 image_iff prefixes_set_ob) have "ioP @ (p_io pt) \ tmax" using \io \ tmax\ \(q,P) \ prs\ \ioP \ remove_proper_prefixes (L_acyclic P)\ \pt \ tps q\ unfolding tmax_def by force then have "io \ ioP @ (p_io pt)" using \io \ tmax\ by blast then have "pt2 \ []" using \io = ioP @ pt1\ unfolding \p_io pt = pt1@pt2\ by auto have "ioP @ (p_io pt) \ test_suite_to_io M T" using \ioP @ (p_io pt) \ tmax\ tmax_sub by blast then show ?thesis unfolding \io = ioP @ pt1\ append.assoc \p_io pt = pt1@pt2\ using \pt2 \ []\ by blast next case b2 then obtain q' A t1 t2 where "q' \ rd_targets (q,pt)" and "(A,t1,t2) \ atcs (target q pt,q')" and "io \ (\ io_atc . ioP @ p_io pt @ io_atc) ` (acyclic_language_intersection (from_FSM M (target q pt)) A)" by blast then obtain ioA where "io = ioP @ p_io pt @ ioA" and "ioA \ acyclic_language_intersection (from_FSM M (target q pt)) A" by blast moreover have "ioA \ (remove_proper_prefixes (acyclic_language_intersection (from_FSM M (target q pt)) A))" proof assume "ioA \ remove_proper_prefixes (acyclic_language_intersection (FSM.from_FSM M (target q pt)) A)" then have "io \ tmax" using \(q,P) \ prs\ \ioP \ remove_proper_prefixes (L_acyclic P)\ \pt \ tps q\ \q' \ rd_targets (q,pt)\ \(A,t1,t2) \ atcs (target q pt,q')\ unfolding tmax_def \io = ioP @ p_io pt @ ioA\ by force then show False using \io \ tmax\ by simp qed ultimately obtain ioA2 where "ioA @ ioA2 \ acyclic_language_intersection (from_FSM M (target q pt)) A" and "ioA2 \ []" unfolding remove_proper_prefixes_def by blast then have "io@ioA2 \ test_suite_to_io M T" using \(q,P) \ prs\ \ioP \ remove_proper_prefixes (L_acyclic P)\ \pt \ tps q\ \q' \ rd_targets (q,pt)\ \(A,t1,t2) \ atcs (target q pt,q')\ \ioA @ ioA2 \ acyclic_language_intersection (from_FSM M (target q pt)) A\ unfolding t_def t1_def \io = ioP @ p_io pt @ ioA\ by force then show ?thesis using \ioA2 \ []\ by blast qed qed qed show ?thesis unfolding t2_def proof show "{io' \ test_suite_to_io M T. \io''. io'' \ [] \ io' @ io'' \ test_suite_to_io M T} \ remove_proper_prefixes tmax" proof fix io assume "io \ {io' \ test_suite_to_io M T. \io''. io'' \ [] \ io' @ io'' \ test_suite_to_io M T}" then have "io \ test_suite_to_io M T" and "\io''. io'' \ [] \ io @ io'' \ test_suite_to_io M T" by blast+ show "io \ remove_proper_prefixes tmax" using \\io''. io'' \ [] \ io @ io'' \ test_suite_to_io M T\ using tmax_sub tmax_max[OF \io \ test_suite_to_io M T\] unfolding remove_proper_prefixes_def by auto qed show "remove_proper_prefixes tmax \ {io' \ test_suite_to_io M T. \io''. io'' \ [] \ io' @ io'' \ test_suite_to_io M T}" proof fix io assume "io \ remove_proper_prefixes tmax" then have "io \ tmax" and "\io''. io'' \ [] \ io @ io'' \ remove_proper_prefixes tmax" unfolding remove_proper_prefixes_def by blast+ then have "io \ test_suite_to_io M T" using tmax_sub by blast moreover have "\io''. io'' \ [] \ io @ io'' \ test_suite_to_io M T" proof assume "\io''. io'' \ [] \ io @ io'' \ test_suite_to_io M T" then obtain io'' where "io'' \ []" and "io @ io'' \ test_suite_to_io M T" by blast then obtain io''' where "(io @ io'') @ io''' \ test_suite_to_io M T" and "(\zs. zs \ [] \ (io @ io'') @ io''' @ zs \ test_suite_to_io M T)" using finite_set_elem_maximal_extension_ex[OF \io @ io'' \ test_suite_to_io M T\ test_suite_to_io_finite[OF assms(1,2)]] by blast have "io @ (io'' @ io''') \ tmax" using tmax_max[OF \(io @ io'') @ io''' \ test_suite_to_io M T\] \(\zs. zs \ [] \ (io @ io'') @ io''' @ zs \ test_suite_to_io M T)\ by force moreover have "io''@io''' \ []" using \io'' \ []\ by auto ultimately show "False" using \\io''. io'' \ [] \ io @ io'' \ remove_proper_prefixes tmax\ by (metis (mono_tags, lifting) \io \ remove_proper_prefixes tmax\ mem_Collect_eq remove_proper_prefixes_def) qed ultimately show "io \ {io' \ test_suite_to_io M T. \io''. io'' \ [] \ io' @ io'' \ test_suite_to_io M T}" by blast qed qed qed lemma test_suite_to_io_pass_maximal : assumes "implies_completeness T M m" and "is_finite_test_suite T" shows "pass_io_set M' (test_suite_to_io M T) = pass_io_set_maximal M' {io' \ (test_suite_to_io M T) . \ (\ io'' . io'' \ [] \ io'@io'' \ (test_suite_to_io M T))}" proof - (* finiteness *) have p1: "finite (test_suite_to_io M T)" using test_suite_to_io_finite[OF assms] by assumption (* prefix closure *) obtain prs tps rd_targets atcs where "T = Test_Suite prs tps rd_targets atcs" by (meson test_suite.exhaust) then obtain repetition_sets where repetition_sets_def: "implies_completeness_for_repetition_sets (Test_Suite prs tps rd_targets atcs) M m repetition_sets" using assms(1) unfolding implies_completeness_def by blast then have "implies_completeness (Test_Suite prs tps rd_targets atcs) M m" unfolding implies_completeness_def by blast then have test_suite_language_prop: "test_suite_to_io M (Test_Suite prs tps rd_targets atcs) \ L M" using test_suite_to_io_language by blast have p2: "\io' io''. io' @ io'' \ test_suite_to_io M T \ io' \ test_suite_to_io M T" unfolding \T = Test_Suite prs tps rd_targets atcs\ proof - fix io' io'' assume "io' @ io'' \ test_suite_to_io M (Test_Suite prs tps rd_targets atcs)" have preamble_prop : "\ io''' q P . (q,P) \ prs \ io'@io''' \ L P \ io' \ test_suite_to_io M (Test_Suite prs tps rd_targets atcs)" proof - fix io''' q P assume "(q,P) \ prs" and "io'@io''' \ L P" have "io' \ (\ (q,P) \ prs . L P)" using \(q,P) \ prs\ language_prefix[OF \io'@io''' \ L P\] by auto then show "io' \ test_suite_to_io M (Test_Suite prs tps rd_targets atcs)" unfolding test_suite_to_io.simps by blast qed have traversal_path_prop : "\ io''' p pt q P . io'@io''' \ (\ io' . p_io p @ io') ` (set (prefixes (p_io pt))) \ (q,P) \ prs \ path P (initial P) p \ target (initial P) p = q \ pt \ tps q \ io' \ test_suite_to_io M (Test_Suite prs tps rd_targets atcs)" proof - fix io''' p pt q P assume "io'@io''' \ (\ io' . p_io p @ io') ` (set (prefixes (p_io pt)))" and "(q,P) \ prs" and "path P (initial P) p" and "target (initial P) p = q" and "pt \ tps q" obtain ioP1 where "io'@io''' = p_io p @ ioP1" and "ioP1 \ (set (prefixes (p_io pt)))" using \io'@io''' \ (\ io' . p_io p @ io') ` (set (prefixes (p_io pt)))\ by auto then obtain ioP2 where "ioP1 @ ioP2 = p_io pt" unfolding prefixes_set by blast show "io' \ test_suite_to_io M (Test_Suite prs tps rd_targets atcs)" proof (cases "length io' \ length (p_io p)") case True then have "io' = (take (length io') (p_io p))" using \io'@io''' = p_io p @ ioP1\ by (metis (no_types, lifting) order_refl take_all take_le) then have "p_io p = io'@(drop (length io') (p_io p))" by (metis (no_types, lifting) append_take_drop_id) then have "io'@(drop (length io') (p_io p)) \ L P" using language_state_containment[OF \path P (initial P) p\] by blast then show ?thesis using preamble_prop[OF \(q,P) \ prs\] by blast next case False then have "io' = p_io p @ (take (length io' - length (p_io p)) ioP1)" using \io'@io''' = p_io p @ ioP1\ by (metis (no_types, lifting) le_cases take_all take_append take_le) moreover have "(take (length io' - length (p_io p)) ioP1) \ (set (prefixes (p_io pt)))" proof - have "ioP1 = (take (length io' - length (p_io p)) ioP1) @ (drop (length io' - length (p_io p)) ioP1)" by auto then have "(take (length io' - length (p_io p)) ioP1) @ ((drop (length io' - length (p_io p)) ioP1) @ ioP2) = p_io pt" using \ioP1 @ ioP2 = p_io pt\ by (metis (mono_tags, lifting) append_assoc) then show ?thesis unfolding prefixes_set by blast qed ultimately have "io' \ (\ io' . p_io p @ io') ` (set (prefixes (p_io pt)))" by blast then have "io' \ (\{(\ io' . p_io p @ io') ` (set (prefixes (p_io pt))) | p pt . \ q P . (q,P) \ prs \ path P (initial P) p \ target (initial P) p = q \ pt \ tps q})" using \(q,P) \ prs\ \path P (initial P) p\ \target (initial P) p = q\ \pt \ tps q\ by blast then show ?thesis unfolding test_suite_to_io.simps by blast qed qed from \io' @ io'' \ test_suite_to_io M (Test_Suite prs tps rd_targets atcs)\ consider (a) "io' @ io'' \ (\ (q,P) \ prs . L P)" | (b) "io' @ io'' \ (\{(\ io' . p_io p @ io') ` (set (prefixes (p_io pt))) | p pt . \ q P . (q,P) \ prs \ path P (initial P) p \ target (initial P) p = q \ pt \ tps q})" | (c) "io' @ io'' \ (\{(\ io_atc . p_io p @ p_io pt @ io_atc) ` (atc_to_io_set (from_FSM M (target q pt)) A) | p pt q A . \ P q' t1 t2 . (q,P) \ prs \ path P (initial P) p \ target (initial P) p = q \ pt \ tps q \ q' \ rd_targets (q,pt) \ (A,t1,t2) \ atcs (target q pt,q') })" unfolding test_suite_to_io.simps by blast then show "io' \ test_suite_to_io M (Test_Suite prs tps rd_targets atcs)" proof cases case a then show ?thesis using preamble_prop by blast next case b then show ?thesis using traversal_path_prop by blast next case c then obtain p pt q A P q' t1 t2 where "io' @ io'' \ (\ io_atc . p_io p @ p_io pt @ io_atc) ` (atc_to_io_set (from_FSM M (target q pt)) A)" and "(q,P) \ prs" and "path P (initial P) p" and "target (initial P) p = q" and "pt \ tps q" and "q' \ rd_targets (q,pt)" and "(A,t1,t2) \ atcs (target q pt,q')" by blast obtain ioA where "io' @ io'' = p_io p @ p_io pt @ ioA" and "ioA \ (atc_to_io_set (from_FSM M (target q pt)) A)" using \io' @ io'' \ (\ io_atc . p_io p @ p_io pt @ io_atc) ` (atc_to_io_set (from_FSM M (target q pt)) A)\ by blast show ?thesis proof (cases "length io' \ length (p_io p @ p_io pt)") case True then have "io' @ (drop (length io') (p_io p @ p_io pt)) = p_io p @ p_io pt" using \io' @ io'' = p_io p @ p_io pt @ ioA\ by (simp add: append_eq_conv_conj) moreover have "p_io p @ p_io pt \ (\ io' . p_io p @ io') ` (set (prefixes (p_io pt)))" unfolding prefixes_set by blast ultimately have "io'@(drop (length io') (p_io p @ p_io pt)) \ (\ io' . p_io p @ io') ` (set (prefixes (p_io pt)))" by simp then show ?thesis using traversal_path_prop[OF _ \(q,P) \ prs\ \path P (initial P) p\ \target (initial P) p = q\ \pt \ tps q\] by blast next case False then have "io' = (p_io p @ p_io pt) @ (take (length io' - length (p_io p @ p_io pt)) ioA)" proof - have "io' = take (length io') (io' @ io'')" by auto then show ?thesis using False \io' @ io'' = p_io p @ p_io pt @ ioA\ by fastforce qed moreover have "(take (length io' - length (p_io p @ p_io pt)) ioA) \ (atc_to_io_set (from_FSM M (target q pt)) A)" proof - have "(take (length io' - length (p_io p @ p_io pt)) ioA)@(drop (length io' - length (p_io p @ p_io pt)) ioA) \ L (from_FSM M (target q pt)) \ L A" using \ioA \ (atc_to_io_set (from_FSM M (target q pt)) A)\ by auto then have *: "(take (length io' - length (p_io p @ p_io pt)) ioA)@(drop (length io' - length (p_io p @ p_io pt)) ioA) \ L (from_FSM M (target q pt))" and **: "(take (length io' - length (p_io p @ p_io pt)) ioA)@(drop (length io' - length (p_io p @ p_io pt)) ioA) \ L A" by blast+ show ?thesis using language_prefix[OF *] language_prefix[OF **] unfolding atc_to_io_set.simps by blast qed ultimately have "io' \ (\ io_atc . p_io p @ p_io pt @ io_atc) ` (atc_to_io_set (from_FSM M (target q pt)) A)" by force then have "io' \ (\{(\ io_atc . p_io p @ p_io pt @ io_atc) ` (atc_to_io_set (from_FSM M (target q pt)) A) | p pt q A . \ P q' t1 t2 . (q,P) \ prs \ path P (initial P) p \ target (initial P) p = q \ pt \ tps q \ q' \ rd_targets (q,pt) \ (A,t1,t2) \ atcs (target q pt,q') })" using \(q,P) \ prs\ \path P (initial P) p\ \target (initial P) p = q\ \pt \ tps q\ \q' \ rd_targets (q,pt)\ \(A,t1,t2) \ atcs (target q pt,q')\ by blast then show ?thesis unfolding test_suite_to_io.simps by blast qed qed qed then show ?thesis using pass_io_set_maximal_from_pass_io_set[OF p1] by blast qed lemma passes_test_suite_as_maximal_sequences_completeness : assumes "implies_completeness T M m" and "is_finite_test_suite T" and "observable M" and "observable M'" and "inputs M' = inputs M" and "inputs M \ {}" and "completely_specified M" and "completely_specified M'" and "size M' \ m" shows "(L M' \ L M) \ pass_io_set_maximal M' (test_suite_to_io_maximal M T)" unfolding passes_test_suite_completeness[OF assms(1,3-9)] unfolding test_suite_to_io_pass[OF assms(1,3-8),symmetric] unfolding test_suite_to_io_pass_maximal[OF assms(1,2)] unfolding test_suite_to_io_maximal_code[OF assms(1,2,3)] by simp lemma test_suite_to_io_maximal_finite : assumes "implies_completeness T M m" and "is_finite_test_suite T" and "observable M" shows "finite (test_suite_to_io_maximal M T)" using test_suite_to_io_finite[OF assms(1,2)] unfolding test_suite_to_io_maximal_code[OF assms, symmetric] by simp end \ No newline at end of file diff --git a/thys/Lambda_Free_RPOs/Extension_Orders.thy b/thys/Lambda_Free_RPOs/Extension_Orders.thy --- a/thys/Lambda_Free_RPOs/Extension_Orders.thy +++ b/thys/Lambda_Free_RPOs/Extension_Orders.thy @@ -1,2055 +1,2054 @@ (* Title: Extension Orders Author: Heiko Becker , 2016 Author: Jasmin Blanchette , 2016 Author: Dmitriy Traytel , 2014 Maintainer: Jasmin Blanchette *) section \Extension Orders\ theory Extension_Orders imports Lambda_Free_Util Infinite_Chain "HOL-Cardinals.Wellorder_Extension" begin text \ This theory defines locales for categorizing extension orders used for orders on \\\-free higher-order terms and defines variants of the lexicographic and multiset orders. \ subsection \Locales\ locale ext = fixes ext :: "('a \ 'a \ bool) \ 'a list \ 'a list \ bool" assumes mono_strong: "(\y \ set ys. \x \ set xs. gt y x \ gt' y x) \ ext gt ys xs \ ext gt' ys xs" and map: "finite A \ ys \ lists A \ xs \ lists A \ (\x \ A. \ gt (f x) (f x)) \ (\z \ A. \y \ A. \x \ A. gt (f z) (f y) \ gt (f y) (f x) \ gt (f z) (f x)) \ (\y \ A. \x \ A. gt y x \ gt (f y) (f x)) \ ext gt ys xs \ ext gt (map f ys) (map f xs)" begin lemma mono[mono]: "gt \ gt' \ ext gt \ ext gt'" using mono_strong by blast end locale ext_irrefl = ext + assumes irrefl: "(\x \ set xs. \ gt x x) \ \ ext gt xs xs" locale ext_trans = ext + assumes trans: "zs \ lists A \ ys \ lists A \ xs \ lists A \ (\z \ A. \y \ A. \x \ A. gt z y \ gt y x \ gt z x) \ ext gt zs ys \ ext gt ys xs \ ext gt zs xs" locale ext_irrefl_before_trans = ext_irrefl + assumes trans_from_irrefl: "finite A \ zs \ lists A \ ys \ lists A \ xs \ lists A \ (\x \ A. \ gt x x) \ (\z \ A. \y \ A. \x \ A. gt z y \ gt y x \ gt z x) \ ext gt zs ys \ ext gt ys xs \ ext gt zs xs" locale ext_trans_before_irrefl = ext_trans + assumes irrefl_from_trans: "(\z \ set xs. \y \ set xs. \x \ set xs. gt z y \ gt y x \ gt z x) \ (\x \ set xs. \ gt x x) \ \ ext gt xs xs" locale ext_irrefl_trans_strong = ext_irrefl + assumes trans_strong: "(\z \ set zs. \y \ set ys. \x \ set xs. gt z y \ gt y x \ gt z x) \ ext gt zs ys \ ext gt ys xs \ ext gt zs xs" sublocale ext_irrefl_trans_strong < ext_irrefl_before_trans by standard (erule irrefl, metis in_listsD trans_strong) sublocale ext_irrefl_trans_strong < ext_trans by standard (metis in_listsD trans_strong) sublocale ext_irrefl_trans_strong < ext_trans_before_irrefl by standard (rule irrefl) locale ext_snoc = ext + assumes snoc: "ext gt (xs @ [x]) xs" locale ext_compat_cons = ext + assumes compat_cons: "ext gt ys xs \ ext gt (x # ys) (x # xs)" begin lemma compat_append_left: "ext gt ys xs \ ext gt (zs @ ys) (zs @ xs)" by (induct zs) (auto intro: compat_cons) end locale ext_compat_snoc = ext + assumes compat_snoc: "ext gt ys xs \ ext gt (ys @ [x]) (xs @ [x])" begin lemma compat_append_right: "ext gt ys xs \ ext gt (ys @ zs) (xs @ zs)" by (induct zs arbitrary: xs ys rule: rev_induct) (auto intro: compat_snoc simp del: append_assoc simp: append_assoc[symmetric]) end locale ext_compat_list = ext + assumes compat_list: "y \ x \ gt y x \ ext gt (xs @ y # xs') (xs @ x # xs')" locale ext_singleton = ext + assumes singleton: "y \ x \ ext gt [y] [x] \ gt y x" locale ext_compat_list_strong = ext_compat_cons + ext_compat_snoc + ext_singleton begin lemma compat_list: "y \ x \ gt y x \ ext gt (xs @ y # xs') (xs @ x # xs')" using compat_append_left[of gt "y # xs'" "x # xs'" xs] compat_append_right[of gt, of "[y]" "[x]" xs'] singleton[of y x gt] by fastforce end sublocale ext_compat_list_strong < ext_compat_list by standard (fact compat_list) locale ext_total = ext + assumes total: "(\y \ A. \x \ A. gt y x \ gt x y \ y = x) \ ys \ lists A \ xs \ lists A \ ext gt ys xs \ ext gt xs ys \ ys = xs" locale ext_wf = ext + assumes wf: "wfP (\x y. gt y x) \ wfP (\xs ys. ext gt ys xs)" locale ext_hd_or_tl = ext + assumes hd_or_tl: "(\z y x. gt z y \ gt y x \ gt z x) \ (\y x. gt y x \ gt x y \ y = x) \ length ys = length xs \ ext gt (y # ys) (x # xs) \ gt y x \ ext gt ys xs" locale ext_wf_bounded = ext_irrefl_before_trans + ext_hd_or_tl begin context fixes gt :: "'a \ 'a \ bool" assumes gt_irrefl: "\z. \ gt z z" and gt_trans: "\z y x. gt z y \ gt y x \ gt z x" and gt_total: "\y x. gt y x \ gt x y \ y = x" and gt_wf: "wfP (\x y. gt y x)" begin lemma irrefl_gt: "\ ext gt xs xs" using irrefl gt_irrefl by simp lemma trans_gt: "ext gt zs ys \ ext gt ys xs \ ext gt zs xs" by (rule trans_from_irrefl[of "set zs \ set ys \ set xs" zs ys xs gt]) (auto intro: gt_trans simp: gt_irrefl) lemma hd_or_tl_gt: "length ys = length xs \ ext gt (y # ys) (x # xs) \ gt y x \ ext gt ys xs" by (rule hd_or_tl) (auto intro: gt_trans simp: gt_total) lemma wf_same_length_if_total: "wfP (\xs ys. length ys = n \ length xs = n \ ext gt ys xs)" proof (induct n) case 0 thus ?case unfolding wfP_def wf_def using irrefl by auto next case (Suc n) note ih = this(1) define gt_hd where "\ys xs. gt_hd ys xs \ gt (hd ys) (hd xs)" define gt_tl where "\ys xs. gt_tl ys xs \ ext gt (tl ys) (tl xs)" have hd_tl: "gt_hd ys xs \ gt_tl ys xs" if len_ys: "length ys = Suc n" and len_xs: "length xs = Suc n" and ys_gt_xs: "ext gt ys xs" for n ys xs using len_ys len_xs ys_gt_xs unfolding gt_hd_def gt_tl_def by (cases xs; cases ys) (auto simp: hd_or_tl_gt) show ?case unfolding wfP_iff_no_inf_chain proof (intro notI) let ?gtsn = "\ys xs. length ys = n \ length xs = n \ ext gt ys xs" let ?gtsSn = "\ys xs. length ys = Suc n \ length xs = Suc n \ ext gt ys xs" let ?gttlSn = "\ys xs. length ys = Suc n \ length xs = Suc n \ gt_tl ys xs" assume "\f. inf_chain ?gtsSn f" then obtain xs where xs_bad: "bad ?gtsSn xs" unfolding inf_chain_def bad_def by blast let ?ff = "worst_chain ?gtsSn gt_hd" have wf_hd: "wf {(xs, ys). gt_hd ys xs}" unfolding gt_hd_def by (rule wfP_app[OF gt_wf, of hd, unfolded wfP_def]) have "inf_chain ?gtsSn ?ff" by (rule worst_chain_bad[OF wf_hd xs_bad]) moreover have "\ gt_hd (?ff i) (?ff (Suc i))" for i by (rule worst_chain_not_gt[OF wf_hd xs_bad]) (blast intro: trans_gt) ultimately have tl_bad: "inf_chain ?gttlSn ?ff" unfolding inf_chain_def using hd_tl by blast have "\ inf_chain ?gtsn (tl \ ?ff)" using wfP_iff_no_inf_chain[THEN iffD1, OF ih] by blast hence tl_good: "\ inf_chain ?gttlSn ?ff" unfolding inf_chain_def gt_tl_def by force show False using tl_bad tl_good by sat qed qed lemma wf_bounded_if_total: "wfP (\xs ys. length ys \ n \ length xs \ n \ ext gt ys xs)" unfolding wfP_iff_no_inf_chain proof (intro notI, induct n rule: less_induct) case (less n) note ih = this(1) and ex_bad = this(2) let ?gtsle = "\ys xs. length ys \ n \ length xs \ n \ ext gt ys xs" obtain xs where xs_bad: "bad ?gtsle xs" using ex_bad unfolding inf_chain_def bad_def by blast let ?ff = "worst_chain ?gtsle (\ys xs. length ys > length xs)" note wf_len = wf_app[OF wellorder_class.wf, of length, simplified] have ff_bad: "inf_chain ?gtsle ?ff" by (rule worst_chain_bad[OF wf_len xs_bad]) have ffi_bad: "\i. bad ?gtsle (?ff i)" by (rule inf_chain_bad[OF ff_bad]) have len_le_n: "\i. length (?ff i) \ n" using worst_chain_pred[OF wf_len xs_bad] by simp have len_le_Suc: "\i. length (?ff i) \ length (?ff (Suc i))" using worst_chain_not_gt[OF wf_len xs_bad] not_le_imp_less by (blast intro: trans_gt) show False proof (cases "\k. length (?ff k) = n") case False hence len_lt_n: "\i. length (?ff i) < n" using len_le_n by (blast intro: le_neq_implies_less) hence nm1_le: "n - 1 < n" by fastforce let ?gtslt = "\ys xs. length ys \ n - 1 \ length xs \ n - 1 \ ext gt ys xs" have "inf_chain ?gtslt ?ff" using ff_bad len_lt_n unfolding inf_chain_def by (metis (no_types, lifting) Suc_diff_1 le_antisym nat_neq_iff not_less0 not_less_eq_eq) thus False using ih[OF nm1_le] by blast next case True then obtain k where len_eq_n: "length (?ff k) = n" by blast let ?gtssl = "\ys xs. length ys = n \ length xs = n \ ext gt ys xs" have len_eq_n: "length (?ff (i + k)) = n" for i by (induct i) (simp add: len_eq_n, metis (lifting) len_le_n len_le_Suc add_Suc dual_order.antisym) have "inf_chain ?gtsle (\i. ?ff (i + k))" by (rule inf_chain_offset[OF ff_bad]) hence "inf_chain ?gtssl (\i. ?ff (i + k))" unfolding inf_chain_def using len_eq_n by presburger hence "\ wfP (\xs ys. ?gtssl ys xs)" using wfP_iff_no_inf_chain by blast thus False using wf_same_length_if_total[of n] by sat qed qed end context fixes gt :: "'a \ 'a \ bool" assumes gt_irrefl: "\z. \ gt z z" and gt_wf: "wfP (\x y. gt y x)" begin lemma wf_bounded: "wfP (\xs ys. length ys \ n \ length xs \ n \ ext gt ys xs)" proof - obtain Ge' where gt_sub_Ge': "{(x, y). gt y x} \ Ge'" and Ge'_wo: "Well_order Ge'" and Ge'_fld: "Field Ge' = UNIV" using total_well_order_extension[OF gt_wf[unfolded wfP_def]] by blast define gt' where "\y x. gt' y x \ y \ x \ (x, y) \ Ge'" have gt_imp_gt': "gt \ gt'" by (auto simp: gt'_def gt_irrefl intro: gt_sub_Ge'[THEN subsetD]) have gt'_irrefl: "\z. \ gt' z z" unfolding gt'_def by simp have gt'_trans: "\z y x. gt' z y \ gt' y x \ gt' z x" using Ge'_wo unfolding gt'_def well_order_on_def linear_order_on_def partial_order_on_def preorder_on_def trans_def antisym_def by blast have "wf {(x, y). (x, y) \ Ge' \ x \ y}" by (rule Ge'_wo[unfolded well_order_on_def set_diff_eq case_prod_eta[symmetric, of "\xy. xy \ Ge' \ xy \ Id"] pair_in_Id_conv, THEN conjunct2]) moreover have "\y x. (x, y) \ Ge' \ x \ y \ y \ x \ (x, y) \ Ge'" by auto ultimately have gt'_wf: "wfP (\x y. gt' y x)" unfolding wfP_def gt'_def by simp have gt'_total: "\x y. gt' y x \ gt' x y \ y = x" using Ge'_wo unfolding gt'_def well_order_on_def linear_order_on_def total_on_def Ge'_fld by blast have "wfP (\xs ys. length ys \ n \ length xs \ n \ ext gt' ys xs)" using wf_bounded_if_total gt'_total gt'_irrefl gt'_trans gt'_wf by blast thus ?thesis by (rule wfP_subset) (auto intro: mono[OF gt_imp_gt', THEN predicate2D]) qed end end subsection \Lexicographic Extension\ inductive lexext :: "('a \ 'a \ bool) \ 'a list \ 'a list \ bool" for gt where lexext_Nil: "lexext gt (y # ys) []" | lexext_Cons: "gt y x \ lexext gt (y # ys) (x # xs)" | lexext_Cons_eq: "lexext gt ys xs \ lexext gt (x # ys) (x # xs)" lemma lexext_simps[simp]: "lexext gt ys [] \ ys \ []" "\ lexext gt [] xs" "lexext gt (y # ys) (x # xs) \ gt y x \ x = y \ lexext gt ys xs" proof show "lexext gt ys [] \ (ys \ [])" by (metis lexext.cases list.distinct(1)) next show "ys \ [] \ lexext gt ys []" by (metis lexext_Nil list.exhaust) next show "\ lexext gt [] xs" using lexext.cases by auto next show "lexext gt (y # ys) (x # xs) = (gt y x \ x = y \ lexext gt ys xs)" proof - have fwdd: "lexext gt (y # ys) (x # xs) \ gt y x \ x = y \ lexext gt ys xs" proof assume "lexext gt (y # ys) (x # xs)" thus "gt y x \ x = y \ lexext gt ys xs" using lexext.cases by blast qed have backd: "gt y x \ x = y \ lexext gt ys xs \ lexext gt (y # ys) (x # xs)" by (simp add: lexext_Cons lexext_Cons_eq) show "lexext gt (y # ys) (x # xs) = (gt y x \ x = y \ lexext gt ys xs)" using fwdd backd by blast qed qed lemma lexext_mono_strong: assumes "\y \ set ys. \x \ set xs. gt y x \ gt' y x" and "lexext gt ys xs" shows "lexext gt' ys xs" using assms by (induct ys xs rule: list_induct2') auto lemma lexext_map_strong: "(\y \ set ys. \x \ set xs. gt y x \ gt (f y) (f x)) \ lexext gt ys xs \ lexext gt (map f ys) (map f xs)" by (induct ys xs rule: list_induct2') auto lemma lexext_irrefl: assumes "\x \ set xs. \ gt x x" shows "\ lexext gt xs xs" using assms by (induct xs) auto lemma lexext_trans_strong: assumes "\z \ set zs. \y \ set ys. \x \ set xs. gt z y \ gt y x \ gt z x" and "lexext gt zs ys" and "lexext gt ys xs" shows "lexext gt zs xs" using assms proof (induct zs arbitrary: ys xs) case (Cons z zs) note zs_trans = this(1) show ?case using Cons(2-4) proof (induct ys arbitrary: xs rule: list.induct) case (Cons y ys) note ys_trans = this(1) and gt_trans = this(2) and zzs_gt_yys = this(3) and yys_gt_xs = this(4) show ?case proof (cases xs) case xs: (Cons x xs) thus ?thesis proof (unfold xs) note yys_gt_xxs = yys_gt_xs[unfolded xs] note gt_trans = gt_trans[unfolded xs] let ?case = "lexext gt (z # zs) (x # xs)" { assume "gt z y" and "gt y x" hence ?case using gt_trans by simp } moreover { assume "gt z y" and "x = y" hence ?case by simp } moreover { assume "y = z" and "gt y x" hence ?case by simp } moreover { assume y_eq_z: "y = z" and zs_gt_ys: "lexext gt zs ys" and x_eq_y: "x = y" and ys_gt_xs: "lexext gt ys xs" have "lexext gt zs xs" by (rule zs_trans[OF _ zs_gt_ys ys_gt_xs]) (meson gt_trans[simplified]) hence ?case by (simp add: x_eq_y y_eq_z) } ultimately show ?case using zzs_gt_yys yys_gt_xxs by force qed qed auto qed auto qed auto lemma lexext_snoc: "lexext gt (xs @ [x]) xs" by (induct xs) auto lemmas lexext_compat_cons = lexext_Cons_eq lemma lexext_compat_snoc_if_same_length: assumes "length ys = length xs" and "lexext gt ys xs" shows "lexext gt (ys @ [x]) (xs @ [x])" using assms(2,1) by (induct rule: lexext.induct) auto lemma lexext_compat_list: "gt y x \ lexext gt (xs @ y # xs') (xs @ x # xs')" by (induct xs) auto lemma lexext_singleton: "lexext gt [y] [x] \ gt y x" by simp lemma lexext_total: "(\y \ B. \x \ A. gt y x \ gt x y \ y = x) \ ys \ lists B \ xs \ lists A \ lexext gt ys xs \ lexext gt xs ys \ ys = xs" by (induct ys xs rule: list_induct2') auto lemma lexext_hd_or_tl: "lexext gt (y # ys) (x # xs) \ gt y x \ lexext gt ys xs" by auto interpretation lexext: ext lexext by standard (fact lexext_mono_strong, rule lexext_map_strong, metis in_listsD) interpretation lexext: ext_irrefl_trans_strong lexext by standard (fact lexext_irrefl, fact lexext_trans_strong) interpretation lexext: ext_snoc lexext by standard (fact lexext_snoc) interpretation lexext: ext_compat_cons lexext by standard (fact lexext_compat_cons) interpretation lexext: ext_compat_list lexext by standard (rule lexext_compat_list) interpretation lexext: ext_singleton lexext by standard (rule lexext_singleton) interpretation lexext: ext_total lexext by standard (fact lexext_total) interpretation lexext: ext_hd_or_tl lexext by standard (rule lexext_hd_or_tl) interpretation lexext: ext_wf_bounded lexext by standard subsection \Reverse (Right-to-Left) Lexicographic Extension\ abbreviation lexext_rev :: "('a \ 'a \ bool) \ 'a list \ 'a list \ bool" where "lexext_rev gt ys xs \ lexext gt (rev ys) (rev xs)" lemma lexext_rev_simps[simp]: "lexext_rev gt ys [] \ ys \ []" "\ lexext_rev gt [] xs" "lexext_rev gt (ys @ [y]) (xs @ [x]) \ gt y x \ x = y \ lexext_rev gt ys xs" by simp+ lemma lexext_rev_cons_cons: assumes "length ys = length xs" shows "lexext_rev gt (y # ys) (x # xs) \ lexext_rev gt ys xs \ ys = xs \ gt y x" using assms proof (induct arbitrary: y x rule: rev_induct2) case Nil thus ?case by simp next case (snoc y' ys x' xs) show ?case using snoc(2) by auto qed lemma lexext_rev_mono_strong: assumes "\y \ set ys. \x \ set xs. gt y x \ gt' y x" and "lexext_rev gt ys xs" shows "lexext_rev gt' ys xs" using assms by (simp add: lexext_mono_strong) lemma lexext_rev_map_strong: "(\y \ set ys. \x \ set xs. gt y x \ gt (f y) (f x)) \ lexext_rev gt ys xs \ lexext_rev gt (map f ys) (map f xs)" by (simp add: lexext_map_strong rev_map) lemma lexext_rev_irrefl: assumes "\x \ set xs. \ gt x x" shows "\ lexext_rev gt xs xs" using assms by (simp add: lexext_irrefl) lemma lexext_rev_trans_strong: assumes "\z \ set zs. \y \ set ys. \x \ set xs. gt z y \ gt y x \ gt z x" and "lexext_rev gt zs ys" and "lexext_rev gt ys xs" shows "lexext_rev gt zs xs" using assms(1) lexext_trans_strong[OF _ assms(2,3), unfolded set_rev] by sat lemma lexext_rev_compat_cons_if_same_length: assumes "length ys = length xs" and "lexext_rev gt ys xs" shows "lexext_rev gt (x # ys) (x # xs)" using assms by (simp add: lexext_compat_snoc_if_same_length) lemma lexext_rev_compat_snoc: "lexext_rev gt ys xs \ lexext_rev gt (ys @ [x]) (xs @ [x])" by (simp add: lexext_compat_cons) lemma lexext_rev_compat_list: "gt y x \ lexext_rev gt (xs @ y # xs') (xs @ x # xs')" by (induct xs' rule: rev_induct) auto lemma lexext_rev_singleton: "lexext_rev gt [y] [x] \ gt y x" by simp lemma lexext_rev_total: "(\y \ B. \x \ A. gt y x \ gt x y \ y = x) \ ys \ lists B \ xs \ lists A \ lexext_rev gt ys xs \ lexext_rev gt xs ys \ ys = xs" by (rule lexext_total[of _ _ _ "rev ys" "rev xs", simplified]) lemma lexext_rev_hd_or_tl: assumes "length ys = length xs" and "lexext_rev gt (y # ys) (x # xs)" shows "gt y x \ lexext_rev gt ys xs" using assms lexext_rev_cons_cons by fastforce interpretation lexext_rev: ext lexext_rev by standard (fact lexext_rev_mono_strong, rule lexext_rev_map_strong, metis in_listsD) interpretation lexext_rev: ext_irrefl_trans_strong lexext_rev by standard (fact lexext_rev_irrefl, fact lexext_rev_trans_strong) interpretation lexext_rev: ext_compat_snoc lexext_rev by standard (fact lexext_rev_compat_snoc) interpretation lexext_rev: ext_compat_list lexext_rev by standard (rule lexext_rev_compat_list) interpretation lexext_rev: ext_singleton lexext_rev by standard (rule lexext_rev_singleton) interpretation lexext_rev: ext_total lexext_rev by standard (fact lexext_rev_total) interpretation lexext_rev: ext_hd_or_tl lexext_rev by standard (rule lexext_rev_hd_or_tl) interpretation lexext_rev: ext_wf_bounded lexext_rev by standard subsection \Generic Length Extension\ definition lenext :: "('a list \ 'a list \ bool) \ 'a list \ 'a list \ bool" where "lenext gts ys xs \ length ys > length xs \ length ys = length xs \ gts ys xs" lemma lenext_mono_strong: "(gts ys xs \ gts' ys xs) \ lenext gts ys xs \ lenext gts' ys xs" and lenext_map_strong: "(length ys = length xs \ gts ys xs \ gts (map f ys) (map f xs)) \ lenext gts ys xs \ lenext gts (map f ys) (map f xs)" and lenext_irrefl: "\ gts xs xs \ \ lenext gts xs xs" and lenext_trans: "(gts zs ys \ gts ys xs \ gts zs xs) \ lenext gts zs ys \ lenext gts ys xs \ lenext gts zs xs" and lenext_snoc: "lenext gts (xs @ [x]) xs" and lenext_compat_cons: "(length ys = length xs \ gts ys xs \ gts (x # ys) (x # xs)) \ lenext gts ys xs \ lenext gts (x # ys) (x # xs)" and lenext_compat_snoc: "(length ys = length xs \ gts ys xs \ gts (ys @ [x]) (xs @ [x])) \ lenext gts ys xs \ lenext gts (ys @ [x]) (xs @ [x])" and lenext_compat_list: "gts (xs @ y # xs') (xs @ x # xs') \ lenext gts (xs @ y # xs') (xs @ x # xs')" and lenext_singleton: "lenext gts [y] [x] \ gts [y] [x]" and lenext_total: "(gts ys xs \ gts xs ys \ ys = xs) \ lenext gts ys xs \ lenext gts xs ys \ ys = xs" and lenext_hd_or_tl: "(length ys = length xs \ gts (y # ys) (x # xs) \ gt y x \ gts ys xs) \ lenext gts (y # ys) (x # xs) \ gt y x \ lenext gts ys xs" unfolding lenext_def by auto subsection \Length-Lexicographic Extension\ abbreviation len_lexext :: "('a \ 'a \ bool) \ 'a list \ 'a list \ bool" where "len_lexext gt \ lenext (lexext gt)" lemma len_lexext_mono_strong: "(\y \ set ys. \x \ set xs. gt y x \ gt' y x) \ len_lexext gt ys xs \ len_lexext gt' ys xs" by (rule lenext_mono_strong[OF lexext_mono_strong]) lemma len_lexext_map_strong: "(\y \ set ys. \x \ set xs. gt y x \ gt (f y) (f x)) \ len_lexext gt ys xs \ len_lexext gt (map f ys) (map f xs)" by (rule lenext_map_strong) (metis lexext_map_strong) lemma len_lexext_irrefl: "(\x \ set xs. \ gt x x) \ \ len_lexext gt xs xs" by (rule lenext_irrefl[OF lexext_irrefl]) lemma len_lexext_trans_strong: "(\z \ set zs. \y \ set ys. \x \ set xs. gt z y \ gt y x \ gt z x) \ len_lexext gt zs ys \ len_lexext gt ys xs \ len_lexext gt zs xs" by (rule lenext_trans[OF lexext_trans_strong]) lemma len_lexext_snoc: "len_lexext gt (xs @ [x]) xs" by (rule lenext_snoc) lemma len_lexext_compat_cons: "len_lexext gt ys xs \ len_lexext gt (x # ys) (x # xs)" by (intro lenext_compat_cons lexext_compat_cons) lemma len_lexext_compat_snoc: "len_lexext gt ys xs \ len_lexext gt (ys @ [x]) (xs @ [x])" by (intro lenext_compat_snoc lexext_compat_snoc_if_same_length) lemma len_lexext_compat_list: "gt y x \ len_lexext gt (xs @ y # xs') (xs @ x # xs')" by (intro lenext_compat_list lexext_compat_list) lemma len_lexext_singleton[simp]: "len_lexext gt [y] [x] \ gt y x" by (simp only: lenext_singleton lexext_singleton) lemma len_lexext_total: "(\y \ B. \x \ A. gt y x \ gt x y \ y = x) \ ys \ lists B \ xs \ lists A \ len_lexext gt ys xs \ len_lexext gt xs ys \ ys = xs" by (rule lenext_total[OF lexext_total]) lemma len_lexext_iff_lenlex: "len_lexext gt ys xs \ (xs, ys) \ lenlex {(x, y). gt y x}" proof - { assume "length xs = length ys" hence "lexext gt ys xs \ (xs, ys) \ lex {(x, y). gt y x}" by (induct xs ys rule: list_induct2) auto } thus ?thesis unfolding lenext_def lenlex_conv by auto qed lemma len_lexext_wf: "wfP (\x y. gt y x) \ wfP (\xs ys. len_lexext gt ys xs)" unfolding wfP_def len_lexext_iff_lenlex by (simp add: wf_lenlex) lemma len_lexext_hd_or_tl: "len_lexext gt (y # ys) (x # xs) \ gt y x \ len_lexext gt ys xs" using lenext_hd_or_tl lexext_hd_or_tl by metis interpretation len_lexext: ext len_lexext by standard (fact len_lexext_mono_strong, rule len_lexext_map_strong, metis in_listsD) interpretation len_lexext: ext_irrefl_trans_strong len_lexext by standard (fact len_lexext_irrefl, fact len_lexext_trans_strong) interpretation len_lexext: ext_snoc len_lexext by standard (fact len_lexext_snoc) interpretation len_lexext: ext_compat_cons len_lexext by standard (fact len_lexext_compat_cons) interpretation len_lexext: ext_compat_snoc len_lexext by standard (fact len_lexext_compat_snoc) interpretation len_lexext: ext_compat_list len_lexext by standard (rule len_lexext_compat_list) interpretation len_lexext: ext_singleton len_lexext by standard (rule len_lexext_singleton) interpretation len_lexext: ext_total len_lexext by standard (fact len_lexext_total) interpretation len_lexext: ext_wf len_lexext by standard (fact len_lexext_wf) interpretation len_lexext: ext_hd_or_tl len_lexext by standard (rule len_lexext_hd_or_tl) interpretation len_lexext: ext_wf_bounded len_lexext by standard subsection \Reverse (Right-to-Left) Length-Lexicographic Extension\ abbreviation len_lexext_rev :: "('a \ 'a \ bool) \ 'a list \ 'a list \ bool" where "len_lexext_rev gt \ lenext (lexext_rev gt)" lemma len_lexext_rev_mono_strong: "(\y \ set ys. \x \ set xs. gt y x \ gt' y x) \ len_lexext_rev gt ys xs \ len_lexext_rev gt' ys xs" by (rule lenext_mono_strong) (rule lexext_rev_mono_strong) lemma len_lexext_rev_map_strong: "(\y \ set ys. \x \ set xs. gt y x \ gt (f y) (f x)) \ len_lexext_rev gt ys xs \ len_lexext_rev gt (map f ys) (map f xs)" by (rule lenext_map_strong) (rule lexext_rev_map_strong) lemma len_lexext_rev_irrefl: "(\x \ set xs. \ gt x x) \ \ len_lexext_rev gt xs xs" by (rule lenext_irrefl) (rule lexext_rev_irrefl) lemma len_lexext_rev_trans_strong: "(\z \ set zs. \y \ set ys. \x \ set xs. gt z y \ gt y x \ gt z x) \ len_lexext_rev gt zs ys \ len_lexext_rev gt ys xs \ len_lexext_rev gt zs xs" by (rule lenext_trans) (rule lexext_rev_trans_strong) lemma len_lexext_rev_snoc: "len_lexext_rev gt (xs @ [x]) xs" by (rule lenext_snoc) lemma len_lexext_rev_compat_cons: "len_lexext_rev gt ys xs \ len_lexext_rev gt (x # ys) (x # xs)" by (intro lenext_compat_cons lexext_rev_compat_cons_if_same_length) lemma len_lexext_rev_compat_snoc: "len_lexext_rev gt ys xs \ len_lexext_rev gt (ys @ [x]) (xs @ [x])" by (intro lenext_compat_snoc lexext_rev_compat_snoc) lemma len_lexext_rev_compat_list: "gt y x \ len_lexext_rev gt (xs @ y # xs') (xs @ x # xs')" by (intro lenext_compat_list lexext_rev_compat_list) lemma len_lexext_rev_singleton[simp]: "len_lexext_rev gt [y] [x] \ gt y x" by (simp only: lenext_singleton lexext_rev_singleton) lemma len_lexext_rev_total: "(\y \ B. \x \ A. gt y x \ gt x y \ y = x) \ ys \ lists B \ xs \ lists A \ len_lexext_rev gt ys xs \ len_lexext_rev gt xs ys \ ys = xs" by (rule lenext_total[OF lexext_rev_total]) lemma len_lexext_rev_iff_len_lexext: "len_lexext_rev gt ys xs \ len_lexext gt (rev ys) (rev xs)" unfolding lenext_def by simp lemma len_lexext_rev_wf: "wfP (\x y. gt y x) \ wfP (\xs ys. len_lexext_rev gt ys xs)" unfolding len_lexext_rev_iff_len_lexext by (rule wfP_app[of "\xs ys. len_lexext gt ys xs" rev, simplified]) (rule len_lexext_wf) lemma len_lexext_rev_hd_or_tl: "len_lexext_rev gt (y # ys) (x # xs) \ gt y x \ len_lexext_rev gt ys xs" using lenext_hd_or_tl lexext_rev_hd_or_tl by metis interpretation len_lexext_rev: ext len_lexext_rev by standard (fact len_lexext_rev_mono_strong, rule len_lexext_rev_map_strong, metis in_listsD) interpretation len_lexext_rev: ext_irrefl_trans_strong len_lexext_rev by standard (fact len_lexext_rev_irrefl, fact len_lexext_rev_trans_strong) interpretation len_lexext_rev: ext_snoc len_lexext_rev by standard (fact len_lexext_rev_snoc) interpretation len_lexext_rev: ext_compat_cons len_lexext_rev by standard (fact len_lexext_rev_compat_cons) interpretation len_lexext_rev: ext_compat_snoc len_lexext_rev by standard (fact len_lexext_rev_compat_snoc) interpretation len_lexext_rev: ext_compat_list len_lexext_rev by standard (rule len_lexext_rev_compat_list) interpretation len_lexext_rev: ext_singleton len_lexext_rev by standard (rule len_lexext_rev_singleton) interpretation len_lexext_rev: ext_total len_lexext_rev by standard (fact len_lexext_rev_total) interpretation len_lexext_rev: ext_wf len_lexext_rev by standard (fact len_lexext_rev_wf) interpretation len_lexext_rev: ext_hd_or_tl len_lexext_rev by standard (rule len_lexext_rev_hd_or_tl) interpretation len_lexext_rev: ext_wf_bounded len_lexext_rev by standard subsection \Dershowitz--Manna Multiset Extension\ definition msetext_dersh where "msetext_dersh gt ys xs = (let N = mset ys; M = mset xs in (\Y X. Y \ {#} \ Y \# N \ M = (N - Y) + X \ (\x. x \# X \ (\y. y \# Y \ gt y x))))" text \ The following proof is based on that of @{thm[source] less_multiset\<^sub>D\<^sub>M_imp_mult}. \ lemma msetext_dersh_imp_mult_rel: assumes ys_a: "ys \ lists A" and xs_a: "xs \ lists A" and ys_gt_xs: "msetext_dersh gt ys xs" shows "(mset xs, mset ys) \ mult {(x, y). x \ A \ y \ A \ gt y x}" proof - obtain Y X where y_nemp: "Y \ {#}" and y_sub_ys: "Y \# mset ys" and xs_eq: "mset xs = mset ys - Y + X" and ex_y: "\x. x \# X \ (\y. y \# Y \ gt y x)" using ys_gt_xs[unfolded msetext_dersh_def Let_def] by blast have ex_y': "\x. x \# X \ (\y. y \# Y \ x \ A \ y \ A \ gt y x)" using ex_y y_sub_ys xs_eq ys_a xs_a by (metis in_listsD mset_subset_eqD set_mset_mset union_iff) hence "(mset ys - Y + X, mset ys - Y + Y) \ mult {(x, y). x \ A \ y \ A \ gt y x}" using y_nemp y_sub_ys by (intro one_step_implies_mult) (auto simp: Bex_def trans_def) thus ?thesis using xs_eq y_sub_ys by (simp add: subset_mset.diff_add) qed lemma msetext_dersh_imp_mult: "msetext_dersh gt ys xs \ (mset xs, mset ys) \ mult {(x, y). gt y x}" using msetext_dersh_imp_mult_rel[of _ UNIV] by auto lemma mult_imp_msetext_dersh_rel: assumes ys_a: "set_mset (mset ys) \ A" and xs_a: "set_mset (mset xs) \ A" and in_mult: "(mset xs, mset ys) \ mult {(x, y). x \ A \ y \ A \ gt y x}" and trans: "\z \ A. \y \ A. \x \ A. gt z y \ gt y x \ gt z x" shows "msetext_dersh gt ys xs" using in_mult ys_a xs_a unfolding mult_def msetext_dersh_def Let_def proof induct case (base Ys) then obtain y M0 X where "Ys = M0 + {#y#}" and "mset xs = M0 + X" and "\a. a \# X \ gt y a" unfolding mult1_def by auto thus ?case by (auto intro: exI[of _ "{#y#}"] exI[of _ X]) next case (step Ys Zs) note ys_zs_in_mult1 = this(2) and ih = this(3) and zs_a = this(4) and xs_a = this(5) have Ys_a: "set_mset Ys \ A" using ys_zs_in_mult1 zs_a unfolding mult1_def by auto obtain Y X where y_nemp: "Y \ {#}" and y_sub_ys: "Y \# Ys" and xs_eq: "mset xs = Ys - Y + X" and ex_y: "\x. x \# X \ (\y. y \# Y \ gt y x)" using ih[OF Ys_a xs_a] by blast obtain z M0 Ya where zs_eq: "Zs = M0 + {#z#}" and ys_eq: "Ys = M0 + Ya" and z_gt: "\y. y \# Ya \ y \ A \ z \ A \ gt z y" using ys_zs_in_mult1[unfolded mult1_def] by auto let ?Za = "Y - Ya + {#z#}" let ?Xa = "X + Ya + (Y - Ya) - Y" have xa_sub_x_ya: "set_mset ?Xa \ set_mset (X + Ya)" by (metis diff_subset_eq_self in_diffD subsetI subset_mset.diff_diff_right) have x_a: "set_mset X \ A" using xs_a xs_eq by auto have ya_a: "set_mset Ya \ A" by (simp add: subsetI z_gt) have ex_y': "\y. y \# Y - Ya + {#z#} \ gt y x" if x_in: "x \# X + Ya" for x proof (cases "x \# X") case True then obtain y where y_in: "y \# Y" and y_gt_x: "gt y x" using ex_y by blast show ?thesis proof (cases "y \# Ya") case False hence "y \# Y - Ya + {#z#}" using y_in by fastforce thus ?thesis using y_gt_x by blast next case True hence "y \ A" and "z \ A" and "gt z y" using z_gt by blast+ hence "gt z x" using trans y_gt_x x_a ya_a x_in by (meson subsetCE union_iff) thus ?thesis by auto qed next case False hence "x \# Ya" using x_in by auto hence "x \ A" and "z \ A" and "gt z x" using z_gt by blast+ thus ?thesis by auto qed show ?case proof (rule exI[of _ ?Za], rule exI[of _ ?Xa], intro conjI) show "Y - Ya + {#z#} \# Zs" using mset_subset_eq_mono_add subset_eq_diff_conv y_sub_ys ys_eq zs_eq by fastforce next show "mset xs = Zs - (Y - Ya + {#z#}) + (X + Ya + (Y - Ya) - Y)" unfolding xs_eq ys_eq zs_eq by (auto simp: multiset_eq_iff) next show "\x. x \# X + Ya + (Y - Ya) - Y \ (\y. y \# Y - Ya + {#z#} \ gt y x)" using ex_y' xa_sub_x_ya by blast qed auto qed lemma msetext_dersh_mono_strong: "(\y \ set ys. \x \ set xs. gt y x \ gt' y x) \ msetext_dersh gt ys xs \ msetext_dersh gt' ys xs" unfolding msetext_dersh_def Let_def by (metis mset_subset_eqD mset_subset_eq_add_right set_mset_mset) lemma msetext_dersh_map_strong: assumes compat_f: "\y \ set ys. \x \ set xs. gt y x \ gt (f y) (f x)" and ys_gt_xs: "msetext_dersh gt ys xs" shows "msetext_dersh gt (map f ys) (map f xs)" proof - obtain Y X where y_nemp: "Y \ {#}" and y_sub_ys: "Y \# mset ys" and xs_eq: "mset xs = mset ys - Y + X" and ex_y: "\x. x \# X \ (\y. y \# Y \ gt y x)" using ys_gt_xs[unfolded msetext_dersh_def Let_def mset_map] by blast have x_sub_xs: "X \# mset xs" using xs_eq by simp let ?fY = "image_mset f Y" let ?fX = "image_mset f X" show ?thesis unfolding msetext_dersh_def Let_def mset_map proof (intro exI conjI) show "image_mset f (mset xs) = image_mset f (mset ys) - ?fY + ?fX" using xs_eq[THEN arg_cong, of "image_mset f"] y_sub_ys by (metis image_mset_Diff image_mset_union) next obtain y where y: "\x. x \# X \ y x \# Y \ gt (y x) x" using ex_y by moura show "\fx. fx \# ?fX \ (\fy. fy \# ?fY \ gt fy fx)" proof (intro allI impI) fix fx assume "fx \# ?fX" then obtain x where fx: "fx = f x" and x_in: "x \# X" by auto hence y_in: "y x \# Y" and y_gt: "gt (y x) x" using y[rule_format, OF x_in] by blast+ hence "f (y x) \# ?fY \ gt (f (y x)) (f x)" using compat_f y_sub_ys x_sub_xs x_in by (metis image_eqI in_image_mset mset_subset_eqD set_mset_mset) thus "\fy. fy \# ?fY \ gt fy fx" unfolding fx by auto qed qed (auto simp: y_nemp y_sub_ys image_mset_subseteq_mono) qed lemma msetext_dersh_trans: assumes zs_a: "zs \ lists A" and ys_a: "ys \ lists A" and xs_a: "xs \ lists A" and trans: "\z \ A. \y \ A. \x \ A. gt z y \ gt y x \ gt z x" and zs_gt_ys: "msetext_dersh gt zs ys" and ys_gt_xs: "msetext_dersh gt ys xs" shows "msetext_dersh gt zs xs" proof (rule mult_imp_msetext_dersh_rel[OF _ _ _ trans]) show "set_mset (mset zs) \ A" using zs_a by auto next show "set_mset (mset xs) \ A" using xs_a by auto next let ?Gt = "{(x, y). x \ A \ y \ A \ gt y x}" have "(mset xs, mset ys) \ mult ?Gt" by (rule msetext_dersh_imp_mult_rel[OF ys_a xs_a ys_gt_xs]) moreover have "(mset ys, mset zs) \ mult ?Gt" by (rule msetext_dersh_imp_mult_rel[OF zs_a ys_a zs_gt_ys]) ultimately show "(mset xs, mset zs) \ mult ?Gt" unfolding mult_def by simp qed lemma msetext_dersh_irrefl_from_trans: assumes trans: "\z \ set xs. \y \ set xs. \x \ set xs. gt z y \ gt y x \ gt z x" and irrefl: "\x \ set xs. \ gt x x" shows "\ msetext_dersh gt xs xs" unfolding msetext_dersh_def Let_def proof clarify fix Y X assume y_nemp: "Y \ {#}" and y_sub_xs: "Y \# mset xs" and xs_eq: "mset xs = mset xs - Y + X" and ex_y: "\x. x \# X \ (\y. y \# Y \ gt y x)" have x_eq_y: "X = Y" using y_sub_xs xs_eq by (metis diff_union_cancelL subset_mset.diff_add) let ?Gt = "{(y, x). y \# Y \ x \# Y \ gt y x}" have "?Gt \ set_mset Y \ set_mset Y" by auto hence fin: "finite ?Gt" by (auto dest!: infinite_super) moreover have "irrefl ?Gt" unfolding irrefl_def using irrefl y_sub_xs by (fastforce dest!: set_mset_mono) moreover have "trans ?Gt" unfolding trans_def using trans y_sub_xs by (fastforce dest!: set_mset_mono) ultimately have acyc: "acyclic ?Gt" by (rule finite_irrefl_trans_imp_wf[THEN wf_acyclic]) have fin_y: "finite (set_mset Y)" using y_sub_xs by simp hence cyc: "\ acyclic ?Gt" proof (rule finite_nonempty_ex_succ_imp_cyclic) show "\x \# Y. \y \# Y. (y, x) \ ?Gt" using ex_y[unfolded x_eq_y] by auto qed (auto simp: y_nemp) show False using acyc cyc by sat qed lemma msetext_dersh_snoc: "msetext_dersh gt (xs @ [x]) xs" unfolding msetext_dersh_def Let_def proof (intro exI conjI) show "mset xs = mset (xs @ [x]) - {#x#} + {#}" by simp qed auto lemma msetext_dersh_compat_cons: assumes ys_gt_xs: "msetext_dersh gt ys xs" shows "msetext_dersh gt (x # ys) (x # xs)" proof - obtain Y X where y_nemp: "Y \ {#}" and y_sub_ys: "Y \# mset ys" and xs_eq: "mset xs = mset ys - Y + X" and ex_y: "\x. x \# X \ (\y. y \# Y \ gt y x)" using ys_gt_xs[unfolded msetext_dersh_def Let_def mset_map] by blast show ?thesis unfolding msetext_dersh_def Let_def proof (intro exI conjI) show "Y \# mset (x # ys)" using y_sub_ys by (metis add_mset_add_single mset.simps(2) mset_subset_eq_add_left subset_mset.add_increasing2) next show "mset (x # xs) = mset (x # ys) - Y + X" proof - have "X + (mset ys - Y) = mset xs" by (simp add: union_commute xs_eq) hence "mset (x # xs) = X + (mset (x # ys) - Y)" by (metis add_mset_add_single mset.simps(2) mset_subset_eq_multiset_union_diff_commute union_mset_add_mset_right y_sub_ys) thus ?thesis by (simp add: union_commute) qed qed (auto simp: y_nemp ex_y) qed lemma msetext_dersh_compat_snoc: "msetext_dersh gt ys xs \ msetext_dersh gt (ys @ [x]) (xs @ [x])" using msetext_dersh_compat_cons[of gt ys xs x] unfolding msetext_dersh_def by simp lemma msetext_dersh_compat_list: assumes y_gt_x: "gt y x" shows "msetext_dersh gt (xs @ y # xs') (xs @ x # xs')" unfolding msetext_dersh_def Let_def proof (intro exI conjI) show "mset (xs @ x # xs') = mset (xs @ y # xs') - {#y#} + {#x#}" by auto qed (auto intro: y_gt_x) lemma msetext_dersh_singleton: "msetext_dersh gt [y] [x] \ gt y x" unfolding msetext_dersh_def Let_def by (auto dest: nonempty_subseteq_mset_eq_singleton simp: nonempty_subseteq_mset_iff_singleton) lemma msetext_dersh_wf: assumes wf_gt: "wfP (\x y. gt y x)" shows "wfP (\xs ys. msetext_dersh gt ys xs)" proof (rule wfP_subset, rule wfP_app[of "\xs ys. (xs, ys) \ mult {(x, y). gt y x}" mset]) show "wfP (\xs ys. (xs, ys) \ mult {(x, y). gt y x})" using wf_gt unfolding wfP_def by (auto intro: wf_mult) next show "(\xs ys. msetext_dersh gt ys xs) \ (\x y. (mset x, mset y) \ mult {(x, y). gt y x})" using msetext_dersh_imp_mult by blast qed interpretation msetext_dersh: ext msetext_dersh by standard (fact msetext_dersh_mono_strong, rule msetext_dersh_map_strong, metis in_listsD) interpretation msetext_dersh: ext_trans_before_irrefl msetext_dersh by standard (fact msetext_dersh_trans, fact msetext_dersh_irrefl_from_trans) interpretation msetext_dersh: ext_snoc msetext_dersh by standard (fact msetext_dersh_snoc) interpretation msetext_dersh: ext_compat_cons msetext_dersh by standard (fact msetext_dersh_compat_cons) interpretation msetext_dersh: ext_compat_snoc msetext_dersh by standard (fact msetext_dersh_compat_snoc) interpretation msetext_dersh: ext_compat_list msetext_dersh by standard (rule msetext_dersh_compat_list) interpretation msetext_dersh: ext_singleton msetext_dersh by standard (rule msetext_dersh_singleton) interpretation msetext_dersh: ext_wf msetext_dersh by standard (fact msetext_dersh_wf) subsection \Huet--Oppen Multiset Extension\ definition msetext_huet where "msetext_huet gt ys xs = (let N = mset ys; M = mset xs in M \ N \ (\x. count M x > count N x \ (\y. gt y x \ count N y > count M y)))" lemma msetext_huet_imp_count_gt: assumes ys_gt_xs: "msetext_huet gt ys xs" shows "\x. count (mset ys) x > count (mset xs) x" proof - obtain x where "count (mset ys) x \ count (mset xs) x" using ys_gt_xs[unfolded msetext_huet_def Let_def] by (fastforce intro: multiset_eqI) moreover { assume "count (mset ys) x < count (mset xs) x" hence ?thesis using ys_gt_xs[unfolded msetext_huet_def Let_def] by blast } moreover { assume "count (mset ys) x > count (mset xs) x" hence ?thesis by fast } ultimately show ?thesis by fastforce qed lemma msetext_huet_imp_dersh: assumes huet: "msetext_huet gt ys xs" shows "msetext_dersh gt ys xs" proof (unfold msetext_dersh_def Let_def, intro exI conjI) let ?X = "mset xs - mset ys" let ?Y = "mset ys - mset xs" show "?Y \ {#}" by (metis msetext_huet_imp_count_gt[OF huet] empty_iff in_diff_count set_mset_empty) show "?Y \# mset ys" by auto show "mset xs = mset ys - ?Y + ?X" by (rule multiset_eqI) simp show "\x. x \# ?X \ (\y. y \# ?Y \ gt y x)" using huet[unfolded msetext_huet_def Let_def, THEN conjunct2] by (meson in_diff_count) qed text \ The following proof is based on that of @{thm[source] mult_imp_less_multiset\<^sub>H\<^sub>O}. \ lemma mult_imp_msetext_huet: assumes irrefl: "irreflp gt" and trans: "transp gt" and in_mult: "(mset xs, mset ys) \ mult {(x, y). gt y x}" shows "msetext_huet gt ys xs" using in_mult unfolding mult_def msetext_huet_def Let_def proof (induct rule: trancl_induct) case (base Ys) thus ?case using irrefl unfolding irreflp_def msetext_huet_def Let_def mult1_def by (auto 0 3 split: if_splits) next case (step Ys Zs) have asym[unfolded antisym_def, simplified]: "antisymp gt" by (rule irreflp_transp_imp_antisymP[OF irrefl trans]) from step(3) have "mset xs \ Ys" and **: "\x. count Ys x < count (mset xs) x \ (\y. gt y x \ count (mset xs) y < count Ys y)" by blast+ from step(2) obtain M0 a K where *: "Zs = M0 + {#a#}" "Ys = M0 + K" "a \# K" "\b. b \# K \ gt a b" using irrefl unfolding mult1_def irreflp_def by force have "mset xs \ Zs" proof (cases "K = {#}") case True thus ?thesis using \mset xs \ Ys\ ** *(1,2) irrefl[unfolded irreflp_def] by (metis One_nat_def add.comm_neutral count_single diff_union_cancelL lessI minus_multiset.rep_eq not_add_less2 plus_multiset.rep_eq union_commute zero_less_diff) next case False thus ?thesis proof - obtain aa :: "'a \ 'a" where f1: "\a. \ count Ys a < count (mset xs) a \ gt (aa a) a \ count (mset xs) (aa a) < count Ys (aa a)" - using "**" by moura + using ** by moura have f2: "K + M0 = Ys" - using "*"(2) union_ac(2) by blast + using *(2) union_ac(2) by blast have f3: "\aa. count Zs aa = count M0 aa + count {#a#} aa" - by (simp add: "*"(1)) + by (simp add: *(1)) have f4: "\a. count Ys a = count K a + count M0 a" using f2 by auto have f5: "count K a = 0" - by (meson "*"(3) count_inI) + by (meson *(3) count_inI) have "Zs - M0 = {#a#}" - using "*"(1) add_diff_cancel_left' by blast + using *(1) add_diff_cancel_left' by blast then have f6: "count M0 a < count Zs a" by (metis in_diff_count union_single_eq_member) have "\m. count m a = 0 + count m a" by simp moreover { assume "aa a \ a" then have "mset xs = Zs \ count Zs (aa a) < count K (aa a) + count M0 (aa a) \ count K (aa a) + count M0 (aa a) < count Zs (aa a)" - using f5 f3 f2 f1 "*"(4) asym by (auto dest!: antisympD) } + using f5 f3 f2 f1 *(4) asym by (auto dest!: antisympD) } ultimately show ?thesis using f6 f5 f4 f1 by (metis less_imp_not_less) qed qed moreover { assume "count Zs a \ count (mset xs) a" with \a \# K\ have "count Ys a < count (mset xs) a" unfolding *(1,2) by (auto simp add: not_in_iff) - with ** obtain z where z: "gt z a" "count (mset xs) z < count Ys z" - by blast - with * have "count Ys z \ count Zs z" - using asym - by (auto simp: intro: count_inI dest: antisympD) - with z have "\z. gt z a \ count (mset xs) z < count Zs z" by auto + with ** obtain z where z: "gt z a" "count (mset xs) z < count Ys z" + by blast + with * have "count Ys z \ count Zs z" + using asym by (auto simp: intro: count_inI dest: antisympD) + with z have "\z. gt z a \ count (mset xs) z < count Zs z" by auto } note count_a = this { fix y assume count_y: "count Zs y < count (mset xs) y" have "\x. gt x y \ count (mset xs) x < count Zs x" proof (cases "y = a") case True with count_y count_a show ?thesis by auto next case False show ?thesis proof (cases "y \# K") case True with *(4) have "gt a y" by simp then show ?thesis by (cases "count Zs a \ count (mset xs) a", blast dest: count_a trans[unfolded transp_def, rule_format], auto dest: count_a) next case False with \y \ a\ have "count Zs y = count Ys y" unfolding *(1,2) by (simp add: not_in_iff) with count_y ** obtain z where z: "gt z y" "count (mset xs) z < count Ys z" by auto show ?thesis proof (cases "z \# K") case True with *(4) have "gt a z" by simp with z(1) show ?thesis by (cases "count Zs a \ count (mset xs) a") (blast dest: count_a not_le_imp_less trans[unfolded transp_def, rule_format])+ next case False with \a \# K\ have "count Ys z \ count Zs z" unfolding * by (auto simp add: not_in_iff) with z show ?thesis by auto qed qed qed } ultimately show ?case unfolding msetext_huet_def Let_def by blast qed theorem msetext_huet_eq_dersh: "irreflp gt \ transp gt \ msetext_dersh gt = msetext_huet gt" using msetext_huet_imp_dersh msetext_dersh_imp_mult mult_imp_msetext_huet by fast lemma msetext_huet_mono_strong: "(\y \ set ys. \x \ set xs. gt y x \ gt' y x) \ msetext_huet gt ys xs \ msetext_huet gt' ys xs" unfolding msetext_huet_def by (metis less_le_trans mem_Collect_eq not_le not_less0 set_mset_mset[unfolded set_mset_def]) lemma msetext_huet_map: assumes fin: "finite A" and ys_a: "ys \ lists A" and xs_a: "xs \ lists A" and irrefl_f: "\x \ A. \ gt (f x) (f x)" and trans_f: "\z \ A. \y \ A. \x \ A. gt (f z) (f y) \ gt (f y) (f x) \ gt (f z) (f x)" and compat_f: "\y \ A. \x \ A. gt y x \ gt (f y) (f x)" and ys_gt_xs: "msetext_huet gt ys xs" shows "msetext_huet gt (map f ys) (map f xs)" (is "msetext_huet _ ?fys ?fxs") proof - have irrefl: "\x \ A. \ gt x x" using irrefl_f compat_f by blast have ms_xs_ne_ys: "mset xs \ mset ys" and ex_gt: "\x. count (mset ys) x < count (mset xs) x \ (\y. gt y x \ count (mset xs) y < count (mset ys) y)" using ys_gt_xs[unfolded msetext_huet_def Let_def] by blast+ have ex_y: "\y. gt (f y) (f x) \ count (mset ?fxs) (f y) < count (mset (map f ys)) (f y)" if cnt_x: "count (mset xs) x > count (mset ys) x" for x proof - have x_in_a: "x \ A" using cnt_x xs_a dual_order.strict_trans2 by fastforce obtain y where y_gt_x: "gt y x" and cnt_y: "count (mset ys) y > count (mset xs) y" using cnt_x ex_gt by blast have y_in_a: "y \ A" using cnt_y ys_a dual_order.strict_trans2 by fastforce have wf_gt_f: "wfP (\y x. y \ A \ x \ A \ gt (f y) (f x))" by (rule finite_irreflp_transp_imp_wfp) (auto elim: trans_f[rule_format] simp: fin irrefl_f Collect_case_prod_Sigma irreflp_def transp_def) obtain yy where fyy_gt_fx: "gt (f yy) (f x)" and cnt_yy: "count (mset ys) yy > count (mset xs) yy" and max_yy: "\y \ A. yy \ A \ gt (f y) (f yy) \ gt (f y) (f x) \ count (mset xs) y \ count (mset ys) y" using wfP_eq_minimal[THEN iffD1, OF wf_gt_f, rule_format, of y "{y. gt (f y) (f x) \ count (mset xs) y < count (mset ys) y}", simplified] y_gt_x cnt_y by (metis compat_f not_less x_in_a y_in_a) have yy_in_a: "yy \ A" using cnt_yy ys_a dual_order.strict_trans2 by fastforce { assume "count (mset ?fxs) (f yy) \ count (mset ?fys) (f yy)" then obtain u where fu_eq_fyy: "f u = f yy" and cnt_u: "count (mset xs) u > count (mset ys) u" using count_image_mset_le_imp_lt cnt_yy mset_map by (metis (mono_tags)) have u_in_a: "u \ A" using cnt_u xs_a dual_order.strict_trans2 by fastforce obtain v where v_gt_u: "gt v u" and cnt_v: "count (mset ys) v > count (mset xs) v" using cnt_u ex_gt by blast have v_in_a: "v \ A" using cnt_v ys_a dual_order.strict_trans2 by fastforce have fv_gt_fu: "gt (f v) (f u)" using v_gt_u compat_f v_in_a u_in_a by blast hence fv_gt_fyy: "gt (f v) (f yy)" by (simp only: fu_eq_fyy) have "gt (f v) (f x)" using fv_gt_fyy fyy_gt_fx v_in_a yy_in_a x_in_a trans_f by blast hence False using max_yy[rule_format, of v] fv_gt_fyy v_in_a yy_in_a cnt_v by linarith } thus ?thesis using fyy_gt_fx leI by blast qed show ?thesis unfolding msetext_huet_def Let_def proof (intro conjI allI impI) { assume len_eq: "length xs = length ys" obtain x where cnt_x: "count (mset xs) x > count (mset ys) x" using len_eq ms_xs_ne_ys by (metis size_eq_ex_count_lt size_mset) hence "mset ?fxs \ mset ?fys" using ex_y by fastforce } thus "mset ?fxs \ mset (map f ys)" by (metis length_map size_mset) next fix fx assume cnt_fx: "count (mset ?fxs) fx > count (mset ?fys) fx" then obtain x where fx: "fx = f x" and cnt_x: "count (mset xs) x > count (mset ys) x" using count_image_mset_lt_imp_lt mset_map by (metis (mono_tags)) thus "\fy. gt fy fx \ count (mset ?fxs) fy < count (mset (map f ys)) fy" using ex_y[OF cnt_x] by blast qed qed lemma msetext_huet_irrefl: "(\x \ set xs. \ gt x x) \ \ msetext_huet gt xs xs" unfolding msetext_huet_def by simp lemma msetext_huet_trans_from_irrefl: assumes fin: "finite A" and zs_a: "zs \ lists A" and ys_a: "ys \ lists A" and xs_a: "xs \ lists A" and irrefl: "\x \ A. \ gt x x" and trans: "\z \ A. \y \ A. \x \ A. gt z y \ gt y x \ gt z x" and zs_gt_ys: "msetext_huet gt zs ys" and ys_gt_xs: "msetext_huet gt ys xs" shows "msetext_huet gt zs xs" proof - have wf_gt: "wfP (\y x. y \ A \ x \ A \ gt y x)" by (rule finite_irreflp_transp_imp_wfp) (auto elim: trans[rule_format] simp: fin irrefl Collect_case_prod_Sigma irreflp_def transp_def) show ?thesis unfolding msetext_huet_def Let_def proof (intro conjI allI impI) obtain x where cnt_x: "count (mset zs) x > count (mset ys) x" using msetext_huet_imp_count_gt[OF zs_gt_ys] by blast have x_in_a: "x \ A" using cnt_x zs_a dual_order.strict_trans2 by fastforce obtain xx where cnt_xx: "count (mset zs) xx > count (mset ys) xx" and max_xx: "\y \ A. xx \ A \ gt y xx \ count (mset ys) y \ count (mset zs) y" using wfP_eq_minimal[THEN iffD1, OF wf_gt, rule_format, of x "{y. count (mset ys) y < count (mset zs) y}", simplified] cnt_x by force have xx_in_a: "xx \ A" using cnt_xx zs_a dual_order.strict_trans2 by fastforce show "mset xs \ mset zs" proof (cases "count (mset ys) xx \ count (mset xs) xx") case True thus ?thesis using cnt_xx by fastforce next case False hence "count (mset ys) xx < count (mset xs) xx" by fastforce then obtain z where z_gt_xx: "gt z xx" and cnt_z: "count (mset ys) z > count (mset xs) z" using ys_gt_xs[unfolded msetext_huet_def Let_def] by blast have z_in_a: "z \ A" using cnt_z ys_a dual_order.strict_trans2 by fastforce have "count (mset zs) z \ count (mset ys) z" using max_xx[rule_format, of z] z_in_a xx_in_a z_gt_xx by blast moreover { assume "count (mset zs) z < count (mset ys) z" then obtain u where u_gt_z: "gt u z" and cnt_u: "count (mset ys) u < count (mset zs) u" using zs_gt_ys[unfolded msetext_huet_def Let_def] by blast have u_in_a: "u \ A" using cnt_u zs_a dual_order.strict_trans2 by fastforce have u_gt_xx: "gt u xx" using trans u_in_a z_in_a xx_in_a u_gt_z z_gt_xx by blast have False using max_xx[rule_format, of u] u_in_a xx_in_a u_gt_xx cnt_u by fastforce } ultimately have "count (mset zs) z = count (mset ys) z" by fastforce thus ?thesis using cnt_z by fastforce qed next fix x assume cnt_x_xz: "count (mset zs) x < count (mset xs) x" have x_in_a: "x \ A" using cnt_x_xz xs_a dual_order.strict_trans2 by fastforce let ?case = "\y. gt y x \ count (mset zs) y > count (mset xs) y" { assume cnt_x: "count (mset zs) x < count (mset ys) x" then obtain y where y_gt_x: "gt y x" and cnt_y: "count (mset zs) y > count (mset ys) y" using zs_gt_ys[unfolded msetext_huet_def Let_def] by blast have y_in_a: "y \ A" using cnt_y zs_a dual_order.strict_trans2 by fastforce obtain yy where yy_gt_x: "gt yy x" and cnt_yy: "count (mset zs) yy > count (mset ys) yy" and max_yy: "\y \ A. yy \ A \ gt y yy \ gt y x \ count (mset ys) y \ count (mset zs) y" using wfP_eq_minimal[THEN iffD1, OF wf_gt, rule_format, of y "{y. gt y x \ count (mset ys) y < count (mset zs) y}", simplified] y_gt_x cnt_y by force have yy_in_a: "yy \ A" using cnt_yy zs_a dual_order.strict_trans2 by fastforce have ?case proof (cases "count (mset ys) yy \ count (mset xs) yy") case True thus ?thesis using yy_gt_x cnt_yy by fastforce next case False hence "count (mset ys) yy < count (mset xs) yy" by fastforce then obtain z where z_gt_yy: "gt z yy" and cnt_z: "count (mset ys) z > count (mset xs) z" using ys_gt_xs[unfolded msetext_huet_def Let_def] by blast have z_in_a: "z \ A" using cnt_z ys_a dual_order.strict_trans2 by fastforce have z_gt_x: "gt z x" using trans z_in_a yy_in_a x_in_a z_gt_yy yy_gt_x by blast have "count (mset zs) z \ count (mset ys) z" using max_yy[rule_format, of z] z_in_a yy_in_a z_gt_yy z_gt_x by blast moreover { assume "count (mset zs) z < count (mset ys) z" then obtain u where u_gt_z: "gt u z" and cnt_u: "count (mset ys) u < count (mset zs) u" using zs_gt_ys[unfolded msetext_huet_def Let_def] by blast have u_in_a: "u \ A" using cnt_u zs_a dual_order.strict_trans2 by fastforce have u_gt_yy: "gt u yy" using trans u_in_a z_in_a yy_in_a u_gt_z z_gt_yy by blast have u_gt_x: "gt u x" using trans u_in_a z_in_a x_in_a u_gt_z z_gt_x by blast have False using max_yy[rule_format, of u] u_in_a yy_in_a u_gt_yy u_gt_x cnt_u by fastforce } ultimately have "count (mset zs) z = count (mset ys) z" by fastforce thus ?thesis using z_gt_x cnt_z by fastforce qed } moreover { assume "count (mset zs) x \ count (mset ys) x" hence "count (mset ys) x < count (mset xs) x" using cnt_x_xz by fastforce then obtain y where y_gt_x: "gt y x" and cnt_y: "count (mset ys) y > count (mset xs) y" using ys_gt_xs[unfolded msetext_huet_def Let_def] by blast have y_in_a: "y \ A" using cnt_y ys_a dual_order.strict_trans2 by fastforce obtain yy where yy_gt_x: "gt yy x" and cnt_yy: "count (mset ys) yy > count (mset xs) yy" and max_yy: "\y \ A. yy \ A \ gt y yy \ gt y x \ count (mset xs) y \ count (mset ys) y" using wfP_eq_minimal[THEN iffD1, OF wf_gt, rule_format, of y "{y. gt y x \ count (mset xs) y < count (mset ys) y}", simplified] y_gt_x cnt_y by force have yy_in_a: "yy \ A" using cnt_yy ys_a dual_order.strict_trans2 by fastforce have ?case proof (cases "count (mset zs) yy \ count (mset ys) yy") case True thus ?thesis using yy_gt_x cnt_yy by fastforce next case False hence "count (mset zs) yy < count (mset ys) yy" by fastforce then obtain z where z_gt_yy: "gt z yy" and cnt_z: "count (mset zs) z > count (mset ys) z" using zs_gt_ys[unfolded msetext_huet_def Let_def] by blast have z_in_a: "z \ A" using cnt_z zs_a dual_order.strict_trans2 by fastforce have z_gt_x: "gt z x" using trans z_in_a yy_in_a x_in_a z_gt_yy yy_gt_x by blast have "count (mset ys) z \ count (mset xs) z" using max_yy[rule_format, of z] z_in_a yy_in_a z_gt_yy z_gt_x by blast moreover { assume "count (mset ys) z < count (mset xs) z" then obtain u where u_gt_z: "gt u z" and cnt_u: "count (mset xs) u < count (mset ys) u" using ys_gt_xs[unfolded msetext_huet_def Let_def] by blast have u_in_a: "u \ A" using cnt_u ys_a dual_order.strict_trans2 by fastforce have u_gt_yy: "gt u yy" using trans u_in_a z_in_a yy_in_a u_gt_z z_gt_yy by blast have u_gt_x: "gt u x" using trans u_in_a z_in_a x_in_a u_gt_z z_gt_x by blast have False using max_yy[rule_format, of u] u_in_a yy_in_a u_gt_yy u_gt_x cnt_u by fastforce } ultimately have "count (mset ys) z = count (mset xs) z" by fastforce thus ?thesis using z_gt_x cnt_z by fastforce qed } ultimately show "\y. gt y x \ count (mset xs) y < count (mset zs) y" by fastforce qed qed lemma msetext_huet_snoc: "msetext_huet gt (xs @ [x]) xs" unfolding msetext_huet_def Let_def by simp lemma msetext_huet_compat_cons: "msetext_huet gt ys xs \ msetext_huet gt (x # ys) (x # xs)" unfolding msetext_huet_def Let_def by auto lemma msetext_huet_compat_snoc: "msetext_huet gt ys xs \ msetext_huet gt (ys @ [x]) (xs @ [x])" unfolding msetext_huet_def Let_def by auto lemma msetext_huet_compat_list: "y \ x \ gt y x \ msetext_huet gt (xs @ y # xs') (xs @ x # xs')" unfolding msetext_huet_def Let_def by auto lemma msetext_huet_singleton: "y \ x \ msetext_huet gt [y] [x] \ gt y x" unfolding msetext_huet_def by simp lemma msetext_huet_wf: "wfP (\x y. gt y x) \ wfP (\xs ys. msetext_huet gt ys xs)" by (erule wfP_subset[OF msetext_dersh_wf]) (auto intro: msetext_huet_imp_dersh) lemma msetext_huet_hd_or_tl: assumes trans: "\z y x. gt z y \ gt y x \ gt z x" and total: "\y x. gt y x \ gt x y \ y = x" and len_eq: "length ys = length xs" and yys_gt_xxs: "msetext_huet gt (y # ys) (x # xs)" shows "gt y x \ msetext_huet gt ys xs" proof - let ?Y = "mset (y # ys)" let ?X = "mset (x # xs)" let ?Ya = "mset ys" let ?Xa = "mset xs" have Y_ne_X: "?Y \ ?X" and ex_gt_Y: "\xa. count ?X xa > count ?Y xa \ \ya. gt ya xa \ count ?Y ya > count ?X ya" using yys_gt_xxs[unfolded msetext_huet_def Let_def] by auto obtain yy where yy: "\xa. count ?X xa > count ?Y xa \ gt (yy xa) xa \ count ?Y (yy xa) > count ?X (yy xa)" using ex_gt_Y by metis have cnt_Y_pres: "count ?Ya xa > count ?Xa xa" if "count ?Y xa > count ?X xa" and "xa \ y" for xa using that by (auto split: if_splits) have cnt_X_pres: "count ?Xa xa > count ?Ya xa" if "count ?X xa > count ?Y xa" and "xa \ x" for xa using that by (auto split: if_splits) { assume y_eq_x: "y = x" have "?Xa \ ?Ya" using y_eq_x Y_ne_X by simp moreover have "\xa. count ?Xa xa > count ?Ya xa \ \ya. gt ya xa \ count ?Ya ya > count ?Xa ya" proof - fix xa :: 'a assume a1: "count (mset ys) xa < count (mset xs) xa" from ex_gt_Y obtain aa :: "'a \ 'a" where f3: "\a. \ count (mset (y # ys)) a < count (mset (x # xs)) a \ gt (aa a) a \ count (mset (x # xs)) (aa a) < count (mset (y # ys)) (aa a)" by (metis (full_types)) then have f4: "\a. count (mset (x # xs)) (aa a) < count (mset (x # ys)) (aa a) \ \ count (mset (x # ys)) a < count (mset (x # xs)) a" using y_eq_x by meson have "\a as aa. count (mset ((a::'a) # as)) aa = count (mset as) aa \ aa = a" by fastforce then have "xa = x \ count (mset (x # xs)) (aa xa) < count (mset (x # ys)) (aa xa)" using f4 a1 by (metis (no_types)) then show "\a. gt a xa \ count (mset xs) a < count (mset ys) a" using f3 y_eq_x a1 by (metis (no_types) Suc_less_eq count_add_mset mset.simps(2)) qed ultimately have "msetext_huet gt ys xs" unfolding msetext_huet_def Let_def by simp } moreover { assume x_gt_y: "gt x y" and y_ngt_x: "\ gt y x" hence y_ne_x: "y \ x" by fast obtain z where z_cnt: "count ?X z > count ?Y z" using size_eq_ex_count_lt[of ?Y ?X] size_mset size_mset len_eq Y_ne_X by auto have Xa_ne_Ya: "?Xa \ ?Ya" proof (cases "z = x") case True hence "yy z \ y" using y_ngt_x yy z_cnt by blast hence "count ?Ya (yy z) > count ?Xa (yy z)" using cnt_Y_pres yy z_cnt by blast thus ?thesis by auto next case False hence "count ?Xa z > count ?Ya z" using z_cnt cnt_X_pres by blast thus ?thesis by auto qed have "\ya. gt ya xa \ count ?Ya ya > count ?Xa ya" if xa_cnta: "count ?Xa xa > count ?Ya xa" for xa proof (cases "xa = y") case xa_eq_y: True { assume "count ?Ya x > count ?Xa x" moreover have "gt x xa" unfolding xa_eq_y by (rule x_gt_y) ultimately have ?thesis by fast } moreover { assume "count ?Xa x \ count ?Ya x" hence x_cnt: "count ?X x > count ?Y x" by (simp add: y_ne_x) hence yyx_gt_x: "gt (yy x) x" and yyx_cnt: "count ?Y (yy x) > count ?X (yy x)" using yy by blast+ have yyx_ne_y: "yy x \ y" using y_ngt_x yyx_gt_x by auto have "gt (yy x) xa" unfolding xa_eq_y using trans yyx_gt_x x_gt_y by blast moreover have "count ?Ya (yy x) > count ?Xa (yy x)" using cnt_Y_pres yyx_cnt yyx_ne_y by blast ultimately have ?thesis by blast } ultimately show ?thesis by fastforce next case False hence xa_cnt: "count ?X xa > count ?Y xa" using xa_cnta by fastforce show ?thesis proof (cases "yy xa = y \ count ?Ya y \ count ?Xa y") case yyxa_ne_y_or: False have yyxa_gt_xa: "gt (yy xa) xa" and yyxa_cnt: "count ?Y (yy xa) > count ?X (yy xa)" using yy[OF xa_cnt] by blast+ have "count ?Ya (yy xa) > count ?Xa (yy xa)" using cnt_Y_pres yyxa_cnt yyxa_ne_y_or by fastforce thus ?thesis using yyxa_gt_xa by blast next case True note yyxa_eq_y = this[THEN conjunct1] and y_cnt = this[THEN conjunct2] { assume "count ?Ya x > count ?Xa x" moreover have "gt x xa" using trans x_gt_y xa_cnt yy yyxa_eq_y by blast ultimately have ?thesis by fast } moreover { assume "count ?Xa x \ count ?Ya x" hence x_cnt: "count ?X x > count ?Y x" by (simp add: y_ne_x) hence yyx_gt_x: "gt (yy x) x" and yyx_cnt: "count ?Y (yy x) > count ?X (yy x)" using yy by blast+ have yyx_ne_y: "yy x \ y" using y_ngt_x yyx_gt_x by auto have "gt (yy x) xa" using trans x_gt_y xa_cnt yy yyx_gt_x yyxa_eq_y by blast moreover have "count ?Ya (yy x) > count ?Xa (yy x)" using cnt_Y_pres yyx_cnt yyx_ne_y by blast ultimately have ?thesis by blast } ultimately show ?thesis by fastforce qed qed hence "msetext_huet gt ys xs" unfolding msetext_huet_def Let_def using Xa_ne_Ya by fast } ultimately show ?thesis using total by blast qed interpretation msetext_huet: ext msetext_huet by standard (fact msetext_huet_mono_strong, fact msetext_huet_map) interpretation msetext_huet: ext_irrefl_before_trans msetext_huet by standard (fact msetext_huet_irrefl, fact msetext_huet_trans_from_irrefl) interpretation msetext_huet: ext_snoc msetext_huet by standard (fact msetext_huet_snoc) interpretation msetext_huet: ext_compat_cons msetext_huet by standard (fact msetext_huet_compat_cons) interpretation msetext_huet: ext_compat_snoc msetext_huet by standard (fact msetext_huet_compat_snoc) interpretation msetext_huet: ext_compat_list msetext_huet by standard (fact msetext_huet_compat_list) interpretation msetext_huet: ext_singleton msetext_huet by standard (fact msetext_huet_singleton) interpretation msetext_huet: ext_wf msetext_huet by standard (fact msetext_huet_wf) interpretation msetext_huet: ext_hd_or_tl msetext_huet by standard (rule msetext_huet_hd_or_tl) interpretation msetext_huet: ext_wf_bounded msetext_huet by standard subsection \Componentwise Extension\ definition cwiseext :: "('a \ 'a \ bool) \ 'a list \ 'a list \ bool" where "cwiseext gt ys xs \ length ys = length xs \ (\i < length ys. gt (ys ! i) (xs ! i) \ ys ! i = xs ! i) \ (\i < length ys. gt (ys ! i) (xs ! i))" lemma cwiseext_imp_len_lexext: assumes cw: "cwiseext gt ys xs" shows "len_lexext gt ys xs" proof - have len_eq: "length ys = length xs" using cw[unfolded cwiseext_def] by sat moreover have "lexext gt ys xs" proof - obtain j where j_len: "j < length ys" and j_gt: "gt (ys ! j) (xs ! j)" using cw[unfolded cwiseext_def] by blast then obtain j0 where j0_len: "j0 < length ys" and j0_gt: "gt (ys ! j0) (xs ! j0)" and j0_min: "\i. i < j0 \ \ gt (ys ! i) (xs ! i)" using wf_eq_minimal[THEN iffD1, OF wf_less, rule_format, of _ "{i. gt (ys ! i) (xs ! i)}", simplified, OF j_gt] by (metis less_trans nat_neq_iff) have j0_eq: "\i. i < j0 \ ys ! i = xs ! i" using cw[unfolded cwiseext_def] by (metis j0_len j0_min less_trans) have "lexext gt (drop j0 ys) (drop j0 xs)" using lexext_Cons[of gt _ _ "drop (Suc j0) ys" "drop (Suc j0) xs", OF j0_gt] by (metis Cons_nth_drop_Suc j0_len len_eq) thus ?thesis using cw len_eq j0_len j0_min proof (induct j0 arbitrary: ys xs) case (Suc k) note ih0 = this(1) and gts_dropSk = this(2) and cw = this(3) and len_eq = this(4) and Sk_len = this(5) and Sk_min = this(6) have Sk_eq: "\i. i < Suc k \ ys ! i = xs ! i" using cw[unfolded cwiseext_def] by (metis Sk_len Sk_min less_trans) have k_len: "k < length ys" using Sk_len by simp have k_min: "\i. i < k \ \ gt (ys ! i) (xs ! i)" using Sk_min by simp have k_eq: "\i. i < k \ ys ! i = xs ! i" using Sk_eq by simp note ih = ih0[OF _ cw len_eq k_len k_min] show ?case proof (cases "k < length ys") case k_lt_ys: True note k_lt_xs = k_lt_ys[unfolded len_eq] obtain x where x: "x = xs ! k" by simp hence y: "x = ys ! k" using Sk_eq[of k] by simp have dropk_xs: "drop k xs = x # drop (Suc k) xs" using k_lt_xs x by (simp add: Cons_nth_drop_Suc) have dropk_ys: "drop k ys = x # drop (Suc k) ys" using k_lt_ys y by (simp add: Cons_nth_drop_Suc) show ?thesis by (rule ih, unfold dropk_xs dropk_ys, rule lexext_Cons_eq[OF gts_dropSk]) next case False hence "drop k xs = []" and "drop k ys = []" using len_eq by simp_all hence "lexext gt [] []" using gts_dropSk by simp hence "lexext gt (drop k ys) (drop k xs)" by simp thus ?thesis by (rule ih) qed qed simp qed ultimately show ?thesis unfolding lenext_def by sat qed lemma cwiseext_mono_strong: "(\y \ set ys. \x \ set xs. gt y x \ gt' y x) \ cwiseext gt ys xs \ cwiseext gt' ys xs" unfolding cwiseext_def by (induct, force, fast) lemma cwiseext_map_strong: "(\y \ set ys. \x \ set xs. gt y x \ gt (f y) (f x)) \ cwiseext gt ys xs \ cwiseext gt (map f ys) (map f xs)" unfolding cwiseext_def by auto lemma cwiseext_irrefl: "(\x \ set xs. \ gt x x) \ \ cwiseext gt xs xs" unfolding cwiseext_def by (blast intro: nth_mem) lemma cwiseext_trans_strong: assumes "\z \ set zs. \y \ set ys. \x \ set xs. gt z y \ gt y x \ gt z x" and "cwiseext gt zs ys" and "cwiseext gt ys xs" shows "cwiseext gt zs xs" using assms unfolding cwiseext_def by (metis (mono_tags) nth_mem) lemma cwiseext_compat_cons: "cwiseext gt ys xs \ cwiseext gt (x # ys) (x # xs)" unfolding cwiseext_def proof (elim conjE, intro conjI) assume "length ys = length xs" and "\i < length ys. gt (ys ! i) (xs ! i) \ ys ! i = xs ! i" thus "\i < length (x # ys). gt ((x # ys) ! i) ((x # xs) ! i) \ (x # ys) ! i = (x # xs) ! i" by (simp add: nth_Cons') next assume "\i < length ys. gt (ys ! i) (xs ! i)" thus "\i < length (x # ys). gt ((x # ys) ! i) ((x # xs) ! i)" by fastforce qed auto lemma cwiseext_compat_snoc: "cwiseext gt ys xs \ cwiseext gt (ys @ [x]) (xs @ [x])" unfolding cwiseext_def proof (elim conjE, intro conjI) assume "length ys = length xs" and "\i < length ys. gt (ys ! i) (xs ! i) \ ys ! i = xs ! i" thus "\i < length (ys @ [x]). gt ((ys @ [x]) ! i) ((xs @ [x]) ! i) \ (ys @ [x]) ! i = (xs @ [x]) ! i" by (simp add: nth_append) next assume "length ys = length xs" and "\i < length ys. gt (ys ! i) (xs ! i)" thus "\i < length (ys @ [x]). gt ((ys @ [x]) ! i) ((xs @ [x]) ! i)" by (metis length_append_singleton less_Suc_eq nth_append) qed auto lemma cwiseext_compat_list: assumes y_gt_x: "gt y x" shows "cwiseext gt (xs @ y # xs') (xs @ x # xs')" unfolding cwiseext_def proof (intro conjI) show "\i < length (xs @ y # xs'). gt ((xs @ y # xs') ! i) ((xs @ x # xs') ! i) \ (xs @ y # xs') ! i = (xs @ x # xs') ! i" using y_gt_x by (simp add: nth_Cons' nth_append) next show "\i < length (xs @ y # xs'). gt ((xs @ y # xs') ! i) ((xs @ x # xs') ! i)" using y_gt_x by (metis add_diff_cancel_right' append_is_Nil_conv diff_less length_append length_greater_0_conv list.simps(3) nth_append_length) qed auto lemma cwiseext_singleton: "cwiseext gt [y] [x] \ gt y x" unfolding cwiseext_def by auto lemma cwiseext_wf: "wfP (\x y. gt y x) \ wfP (\xs ys. cwiseext gt ys xs)" by (auto intro: cwiseext_imp_len_lexext wfP_subset[OF len_lexext_wf]) lemma cwiseext_hd_or_tl: "cwiseext gt (y # ys) (x # xs) \ gt y x \ cwiseext gt ys xs" unfolding cwiseext_def proof (elim conjE, intro disj_imp[THEN iffD2, rule_format] conjI) assume "\i < length (y # ys). gt ((y # ys) ! i) ((x # xs) ! i)" and "\ gt y x" thus "\i < length ys. gt (ys ! i) (xs ! i)" by (metis (no_types) One_nat_def diff_le_self diff_less dual_order.strict_trans2 length_Cons less_Suc_eq linorder_neqE_nat not_less0 nth_Cons') qed auto locale ext_cwiseext = ext_compat_list + ext_compat_cons begin context fixes gt :: "'a \ 'a \ bool" assumes gt_irrefl: "\ gt x x" and trans_gt: "ext gt zs ys \ ext gt ys xs \ ext gt zs xs" begin lemma assumes ys_gtcw_xs: "cwiseext gt ys xs" shows "ext gt ys xs" proof - have "length ys = length xs" by (rule ys_gtcw_xs[unfolded cwiseext_def, THEN conjunct1]) thus ?thesis using ys_gtcw_xs proof (induct rule: list_induct2) case Nil thus ?case unfolding cwiseext_def by simp next case (Cons y ys x xs) note len_ys_eq_xs = this(1) and ih = this(2) and yys_gtcw_xxs = this(3) have xys_gts_xxs: "ext gt (x # ys) (x # xs)" if ys_ne_xs: "ys \ xs" proof - have ys_gtcw_xs: "cwiseext gt ys xs" using yys_gtcw_xxs unfolding cwiseext_def proof (elim conjE, intro conjI) assume "\i < length (y # ys). gt ((y # ys) ! i) ((x # xs) ! i) \ (y # ys) ! i = (x # xs) ! i" hence ge: "\i < length ys. gt (ys ! i) (xs ! i) \ ys ! i = xs ! i" by auto thus "\i < length ys. gt (ys ! i) (xs ! i)" using ys_ne_xs len_ys_eq_xs nth_equalityI by blast qed auto hence "ext gt ys xs" by (rule ih) thus "ext gt (x # ys) (x # xs)" by (rule compat_cons) qed have "gt y x \ y = x" using yys_gtcw_xxs unfolding cwiseext_def by fastforce moreover { assume y_eq_x: "y = x" have ?case proof (cases "ys = xs") case True hence False using y_eq_x gt_irrefl yys_gtcw_xxs unfolding cwiseext_def by presburger thus ?thesis by sat next case False thus ?thesis using y_eq_x xys_gts_xxs by simp qed } moreover { assume "y \ x" and "gt y x" hence yys_gts_xys: "ext gt (y # ys) (x # ys)" using compat_list[of _ _ gt "[]"] by simp have ?case proof (cases "ys = xs") case ys_eq_xs: True thus ?thesis using yys_gts_xys by simp next case False thus ?thesis using yys_gts_xys xys_gts_xxs trans_gt by blast qed } ultimately show ?case by sat qed qed end end interpretation cwiseext: ext cwiseext by standard (fact cwiseext_mono_strong, rule cwiseext_map_strong, metis in_listsD) interpretation cwiseext: ext_irrefl_trans_strong cwiseext by standard (fact cwiseext_irrefl, fact cwiseext_trans_strong) interpretation cwiseext: ext_compat_cons cwiseext by standard (fact cwiseext_compat_cons) interpretation cwiseext: ext_compat_snoc cwiseext by standard (fact cwiseext_compat_snoc) interpretation cwiseext: ext_compat_list cwiseext by standard (rule cwiseext_compat_list) interpretation cwiseext: ext_singleton cwiseext by standard (rule cwiseext_singleton) interpretation cwiseext: ext_wf cwiseext by standard (rule cwiseext_wf) interpretation cwiseext: ext_hd_or_tl cwiseext by standard (rule cwiseext_hd_or_tl) interpretation cwiseext: ext_wf_bounded cwiseext by standard end diff --git a/thys/Refine_Monadic/Refine_Leof.thy b/thys/Refine_Monadic/Refine_Leof.thy --- a/thys/Refine_Monadic/Refine_Leof.thy +++ b/thys/Refine_Monadic/Refine_Leof.thy @@ -1,224 +1,224 @@ section \Less-Equal or Fail\ (* TODO: Move to Refinement Framework *) theory Refine_Leof imports Refine_Basic begin text \A predicate that states refinement or that the LHS fails.\ definition le_or_fail :: "'a nres \ 'a nres \ bool" (infix "\\<^sub>n" 50) where "m \\<^sub>n m' \ nofail m \ m \ m'" lemma leofI[intro?]: assumes "nofail m \ m \ m'" shows "m \\<^sub>n m'" using assms unfolding le_or_fail_def by auto lemma leofD: assumes "nofail m" assumes "m \\<^sub>n m'" shows "m \ m'" using assms unfolding le_or_fail_def by blast lemma pw_leof_iff: "m \\<^sub>n m' \ (nofail m \ (\x. inres m x \ inres m' x))" unfolding le_or_fail_def by (auto simp add: pw_le_iff refine_pw_simps) lemma le_by_leofI: "\ nofail m' \ nofail m; m\\<^sub>nm' \ \ m \ m'" by (auto simp: pw_le_iff pw_leof_iff) lemma inres_leof_mono: "m\\<^sub>nm' \ nofail m \ inres m x \ inres m' x" by (auto simp: pw_leof_iff) lemma leof_trans[trans]: "\a \\<^sub>n RES X; RES X \\<^sub>n c\ \ a \\<^sub>n c" by (auto simp: pw_leof_iff) lemma leof_trans_nofail: "\ a\\<^sub>nb; nofail b; b\\<^sub>nc \ \ a \\<^sub>n c" by (auto simp: pw_leof_iff) lemma leof_refl[simp]: "a \\<^sub>n a" by (auto simp: pw_leof_iff) lemma leof_RES_UNIV[simp, intro!]: "m \\<^sub>n RES UNIV" by (auto simp: pw_leof_iff) lemma leof_FAIL[simp, intro!]: "m \\<^sub>n FAIL" by (auto simp: pw_leof_iff) lemma FAIL_leof[simp, intro!]: "FAIL \\<^sub>n m" by (auto simp: le_or_fail_def) lemma leof_lift: "m \ F \ m \\<^sub>n F" by (auto simp add: pw_leof_iff pw_le_iff) lemma leof_RETURN_rule[refine_vcg]: "\ m \ RETURN m \\<^sub>n SPEC \" by (simp add: pw_leof_iff) lemma leof_bind_rule[refine_vcg]: "\ m \\<^sub>n SPEC (\x. f x \\<^sub>n SPEC \) \ \ m\f \\<^sub>n SPEC \" by (auto simp add: pw_leof_iff refine_pw_simps) lemma RETURN_leof_RES_iff[simp]: "RETURN x \\<^sub>n RES Y \ x\Y" by (auto simp add: pw_leof_iff refine_pw_simps) lemma RES_leof_RES_iff[simp]: "RES X \\<^sub>n RES Y \ X \ Y" by (auto simp add: pw_leof_iff refine_pw_simps) lemma leof_Let_rule[refine_vcg]: "f x \\<^sub>n SPEC \ \ Let x f \\<^sub>n SPEC \" by simp lemma leof_If_rule[refine_vcg]: "\c \ t \\<^sub>n SPEC \; \c \ e \\<^sub>n SPEC \\ \ If c t e \\<^sub>n SPEC \" by simp lemma leof_RES_rule[refine_vcg]: "\\x. \ x \ \ x\ \ SPEC \ \\<^sub>n SPEC \" "\\x. x\X \ \ x\ \ RES X \\<^sub>n SPEC \" by auto lemma leof_True_rule: "\\x. \ x\ \ m \\<^sub>n SPEC \" by (auto simp add: pw_leof_iff refine_pw_simps) lemma sup_leof_iff: "(sup a b \\<^sub>n m) \ (nofail a \ nofail b \ a\\<^sub>nm \ b\\<^sub>nm)" by (auto simp: pw_leof_iff refine_pw_simps) lemma sup_leof_rule[refine_vcg]: assumes "\nofail a; nofail b\ \ a\\<^sub>nm" assumes "\nofail a; nofail b\ \ b\\<^sub>nm" shows "sup a b \\<^sub>n m" using assms by (auto simp: pw_leof_iff refine_pw_simps) lemma leof_option_rule[refine_vcg]: "\v = None \ S1 \\<^sub>n SPEC \; \x. v = Some x \ f2 x \\<^sub>n SPEC \\ \ (case v of None \ S1 | Some x \ f2 x) \\<^sub>n SPEC \" by (cases v) auto lemma ASSERT_leof_rule[refine_vcg]: assumes "\ \ m \\<^sub>n m'" shows "do {ASSERT \; m} \\<^sub>n m'" using assms by (cases \, auto simp: pw_leof_iff) lemma leof_ASSERT_rule[refine_vcg]: "\\ \ m \\<^sub>n m'\ \ m \\<^sub>n ASSERT \ \ m'" by (auto simp: pw_leof_iff refine_pw_simps) lemma leof_ASSERT_refine_rule[refine]: "\\ \ m \\<^sub>n \R m'\ \ m \\<^sub>n \R (ASSERT \ \ m')" by (auto simp: pw_leof_iff refine_pw_simps) lemma ASSUME_leof_iff: "(ASSUME \ \\<^sub>n SPEC \) \ (\ \ \ ())" by (auto simp: pw_leof_iff) lemma ASSUME_leof_rule[refine_vcg]: assumes "\ \ \ ()" shows "ASSUME \ \\<^sub>n SPEC \" using assms by (auto simp: ASSUME_leof_iff) lemma SPEC_rule_conj_leofI1: assumes "m \ SPEC \" assumes "m \\<^sub>n SPEC \" shows "m \ SPEC (\s. \ s \ \ s)" using assms by (auto simp: pw_le_iff pw_leof_iff) lemma SPEC_rule_conj_leofI2: assumes "m \\<^sub>n SPEC \" assumes "m \ SPEC \" shows "m \ SPEC (\s. \ s \ \ s)" using assms by (auto simp: pw_le_iff pw_leof_iff) lemma SPEC_rule_leof_conjI: assumes "m \\<^sub>n SPEC \" "m \\<^sub>n SPEC \" shows "m \\<^sub>n SPEC (\x. \ x \ \ x)" using assms by (auto simp: pw_leof_iff) lemma leof_use_spec_rule: assumes "m \\<^sub>n SPEC \" assumes "m \\<^sub>n SPEC (\s. \ s \ \ s)" shows "m \\<^sub>n SPEC \" using assms by (auto simp: pw_leof_iff refine_pw_simps) lemma use_spec_leof_rule: assumes "m \\<^sub>n SPEC \" assumes "m \ SPEC (\s. \ s \ \ s)" shows "m \ SPEC \" using assms by (auto simp: pw_leof_iff pw_le_iff refine_pw_simps) lemma leof_strengthen_SPEC: "m \\<^sub>n SPEC \ \ m \\<^sub>n SPEC (\x. inres m x \ \ x)" by (auto simp: pw_leof_iff) lemma build_rel_SPEC_leof: assumes "m \\<^sub>n SPEC (\x. I x \ \ (\ x))" shows "m \\<^sub>n \(br \ I) (SPEC \)" using assms by (auto simp: build_rel_SPEC_conv) lemma RETURN_as_SPEC_refine_leof[refine2]: assumes "M \\<^sub>n SPEC (\c. (c,a)\R)" shows "M \\<^sub>n \R (RETURN a)" using assms by (simp add: pw_leof_iff refine_pw_simps) lemma ASSERT_leof_defI: assumes "c \ do { ASSERT \; m'}" assumes "\ \ m' \\<^sub>n m" shows "c \\<^sub>n m" using assms by (auto simp: pw_leof_iff refine_pw_simps) lemma leof_fun_conv_le: "(f x \\<^sub>n M x) \ (f x \ (if nofail (f x) then M x else FAIL))" by (auto simp: pw_le_iff pw_leof_iff) lemma leof_add_nofailI: "\ nofail m \ m\\<^sub>nm' \ \ m\\<^sub>nm'" by (auto simp: pw_le_iff pw_leof_iff) lemma leof_cons_rule[refine_vcg_cons]: assumes "m \\<^sub>n SPEC Q" assumes "\x. Q x \ P x" shows "m \\<^sub>n SPEC P" using assms by (auto simp: pw_le_iff pw_leof_iff) lemma RECT_rule_leof: assumes WF: "wf (V::('x\'x) set)" assumes I0: "pre (x::'x)" assumes IS: "\f x. \ \x'. \pre x'; (x',x)\V\ \ f x' \\<^sub>n M x'; pre x; RECT body = f \ \ body f x \\<^sub>n M x" shows "RECT body x \\<^sub>n M x" apply (cases "\trimono body") apply (simp add: RECT_def) using assms unfolding leof_fun_conv_le apply - apply (rule RECT_rule[where pre=pre and V=V]) - apply clarsimp_all + apply clarsimp_all proof - fix xa :: 'x assume a1: "\x'. \pre x'; (x', xa) \ V\ \ REC\<^sub>T body x' \ (if nofail (REC\<^sub>T body x') then M x' else FAIL)" assume a2: "\x f. \\x'. \pre x'; (x', x) \ V\ \ f x' \ (if nofail (f x') then M x' else FAIL); pre x; REC\<^sub>T body = f\ \ body f x \ (if nofail (body f x) then M x else FAIL)" assume a3: "pre xa" assume a4: "nofail (REC\<^sub>T body xa)" assume a5: "trimono body" have f6: "\x. \ pre x \ (x, xa) \ V \ (if nofail (REC\<^sub>T body x) then REC\<^sub>T body x \ M x else REC\<^sub>T body x \ FAIL)" using a1 by presburger have f7: "\x f. ((\xa. (pre xa \ (xa, x) \ V) \ \ f xa \ (if nofail (f xa) then M xa else FAIL)) \ \ pre x \ REC\<^sub>T body \ f) \ body f x \ (if nofail (body f x) then M x else FAIL)" using a2 by blast obtain xx :: "('x \ 'a nres) \ 'x \ 'x" where f8: "\x0 x1. (\v2. (pre v2 \ (v2, x1) \ V) \ \ x0 v2 \ (if nofail (x0 v2) then M v2 else FAIL)) = ((pre (xx x0 x1) \ (xx x0 x1, x1) \ V) \ \ x0 (xx x0 x1) \ (if nofail (x0 (xx x0 x1)) then M (xx x0 x1) else FAIL))" by moura have f9: "\x0 x1. (x0 (xx x0 x1) \ (if nofail (x0 (xx x0 x1)) then M (xx x0 x1) else FAIL)) = (if nofail (x0 (xx x0 x1)) then x0 (xx x0 x1) \ M (xx x0 x1) else x0 (xx x0 x1) \ FAIL)" by presburger have "nofail (body (REC\<^sub>T body) xa)" using a5 a4 by (metis (no_types) RECT_unfold) then show "body (REC\<^sub>T body) xa \ M xa" using f9 f8 f7 f6 a3 by fastforce qed (* TODO: REC_rule_leof! (However, this may require some fix to the domain theory model of Refine_Monadic!) *) end diff --git a/thys/Subset_Boolean_Algebras/Subset_Boolean_Algebras.thy b/thys/Subset_Boolean_Algebras/Subset_Boolean_Algebras.thy --- a/thys/Subset_Boolean_Algebras/Subset_Boolean_Algebras.thy +++ b/thys/Subset_Boolean_Algebras/Subset_Boolean_Algebras.thy @@ -1,3608 +1,3608 @@ (* Title: Subset Boolean Algebras Authors: Walter Guttmann, Bernhard Möller Maintainer: Walter Guttmann *) theory Subset_Boolean_Algebras imports Stone_Algebras.P_Algebras begin section \Boolean Algebras\ text \ We show that Isabelle/HOL's \boolean_algebra\ class is equivalent to Huntington's axioms \<^cite>\"Huntington1933"\. See \<^cite>\"WamplerDoty2016"\ for related results. \ subsection \Huntington's Axioms\ text \Definition 1\ class huntington = sup + uminus + assumes associative: "x \ (y \ z) = (x \ y) \ z" assumes commutative: "x \ y = y \ x" assumes huntington: "x = -(-x \ y) \ -(-x \ -y)" begin lemma top_unique: "x \ -x = y \ -y" proof - have "x \ -x = y \ -(--y \ -x) \ -(--y \ --x)" by (smt associative commutative huntington) thus ?thesis by (metis associative huntington) qed end subsection \Equivalence to \boolean_algebra\ Class\ text \Definition 2\ class extended = sup + inf + minus + uminus + bot + top + ord + assumes top_def: "top = (THE x . \y . x = y \ -y)" (* define without imposing uniqueness *) assumes bot_def: "bot = -(THE x . \y . x = y \ -y)" assumes inf_def: "x \ y = -(-x \ -y)" assumes minus_def: "x - y = -(-x \ y)" assumes less_eq_def: "x \ y \ x \ y = y" assumes less_def: "x < y \ x \ y = y \ \ (y \ x = x)" class huntington_extended = huntington + extended begin lemma top_char: "top = x \ -x" using top_def top_unique by auto lemma bot_char: "bot = -top" by (simp add: bot_def top_def) subclass boolean_algebra proof show 1: "\x y. (x < y) = (x \ y \ \ y \ x)" by (simp add: less_def less_eq_def) show 2: "\x. x \ x" proof - fix x have "x \ top = top \ --x" by (metis (full_types) associative top_char) thus "x \ x" by (metis (no_types) associative huntington less_eq_def top_char) qed show 3: "\x y z. x \ y \ y \ z \ x \ z" by (metis associative less_eq_def) show 4: "\x y. x \ y \ y \ x \ x = y" by (simp add: commutative less_eq_def) show 5: "\x y. x \ y \ x" using 2 by (metis associative huntington inf_def less_eq_def) show 6: "\x y. x \ y \ y" using 5 commutative inf_def by fastforce show 8: "\x y. x \ x \ y" using 2 associative less_eq_def by auto show 9: "\y x. y \ x \ y" using 8 commutative by fastforce show 10: "\y x z. y \ x \ z \ x \ y \ z \ x" by (metis associative less_eq_def) show 11: "\x. bot \ x" using 8 by (metis bot_char huntington top_char) show 12: "\x. x \ top" using 6 11 by (metis huntington bot_def inf_def less_eq_def top_def) show 13: "\x y z. x \ y \ z = (x \ y) \ (x \ z)" proof - have 2: "\x y z . x \ (y \ z) = (x \ y) \ z" by (simp add: associative) have 3: "\x y z . (x \ y) \ z = x \ (y \ z)" using 2 by metis have 4: "\x y . x \ y = y \ x" by (simp add: commutative) have 5: "\x y . x = - (- x \ y) \ - (- x \ - y)" by (simp add: huntington) have 6: "\x y . - (- x \ y) \ - (- x \ - y) = x" using 5 by metis have 7: "\x y . x \ y = - (- x \ - y)" by (simp add: inf_def) have 10: "\x y z . x \ (y \ z) = y \ (x \ z)" using 3 4 by metis have 11: "\x y z . - (- x \ y) \ (- (- x \ - y) \ z) = x \ z" using 3 6 by metis have 12: "\x y . - (x \ - y) \ - (- y \ - x) = y" using 4 6 by metis have 13: "\x y . - (- x \ y) \ - (- y \ - x) = x" using 4 6 by metis have 14: "\x y . - x \ - (- (- x \ y) \ - - (- x \ - y)) = - x \ y" using 6 by metis have 18: "\x y z . - (x \ - y) \ (- (- y \ - x) \ z) = y \ z" using 3 12 by metis have 20: "\x y . - (- x \ - y) \ - (y \ - x) = x" using 4 12 by metis have 21: "\x y . - (x \ - y) \ - (- x \ - y) = y" using 4 12 by metis have 22: "\x y . - x \ - (- (y \ - x) \ - - (- x \ - y)) = y \ - x" using 6 12 by metis have 23: "\x y . - x \ - (- x \ (- y \ - (y \ - x))) = y \ - x" using 3 4 6 12 by metis have 24: "\x y . - x \ - (- (- x \ - y) \ - - (- x \ y)) = - x \ - y" using 6 12 by metis have 28: "\x y . - (- x \ - y) \ - (- y \ x) = y" using 4 13 by metis have 30: "\x y . - x \ - (- y \ (- x \ - (- x \ y))) = - x \ y" using 3 4 6 13 by metis have 32: "\x y z . - (- x \ y) \ (z \ - (- y \ - x)) = z \ x" using 10 13 by metis have 37: "\x y z . - (- x \ - y) \ (- (y \ - x) \ z) = x \ z" using 3 20 by metis have 39: "\x y z . - (- x \ - y) \ (z \ - (y \ - x)) = z \ x" using 10 20 by metis have 40: "\x y z . - (x \ - y) \ (- (- x \ - y) \ z) = y \ z" using 3 21 by metis have 43: "\x y . - x \ - (- y \ (- x \ - (y \ - x))) = y \ - x" using 3 4 6 21 by metis have 47: "\x y z . - (x \ y) \ - (- (- x \ z) \ - (- (- x \ - z) \ y)) = - x \ z" using 6 11 by metis have 55: "\x y . x \ - (- y \ - - x) = y \ - (- x \ y)" using 4 11 12 by metis have 58: "\x y . x \ - (- - y \ - x) = x \ - (- x \ y)" using 4 11 13 by metis have 63: "\x y . x \ - (- - x \ - y) = y \ - (- x \ y)" using 4 11 21 by metis have 71: "\x y . x \ - (- y \ x) = y \ - (- x \ y)" using 4 11 28 by metis have 75: "\x y . x \ - (- y \ x) = y \ - (y \ - x)" using 4 71 by metis have 78: "\x y . - x \ (y \ - (- x \ (y \ - - (- x \ - y)))) = - x \ - (- x \ - y)" using 3 4 6 71 by metis have 86: "\x y . - (- x \ - (- y \ x)) \ - (y \ - (- x \ y)) = - y \ x" using 4 20 71 by metis have 172: "\x y . - x \ - (- x \ - y) = y \ - (- - x \ y)" using 14 75 by metis have 201: "\x y . x \ - (- y \ - - x) = y \ - (y \ - x)" using 4 55 by metis have 236: "\x y . x \ - (- - y \ - x) = x \ - (y \ - x)" using 4 58 by metis have 266: "\x y . - x \ - (- (- x \ - (y \ - - x)) \ - - (- x \ - - (- - x \ y))) = - x \ - (- - x \ y)" using 14 58 236 by metis have 678: "\x y z . - (- x \ - (- y \ x)) \ (- (y \ - (- x \ y)) \ z) = - y \ (x \ z)" using 3 4 37 71 by smt have 745: "\x y z . - (- x \ - (- y \ x)) \ (z \ - (y \ - (- x \ y))) = z \ (- y \ x)" using 4 39 71 by metis have 800: "\x y . - - x \ (- y \ (- (y \ - - x) \ - (- x \ (- - x \ (- y \ - (y \ - - x)))))) = x \ - (y \ - - x)" using 3 23 63 by metis have 944: "\x y . x \ - (x \ - - (- (- x \ - y) \ - - (- x \ y))) = - (- x \ - y) \ - (- (- x \ - y) \ - - (- x \ y))" using 4 24 71 by metis have 948: "\x y . - x \ - (- (y \ - (y \ - - x)) \ - - (- x \ (- y \ - x))) = - x \ - (- y \ - x)" using 24 75 by metis have 950: "\x y . - x \ - (- (y \ - (- - x \ y)) \ - - (- x \ (- x \ - y))) = - x \ - (- x \ - y)" using 24 75 by metis have 961: "\x y . - x \ - (- (y \ - (- - x \ y)) \ - - (- x \ (- - - x \ - y))) = y \ - (- - x \ y)" using 24 63 by metis have 966: "\x y . - x \ - (- (y \ - (y \ - - x)) \ - - (- x \ (- y \ - - - x))) = y \ - (y \ - - x)" using 24 201 by metis have 969: "\x y . - x \ - (- (- x \ - (y \ - - x)) \ - - (- x \ (- - y \ - - x))) = - x \ - (y \ - - x)" using 24 236 by metis have 1096: "\x y z . - x \ (- (- x \ - y) \ z) = y \ (- (- - x \ y) \ z)" using 3 172 by metis have 1098: "\x y z . - x \ (y \ - (- x \ - z)) = y \ (z \ - (- - x \ z))" using 10 172 by metis have 1105: "\x y . x \ - x = y \ - y" using 4 10 12 32 172 by metis have 1109: "\x y z . x \ (- x \ y) = z \ (- z \ y)" using 3 1105 by metis have 1110: "\x y z . x \ - x = y \ (z \ - (y \ z))" using 3 1105 by metis have 1114: "\x y . - (- x \ - - x) = - (y \ - y)" using 7 1105 by metis have 1115: "\x y z . x \ (y \ - y) = z \ (x \ - z)" using 10 1105 by metis have 1117: "\x y . - (x \ - - x) \ - (y \ - y) = - x" using 4 13 1105 by metis have 1121: "\x y . - (x \ - x) \ - (y \ - - y) = - y" using 4 28 1105 by metis have 1122: "\x . - - x = x" using 4 28 1105 1117 by metis have 1134: "\x y z . - (x \ - y) \ (z \ - z) = y \ (- y \ - x)" using 18 1105 1122 by metis have 1140: "\x . - x \ - (x \ (x \ - x)) = - x \ - x" using 4 22 1105 1122 1134 by metis have 1143: "\x y . x \ (- x \ y) = y \ (x \ - y)" using 37 1105 1122 1134 by metis have 1155: "\x y . - (x \ - x) \ - (y \ y) = - y" using 1121 1122 by metis have 1156: "\x y . - (x \ x) \ - (y \ - y) = - x" using 1117 1122 by metis have 1157: "\x y . - (x \ - x) = - (y \ - y)" using 4 1114 1122 by metis have 1167: "\x y z . - x \ (y \ - (- x \ - z)) = y \ (z \ - (x \ z))" using 1098 1122 by metis have 1169: "\x y z . - x \ (- (- x \ - y) \ z) = y \ (- (x \ y) \ z)" using 1096 1122 by metis have 1227: "\x y . - x \ - (- x \ (y \ (x \ - (- x \ - (y \ x))))) = - x \ - (y \ x)" using 3 4 969 1122 by smt have 1230: "\x y . - x \ - (- x \ (- y \ (- x \ - (y \ - (y \ x))))) = y \ - (y \ x)" using 3 4 966 1122 by smt have 1234: "\x y . - x \ - (- x \ (- x \ (- y \ - (y \ - (x \ y))))) = y \ - (x \ y)" using 3 4 961 1122 by metis have 1239: "\x y . - x \ - (- x \ - y) = y \ - (x \ y)" using 3 4 950 1122 1234 by metis have 1240: "\x y . - x \ - (- y \ - x) = y \ - (y \ x)" using 3 4 948 1122 1230 by metis have 1244: "\x y . x \ - (x \ (y \ (y \ - (x \ y)))) = - (- x \ - y) \ - (y \ (y \ - (x \ y)))" using 3 4 944 1122 1167 by metis have 1275: "\x y . x \ (- y \ (- (y \ x) \ - (x \ (- x \ (- y \ - (y \ x)))))) = x \ - (y \ x)" using 10 800 1122 by metis have 1346: "\x y . - x \ - (x \ (y \ (y \ (x \ - (x \ (y \ x)))))) = - x \ - (x \ y)" using 3 4 10 266 1122 1167 by smt have 1377: "\x y . - x \ (y \ - (- x \ (y \ (- x \ - y)))) = y \ - (x \ y)" using 78 1122 1239 by metis have 1394: "\x y . - (- x \ - y) \ - (y \ (y \ (- x \ - (x \ y)))) = x" using 3 4 10 20 30 1122 1239 by smt have 1427: "\x y . - (- x \ - y) \ - (y \ - (x \ (x \ - (x \ y)))) = x \ (x \ - (x \ y))" using 3 4 30 40 1240 by smt have 1436: "\x . - x \ - (x \ (x \ (- x \ - x))) = - x \ (- x \ - (x \ - x))" using 3 4 30 1140 1239 by smt have 1437: "\x y . - (x \ y) \ - (x \ - y) = - x" using 6 1122 by metis have 1438: "\x y . - (x \ y) \ - (y \ - x) = - y" using 12 1122 by metis have 1439: "\x y . - (x \ y) \ - (- y \ x) = - x" using 13 1122 by metis have 1440: "\x y . - (x \ - y) \ - (y \ x) = - x" using 20 1122 by metis have 1441: "\x y . - (x \ y) \ - (- x \ y) = - y" using 21 1122 by metis have 1568: "\x y . x \ (- y \ - x) = y \ (- y \ x)" using 10 1122 1143 by metis have 1598: "\x . - x \ - (x \ (x \ (x \ - x))) = - x \ (- x \ - (x \ - x))" using 4 1436 1568 by metis have 1599: "\x y . - x \ (y \ - (x \ (- x \ (- x \ y)))) = y \ - (x \ y)" using 10 1377 1568 by smt have 1617: "\x . x \ (- x \ (- x \ - (x \ - x))) = x \ - x" using 3 4 10 71 1122 1155 1568 1598 by metis have 1632: "\x y z . - (x \ - x) \ - (- y \ (- (z \ - z) \ - (y \ - (x \ - x)))) = y \ - (x \ - x)" using 43 1157 by metis have 1633: "\x y z . - (x \ - x) \ - (- y \ (- (x \ - x) \ - (y \ - (z \ - z)))) = y \ - (x \ - x)" using 43 1157 by metis have 1636: "\x y . x \ - (y \ (- y \ - (x \ x))) = x \ x" using 43 1109 1122 by metis have 1645: "\x y . x \ - x = y \ (y \ - y)" using 3 1110 1156 by metis have 1648: "\x y z . - (x \ (y \ (- y \ - x))) \ - (z \ - z) = - (y \ - y)" using 3 1115 1156 by metis have 1657: "\x y z . x \ - x = y \ (z \ - z)" using 1105 1645 by metis have 1664: "\x y z . x \ - x = y \ (z \ - y)" using 1115 1645 by metis have 1672: "\x y z . x \ - x = y \ (- y \ z)" using 3 4 1657 by metis have 1697: "\x y z . - x \ (y \ x) = z \ - z" using 1122 1664 by metis have 1733: "\x y z . - (x \ y) \ - (- (z \ - z) \ - (- (- x \ - x) \ y)) = x \ - x" using 4 47 1105 1122 by metis have 1791: "\x y z . x \ - (y \ (- y \ z)) = x \ - (x \ - x)" using 4 71 1122 1672 by metis have 1818: "\x y z . x \ - (- y \ (z \ y)) = x \ - (x \ - x)" using 4 71 1122 1697 by metis have 1861: "\x y z . - (x \ - x) \ - (y \ - (z \ - z)) = - y" using 1437 1657 by metis have 1867: "\x y z . - (x \ - x) \ - (- y \ - (z \ y)) = y" using 1122 1437 1697 by metis have 1868: "\x y . x \ - (y \ - y) = x" using 1122 1155 1633 1861 by metis have 1869: "\x y z . - (x \ - x) \ - (- y \ (- (z \ - z) \ - y)) = y" using 1632 1868 by metis have 1870: "\x y . - (x \ - x) \ - y = - y" using 1861 1868 by metis have 1872: "\x y z . x \ - (- y \ (z \ y)) = x" using 1818 1868 by metis have 1875: "\x y z . x \ - (y \ (- y \ z)) = x" using 1791 1868 by metis have 1883: "\x y . - (x \ (y \ (- y \ - x))) = - (y \ - y)" using 1648 1868 by metis have 1885: "\x . x \ (x \ - x) = x \ - x" using 4 1568 1617 1868 by metis have 1886: "\x . - x \ - x = - x" using 1598 1868 1885 by metis have 1890: "\x . - (x \ x) = - x" using 1156 1868 by metis have 1892: "\x y . - (x \ - x) \ y = y" using 1122 1869 1870 1886 by metis have 1893: "\x y . - (- x \ - (y \ x)) = x" using 1867 1892 by metis have 1902: "\x y . x \ (y \ - (x \ y)) = x \ - x" using 3 4 1122 1733 1886 1892 by metis have 1908: "\x . x \ x = x" using 1636 1875 1890 by metis have 1910: "\x y . x \ - (y \ x) = - y \ x" using 1599 1875 by metis have 1921: "\x y . x \ (- y \ - (y \ x)) = - y \ x" using 1275 1875 1910 by metis have 1951: "\x y . - x \ - (y \ x) = - x" using 1227 1872 1893 1908 by metis have 1954: "\x y z . x \ (y \ - (x \ z)) = y \ (- z \ x)" using 745 1122 1910 1951 by metis have 1956: "\x y z . x \ (- (x \ y) \ z) = - y \ (x \ z)" using 678 1122 1910 1951 by metis have 1959: "\x y . x \ - (x \ y) = - y \ x" using 86 1122 1910 1951 by metis have 1972: "\x y . x \ (- x \ y) = x \ - x" using 1902 1910 by metis have 2000: "\x y . - (- x \ - y) \ - (y \ (- x \ y)) = x \ - (y \ (- x \ y))" using 4 1244 1910 1959 by metis have 2054: "\x y . x \ - (y \ (- x \ y)) = x" using 1394 1921 2000 by metis have 2057: "\x y . - (x \ (y \ - y)) = - (y \ - y)" using 1883 1972 by metis have 2061: "\x y . x \ (- y \ x) = x \ - y" using 4 1122 1427 1910 1959 2054 by metis have 2090: "\x y z . x \ (- (y \ x) \ z) = x \ (- y \ z)" using 1122 1169 1956 by metis have 2100: "\x y . - x \ - (x \ y) = - x" using 4 1346 1868 1885 1910 1959 1972 2057 by metis have 2144: "\x y . x \ - (y \ - x) = x" using 1122 1440 2000 2061 by metis have 2199: "\x y . x \ (x \ y) = x \ y" using 3 1908 by metis have 2208: "\x y z . x \ (- (y \ - x) \ z) = x \ z" using 3 2144 by metis have 2349: "\x y z . - (x \ y) \ - (x \ (y \ z)) = - (x \ y)" using 3 2100 by metis have 2432: "\x y z . - (x \ (y \ z)) \ - (y \ (z \ - x)) = - (y \ z)" using 3 1438 by metis have 2530: "\x y z . - (- (x \ y) \ z) = - (y \ (- x \ z)) \ - (- y \ z)" using 4 1122 1439 2090 2208 by smt have 3364: "\x y z . - (- x \ y) \ (z \ - (x \ y)) = z \ - y" using 3 4 1122 1441 1910 1954 2199 by metis have 5763: "\x y z . - (x \ y) \ - (- x \ (y \ z)) = - (x \ y) \ - (y \ z)" using 4 2349 3364 by metis have 6113: "\x y z . - (x \ (y \ z)) \ - (z \ - x) = - (y \ z) \ - (z \ - x)" using 4 2432 3364 5763 by metis show "\x y z. x \ y \ z = (x \ y) \ (x \ z)" proof - fix x y z have "- (y \ z \ x) = - (- (- y \ z) \ - (- y \ - z) \ x) \ - (x \ - - z)" using 1437 2530 6113 by (smt commutative inf_def) thus "x \ y \ z = (x \ y) \ (x \ z)" using 12 1122 by (metis commutative inf_def) qed qed show 14: "\x. x \ - x = bot" proof - fix x have "(bot \ x) \ (bot \ -x) = bot" using huntington bot_def inf_def by auto thus "x \ -x = bot" using 11 less_eq_def by force qed show 15: "\x. x \ - x = top" using 5 14 by (metis (no_types, lifting) huntington bot_def less_eq_def top_def) show 16: "\x y. x - y = x \ - y" using 15 by (metis commutative huntington inf_def minus_def) show 7: "\x y z. x \ y \ x \ z \ x \ y \ z" by (simp add: 13 less_eq_def) qed end context boolean_algebra begin sublocale ba_he: huntington_extended proof show "\x y z. x \ (y \ z) = x \ y \ z" by (simp add: sup_assoc) show "\x y. x \ y = y \ x" by (simp add: sup_commute) show "\x y. x = - (- x \ y) \ - (- x \ - y)" by simp show "top = (THE x. \y. x = y \ - y)" by auto show "bot = - (THE x. \y. x = y \ - y)" by auto show "\x y. x \ y = - (- x \ - y)" by simp show "\x y. x - y = - (- x \ y)" by (simp add: diff_eq) show "\x y. (x \ y) = (x \ y = y)" by (simp add: le_iff_sup) show "\x y. (x < y) = (x \ y = y \ y \ x \ x)" using sup.strict_order_iff sup_commute by auto qed end subsection \Stone Algebras\ text \ We relate Stone algebras to Boolean algebras. \ class stone_algebra_extended = stone_algebra + minus + assumes stone_minus_def[simp]: "x - y = x \ -y" class regular_stone_algebra = stone_algebra_extended + assumes double_complement[simp]: "--x = x" begin subclass boolean_algebra proof show "\x. x \ - x = bot" by simp show "\x. x \ - x = top" using regular_dense_top by fastforce show "\x y. x - y = x \ - y" by simp qed end context boolean_algebra begin sublocale ba_rsa: regular_stone_algebra proof show "\x y. x - y = x \ - y" by (simp add: diff_eq) show "\x. - - x = x" by simp qed end section \Alternative Axiomatisations of Boolean Algebras\ text \ We consider four axiomatisations of Boolean algebras based only on join and complement. The first three are from the literature and the fourth, a version using equational axioms, is new. The motivation for Byrne's and the new axiomatisation is that the axioms are easier to understand than Huntington's third axiom. We also include Meredith's axiomatisation. \ subsection \Lee Byrne's Formulation A\ text \ The following axiomatisation is from \<^cite>\\Formulation A\ in "Byrne1946"\; see also \<^cite>\"Frink1941"\. \ text \Theorem 3\ class boolean_algebra_1 = sup + uminus + assumes ba1_associative: "x \ (y \ z) = (x \ y) \ z" assumes ba1_commutative: "x \ y = y \ x" assumes ba1_complement: "x \ -y = z \ -z \ x \ y = x" begin subclass huntington proof show 1: "\x y z. x \ (y \ z) = x \ y \ z" by (simp add: ba1_associative) show "\x y. x \ y = y \ x" by (simp add: ba1_commutative) show "\x y. x = - (- x \ y) \ - (- x \ - y)" proof - have 2: "\x y. y \ (y \ x) = y \ x" using 1 by (metis ba1_complement) hence "\x. --x = x" by (smt ba1_associative ba1_commutative ba1_complement) hence "\x y. y \ -(y \ -x) = y \ x" by (smt ba1_associative ba1_commutative ba1_complement) thus "\x y. x = -(-x \ y) \ -(-x \ - y)" using 2 by (smt ba1_commutative ba1_complement) qed qed end context huntington begin sublocale h_ba1: boolean_algebra_1 proof show "\x y z. x \ (y \ z) = x \ y \ z" by (simp add: associative) show "\x y. x \ y = y \ x" by (simp add: commutative) show "\x y z. (x \ - y = z \ - z) = (x \ y = x)" proof fix x y z have 1: "\x y z. -(-x \ y) \ (-(-x \ -y) \ z) = x \ z" using associative huntington by force have 2: "\x y. -(x \ -y) \ -(-y \ -x) = y" by (metis commutative huntington) show "x \ - y = z \ - z \ x \ y = x" by (metis 1 2 associative commutative top_unique) show "x \ y = x \ x \ - y = z \ - z" by (metis associative huntington commutative top_unique) qed qed end subsection \Lee Byrne's Formulation B\ text \ The following axiomatisation is from \<^cite>\\Formulation B\ in "Byrne1946"\. \ text \Theorem 4\ class boolean_algebra_2 = sup + uminus + assumes ba2_associative_commutative: "(x \ y) \ z = (y \ z) \ x" assumes ba2_complement: "x \ -y = z \ -z \ x \ y = x" begin subclass boolean_algebra_1 proof show "\x y z. x \ (y \ z) = x \ y \ z" by (smt ba2_associative_commutative ba2_complement) show "\x y. x \ y = y \ x" by (metis ba2_associative_commutative ba2_complement) show "\x y z. (x \ - y = z \ - z) = (x \ y = x)" by (simp add: ba2_complement) qed end context boolean_algebra_1 begin sublocale ba1_ba2: boolean_algebra_2 proof show "\x y z. x \ y \ z = y \ z \ x" using ba1_associative commutative by force show "\x y z. (x \ - y = z \ - z) = (x \ y = x)" by (simp add: ba1_complement) qed end subsection \Meredith's Equational Axioms\ text \ The following axiomatisation is from \<^cite>\\page 221 (1) \{A,N\}\ in "MeredithPrior1968"\. \ class boolean_algebra_mp = sup + uminus + assumes ba_mp_1: "-(-x \ y) \ x = x" assumes ba_mp_2: "-(-x \ y) \ (z \ y) = y \ (z \ x)" begin subclass huntington proof show "\x y z. x \ (y \ z) = x \ y \ z" by (metis ba_mp_1 ba_mp_2) show "\x y. x \ y = y \ x" by (metis ba_mp_1 ba_mp_2) show "\x y. x = - (- x \ y) \ - (- x \ - y)" by (metis ba_mp_1 ba_mp_2) qed end context huntington begin sublocale mp_h: boolean_algebra_mp proof show 1: "\x y. - (- x \ y) \ x = x" by (metis h_ba1.ba1_associative h_ba1.ba1_complement huntington) show "\x y z. - (- x \ y) \ (z \ y) = y \ (z \ x)" proof - fix x y z have "y = -(-x \ -y) \ y" using 1 h_ba1.ba1_commutative by auto thus "-(-x \ y) \ (z \ y) = y \ (z \ x)" by (metis h_ba1.ba1_associative h_ba1.ba1_commutative huntington) qed qed end subsection \An Equational Axiomatisation based on Semilattices\ text \ The following version is an equational axiomatisation based on semilattices. We add the double complement rule and that \top\ is unique. The final axiom \ba3_export\ encodes the logical statement $P \vee Q = P \vee (\neg P \wedge Q)$. Its dual appears in \<^cite>\"BalbesHorn1970"\. \ text \Theorem 5\ class boolean_algebra_3 = sup + uminus + assumes ba3_associative: "x \ (y \ z) = (x \ y) \ z" assumes ba3_commutative: "x \ y = y \ x" assumes ba3_idempotent[simp]: "x \ x = x" assumes ba3_double_complement[simp]: "--x = x" assumes ba3_top_unique: "x \ -x = y \ -y" assumes ba3_export: "x \ -(x \ y) = x \ -y" begin subclass huntington proof show "\x y z. x \ (y \ z) = x \ y \ z" by (simp add: ba3_associative) show "\x y. x \ y = y \ x" by (simp add: ba3_commutative) show "\x y. x = - (- x \ y) \ - (- x \ - y)" by (metis ba3_commutative ba3_double_complement ba3_export ba3_idempotent ba3_top_unique) qed end context huntington begin sublocale h_ba3: boolean_algebra_3 proof show "\x y z. x \ (y \ z) = x \ y \ z" by (simp add: h_ba1.ba1_associative) show "\x y. x \ y = y \ x" by (simp add: h_ba1.ba1_commutative) show 3: "\x. x \ x = x" using h_ba1.ba1_complement by blast show 4: "\x. - - x = x" by (metis h_ba1.ba1_commutative huntington top_unique) show "\x y. x \ - x = y \ - y" by (simp add: top_unique) show "\x y. x \ - (x \ y) = x \ - y" using 3 4 by (smt h_ba1.ba1_ba2.ba2_associative_commutative h_ba1.ba1_complement) qed end section \Subset Boolean Algebras\ text \ We apply Huntington's axioms to the range of a unary operation, which serves as complement on the range. This gives a Boolean algebra structure on the range without imposing any further constraints on the set. The obtained structure is used as a reference in the subsequent development and to inherit the results proved here. This is taken from \<^cite>\"Guttmann2012c" and "GuttmannStruthWeber2011b"\ and follows the development of Boolean algebras in \<^cite>\"Maddux1996"\. \ text \Definition 6\ class subset_boolean_algebra = sup + uminus + assumes sub_associative: "-x \ (-y \ -z) = (-x \ -y) \ -z" assumes sub_commutative: "-x \ -y = -y \ -x" assumes sub_complement: "-x = -(--x \ -y) \ -(--x \ --y)" assumes sub_sup_closed: "-x \ -y = --(-x \ -y)" begin text \uniqueness of \top\, resulting in the lemma \top_def\ to replace the assumption \sub_top_def\\ lemma top_unique: "-x \ --x = -y \ --y" by (metis sub_associative sub_commutative sub_complement) text \consequences for join and complement\ lemma double_negation[simp]: "---x = -x" by (metis sub_complement sub_sup_closed) lemma complement_1: "--x = -(-x \ -y) \ -(-x \ --y)" by (metis double_negation sub_complement) lemma sup_right_zero_var: "-x \ (-y \ --y) = -z \ --z" by (smt complement_1 sub_associative sub_sup_closed top_unique) lemma sup_right_unit_idempotent: "-x \ -x = -x \ -(-y \ --y)" by (metis complement_1 double_negation sub_sup_closed sup_right_zero_var) lemma sup_idempotent[simp]: "-x \ -x = -x" by (smt complement_1 double_negation sub_associative sup_right_unit_idempotent) lemma complement_2: "-x = -(-(-x \ -y) \ -(-x \ --y))" using complement_1 by auto lemma sup_eq_cases: "-x \ -y = -x \ -z \ --x \ -y = --x \ -z \ -y = -z" by (metis complement_2 sub_commutative) lemma sup_eq_cases_2: "-y \ -x = -z \ -x \ -y \ --x = -z \ --x \ -y = -z" using sub_commutative sup_eq_cases by auto end text \Definition 7\ class subset_extended = sup + inf + minus + uminus + bot + top + ord + assumes sub_top_def: "top = (THE x . \y . x = -y \ --y)" (* define without imposing uniqueness *) assumes sub_bot_def: "bot = -(THE x . \y . x = -y \ --y)" assumes sub_inf_def: "-x \ -y = -(--x \ --y)" assumes sub_minus_def: "-x - -y = -(--x \ -y)" assumes sub_less_eq_def: "-x \ -y \ -x \ -y = -y" assumes sub_less_def: "-x < -y \ -x \ -y = -y \ \ (-y \ -x = -x)" class subset_boolean_algebra_extended = subset_boolean_algebra + subset_extended begin lemma top_def: "top = -x \ --x" using sub_top_def top_unique by blast text \consequences for meet\ lemma inf_closed: "-x \ -y = --(-x \ -y)" by (simp add: sub_inf_def) lemma inf_associative: "-x \ (-y \ -z) = (-x \ -y) \ -z" using sub_associative sub_inf_def sub_sup_closed by auto lemma inf_commutative: "-x \ -y = -y \ -x" by (simp add: sub_commutative sub_inf_def) lemma inf_idempotent[simp]: "-x \ -x = -x" by (simp add: sub_inf_def) lemma inf_absorb[simp]: "(-x \ -y) \ -x = -x" by (metis complement_1 sup_idempotent sub_inf_def sub_associative sub_sup_closed) lemma sup_absorb[simp]: "-x \ (-x \ -y) = -x" by (metis sub_associative sub_complement sub_inf_def sup_idempotent) lemma inf_demorgan: "-(-x \ -y) = --x \ --y" using sub_inf_def sub_sup_closed by auto lemma sub_sup_demorgan: "-(-x \ -y) = --x \ --y" by (simp add: sub_inf_def) lemma sup_cases: "-x = (-x \ -y) \ (-x \ --y)" by (metis inf_closed inf_demorgan sub_complement) lemma inf_cases: "-x = (-x \ -y) \ (-x \ --y)" by (metis complement_2 sub_sup_closed sub_sup_demorgan) lemma inf_complement_intro: "(-x \ -y) \ --x = -y \ --x" proof - have "(-x \ -y) \ --x = (-x \ -y) \ (--x \ -y) \ --x" by (metis inf_absorb inf_associative sub_sup_closed) also have "... = -y \ --x" by (metis inf_cases sub_commutative) finally show ?thesis . qed lemma sup_complement_intro: "-x \ -y = -x \ (--x \ -y)" by (metis inf_absorb inf_commutative inf_complement_intro sub_sup_closed sup_cases) lemma inf_left_dist_sup: "-x \ (-y \ -z) = (-x \ -y) \ (-x \ -z)" proof - have "-x \ (-y \ -z) = (-x \ (-y \ -z) \ -y) \ (-x \ (-y \ -z) \ --y)" by (metis sub_inf_def sub_sup_closed sup_cases) also have "... = (-x \ -y) \ (-x \ -z \ --y)" by (metis inf_absorb inf_associative inf_complement_intro sub_sup_closed) also have "... = (-x \ -y) \ ((-x \ -y \ -z) \ (-x \ -z \ --y))" using sub_associative sub_inf_def sup_absorb by auto also have "... = (-x \ -y) \ ((-x \ -z \ -y) \ (-x \ -z \ --y))" by (metis inf_associative inf_commutative) also have "... = (-x \ -y) \ (-x \ -z)" by (metis sub_inf_def sup_cases) finally show ?thesis . qed lemma sup_left_dist_inf: "-x \ (-y \ -z) = (-x \ -y) \ (-x \ -z)" proof - have "-x \ (-y \ -z) = -(--x \ (--y \ --z))" by (metis sub_inf_def sub_sup_closed sub_sup_demorgan) also have "... = (-x \ -y) \ (-x \ -z)" by (metis inf_left_dist_sup sub_sup_closed sub_sup_demorgan) finally show ?thesis . qed lemma sup_right_dist_inf: "(-y \ -z) \ -x = (-y \ -x) \ (-z \ -x)" using sub_commutative sub_inf_def sup_left_dist_inf by auto lemma inf_right_dist_sup: "(-y \ -z) \ -x = (-y \ -x) \ (-z \ -x)" by (metis inf_commutative inf_left_dist_sup sub_sup_closed) lemma case_duality: "(--x \ -y) \ (-x \ -z) = (-x \ -y) \ (--x \ -z)" proof - have 1: "-(--x \ --y) \ ----x = --x \ -y" using inf_commutative inf_complement_intro sub_sup_closed sub_sup_demorgan by auto have 2: "-(----x \ -(--x \ -z)) = -----x \ ---z" by (metis (no_types) double_negation sup_complement_intro sub_sup_demorgan) have 3: "-(--x \ --y) \ -x = -x" using inf_commutative inf_left_dist_sup sub_sup_closed sub_sup_demorgan by auto hence "-(--x \ --y) = -x \ -y" using sub_sup_closed sub_sup_demorgan by auto thus ?thesis by (metis double_negation 1 2 3 inf_associative inf_left_dist_sup sup_complement_intro) qed lemma case_duality_2: "(-x \ -y) \ (--x \ -z) = (-x \ -z) \ (--x \ -y)" using case_duality sub_commutative sub_inf_def by auto lemma complement_cases: "((-v \ -w) \ (--v \ -x)) \ -((-v \ -y) \ (--v \ -z)) = (-v \ -w \ --y) \ (--v \ -x \ --z)" proof - have 1: "(--v \ -w) = --(--v \ -w) \ (-v \ -x) = --(-v \ -x) \ (--v \ --y) = --(--v \ --y) \ (-v \ --z) = --(-v \ --z)" using sub_inf_def sub_sup_closed by auto have 2: "(-v \ (-x \ --z)) = --(-v \ (-x \ --z))" using sub_inf_def sub_sup_closed by auto have "((-v \ -w) \ (--v \ -x)) \ -((-v \ -y) \ (--v \ -z)) = ((-v \ -w) \ (--v \ -x)) \ (-(-v \ -y) \ -(--v \ -z))" using sub_inf_def by auto also have "... = ((-v \ -w) \ (--v \ -x)) \ ((--v \ --y) \ (-v \ --z))" using inf_demorgan by auto also have "... = (--v \ -w) \ (-v \ -x) \ ((--v \ --y) \ (-v \ --z))" by (metis case_duality double_negation) also have "... = (--v \ -w) \ ((-v \ -x) \ ((--v \ --y) \ (-v \ --z)))" by (metis 1 inf_associative sub_inf_def) also have "... = (--v \ -w) \ ((-v \ -x) \ (--v \ --y) \ (-v \ --z))" by (metis 1 inf_associative) also have "... = (--v \ -w) \ ((--v \ --y) \ (-v \ -x) \ (-v \ --z))" by (metis 1 inf_commutative) also have "... = (--v \ -w) \ ((--v \ --y) \ ((-v \ -x) \ (-v \ --z)))" by (metis 1 inf_associative) also have "... = (--v \ -w) \ ((--v \ --y) \ (-v \ (-x \ --z)))" by (simp add: sup_left_dist_inf) also have "... = (--v \ -w) \ (--v \ --y) \ (-v \ (-x \ --z))" using 1 2 by (metis inf_associative) also have "... = (--v \ (-w \ --y)) \ (-v \ (-x \ --z))" by (simp add: sup_left_dist_inf) also have "... = (-v \ (-w \ --y)) \ (--v \ (-x \ --z))" by (metis case_duality complement_1 complement_2 sub_inf_def) also have "... = (-v \ -w \ --y) \ (--v \ -x \ --z)" by (simp add: inf_associative) finally show ?thesis . qed lemma inf_cases_2: "--x = -(-x \ -y) \ -(-x \ --y)" using sub_inf_def sup_cases by auto text \consequences for \top\ and \bot\\ lemma sup_complement[simp]: "-x \ --x = top" using top_def by auto lemma inf_complement[simp]: "-x \ --x = bot" by (metis sub_bot_def sub_inf_def sub_top_def top_def) lemma complement_bot[simp]: "-bot = top" using inf_complement inf_demorgan sup_complement by fastforce lemma complement_top[simp]: "-top = bot" using sub_bot_def sub_top_def by blast lemma sup_right_zero[simp]: "-x \ top = top" using sup_right_zero_var by auto lemma sup_left_zero[simp]: "top \ -x = top" by (metis complement_bot sub_commutative sup_right_zero) lemma inf_right_unit[simp]: "-x \ bot = bot" by (metis complement_bot complement_top double_negation sub_sup_demorgan sup_right_zero) lemma inf_left_unit[simp]: "bot \ -x = bot" by (metis complement_top inf_commutative inf_right_unit) lemma sup_right_unit[simp]: "-x \ bot = -x" using sup_right_unit_idempotent by auto lemma sup_left_unit[simp]: "bot \ -x = -x" by (metis complement_top sub_commutative sup_right_unit) lemma inf_right_zero[simp]: "-x \ top = -x" by (metis inf_left_dist_sup sup_cases top_def) lemma sub_inf_left_zero[simp]: "top \ -x = -x" using inf_absorb top_def by fastforce lemma bot_double_complement[simp]: "--bot = bot" by simp lemma top_double_complement[simp]: "--top = top" by simp text \consequences for the order\ lemma reflexive: "-x \ -x" by (simp add: sub_less_eq_def) lemma transitive: "-x \ -y \ -y \ -z \ -x \ -z" by (metis sub_associative sub_less_eq_def) lemma antisymmetric: "-x \ -y \ -y \ -x \ -x = -y" by (simp add: sub_commutative sub_less_eq_def) lemma sub_bot_least: "bot \ -x" using sup_left_unit complement_top sub_less_eq_def by blast lemma top_greatest: "-x \ top" using complement_bot sub_less_eq_def sup_right_zero by blast lemma upper_bound_left: "-x \ -x \ -y" by (metis sub_associative sub_less_eq_def sub_sup_closed sup_idempotent) lemma upper_bound_right: "-y \ -x \ -y" using sub_commutative upper_bound_left by fastforce lemma sub_sup_left_isotone: assumes "-x \ -y" shows "-x \ -z \ -y \ -z" proof - have "-x \ -y = -y" by (meson assms sub_less_eq_def) thus ?thesis by (metis (full_types) sub_associative sub_commutative sub_sup_closed upper_bound_left) qed lemma sub_sup_right_isotone: "-x \ -y \ -z \ -x \ -z \ -y" by (simp add: sub_commutative sub_sup_left_isotone) lemma sup_isotone: assumes "-p \ -q" and "-r \ -s" shows "-p \ -r \ -q \ -s" proof - have "\x y. \ -x \ -y \ -r \ -x \ -y \ -s" by (metis (full_types) assms(2) sub_sup_closed sub_sup_right_isotone transitive) thus ?thesis by (metis (no_types) assms(1) sub_sup_closed sub_sup_left_isotone) qed lemma sub_complement_antitone: "-x \ -y \ --y \ --x" by (metis inf_absorb inf_demorgan sub_less_eq_def) lemma less_eq_inf: "-x \ -y \ -x \ -y = -x" by (metis inf_absorb inf_commutative sub_less_eq_def upper_bound_right sup_absorb) lemma inf_complement_left_antitone: "-x \ -y \ -(-y \ -z) \ -(-x \ -z)" by (simp add: sub_complement_antitone inf_demorgan sub_sup_left_isotone) lemma sub_inf_left_isotone: "-x \ -y \ -x \ -z \ -y \ -z" using sub_complement_antitone inf_closed inf_complement_left_antitone by fastforce lemma sub_inf_right_isotone: "-x \ -y \ -z \ -x \ -z \ -y" by (simp add: inf_commutative sub_inf_left_isotone) lemma inf_isotone: assumes "-p \ -q" and "-r \ -s" shows "-p \ -r \ -q \ -s" proof - have "\w x y z. (-w \ -x \ -y \ \ -w \ -x \ -z) \ \ -z \ -y" by (metis (no_types) inf_closed sub_inf_right_isotone transitive) thus ?thesis by (metis (no_types) assms inf_closed sub_inf_left_isotone) qed lemma least_upper_bound: "-x \ -z \ -y \ -z \ -x \ -y \ -z" by (metis sub_sup_closed transitive upper_bound_right sup_idempotent sup_isotone upper_bound_left) lemma lower_bound_left: "-x \ -y \ -x" by (metis sub_inf_def upper_bound_right sup_absorb) lemma lower_bound_right: "-x \ -y \ -y" using inf_commutative lower_bound_left by fastforce lemma greatest_lower_bound: "-x \ -y \ -x \ -z \ -x \ -y \ -z" by (metis inf_closed sub_inf_left_isotone less_eq_inf transitive lower_bound_left lower_bound_right) lemma less_eq_sup_top: "-x \ -y \ --x \ -y = top" by (metis complement_1 inf_commutative inf_complement_intro sub_inf_left_zero less_eq_inf sub_complement sup_complement_intro top_def) lemma less_eq_inf_bot: "-x \ -y \ -x \ --y = bot" by (metis complement_bot complement_top double_negation inf_demorgan less_eq_sup_top sub_inf_def) lemma shunting: "-x \ -y \ -z \ -y \ --x \ -z" proof (cases "--x \ (-z \ --y) = top") case True have "\v w. -v \ -w \ -w \ --v \ top" using less_eq_sup_top sub_commutative by blast thus ?thesis by (metis True sub_associative sub_commutative sub_inf_def sub_sup_closed) next case False hence "--x \ (-z \ --y) \ top \ \ -y \ -z \ --x" by (metis (no_types) less_eq_sup_top sub_associative sub_commutative sub_sup_closed) thus ?thesis using less_eq_sup_top sub_associative sub_commutative sub_inf_def sub_sup_closed by auto qed lemma shunting_right: "-x \ -y \ -z \ -x \ -z \ --y" by (metis inf_commutative sub_commutative shunting) lemma sup_less_eq_cases: assumes "-z \ -x \ -y" and "-z \ --x \ -y" shows "-z \ -y" proof - have "-z \ (-x \ -y) \ (--x \ -y)" by (metis assms greatest_lower_bound sub_sup_closed) also have "... = -y" by (metis inf_cases sub_commutative) finally show ?thesis . qed lemma sup_less_eq_cases_2: "-x \ -y \ -x \ -z \ --x \ -y \ --x \ -z \ -y \ -z" by (metis least_upper_bound sup_less_eq_cases sub_sup_closed) lemma sup_less_eq_cases_3: "-y \ -x \ -z \ -x \ -y \ --x \ -z \ --x \ -y \ -z" by (simp add: sup_less_eq_cases_2 sub_commutative) lemma inf_less_eq_cases: "-x \ -y \ -z \ --x \ -y \ -z \ -y \ -z" by (simp add: shunting sup_less_eq_cases) lemma inf_less_eq_cases_2: "-x \ -y \ -x \ -z \ --x \ -y \ --x \ -z \ -y \ -z" by (metis greatest_lower_bound inf_closed inf_less_eq_cases) lemma inf_less_eq_cases_3: "-y \ -x \ -z \ -x \ -y \ --x \ -z \ --x \ -y \ -z" by (simp add: inf_commutative inf_less_eq_cases_2) lemma inf_eq_cases: "-x \ -y = -x \ -z \ --x \ -y = --x \ -z \ -y = -z" by (metis inf_commutative sup_cases) lemma inf_eq_cases_2: "-y \ -x = -z \ -x \ -y \ --x = -z \ --x \ -y = -z" using inf_commutative inf_eq_cases by auto lemma wnf_lemma_1: "((-x \ -y) \ (--x \ -z)) \ -x = -x \ -y" proof - have "\u v w. (-u \ (-v \ --w)) \ -w = -u \ -w" by (metis inf_right_zero sub_associative sub_sup_closed sup_complement sup_idempotent sup_right_dist_inf) thus ?thesis by (metis (no_types) sub_associative sub_commutative sub_sup_closed sup_idempotent) qed lemma wnf_lemma_2: "((-x \ -y) \ (-z \ --y)) \ -y = -x \ -y" using sub_commutative wnf_lemma_1 by fastforce lemma wnf_lemma_3: "((-x \ -z) \ (--x \ -y)) \ --x = --x \ -y" by (metis case_duality case_duality_2 double_negation sub_commutative wnf_lemma_2) lemma wnf_lemma_4: "((-z \ -y) \ (-x \ --y)) \ --y = -x \ --y" using sub_commutative wnf_lemma_3 by auto end class subset_boolean_algebra' = sup + uminus + assumes sub_associative': "-x \ (-y \ -z) = (-x \ -y) \ -z" assumes sub_commutative': "-x \ -y = -y \ -x" assumes sub_complement': "-x = -(--x \ -y) \ -(--x \ --y)" assumes sub_sup_closed': "\z . -x \ -y = -z" begin subclass subset_boolean_algebra proof show "\x y z. - x \ (- y \ - z) = - x \ - y \ - z" by (simp add: sub_associative') show "\x y. - x \ - y = - y \ - x" by (simp add: sub_commutative') show "\x y. - x = - (- - x \ - y) \ - (- - x \ - - y)" by (simp add: sub_complement') show "\x y. - x \ - y = - - (- x \ - y)" proof - fix x y have "\x y. -y \ (-(--y \ -x) \ -(---x \ -y)) = -y \ --x" by (metis (no_types) sub_associative' sub_commutative' sub_complement') hence "\x. ---x = -x" by (metis (no_types) sub_commutative' sub_complement') thus "-x \ -y = --(-x \ -y)" by (metis sub_sup_closed') qed qed end text \ We introduce a type for the range of complement and show that it is an instance of \boolean_algebra\. \ typedef (overloaded) 'a boolean_subset = "{ x::'a::uminus . \y . x = -y }" by auto lemma simp_boolean_subset[simp]: "\y . Rep_boolean_subset x = -y" using Rep_boolean_subset by simp setup_lifting type_definition_boolean_subset text \Theorem 8.1\ instantiation boolean_subset :: (subset_boolean_algebra) huntington begin lift_definition sup_boolean_subset :: "'a boolean_subset \ 'a boolean_subset \ 'a boolean_subset" is sup using sub_sup_closed by auto lift_definition uminus_boolean_subset :: "'a boolean_subset \ 'a boolean_subset" is uminus by auto instance proof show "\x y z::'a boolean_subset. x \ (y \ z) = x \ y \ z" apply transfer using sub_associative by blast show "\x y::'a boolean_subset. x \ y = y \ x" apply transfer using sub_commutative by blast show "\x y::'a boolean_subset. x = - (- x \ y) \ - (- x \ - y)" apply transfer using sub_complement by blast qed end text \Theorem 8.2\ instantiation boolean_subset :: (subset_boolean_algebra_extended) huntington_extended begin lift_definition inf_boolean_subset :: "'a boolean_subset \ 'a boolean_subset \ 'a boolean_subset" is inf using inf_closed by auto lift_definition minus_boolean_subset :: "'a boolean_subset \ 'a boolean_subset \ 'a boolean_subset" is minus using sub_minus_def by auto lift_definition bot_boolean_subset :: "'a boolean_subset" is bot by (metis complement_top) lift_definition top_boolean_subset :: "'a boolean_subset" is top by (metis complement_bot) lift_definition less_eq_boolean_subset :: "'a boolean_subset \ 'a boolean_subset \ bool" is less_eq . lift_definition less_boolean_subset :: "'a boolean_subset \ 'a boolean_subset \ bool" is less . instance proof show 1: "top = (THE x. \y::'a boolean_subset. x = y \ - y)" proof (rule the_equality[symmetric]) show "\y::'a boolean_subset. top = y \ - y" apply transfer by auto show "\x::'a boolean_subset. \y. x = y \ - y \ x = top" apply transfer by force qed have "(bot::'a boolean_subset) = - top" apply transfer by simp thus "bot = - (THE x. \y::'a boolean_subset. x = y \ - y)" using 1 by simp show "\x y::'a boolean_subset. x \ y = - (- x \ - y)" apply transfer using sub_inf_def by blast show "\x y::'a boolean_subset. x - y = - (- x \ y)" apply transfer using sub_minus_def by blast show "\x y::'a boolean_subset. (x \ y) = (x \ y = y)" apply transfer using sub_less_eq_def by blast show "\x y::'a boolean_subset. (x < y) = (x \ y = y \ y \ x \ x)" apply transfer using sub_less_def by blast qed end section \Subset Boolean algebras with Additional Structure\ text \ We now discuss axioms that make the range of a unary operation a Boolean algebra, but add further properties that are common to the intended models. In the intended models, the unary operation can be a complement, a pseudocomplement or the antidomain operation. For simplicity, we mostly call the unary operation `complement'. We first look at structures based only on join and complement, and then add axioms for the remaining operations of Boolean algebras. In the intended models, the operation that is meet on the range of the complement can be a meet in the whole algebra or composition. \ subsection \Axioms Derived from the New Axiomatisation\ text \ The axioms of the first algebra are based on \boolean_algebra_3\. \ text \Definition 9\ class subset_boolean_algebra_1 = sup + uminus + assumes sba1_associative: "x \ (y \ z) = (x \ y) \ z" assumes sba1_commutative: "x \ y = y \ x" assumes sba1_idempotent[simp]: "x \ x = x" assumes sba1_double_complement[simp]: "---x = -x" assumes sba1_bot_unique: "-(x \ -x) = -(y \ -y)" assumes sba1_export: "-x \ -(-x \ y) = -x \ -y" begin text \Theorem 11.1\ subclass subset_boolean_algebra proof show "\x y z. - x \ (- y \ - z) = - x \ - y \ - z" by (simp add: sba1_associative) show "\x y. - x \ - y = - y \ - x" by (simp add: sba1_commutative) show "\x y. - x = - (- - x \ - y) \ - (- - x \ - - y)" by (smt sba1_bot_unique sba1_commutative sba1_double_complement sba1_export sba1_idempotent) thus "\x y. - x \ - y = - - (- x \ - y)" by (metis sba1_double_complement sba1_export) qed definition "sba1_bot \ THE x . \z . x = -(z \ -z)" lemma sba1_bot: "sba1_bot = -(z \ -z)" using sba1_bot_def sba1_bot_unique by auto end text \Boolean algebra operations based on join and complement\ text \Definition 10\ class subset_extended_1 = sup + inf + minus + uminus + bot + top + ord + assumes ba_bot: "bot = (THE x . \z . x = -(z \ -z))" assumes ba_top: "top = -(THE x . \z . x = -(z \ -z))" assumes ba_inf: "-x \ -y = -(--x \ --y)" assumes ba_minus: "-x - -y = -(--x \ -y)" assumes ba_less_eq: "x \ y \ x \ y = y" assumes ba_less: "x < y \ x \ y = y \ \ (y \ x = x)" class subset_extended_2 = subset_extended_1 + assumes ba_bot_unique: "-(x \ -x) = -(y \ -y)" begin lemma ba_bot_def: "bot = -(z \ -z)" using ba_bot ba_bot_unique by auto lemma ba_top_def: "top = --(z \ -z)" using ba_bot_def ba_top by simp end text \Subset forms Boolean Algebra, extended by Boolean algebra operations\ class subset_boolean_algebra_1_extended = subset_boolean_algebra_1 + subset_extended_1 begin subclass subset_extended_2 proof show "\x y. - (x \ - x) = - (y \ - y)" by (simp add: sba1_bot_unique) qed subclass semilattice_sup proof show "\x y. (x < y) = (x \ y \ \ y \ x)" by (simp add: ba_less ba_less_eq) show "\x. x \ x" by (simp add: ba_less_eq) show "\x y z. x \ y \ y \ z \ x \ z" by (metis sba1_associative ba_less_eq) show "\x y. x \ y \ y \ x \ x = y" by (simp add: sba1_commutative ba_less_eq) show "\x y. x \ x \ y" by (simp add: sba1_associative ba_less_eq) thus "\y x. y \ x \ y" by (simp add: sba1_commutative) show "\y x z. y \ x \ z \ x \ y \ z \ x" by (metis sba1_associative ba_less_eq) qed text \Theorem 11.2\ subclass subset_boolean_algebra_extended proof show "top = (THE x. \y. x = - y \ - - y)" by (smt ba_bot ba_bot_def ba_top sub_sup_closed the_equality) thus "bot = - (THE x. \y. x = - y \ - - y)" using ba_bot_def ba_top_def by force show "\x y. - x \ - y = - (- - x \ - - y)" by (simp add: ba_inf) show "\x y. - x - - y = - (- - x \ - y)" by (simp add: ba_minus) show "\x y. (- x \ - y) = (- x \ - y = - y)" using le_iff_sup by auto show "\x y. (- x < - y) = (- x \ - y = - y \ - y \ - x \ - x)" by (simp add: ba_less) qed end subsection \Stronger Assumptions based on Join and Complement\ text \ We add further axioms covering properties common to the antidomain and (pseudo)complement instances. \ text \Definition 12\ class subset_boolean_algebra_2 = sup + uminus + assumes sba2_associative: "x \ (y \ z) = (x \ y) \ z" assumes sba2_commutative: "x \ y = y \ x" assumes sba2_idempotent[simp]: "x \ x = x" assumes sba2_bot_unit: "x \ -(y \ -y) = x" assumes sba2_sub_sup_demorgan: "-(x \ y) = -(--x \ --y)" assumes sba2_export: "-x \ -(-x \ y) = -x \ -y" begin text \Theorem 13.1\ subclass subset_boolean_algebra_1 proof show "\x y z. x \ (y \ z) = x \ y \ z" by (simp add: sba2_associative) show "\x y. x \ y = y \ x" by (simp add: sba2_commutative) show "\x. x \ x = x" by simp show "\x. - - - x = - x" by (metis sba2_idempotent sba2_sub_sup_demorgan) show "\x y. - (x \ - x) = - (y \ - y)" by (metis sba2_bot_unit sba2_commutative) show "\x y. - x \ - (- x \ y) = - x \ - y" by (simp add: sba2_export) qed text \Theorem 13.2\ lemma double_complement_dist_sup: "--(x \ y) = --x \ --y" by (metis sba2_commutative sba2_export sba2_idempotent sba2_sub_sup_demorgan) lemma maddux_3_3[simp]: "-(x \ y) \ -(x \ -y) = -x" by (metis double_complement_dist_sup sba1_double_complement sba2_commutative sub_complement) lemma huntington_3_pp[simp]: "-(-x \ -y) \ -(-x \ y) = --x" using sba2_commutative maddux_3_3 by fastforce end class subset_boolean_algebra_2_extended = subset_boolean_algebra_2 + subset_extended_1 begin subclass subset_boolean_algebra_1_extended .. subclass bounded_semilattice_sup_bot proof show "\x. bot \ x" using sba2_bot_unit ba_bot_def sup_right_divisibility by auto qed text \Theorem 13.3\ lemma complement_antitone: "x \ y \ -y \ -x" by (metis le_iff_sup maddux_3_3 sba2_export sup_monoid.add_commute) lemma double_complement_isotone: "x \ y \ --x \ --y" by (simp add: complement_antitone) lemma sup_demorgan: "-(x \ y) = -x \ -y" using sba2_sub_sup_demorgan ba_inf by auto end subsection \Axioms for Meet\ text \ We add further axioms of \inf\ covering properties common to the antidomain and pseudocomplement instances. We omit the left distributivity rule and the right zero rule as they do not hold in some models. In particular, the operation \inf\ does not have to be commutative. \ text \Definition 14\ class subset_boolean_algebra_3_extended = subset_boolean_algebra_2_extended + assumes sba3_inf_associative: "x \ (y \ z) = (x \ y) \ z" assumes sba3_inf_right_dist_sup: "(x \ y) \ z = (x \ z) \ (y \ z)" assumes sba3_inf_complement_bot: "-x \ x = bot" assumes sba3_inf_left_unit[simp]: "top \ x = x" assumes sba3_complement_inf_double_complement: "-(x \ --y) = -(x \ y)" begin text \Theorem 15\ lemma inf_left_zero: "bot \ x = bot" by (metis inf_right_unit sba3_inf_associative sba3_inf_complement_bot) lemma inf_double_complement_export: "--(--x \ y) = --x \ --y" by (metis inf_closed sba3_complement_inf_double_complement) lemma inf_left_isotone: "x \ y \ x \ z \ y \ z" using sba3_inf_right_dist_sup sup_right_divisibility by auto lemma inf_complement_export: "--(-x \ y) = -x \ --y" by (metis inf_double_complement_export sba1_double_complement) lemma double_complement_above: "--x \ x = x" by (metis sup_monoid.add_0_right complement_bot inf_demorgan sba1_double_complement sba3_inf_complement_bot sba3_inf_right_dist_sup sba3_inf_left_unit) lemma "x \ y \ z \ x \ z \ y" nitpick [expect=genuine] oops lemma "x \ top = x" nitpick [expect=genuine] oops lemma "x \ y = y \ x" nitpick [expect=genuine] oops end subsection \Stronger Assumptions for Meet\ text \ The following axioms also hold in both models, but follow from the axioms of \subset_boolean_algebra_5_operations\. \ text \Definition 16\ class subset_boolean_algebra_4_extended = subset_boolean_algebra_3_extended + assumes sba4_inf_right_unit[simp]: "x \ top = x" assumes inf_right_isotone: "x \ y \ z \ x \ z \ y" begin lemma "x \ top = top" nitpick [expect=genuine] oops lemma "x \ bot = bot" nitpick [expect=genuine] oops lemma "x \ (y \ z) = (x \ y) \ (x \ z)" nitpick [expect=genuine] oops lemma "(x \ y = bot) = (x \ - y)" nitpick [expect=genuine] oops end section \Boolean Algebras in Stone Algebras\ text \ We specialise \inf\ to meet and complement to pseudocomplement. This puts Stone algebras into the picture; for these it is well known that regular elements form a Boolean subalgebra \<^cite>\"Graetzer1971"\. \ text \Definition 17\ class subset_boolean_algebra_5_extended = subset_boolean_algebra_3_extended + assumes sba5_inf_commutative: "x \ y = y \ x" assumes sba5_inf_absorb: "x \ (x \ y) = x" begin subclass distrib_lattice_bot proof show "\x y. x \ y \ x" by (metis sba5_inf_commutative sba3_inf_right_dist_sup sba5_inf_absorb sup_right_divisibility) show "\x y. x \ y \ y" by (metis inf_left_isotone sba5_inf_absorb sba5_inf_commutative sup_ge2) show "\x y z. x \ y \ x \ z \ x \ y \ z" by (metis inf_left_isotone sba5_inf_absorb sup.orderE sup_monoid.add_commute) show "\x y z. x \ y \ z = (x \ y) \ (x \ z) " by (metis sba3_inf_right_dist_sup sba5_inf_absorb sba5_inf_commutative sup_assoc) qed lemma inf_demorgan_2: "-(x \ y) = -x \ -y" using sba3_complement_inf_double_complement sba5_inf_commutative sub_sup_closed sub_sup_demorgan by auto lemma inf_export: "x \ -(x \ y) = x \ -y" using inf_demorgan_2 sba3_inf_complement_bot sba3_inf_right_dist_sup sba5_inf_commutative by auto lemma complement_inf[simp]: "x \ -x = bot" using sba3_inf_complement_bot sba5_inf_commutative by auto text \Theorem 18.2\ subclass stone_algebra proof show "\x. x \ top" by (simp add: inf.absorb_iff2) show "\x y. (x \ y = bot) = (x \ - y)" by (metis (full_types) complement_bot complement_inf inf.cobounded1 inf.order_iff inf_export sba3_complement_inf_double_complement sba3_inf_left_unit) show "\x. - x \ - - x = top" by simp qed text \Theorem 18.1\ subclass subset_boolean_algebra_4_extended proof show "\x. x \ top = x" by simp show "\x y z. x \ y \ z \ x \ z \ y" using inf.sup_right_isotone by blast qed end context stone_algebra_extended begin text \Theorem 18.3\ subclass subset_boolean_algebra_5_extended proof show "\x y z. x \ (y \ z) = x \ y \ z" using sup_assoc by auto show "\x y. x \ y = y \ x" by (simp add: sup_commute) show "\x. x \ x = x" by simp show "\x y. x \ - (y \ - y) = x" by simp show "\x y. - (x \ y) = - (- - x \ - - y)" by auto show "\x y. - x \ - (- x \ y) = - x \ - y" by (metis maddux_3_21_pp p_dist_sup regular_closed_p) show "bot = (THE x. \z. x = - (z \ - z))" by simp thus "top = - (THE x. \z. x = - (z \ - z))" using p_bot by blast show "\x y. - x \ - y = - (- - x \ - - y)" by simp show "\x y. - x - - y = - (- - x \ - y)" by auto show "\x y. (x \ y) = (x \ y = y)" by (simp add: le_iff_sup) thus "\x y. (x < y) = (x \ y = y \ y \ x \ x)" by (simp add: less_le_not_le) show "\x y z. x \ (y \ z) = x \ y \ z" by (simp add: inf.sup_monoid.add_assoc) show "\x y z. (x \ y) \ z = x \ z \ y \ z" by (simp add: inf_sup_distrib2) show "\x. - x \ x = bot" by simp show "\x. top \ x = x" by simp show "\x y. - (x \ - - y) = - (x \ y)" by simp show "\x y. x \ y = y \ x" by (simp add: inf_commute) show "\x y. x \ (x \ y) = x" by simp qed end section \Domain Semirings\ text \ The following development of tests in IL-semirings, prepredomain semirings, predomain semirings and domain semirings is mostly based on \<^cite>\"MoellerDesharnais2019"\; see also \<^cite>\"DesharnaisMoeller2014"\. See \<^cite>\"DesharnaisMoellerStruth2006b"\ for domain axioms in idempotent semirings. See \<^cite>\"DesharnaisJipsenStruth2009" and "JacksonStokes2004"\ for domain axioms in semigroups and monoids. Some variants have been implemented in \<^cite>\"GomesGuttmannHoefnerStruthWeber2016"\. \ subsection \Idempotent Left Semirings\ text \Definition 19\ class il_semiring = sup + inf + bot + top + ord + assumes il_associative: "x \ (y \ z) = (x \ y) \ z" assumes il_commutative: "x \ y = y \ x" assumes il_idempotent[simp]: "x \ x = x" assumes il_bot_unit: "x \ bot = x" assumes il_inf_associative: "x \ (y \ z) = (x \ y) \ z" assumes il_inf_right_dist_sup: "(x \ y) \ z = (x \ z) \ (y \ z)" assumes il_inf_left_unit[simp]: "top \ x = x" assumes il_inf_right_unit[simp]: "x \ top = x" assumes il_sub_inf_left_zero[simp]: "bot \ x = bot" assumes il_sub_inf_right_isotone: "x \ y \ z \ x \ z \ y" assumes il_less_eq: "x \ y \ x \ y = y" assumes il_less_def: "x < y \ x \ y \ \(y \ x)" begin lemma il_unit_bot: "bot \ x = x" using il_bot_unit il_commutative by fastforce subclass order proof show "\x y. (x < y) = (x \ y \ \ y \ x)" by (simp add: il_less_def) show "\x. x \ x" by (simp add: il_less_eq) show "\x y z. x \ y \ y \ z \ x \ z" by (metis il_associative il_less_eq) show "\x y. x \ y \ y \ x \ x = y" by (simp add: il_commutative il_less_eq) qed lemma il_sub_inf_right_isotone_var: "(x \ y) \ (x \ z) \ x \ (y \ z)" by (smt il_associative il_commutative il_idempotent il_less_eq il_sub_inf_right_isotone) lemma il_sub_inf_left_isotone: "x \ y \ x \ z \ y \ z" by (metis il_inf_right_dist_sup il_less_eq) lemma il_sub_inf_left_isotone_var: "(y \ x) \ (z \ x) \ (y \ z) \ x" by (simp add: il_inf_right_dist_sup) lemma sup_left_isotone: "x \ y \ x \ z \ y \ z" by (smt il_associative il_commutative il_idempotent il_less_eq) lemma sup_right_isotone: "x \ y \ z \ x \ z \ y" by (simp add: il_commutative sup_left_isotone) lemma bot_least: "bot \ x" by (simp add: il_less_eq il_unit_bot) lemma less_eq_bot: "x \ bot \ x = bot" by (simp add: il_bot_unit il_less_eq) abbreviation are_complementary :: "'a \ 'a \ bool" where "are_complementary x y \ x \ y = top \ x \ y = bot \ y \ x = bot" abbreviation test :: "'a \ bool" where "test x \ \y . are_complementary x y" definition tests :: "'a set" where "tests = { x . test x }" lemma bot_test: "test bot" by (simp add: il_unit_bot) lemma top_test: "test top" by (simp add: il_bot_unit) lemma test_sub_identity: "test x \ x \ top" using il_associative il_less_eq by auto lemma neg_unique: "are_complementary x y \ are_complementary x z \ y = z" by (metis order.antisym il_inf_left_unit il_inf_right_dist_sup il_inf_right_unit il_sub_inf_right_isotone_var) definition neg :: "'a \ 'a" ("!") where "!x \ THE y . are_complementary x y" lemma neg_char: assumes "test x" shows "are_complementary x (!x)" proof (unfold neg_def) from assms obtain y where 1: "are_complementary x y" by auto show "are_complementary x (THE y. are_complementary x y)" proof (rule theI) show "are_complementary x y" using 1 by simp show "\z. are_complementary x z \ z = y" using 1 neg_unique by blast qed qed lemma are_complementary_symmetric: "are_complementary x y \ are_complementary y x" using il_commutative by auto lemma neg_test: "test x \ test (!x)" using are_complementary_symmetric neg_char by blast lemma are_complementary_test: "test x \ are_complementary x y \ test y" using il_commutative by auto lemma neg_involutive: "test x \ !(!x) = x" using are_complementary_symmetric neg_char neg_unique by blast lemma test_inf_left_below: "test x \ x \ y \ y" by (metis il_associative il_idempotent il_inf_left_unit il_inf_right_dist_sup il_less_eq) lemma test_inf_right_below: "test x \ y \ x \ y" by (metis il_inf_right_unit il_sub_inf_right_isotone test_sub_identity) lemma neg_bot: "!bot = top" using il_unit_bot neg_char by fastforce lemma neg_top: "!top = bot" using bot_test neg_bot neg_involutive by fastforce lemma test_inf_idempotent: "test x \ x \ x = x" by (metis il_bot_unit il_inf_left_unit il_inf_right_dist_sup) lemma test_inf_semicommutative: assumes "test x" and "test y" shows "x \ y \ y \ x" proof - have "x \ y = (y \ x \ y) \ (!y \ x \ y)" by (metis assms(2) il_inf_left_unit il_inf_right_dist_sup neg_char) also have "... \ (y \ x \ y) \ (!y \ y)" proof - obtain z where "are_complementary y z" using assms(2) by blast hence "y \ (x \ y) \ !y \ (x \ y) \ y \ (x \ y)" by (metis assms(1) calculation il_sub_inf_left_isotone il_bot_unit il_idempotent il_inf_associative il_less_eq neg_char test_inf_right_below) thus ?thesis by (simp add: il_associative il_inf_associative il_less_eq) qed also have "... \ (y \ x) \ (!y \ y)" by (metis assms(2) il_bot_unit il_inf_right_unit il_sub_inf_right_isotone neg_char test_sub_identity) also have "... = y \ x" by (simp add: assms(2) il_bot_unit neg_char) finally show ?thesis . qed lemma test_inf_commutative: "test x \ test y \ x \ y = y \ x" by (simp add: order.antisym test_inf_semicommutative) lemma test_inf_bot: "test x \ x \ bot = bot" using il_inf_associative test_inf_idempotent by fastforce lemma test_absorb_1: "test x \ test y \ x \ (x \ y) = x" using il_commutative il_less_eq test_inf_right_below by auto lemma test_absorb_2: "test x \ test y \ x \ (y \ x) = x" by (metis test_absorb_1 test_inf_commutative) lemma test_absorb_3: "test x \ test y \ x \ (x \ y) = x" apply (rule order.antisym) apply (metis il_associative il_inf_right_unit il_less_eq il_sub_inf_right_isotone test_sub_identity) by (metis il_sub_inf_right_isotone_var test_absorb_1 test_inf_idempotent) lemma test_absorb_4: "test x \ test y \ (x \ y) \ x = x" by (smt il_inf_right_dist_sup test_inf_idempotent il_commutative il_less_eq test_inf_left_below) lemma test_import_1: assumes "test x" and "test y" shows "x \ (!x \ y) = x \ y" proof - have "x \ (!x \ y) = x \ ((y \ !y) \ x) \ (!x \ y)" by (simp add: assms(2) neg_char) also have "... = x \ (!y \ x) \ (x \ y) \ (!x \ y)" by (smt assms il_associative il_commutative il_inf_right_dist_sup test_inf_commutative) also have "... = x \ ((x \ !x) \ y)" by (smt calculation il_associative il_commutative il_idempotent il_inf_right_dist_sup) also have "... = x \ y" by (simp add: assms(1) neg_char) finally show ?thesis . qed lemma test_import_2: assumes "test x" and "test y" shows "x \ (y \ !x) = x \ y" proof - obtain z where 1: "are_complementary y z" - using assms(2) by moura + using assms(2) by blast obtain w where 2: "are_complementary x w" using assms(1) by auto hence "x \ !x = bot" using neg_char by blast hence "!x \ y = y \ !x" using 1 2 by (metis il_commutative neg_char test_inf_commutative) thus ?thesis using 1 2 by (metis test_import_1) qed lemma test_import_3: assumes "test x" shows "(!x \ y) \ x = y \ x" by (simp add: assms(1) il_inf_right_dist_sup il_unit_bot neg_char) lemma test_import_4: assumes "test x" and "test y" shows "(!x \ y) \ x = x \ y" by (metis assms test_import_3 test_inf_commutative) lemma test_inf: "test x \ test y \ test z \ z \ x \ y \ z \ x \ z \ y" apply (rule iffI) using dual_order.trans test_inf_left_below test_inf_right_below apply blast by (smt il_less_eq il_sub_inf_right_isotone test_absorb_4) lemma test_shunting: assumes "test x" and "test y" shows "x \ y \ z \ x \ !y \ z" proof assume 1: "x \ y \ z" have "x = (!y \ x) \ (y \ x)" by (metis assms(2) il_commutative il_inf_left_unit il_inf_right_dist_sup neg_char) also have "... \ !y \ (y \ x)" by (simp add: assms(1) sup_left_isotone test_inf_right_below) also have "... \ !y \ z" using 1 by (simp add: assms sup_right_isotone test_inf_commutative) finally show "x \ !y \ z" . next assume "x \ !y \ z" hence "x \ y \ (!y \ z) \ y" using il_sub_inf_left_isotone by blast also have "... = z \ y" by (simp add: assms(2) test_import_3) also have "... \ z" by (simp add: assms(2) test_inf_right_below) finally show "x \ y \ z" . qed lemma test_shunting_bot: assumes "test x" and "test y" shows "x \ y \ x \ !y \ bot" by (simp add: assms il_bot_unit neg_involutive neg_test test_shunting) lemma test_shunting_bot_eq: assumes "test x" and "test y" shows "x \ y \ x \ !y = bot" by (simp add: assms test_shunting_bot less_eq_bot) lemma neg_antitone: assumes "test x" and "test y" and "x \ y" shows "!y \ !x" proof - have 1: "x \ !y = bot" using assms test_shunting_bot_eq by blast have 2: "x \ !x = top" by (simp add: assms(1) neg_char) have "are_complementary y (!y)" by (simp add: assms(2) neg_char) thus ?thesis using 1 2 by (metis il_unit_bot il_commutative il_inf_left_unit il_inf_right_dist_sup il_inf_right_unit il_sub_inf_right_isotone test_sub_identity) qed lemma test_sup_neg_1: assumes "test x" and "test y" shows "(x \ y) \ (!x \ !y) = top" proof - have "x \ !x = top" by (simp add: assms(1) neg_char) hence "x \ (y \ !x) = top" by (metis assms(2) il_associative il_commutative il_idempotent) hence "x \ (y \ !x \ !y) = top" by (simp add: assms neg_test test_import_2) thus ?thesis by (simp add: il_associative) qed lemma test_sup_neg_2: assumes "test x" and "test y" shows "(x \ y) \ (!x \ !y) = bot" proof - have 1: "are_complementary y (!y)" by (simp add: assms(2) neg_char) obtain z where 2: "are_complementary x z" using assms(1) by auto hence "!x = z" using neg_char neg_unique by blast thus ?thesis using 1 2 by (metis are_complementary_symmetric il_inf_associative neg_involutive test_import_3 test_inf_bot test_inf_commutative) qed lemma de_morgan_1: assumes "test x" and "test y" and "test (x \ y)" shows "!(x \ y) = !x \ !y" proof (rule order.antisym) have 1: "test (!(x \ y))" by (simp add: assms neg_test) have "x \ (x \ y) \ !y" by (metis (full_types) assms il_commutative neg_char test_shunting test_shunting_bot_eq) hence "x \ !(x \ y) \ !y" using 1 by (simp add: assms(1,3) neg_involutive test_shunting) hence "!(x \ y) \ x \ !y" using 1 by (metis assms(1) test_inf_commutative) thus "!(x \ y) \ !x \ !y" using 1 assms(1) test_shunting by blast have 2: "!x \ !(x \ y)" by (simp add: assms neg_antitone test_inf_right_below) have "!y \ !(x \ y)" by (simp add: assms neg_antitone test_inf_left_below) thus "!x \ !y \ !(x \ y)" using 2 by (metis il_associative il_less_eq) qed lemma de_morgan_2: assumes "test x" and "test y" and "test (x \ y)" shows "!(x \ y) = !x \ !y" proof (rule order.antisym) have 1: "!(x \ y) \ !x" by (metis assms il_inf_left_unit il_sub_inf_left_isotone neg_antitone test_absorb_3 test_sub_identity) have "!(x \ y) \ !y" by (metis assms il_commutative il_inf_left_unit il_sub_inf_left_isotone neg_antitone test_absorb_3 test_sub_identity) thus "!(x \ y) \ !x \ !y" using 1 by (simp add: assms neg_test test_inf) have "top \ x \ y \ !(x \ y)" by (simp add: assms(3) neg_char) hence "top \ !x \ y \ !(x \ y)" by (smt assms(1) assms(3) il_commutative il_inf_right_dist_sup il_inf_right_unit il_sub_inf_right_isotone il_unit_bot neg_char test_sub_identity) thus "!x \ !y \ !(x \ y)" by (simp add: assms(1) assms(2) neg_involutive neg_test test_shunting) qed lemma test_inf_closed_sup_complement: assumes "test x" and "test y" and "\u v . test u \ test v \ test (u \ v)" shows "!x \ !y \ (x \ y) = bot" proof - have 1: "!(!x \ !y) = x \ y" by (simp add: assms de_morgan_1 neg_involutive neg_test) have "test (!(!x \ !y))" by (metis assms neg_test) thus ?thesis using 1 by (metis assms(1,2) de_morgan_2 neg_char) qed lemma test_sup_complement_sup_closed: assumes "test x" and "test y" and "\u v . test u \ test v \ !u \ !v \ (u \ v) = bot" shows "test (x \ y)" by (meson assms test_sup_neg_1 test_sup_neg_2) lemma test_inf_closed_sup_closed: assumes "test x" and "test y" and "\u v . test u \ test v \ test (u \ v)" shows "test (x \ y)" using assms test_inf_closed_sup_complement test_sup_complement_sup_closed by simp end subsection \Prepredomain Semirings\ class dom = fixes d :: "'a \ 'a" class ppd_semiring = il_semiring + dom + assumes d_closed: "test (d x)" assumes d1: "x \ d x \ x" begin lemma d_sub_identity: "d x \ top" using d_closed test_sub_identity by blast lemma d1_eq: "x = d x \ x" proof - have "x = (d x \ top) \ x" using d_sub_identity il_less_eq by auto thus ?thesis using d1 il_commutative il_inf_right_dist_sup il_less_eq by force qed lemma d_increasing_sub_identity: "x \ top \ x \ d x" by (metis d1_eq il_inf_right_unit il_sub_inf_right_isotone) lemma d_top: "d top = top" by (simp add: d_increasing_sub_identity d_sub_identity dual_order.antisym) lemma d_bot_only: "d x = bot \ x = bot" by (metis d1_eq il_sub_inf_left_zero) lemma d_strict: "d bot \ bot" nitpick [expect=genuine] oops lemma d_isotone_var: "d x \ d (x \ y)" nitpick [expect=genuine] oops lemma d_fully_strict: "d x = bot \ x = bot" nitpick [expect=genuine] oops lemma test_d_fixpoint: "test x \ d x = x" nitpick [expect=genuine] oops end subsection \Predomain Semirings\ class pd_semiring = ppd_semiring + assumes d2: "test p \ d (p \ x) \ p" begin lemma d_strict: "d bot \ bot" using bot_test d2 by fastforce lemma d_strict_eq: "d bot = bot" using d_strict il_bot_unit il_less_eq by auto lemma test_d_fixpoint: "test x \ d x = x" by (metis order.antisym d1_eq d2 test_inf_idempotent test_inf_right_below) lemma d_surjective: "test x \ \y . d y = x" using test_d_fixpoint by blast lemma test_d_fixpoint_iff: "test x \ d x = x" by (metis d_closed test_d_fixpoint) lemma d_surjective_iff: "test x \ (\y . d y = x)" using d_surjective d_closed by blast lemma tests_d_range: "tests = range d" using tests_def image_def d_surjective_iff by auto lemma llp: assumes "test y" shows "d x \ y \ x \ y \ x" by (metis assms d1_eq d2 order.eq_iff il_sub_inf_left_isotone test_inf_left_below) lemma gla: assumes "test y" shows "y \ !(d x) \ y \ x \ bot" proof - obtain ad where 1: "\x. are_complementary (d x) (ad x)" using d_closed by moura hence 2: "\x y. d (d y \ x) \ d y" using d2 by blast have 3: "\x. ad x \ x = bot" using 1 by (metis d1_eq il_inf_associative il_sub_inf_left_zero) have 4: "\x y. d y \ x \ ad y \ x = top \ x" using 1 by (metis il_inf_right_dist_sup) have 5: "\x y z. z \ y \ x \ y \ (z \ x) \ y \ x \ y" by (simp add: il_inf_right_dist_sup il_less_eq) have 6: "\x. !(d x) = ad x" using 1 neg_char neg_unique by blast have 7: "\x. top \ x = x" by auto hence "\x. y \ x \ !y \ x = x" by (metis assms il_inf_right_dist_sup neg_char) thus ?thesis using 1 2 3 4 5 6 7 by (metis assms d1_eq il_commutative il_less_eq test_d_fixpoint) qed lemma gla_var: "test y \ y \ d x \ bot \ y \ x \ bot" using gla d_closed il_bot_unit test_shunting by auto lemma llp_var: assumes "test y" shows "y \ !(d x) \ x \ !y \ x" apply (rule iffI) apply (metis (no_types, opaque_lifting) assms gla Least_equality il_inf_left_unit il_inf_right_dist_sup il_less_eq il_unit_bot order.refl neg_char) by (metis assms gla gla_var llp il_commutative il_sub_inf_right_isotone neg_char) lemma d_idempotent: "d (d x) = d x" using d_closed test_d_fixpoint_iff by auto lemma d_neg: "test x \ d (!x) = !x" using il_commutative neg_char test_d_fixpoint_iff by fastforce lemma d_fully_strict: "d x = bot \ x = bot" using d_strict_eq d_bot_only by blast lemma d_ad_comp: "!(d x) \ x = bot" proof - have "\x. !(d x) \ d x = bot" by (simp add: d_closed neg_char) thus ?thesis by (metis d1_eq il_inf_associative il_sub_inf_left_zero) qed lemma d_isotone: assumes "x \ y" shows "d x \ d y" proof - obtain ad where 1: "\x. are_complementary (d x) (ad x)" using d_closed by moura hence "ad y \ x \ bot" by (metis assms d1_eq il_inf_associative il_sub_inf_left_zero il_sub_inf_right_isotone) thus ?thesis using 1 by (metis d2 il_bot_unit il_inf_left_unit il_inf_right_dist_sup il_less_eq) qed lemma d_isotone_var: "d x \ d (x \ y)" using d_isotone il_associative il_less_eq by auto lemma d3_conv: "d (x \ y) \ d (x \ d y)" by (metis (mono_tags, opaque_lifting) d1_eq d2 d_closed il_inf_associative) lemma d_test_inf_idempotent: "d x \ d x = d x" by (metis d_idempotent d1_eq) lemma d_test_inf_closed: assumes "test x" and "test y" shows "d (x \ y) = x \ y" proof (rule order.antisym) have "d (x \ y) = d (x \ y) \ d (x \ y)" by (simp add: d_test_inf_idempotent) also have "... \ x \ d (x \ y)" by (simp add: assms(1) d2 il_sub_inf_left_isotone) also have "... \ x \ y" by (metis assms d_isotone il_sub_inf_right_isotone test_inf_left_below test_d_fixpoint) finally show "d (x \ y) \ x \ y" . show "x \ y \ d (x \ y)" using assms d_increasing_sub_identity dual_order.trans test_inf_left_below test_sub_identity by blast qed lemma test_inf_closed: "test x \ test y \ test (x \ y)" using d_test_inf_closed test_d_fixpoint_iff by simp lemma test_sup_closed: "test x \ test y \ test (x \ y)" using test_inf_closed test_inf_closed_sup_closed by simp lemma d_export: assumes "test x" shows "d (x \ y) = x \ d y" proof (rule order.antisym) have 1: "d (x \ y) \ x" by (simp add: assms d2) have "d (x \ y) \ d y" by (metis assms d_isotone_var il_inf_left_unit il_inf_right_dist_sup) thus "d (x \ y) \ x \ d y" using 1 by (metis assms d_idempotent llp dual_order.trans il_sub_inf_right_isotone) have "y = (!x \ y) \ (x \ y)" by (metis assms il_commutative il_inf_left_unit il_inf_right_dist_sup neg_char) also have "... = (!x \ y) \ (d (x \ y) \ x \ y)" by (metis d1_eq il_inf_associative) also have "... = (!x \ y) \ (d (x \ y) \ y)" using 1 by (smt calculation d1_eq il_associative il_commutative il_inf_associative il_inf_right_dist_sup il_less_eq il_sub_inf_right_isotone_var) also have "... = (!x \ d (x \ y)) \ y" by (simp add: il_inf_right_dist_sup) finally have "y \ (!x \ d (x \ y)) \ y" by simp hence "d y \ !x \ d (x \ y)" using assms llp test_sup_closed neg_test d_closed by simp hence "d y \ x \ d (x \ y)" by (simp add: assms d_closed test_shunting) thus "x \ d y \ d (x \ y)" by (metis assms d_closed test_inf_commutative) qed lemma test_inf_left_dist_sup: assumes "test x" and "test y" and "test z" shows "x \ (y \ z) = (x \ y) \ (x \ z)" proof - have "x \ (y \ z) = (y \ z) \ x" using assms test_sup_closed test_inf_commutative by smt also have "... = (y \ x) \ (z \ x)" using il_inf_right_dist_sup by simp also have "... = (x \ y) \ (x \ z)" using assms test_sup_closed test_inf_commutative by smt finally show ?thesis . qed lemma "!x \ !y = !(!(!x \ !y))" nitpick [expect=genuine] oops lemma "d x = !(!x)" nitpick [expect=genuine] oops sublocale subset_boolean_algebra where uminus = "\ x . !(d x)" proof show "\x y z. !(d x) \ (!(d y) \ !(d z)) = !(d x) \ !(d y) \ !(d z)" using il_associative by blast show "\x y. !(d x) \ !(d y) = !(d y) \ !(d x)" by (simp add: il_commutative) show "\x y. !(d x) \ !(d y) = !(d (!(d (!(d x) \ !(d y)))))" proof - fix x y have "test (!(d x)) \ test (!(d y))" by (simp add: d_closed neg_test) hence "test (!(d x) \ !(d y))" by (simp add: test_sup_closed) thus "!(d x) \ !(d y) = !(d (!(d (!(d x) \ !(d y)))))" by (simp add: d_neg neg_involutive test_d_fixpoint) qed show "\x y. !(d x) = !(d (!(d (!(d x))) \ !(d y))) \ !(d (!(d (!(d x))) \ !(d (!(d y)))))" proof - fix x y have "!(d (!(d (!(d x))) \ !(d y))) \ !(d (!(d (!(d x))) \ !(d (!(d y))))) = !(d x \ !(d y)) \ !(d x \ d y)" using d_closed neg_test test_sup_closed neg_involutive test_d_fixpoint by auto also have "... = (!(d x) \ d y) \ (!(d x) \ !(d y))" using d_closed neg_test test_sup_closed neg_involutive de_morgan_2 by auto also have "... = !(d x) \ (d y \ !(d y))" using d_closed neg_test test_inf_left_dist_sup by auto also have "... = !(d x) \ top" by (simp add: neg_char d_closed) finally show "!(d x) = !(d (!(d (!(d x))) \ !(d y))) \ !(d (!(d (!(d x))) \ !(d (!(d y)))))" by simp qed qed lemma d_dist_sup: "d (x \ y) = d x \ d y" proof (rule order.antisym) have "x \ d x \ x" by (simp add: d1) also have "... \ (d x \ d y) \ (x \ y)" using il_associative il_inf_right_dist_sup il_less_eq il_sub_inf_right_isotone by auto finally have 1: "x \ (d x \ d y) \ (x \ y)" . have "y \ d y \ y" by (simp add: d1) also have "... \ (d y \ d x) \ (y \ x)" using il_associative il_idempotent il_inf_right_dist_sup il_less_eq il_sub_inf_right_isotone by simp finally have "y \ (d x \ d y) \ (x \ y)" using il_commutative by auto hence "x \ y \ (d x \ d y) \ (x \ y)" using 1 by (metis il_associative il_less_eq) thus "d (x \ y) \ d x \ d y" using llp test_sup_closed neg_test d_closed by simp show "d x \ d y \ d (x \ y)" using d_isotone_var il_associative il_commutative il_less_eq by fastforce qed end class pd_semiring_extended = pd_semiring + uminus + assumes uminus_def: "-x = !(d x)" begin subclass subset_boolean_algebra by (metis subset_boolean_algebra_axioms uminus_def ext) end subsection \Domain Semirings\ class d_semiring = pd_semiring + assumes d3: "d (x \ d y) \ d (x \ y)" begin lemma d3_eq: "d (x \ d y) = d (x \ y)" by (simp add: order.antisym d3 d3_conv) end text \ Axioms (d1), (d2) and (d3) are independent in IL-semirings. \ context il_semiring begin context fixes d :: "'a \ 'a" assumes d_closed: "test (d x)" begin context assumes d1: "x \ d x \ x" assumes d2: "test p \ d (p \ x) \ p" begin lemma d3: "d (x \ d y) \ d (x \ y)" nitpick [expect=genuine] oops end context assumes d1: "x \ d x \ x" assumes d3: "d (x \ d y) \ d (x \ y)" begin lemma d2: "test p \ d (p \ x) \ p" nitpick [expect=genuine] oops end context assumes d2: "test p \ d (p \ x) \ p" assumes d3: "d (x \ d y) \ d (x \ y)" begin lemma d1: "x \ d x \ x" nitpick [expect=genuine] oops end end end class d_semiring_var = ppd_semiring + assumes d3_var: "d (x \ d y) \ d (x \ y)" assumes d_strict_eq_var: "d bot = bot" begin lemma d2_var: assumes "test p" shows "d (p \ x) \ p" proof - have "!p \ p \ x = bot" by (simp add: assms neg_char) hence "d (!p \ p \ x) = bot" by (simp add: d_strict_eq_var) hence "d (!p \ d (p \ x)) = bot" by (metis d3_var il_inf_associative less_eq_bot) hence "!p \ d (p \ x) = bot" using d_bot_only by blast thus ?thesis by (metis (no_types, opaque_lifting) assms d_sub_identity il_bot_unit il_inf_left_unit il_inf_right_dist_sup il_inf_right_unit il_sub_inf_right_isotone neg_char) qed subclass d_semiring proof show "\p x. test p \ d (p \ x) \ p" by (simp add: d2_var) show "\x y. d (x \ d y) \ d (x \ y)" by (simp add: d3_var) qed end section \Antidomain Semirings\ text \ We now develop prepreantidomain semirings, preantidomain semirings and antidomain semirings. See \<^cite>\"DesharnaisStruth2008b" and "DesharnaisStruth2008a" and "DesharnaisStruth2011"\ for related work on internal axioms for antidomain. \ subsection \Prepreantidomain Semirings\ text \Definition 20\ class ppa_semiring = il_semiring + uminus + assumes a_inf_complement_bot: "-x \ x = bot" assumes a_stone[simp]: "-x \ --x = top" begin text \Theorem 21\ lemma l1: "-top = bot" by (metis a_inf_complement_bot il_inf_right_unit) lemma l2: "-bot = top" by (metis l1 a_stone il_unit_bot) lemma l3: "-x \ -y \ -x \ y = bot" by (metis a_inf_complement_bot il_bot_unit il_inf_right_dist_sup il_less_eq) lemma l5: "--x \ --y \ -y \ -x" by (metis (mono_tags, opaque_lifting) l3 a_stone bot_least il_bot_unit il_inf_left_unit il_inf_right_dist_sup il_inf_right_unit il_sub_inf_right_isotone sup_right_isotone) lemma l4: "---x = -x" by (metis l5 a_inf_complement_bot a_stone order.antisym bot_least il_inf_left_unit il_inf_right_dist_sup il_inf_right_unit il_sub_inf_right_isotone il_unit_bot) lemma l6: "-x \ --x = bot" by (metis l3 l5 a_inf_complement_bot a_stone il_inf_left_unit il_inf_right_dist_sup il_inf_right_unit il_less_eq il_sub_inf_right_isotone il_unit_bot) lemma l7: "-x \ -y = -y \ -x" using l6 a_inf_complement_bot a_stone test_inf_commutative by blast lemma l8: "x \ --x \ x" by (metis a_inf_complement_bot a_stone il_idempotent il_inf_left_unit il_inf_right_dist_sup il_less_eq il_unit_bot) sublocale ppa_ppd: ppd_semiring where d = "\x . --x" proof show "\x. test (- - x)" using l4 l6 by force show "\x. x \ - - x \ x" by (simp add: l8) qed (* The following statements have counterexamples, but they take a while to find. lemma "- x = - (- - x \ - y) \ - (- - x \ - - y)" nitpick [card=8, expect=genuine] oops lemma "- x \ - y = - - (- x \ - y)" nitpick [card=8, expect=genuine] oops *) end subsection \Preantidomain Semirings\ text \Definition 22\ class pa_semiring = ppa_semiring + assumes pad2: "--x \ -(-x \ y)" begin text \Theorem 23\ lemma l10: "-x \ y = bot \ -x \ -y" by (metis a_stone il_inf_left_unit il_inf_right_dist_sup il_unit_bot l4 pad2) lemma l10_iff: "-x \ y = bot \ -x \ -y" using l10 l3 by blast lemma l13: "--(--x \ y) \ --x" by (metis l4 l5 pad2) lemma l14: "-(x \ --y) \ -(x \ y)" by (metis il_inf_associative l4 pad2 ppa_ppd.d1_eq) lemma l9: "x \ y \ -y \ -x" by (metis l10 a_inf_complement_bot il_commutative il_less_eq il_sub_inf_right_isotone il_unit_bot) lemma l11: "- x \ - y = - (- - x \ - - y)" proof - have 1: "\x y . x \ y \ x \ y = y" by (simp add: il_less_eq) have 4: "\x y . \(x \ y) \ x \ y = y" using 1 by metis have 5: "\x y z . (x \ y) \ (x \ z) \ x \ (y \ z)" by (simp add: il_sub_inf_right_isotone_var) have 6: "\x y . - - x \ - (- x \ y)" by (simp add: pad2) have 7: "\x y z . x \ (y \ z) = (x \ y) \ z" by (simp add: il_associative) have 8: "\x y z . (x \ y) \ z = x \ (y \ z)" using 7 by metis have 9: "\x y . x \ y = y \ x" by (simp add: il_commutative) have 10: "\x . x \ bot = x" by (simp add: il_bot_unit) have 11: "\x . x \ x = x" by simp have 12: "\x y z . x \ (y \ z) = (x \ y) \ z" by (simp add: il_inf_associative) have 13: "\x y z . (x \ y) \ z = x \ (y \ z)" using 12 by metis have 14: "\x . top \ x = x" by simp have 15: "\x . x \ top = x" by simp have 16: "\x y z . (x \ y) \ z = (x \ z) \ (y \ z)" by (simp add: il_inf_right_dist_sup) have 17: "\x y z . (x \ y) \ (z \ y) = (x \ z) \ y" using 16 by metis have 18: "\x . bot \ x = bot" by simp have 19: "\x . - x \ - - x = top" by simp have 20: "\x . - x \ x = bot" by (simp add: a_inf_complement_bot) have 23: "\x y z . ((x \ y) \ (x \ z)) \ (x \ (y \ z)) = x \ (y \ z)" using 4 5 by metis have 24: "\x y z . (x \ (y \ z)) \ ((x \ y) \ (x \ z)) = x \ (y \ z)" using 9 23 by metis have 25: "\x y . - - x \ - (- x \ y) = - (- x \ y)" using 4 6 by metis have 26: "\x y z . x \ (y \ z) = y \ (x \ z)" using 8 9 by metis have 27: "\x y z . (x \ y) \ ((x \ z) \ (x \ (y \ z))) = x \ (y \ z)" using 9 24 26 by metis have 30: "\x . bot \ x = x" using 9 10 by metis have 31: "\x y . x \ (x \ y) = x \ y" using 8 11 by metis have 34: "\u x y z . ((x \ y) \ z) \ u = (x \ z) \ ((y \ z) \ u)" using 8 17 by metis have 35: "\u x y z . (x \ (y \ z)) \ (u \ z) = ((x \ y) \ u) \ z" using 13 17 by metis have 36: "\u x y z . (x \ y) \ (z \ (u \ y)) = (x \ (z \ u)) \ y" using 13 17 by metis have 39: "\x y . - x \ (- - x \ y) = top \ y" using 8 19 by metis have 41: "\x y . - x \ (x \ y) = bot" using 13 18 20 by metis have 42: "- top = bot" using 15 20 by metis have 43: "\x y . (- x \ y) \ x = y \ x" using 17 20 30 by metis have 44: "\x y . (x \ - y) \ y = x \ y" using 9 17 20 30 by metis have 46: "\x . - bot \ - - x = - bot" using 9 20 25 by metis have 50: "- bot = top" using 19 30 42 by metis have 51: "\x . top \ - - x = top" using 46 50 by metis have 63: "\x y . x \ ((x \ - y) \ (x \ - - y)) = x" using 9 15 19 26 27 by metis have 66: "\x y . (- (x \ y) \ x) \ (- (x \ y) \ y) = bot" using 9 20 27 30 by metis have 67: "\x y z . (x \ - - y) \ (x \ - (- y \ z)) = x \ - (- y \ z)" using 11 25 27 by metis have 70: "\x y . x \ (x \ - - y) = x" using 9 15 27 31 51 by metis have 82: "\x . top \ - x = top" using 9 19 31 by metis have 89: "\x y . x \ (- y \ x) = x" using 14 17 82 by metis have 102: "\x y z . x \ (y \ (x \ - - z)) = y \ x" using 26 70 by metis have 104: "\x y . x \ (x \ - y) = x" using 9 63 102 by metis have 112: "\x y z . (- x \ y) \ ((- - x \ y) \ z) = y \ z" using 14 19 34 by metis have 117: "\x y z . x \ ((x \ - y) \ z) = x \ z" using 8 104 by metis have 120: "\x y z . x \ (y \ (x \ - z)) = y \ x" using 26 104 by metis have 124: "\x . - - x \ x = x" using 14 19 43 by metis have 128: "\x y . - - x \ (x \ y) = x \ y" using 13 124 by metis have 131: "\x . - x \ - - - x = - x" using 9 25 124 by metis have 133: "\x . - - - x = - x" using 9 104 124 131 by metis have 135: "\x y . - x \ - (- - x \ y) = - (- - x \ y)" using 25 133 by metis have 137: "\x y . (- x \ y) \ - - x = y \ - - x" using 43 133 by metis have 145: "\x y z . ((- (x \ y) \ x) \ z) \ y = z \ y" using 20 30 35 by metis have 183: "\x y z . (x \ (- - (y \ z) \ y)) \ z = (x \ y) \ z" using 17 36 124 by metis have 289: "\x y . - x \ - (- x \ y) = top" using 25 39 82 by metis have 316: "\x y . - (- x \ y) \ x = x" using 14 43 289 by metis have 317: "\x y z . - (- x \ y) \ (x \ z) = x \ z" using 13 316 by metis have 320: "\x y . - x \ - - (- x \ y) = - x" using 9 25 316 by metis have 321: "\x y . - - (- x \ y) \ x = bot" using 41 316 by metis have 374: "\x y . - x \ - (x \ y) = - (x \ y)" using 25 128 133 by metis have 388: "\x y . - (x \ y) \ - x = - x" using 128 316 by metis have 389: "\x y . - - (x \ y) \ - x = bot" using 128 321 by metis have 405: "\x y z . - (x \ y) \ (- x \ z) = - x \ z" using 13 388 by metis have 406: "\x y z . - (x \ (y \ z)) \ - (x \ y) = - (x \ y)" using 13 388 by metis have 420: "\x y . - x \ - - (- x \ y) = - - (- x \ y)" using 316 388 by metis have 422: "\x y z . - - (x \ y) \ (- x \ z) = bot" using 13 18 389 by metis have 758: "\x y z . x \ (x \ (- y \ - z)) = x" using 13 104 117 by metis have 1092: "\x y . - (x \ y) \ x = bot" using 9 30 31 66 by metis have 1130: "\x y z . (- (x \ y) \ z) \ x = z \ x" using 17 30 1092 by metis have 1156: "\x y . - - x \ - (- x \ y) = - - x" using 67 104 124 133 by metis have 2098: "\x y . - - (x \ y) \ x = x" using 14 19 1130 by metis have 2125: "\x y . - - (x \ y) \ y = y" using 9 2098 by metis have 2138: "\x y . - x \ - - (x \ y) = top" using 9 289 2098 by metis have 2139: "\x y . - x \ - (x \ y) = - (x \ y)" using 316 2098 by metis have 2192: "\x y . - - x \ (- y \ x) = - y \ x" using 89 2125 by metis have 2202: "\x y . - x \ - - (y \ x) = top" using 9 289 2125 by metis have 2344: "\x y . - (- x \ y) \ - - y = top" using 89 2202 by metis have 2547: "\x y z . - x \ ((- - x \ - y) \ z) = - x \ (- y \ z)" using 112 117 by metis have 3023: "\x y . - x \ - (- y \ - x) = top" using 9 133 2344 by metis have 3134: "\x y . - (- x \ - y) \ y = y" using 14 43 3023 by metis have 3135: "\x y . - x \ (- y \ - x) = - y \ - x" using 14 44 3023 by metis have 3962: "\x y . - - (x \ y) \ - - x = - - x" using 14 137 2138 by metis have 5496: "\x y z . - - (x \ y) \ - (x \ z) = bot" using 422 2139 by metis have 9414: "\x y . - - (- x \ y) \ y = - x \ y" using 9 104 183 320 by metis have 9520: "\x y z . - - (- x \ y) \ - - (x \ z) = bot" using 374 5496 by metis have 11070: "\x y z . - (- - x \ y) \ (- x \ - z) = - (- - x \ y)" using 317 758 by metis have 12371: "\x y . - x \ - (- - x \ y) = - x" using 133 1156 by metis have 12377: "\x y . - x \ - (x \ y) = - x" using 128 133 1156 by metis have 12384: "\x y . - (x \ y) \ - y = - (x \ y)" using 133 1156 2125 by metis have 12394: "\x y . - - (- x \ - y) = - x \ - y" using 1156 3134 9414 by metis have 12640: "\x y . - x \ - (- y \ x) = - x" using 89 12384 by metis have 24648: "\x y . (- x \ - y) \ - (- x \ - y) = top" using 19 12394 by metis have 28270: "\x y z . - - (x \ y) \ - (- x \ z) = - (- x \ z)" using 374 405 by metis have 28339: "\x y . - (- - (x \ y) \ x) = - (x \ y)" using 124 406 12371 by metis have 28423: "\x y . - (- x \ - y) = - (- y \ - x)" using 13 3135 12394 28339 by metis have 28487: "\x y . - x \ - y = - y \ - x" using 2098 3962 12394 28423 by metis have 52423: "\x y . - (- x \ - (- x \ y)) \ y = y" using 14 145 24648 28487 by metis have 52522: "\x y . - x \ - (- x \ y) = - x \ - y" using 13 12377 12394 12640 28487 52423 by metis have 61103: "\x y z . - (- - x \ y) \ z = - x \ (- y \ z)" using 112 2547 12371 52522 by metis have 61158: "\x y . - - (- x \ y) = - x \ - - y" using 420 52522 by metis have 61231: "\x y z . - x \ (- - y \ - (x \ z)) = - x \ - - y" using 13 15 50 133 9520 52522 61158 by metis have 61313: "\x y . - x \ - y = - (- - y \ x)" using 120 11070 61103 by metis have 61393: "\x y . - (- x \ - - y) = - (- x \ y)" using 13 28270 61158 61231 61313 by metis have 61422: "\x y . - (- - x \ y) = - (- - y \ x)" using 13 135 2192 61158 61313 by metis show ?thesis using 61313 61393 61422 by metis qed lemma l12: "- x \ - y = - (x \ y)" proof - have 1: "\x y . x \ y \ x \ y = y" by (simp add: il_less_eq) have 4: "\x y . \(x \ y) \ x \ y = y" using 1 by metis have 5: "\x y z . (x \ y) \ (x \ z) \ x \ (y \ z)" by (simp add: il_sub_inf_right_isotone_var) have 6: "\x y . - - x \ - (- x \ y)" by (simp add: pad2) have 7: "\x y z . x \ (y \ z) = (x \ y) \ z" by (simp add: il_associative) have 8: "\x y z . (x \ y) \ z = x \ (y \ z)" using 7 by metis have 9: "\x y . x \ y = y \ x" by (simp add: il_commutative) have 10: "\x . x \ bot = x" by (simp add: il_bot_unit) have 11: "\x . x \ x = x" by simp have 12: "\x y z . x \ (y \ z) = (x \ y) \ z" by (simp add: il_inf_associative) have 13: "\x y z . (x \ y) \ z = x \ (y \ z)" using 12 by metis have 14: "\x . top \ x = x" by simp have 15: "\x . x \ top = x" by simp have 16: "\x y z . (x \ y) \ z = (x \ z) \ (y \ z)" by (simp add: il_inf_right_dist_sup) have 17: "\x y z . (x \ y) \ (z \ y) = (x \ z) \ y" using 16 by metis have 18: "\x . bot \ x = bot" by simp have 19: "\x . - x \ - - x = top" by simp have 20: "\x . - x \ x = bot" by (simp add: a_inf_complement_bot) have 22: "\x y z . ((x \ y) \ (x \ z)) \ (x \ (y \ z)) = x \ (y \ z)" using 4 5 by metis have 23: "\x y z . (x \ (y \ z)) \ ((x \ y) \ (x \ z)) = x \ (y \ z)" using 9 22 by metis have 24: "\x y . - - x \ - (- x \ y) = - (- x \ y)" using 4 6 by metis have 25: "\x y z . x \ (y \ z) = y \ (x \ z)" using 8 9 by metis have 26: "\x y z . (x \ y) \ ((x \ z) \ (x \ (y \ z))) = x \ (y \ z)" using 9 23 25 by metis have 29: "\x . bot \ x = x" using 9 10 by metis have 30: "\x y . x \ (x \ y) = x \ y" using 8 11 by metis have 32: "\x y . x \ (y \ x) = y \ x" using 8 9 11 by metis have 33: "\u x y z . ((x \ y) \ z) \ u = (x \ z) \ ((y \ z) \ u)" using 8 17 by metis have 34: "\u x y z . (x \ (y \ z)) \ (u \ z) = ((x \ y) \ u) \ z" using 13 17 by metis have 35: "\u x y z . (x \ y) \ (z \ (u \ y)) = (x \ (z \ u)) \ y" using 13 17 by metis have 36: "\x y . (top \ x) \ y = y \ (x \ y)" using 14 17 by metis have 37: "\x y . (x \ top) \ y = y \ (x \ y)" using 9 14 17 by metis have 38: "\x y . - x \ (- - x \ y) = top \ y" using 8 19 by metis have 40: "\x y . - x \ (x \ y) = bot" using 13 18 20 by metis have 41: "- top = bot" using 15 20 by metis have 42: "\x y . (- x \ y) \ x = y \ x" using 17 20 29 by metis have 43: "\x y . (x \ - y) \ y = x \ y" using 9 17 20 29 by metis have 45: "\x . - bot \ - - x = - bot" using 9 20 24 by metis have 46: "\u x y z . (x \ y) \ (z \ (u \ y)) = z \ ((x \ u) \ y)" using 17 25 by metis have 47: "\x y . - x \ (y \ - - x) = y \ top" using 19 25 by metis have 49: "- bot = top" using 19 29 41 by metis have 50: "\x . top \ - - x = top" using 45 49 by metis have 54: "\u x y z . (x \ y) \ ((x \ z) \ ((x \ (y \ z)) \ u)) = (x \ (y \ z)) \ u" using 8 26 by metis have 58: "\u x y z . (x \ (y \ z)) \ ((x \ (y \ u)) \ (x \ (y \ (z \ u)))) = x \ (y \ (z \ u))" using 13 26 by metis have 60: "\x y . x \ ((x \ y) \ (x \ (y \ top))) = x \ (y \ top)" using 15 25 26 by metis have 62: "\x y . x \ ((x \ - y) \ (x \ - - y)) = x" using 9 15 19 25 26 by metis have 65: "\x y . (- (x \ y) \ x) \ (- (x \ y) \ y) = bot" using 9 20 26 29 by metis have 66: "\x y z . (x \ - - y) \ (x \ - (- y \ z)) = x \ - (- y \ z)" using 11 24 26 by metis have 69: "\x y . x \ (x \ - - y) = x" using 9 15 26 30 50 by metis have 81: "\x . top \ - x = top" using 9 19 30 by metis have 82: "\x y z . (x \ y) \ (x \ (y \ z)) = x \ (y \ z)" using 11 26 30 by metis have 83: "\x y . x \ (x \ (y \ top)) = x \ (y \ top)" using 60 82 by metis have 88: "\x y . x \ (- y \ x) = x" using 14 17 81 by metis have 89: "\x y . top \ (x \ - y) = x \ top" using 25 81 by metis have 91: "\x y z . x \ (y \ (z \ x)) = y \ (z \ x)" using 8 32 by metis have 94: "\x y z . x \ (y \ (- z \ x)) = y \ x" using 25 88 by metis have 101: "\x y z . x \ (y \ (x \ - - z)) = y \ x" using 25 69 by metis have 102: "\x . x \ (x \ bot) = x" using 41 49 69 by metis have 103: "\x y . x \ (x \ - y) = x" using 9 62 101 by metis have 109: "\x y . x \ (y \ (x \ bot)) = y \ x" using 25 102 by metis have 111: "\x y z . (- x \ y) \ ((- - x \ y) \ z) = y \ z" using 14 19 33 by metis have 116: "\x y z . x \ ((x \ - y) \ z) = x \ z" using 8 103 by metis have 119: "\x y z . x \ (y \ (x \ - z)) = y \ x" using 25 103 by metis have 123: "\x . - - x \ x = x" using 14 19 42 by metis have 127: "\x y . - - x \ (x \ y) = x \ y" using 13 123 by metis have 130: "\x . - x \ - - - x = - x" using 9 24 123 by metis have 132: "\x . - - - x = - x" using 9 103 123 130 by metis have 134: "\x y . - x \ - (- - x \ y) = - (- - x \ y)" using 24 132 by metis have 136: "\x y . (- x \ y) \ - - x = y \ - - x" using 42 132 by metis have 138: "\x . - x \ - x = - x" using 123 132 by metis have 144: "\x y z . ((- (x \ y) \ x) \ z) \ y = z \ y" using 20 29 34 by metis have 157: "\x y . (- x \ y) \ - x = (top \ y) \ - x" using 17 36 138 by metis have 182: "\x y z . (x \ (- - (y \ z) \ y)) \ z = (x \ y) \ z" using 17 35 123 by metis have 288: "\x y . - x \ - (- x \ y) = top" using 24 38 81 by metis have 315: "\x y . - (- x \ y) \ x = x" using 14 42 288 by metis have 316: "\x y z . - (- x \ y) \ (x \ z) = x \ z" using 13 315 by metis have 319: "\x y . - x \ - - (- x \ y) = - x" using 9 24 315 by metis have 320: "\x y . - - (- x \ y) \ x = bot" using 40 315 by metis have 373: "\x y . - x \ - (x \ y) = - (x \ y)" using 24 127 132 by metis have 387: "\x y . - (x \ y) \ - x = - x" using 127 315 by metis have 388: "\x y . - - (x \ y) \ - x = bot" using 127 320 by metis have 404: "\x y z . - (x \ y) \ (- x \ z) = - x \ z" using 13 387 by metis have 405: "\x y z . - (x \ (y \ z)) \ - (x \ y) = - (x \ y)" using 13 387 by metis have 419: "\x y . - x \ - - (- x \ y) = - - (- x \ y)" using 315 387 by metis have 420: "\x y . - - x \ - - (x \ y) = - - (x \ y)" using 387 by metis have 421: "\x y z . - - (x \ y) \ (- x \ z) = bot" using 13 18 388 by metis have 536: "\x y . (x \ - - y) \ y = (x \ top) \ y" using 42 47 by metis have 662: "\u x y z . (x \ y) \ ((x \ (z \ y)) \ u) = (x \ (z \ y)) \ u" using 9 32 54 by metis have 705: "\u x y z . (x \ (y \ z)) \ ((x \ (y \ (z \ bot))) \ u) = (x \ (y \ z)) \ u" using 25 54 109 662 by metis have 755: "\x y z . (x \ - y) \ (z \ x) = z \ x" using 32 91 116 by metis have 757: "\x y z . x \ (x \ (- y \ - z)) = x" using 13 103 116 by metis have 930: "\x y z . (- (x \ (y \ z)) \ (x \ y)) \ (- (x \ (y \ z)) \ (x \ z)) = bot" using 9 20 29 58 by metis have 1091: "\x y . - (x \ y) \ x = bot" using 9 29 30 65 by metis have 1092: "\x y . - (x \ y) \ y = bot" using 29 30 65 1091 by metis have 1113: "\u x y z . - (x \ ((y \ z) \ u)) \ (x \ (z \ u)) = bot" using 29 46 65 1091 by metis have 1117: "\x y z . - (x \ y) \ (x \ (- z \ y)) = bot" using 29 65 94 1092 by metis have 1128: "\x y z . - (x \ (y \ z)) \ (x \ y) = bot" using 8 1091 by metis have 1129: "\x y z . (- (x \ y) \ z) \ x = z \ x" using 17 29 1091 by metis have 1155: "\x y . - - x \ - (- x \ y) = - - x" using 66 103 123 132 by metis have 1578: "\x y z . - (x \ (y \ z)) \ (x \ y) = bot" using 82 1091 by metis have 1594: "\x y z . - (x \ (y \ z)) \ (x \ z) = bot" using 29 930 1578 by metis have 2094: "\x y z . - (x \ (y \ (z \ top))) \ (x \ y) = bot" using 83 1128 by metis have 2097: "\x y . - - (x \ y) \ x = x" using 14 19 1129 by metis have 2124: "\x y . - - (x \ y) \ y = y" using 9 2097 by metis have 2135: "\x y . - - ((top \ x) \ y) \ y = y" using 36 2097 by metis have 2136: "\x y . - - ((x \ top) \ y) \ y = y" using 37 2097 by metis have 2137: "\x y . - x \ - - (x \ y) = top" using 9 288 2097 by metis have 2138: "\x y . - x \ - (x \ y) = - (x \ y)" using 315 2097 by metis have 2151: "\x y . - x \ - (x \ y) = - x" using 9 132 373 2097 by metis have 2191: "\x y . - - x \ (- y \ x) = - y \ x" using 88 2124 by metis have 2201: "\x y . - x \ - - (y \ x) = top" using 9 288 2124 by metis have 2202: "\x y . - x \ - (y \ x) = - (y \ x)" using 315 2124 by metis have 2320: "\x y . - (x \ (y \ top)) = - x" using 83 373 2151 by metis have 2343: "\x y . - (- x \ y) \ - - y = top" using 88 2201 by metis have 2546: "\x y z . - x \ ((- - x \ - y) \ z) = - x \ (- y \ z)" using 111 116 by metis have 2706: "\x y z . - x \ (y \ - - ((top \ z) \ - x)) = y \ - - ((top \ z) \ - x)" using 755 2135 by metis have 2810: "\x y . - x \ - ((y \ top) \ x) = - ((y \ top) \ x)" using 315 2136 by metis have 3022: "\x y . - x \ - (- y \ - x) = top" using 9 132 2343 by metis have 3133: "\x y . - (- x \ - y) \ y = y" using 14 42 3022 by metis have 3134: "\x y . - x \ (- y \ - x) = - y \ - x" using 14 43 3022 by metis have 3961: "\x y . - - (x \ y) \ - - x = - - x" using 14 136 2137 by metis have 4644: "\x y z . - (x \ - y) \ (x \ - (y \ z)) = bot" using 1594 2151 by metis have 5495: "\x y z . - - (x \ y) \ - (x \ z) = bot" using 421 2138 by metis have 9413: "\x y . - - (- x \ y) \ y = - x \ y" using 9 103 182 319 by metis have 9519: "\x y z . - - (- x \ y) \ - - (x \ z) = bot" using 373 5495 by metis have 11069: "\x y z . - (- - x \ y) \ (- x \ - z) = - (- - x \ y)" using 316 757 by metis have 12370: "\x y . - x \ - (- - x \ y) = - x" using 132 1155 by metis have 12376: "\x y . - x \ - (x \ y) = - x" using 127 132 1155 by metis have 12383: "\x y . - (x \ y) \ - y = - (x \ y)" using 132 1155 2124 by metis have 12393: "\x y . - - (- x \ - y) = - x \ - y" using 1155 3133 9413 by metis have 12407: "\x y . - - x \ - - (x \ y) = - - x" using 1155 2138 by metis have 12639: "\x y . - x \ - (- y \ x) = - x" using 88 12383 by metis have 24647: "\x y . (- x \ - y) \ - (- x \ - y) = top" using 19 12393 by metis have 28269: "\x y z . - - (x \ y) \ - (- x \ z) = - (- x \ z)" using 373 404 by metis have 28338: "\x y . - (- - (x \ y) \ x) = - (x \ y)" using 123 405 12370 by metis have 28422: "\x y . - (- x \ - y) = - (- y \ - x)" using 13 3134 12393 28338 by metis have 28485: "\x y . - x \ - y = - y \ - x" using 2097 3961 12393 28422 by metis have 30411: "\x y . - x \ (x \ (x \ y)) = bot" using 9 82 2094 2320 by metis have 30469: "\x . - x \ (x \ - - x) = bot" using 9 123 132 30411 by metis have 37513: "\x y . - (- x \ - y) \ - (y \ x) = bot" using 2202 4644 by metis have 52421: "\x y . - (- x \ - (- x \ y)) \ y = y" using 14 144 24647 28485 by metis have 52520: "\x y . - x \ - (- x \ y) = - x \ - y" using 13 12376 12393 12639 28485 52421 by metis have 52533: "\x y z . - - (x \ (y \ (z \ top))) \ (x \ y) = x \ y" using 15 49 2094 52421 by metis have 61101: "\x y z . - (- - x \ y) \ z = - x \ (- y \ z)" using 111 2546 12370 52520 by metis have 61156: "\x y . - - (- x \ y) = - x \ - - y" using 419 52520 by metis have 61162: "\x y . - (x \ (x \ y)) = - x" using 15 49 2138 30411 52520 by metis have 61163: "\x . - (x \ - - x) = - x" using 15 49 2138 30469 52520 by metis have 61229: "\x y z . - x \ (- - y \ - (x \ z)) = - x \ - - y" using 13 15 49 132 9519 52520 61156 by metis have 61311: "\x y . - x \ - y = - (- - y \ x)" using 119 11069 61101 by metis have 61391: "\x y . - (- x \ - - y) = - (- x \ y)" using 13 28269 61156 61229 61311 by metis have 61420: "\x y . - (- - x \ y) = - (- - y \ x)" using 13 134 2191 61156 61311 by metis have 61454: "\x y . - (x \ - (- y \ - x)) = - y \ - x" using 9 132 3133 61156 61162 by metis have 61648: "\x y . - x \ (x \ (- y \ - - x)) = bot" using 1117 61163 by metis have 62434: "\x y . - (- - x \ y) \ x = - y \ x" using 43 61311 by metis have 63947: "\x y . - (- x \ y) \ - (- y \ x) = bot" using 37513 61391 by metis have 64227: "\x y . - (x \ (- y \ - - x)) = - x" using 15 49 2138 52520 61648 by metis have 64239: "\x y . - (x \ (- - x \ y)) = - (x \ y)" using 9 25 12407 64227 by metis have 64241: "\x y . - (x \ (- - x \ - y)) = - x" using 28485 64227 by metis have 64260: "\x y . - (x \ - - (x \ y)) = - x" using 420 64241 by metis have 64271: "\x y . - (- x \ (y \ - - (y \ x))) = - (- x \ y)" using 9 25 42 64260 by metis have 64281: "\x y . - (- x \ y) = - (y \ - - ((top \ y) \ - x))" using 9 25 157 2706 64260 by metis have 64282: "\x y . - (x \ - - ((x \ top) \ y)) = - (x \ - - y)" using 9 25 132 536 2810 28485 61311 64260 by metis have 65110: "\x y . - ((- x \ y) \ (- y \ x)) = bot" using 9 14 49 37513 63947 by metis have 65231: "\x y . - (x \ ((- x \ y) \ - y)) = bot" using 9 25 65110 by metis have 65585: "\x y . - (x \ - y) = - - y \ - x" using 61311 61454 64239 by metis have 65615: "\x y . - x \ - ((x \ top) \ y) = - y \ - x" using 132 28485 64282 65585 by metis have 65616: "\x y . - (- x \ y) = - y \ - ((top \ y) \ - x)" using 132 28485 64281 65585 by metis have 65791: "\x y . - x \ - ((top \ x) \ - y) = - - y \ - x" using 89 132 12376 28485 64271 65585 65615 65616 by metis have 65933: "\x y . - (- x \ y) = - - x \ - y" using 65616 65791 by metis have 66082: "\x y z . - (x \ (y \ - z)) = - - z \ - (x \ y)" using 8 65585 by metis have 66204: "\x y . - - x \ - (y \ (- y \ x)) = bot" using 65231 66082 by metis have 66281: "\x y z . - (x \ (- y \ z)) = - - y \ - (x \ z)" using 25 65933 by metis have 67527: "\x y . - - (x \ (- x \ y)) \ y = y" using 14 49 62434 66204 by metis have 67762: "\x y . - (- - x \ (y \ (- y \ x))) = - x" using 61420 67527 by metis have 68018: "\x y z . - (x \ y) \ (x \ (y \ (z \ top))) = bot" using 8 83 1113 2320 by metis have 71989: "\x y z . - (x \ (y \ (z \ top))) = - (x \ y)" using 9 29 52533 67762 68018 by metis have 71997: "\x y z . - ((x \ (y \ top)) \ z) = - (x \ z)" using 17 2320 71989 by metis have 72090: "\x y z . - (x \ ((x \ y) \ z)) = - (x \ z)" using 10 14 705 71997 by metis have 72139: "\x y . - (x \ y) = - x \ - y" using 25 123 132 2138 65933 66281 72090 by metis show ?thesis using 72139 by metis qed lemma l15: "--(x \ y) = --x \ --y" by (simp add: l11 l12 l4) lemma l13_var: "- - (- x \ y) = - x \ - - y" proof - have 1: "\x y . x \ y \ x \ y = y" by (simp add: il_less_eq) have 4: "\x y . \(x \ y) \ x \ y = y" using 1 by metis have 5: "\x y z . (x \ y) \ (x \ z) \ x \ (y \ z)" by (simp add: il_sub_inf_right_isotone_var) have 6: "\x y . - - x \ - (- x \ y)" by (simp add: pad2) have 7: "\x y z . x \ (y \ z) = (x \ y) \ z" by (simp add: il_associative) have 8: "\x y z . (x \ y) \ z = x \ (y \ z)" using 7 by metis have 9: "\x y . x \ y = y \ x" by (simp add: il_commutative) have 10: "\x . x \ bot = x" by (simp add: il_bot_unit) have 11: "\x . x \ x = x" by simp have 12: "\x y z . x \ (y \ z) = (x \ y) \ z" by (simp add: il_inf_associative) have 13: "\x y z . (x \ y) \ z = x \ (y \ z)" using 12 by metis have 14: "\x . top \ x = x" by simp have 15: "\x . x \ top = x" by simp have 16: "\x y z . (x \ y) \ z = (x \ z) \ (y \ z)" by (simp add: il_inf_right_dist_sup) have 17: "\x y z . (x \ y) \ (z \ y) = (x \ z) \ y" using 16 by metis have 19: "\x . - x \ - - x = top" by simp have 20: "\x . - x \ x = bot" by (simp add: a_inf_complement_bot) have 22: "\x y z . ((x \ y) \ (x \ z)) \ (x \ (y \ z)) = x \ (y \ z)" using 4 5 by metis have 23: "\x y z . (x \ (y \ z)) \ ((x \ y) \ (x \ z)) = x \ (y \ z)" using 9 22 by metis have 24: "\x y . - - x \ - (- x \ y) = - (- x \ y)" using 4 6 by metis have 25: "\x y z . x \ (y \ z) = y \ (x \ z)" using 8 9 by metis have 26: "\x y z . (x \ y) \ ((x \ z) \ (x \ (y \ z))) = x \ (y \ z)" using 9 23 25 by metis have 29: "\x . bot \ x = x" using 9 10 by metis have 30: "\x y . x \ (x \ y) = x \ y" using 8 11 by metis have 34: "\u x y z . (x \ (y \ z)) \ (u \ z) = ((x \ y) \ u) \ z" using 13 17 by metis have 35: "\u x y z . (x \ y) \ (z \ (u \ y)) = (x \ (z \ u)) \ y" using 13 17 by metis have 38: "\x y . - x \ (- - x \ y) = top \ y" using 8 19 by metis have 41: "- top = bot" using 15 20 by metis have 42: "\x y . (- x \ y) \ x = y \ x" using 17 20 29 by metis have 43: "\x y . (x \ - y) \ y = x \ y" using 9 17 20 29 by metis have 45: "\x . - bot \ - - x = - bot" using 9 20 24 by metis have 49: "- bot = top" using 19 29 41 by metis have 50: "\x . top \ - - x = top" using 45 49 by metis have 62: "\x y . x \ ((x \ - y) \ (x \ - - y)) = x" using 9 15 19 25 26 by metis have 65: "\x y . (- (x \ y) \ x) \ (- (x \ y) \ y) = bot" using 9 20 26 29 by metis have 66: "\x y z . (x \ - - y) \ (x \ - (- y \ z)) = x \ - (- y \ z)" using 11 24 26 by metis have 69: "\x y . x \ (x \ - - y) = x" using 9 15 26 30 50 by metis have 81: "\x . top \ - x = top" using 9 19 30 by metis have 88: "\x y . x \ (- y \ x) = x" using 14 17 81 by metis have 101: "\x y z . x \ (y \ (x \ - - z)) = y \ x" using 25 69 by metis have 103: "\x y . x \ (x \ - y) = x" using 9 62 101 by metis have 123: "\x . - - x \ x = x" using 14 19 42 by metis have 127: "\x y . - - x \ (x \ y) = x \ y" using 13 123 by metis have 130: "\x . - x \ - - - x = - x" using 9 24 123 by metis have 132: "\x . - - - x = - x" using 9 103 123 130 by metis have 136: "\x y . (- x \ y) \ - - x = y \ - - x" using 42 132 by metis have 144: "\x y z . ((- (x \ y) \ x) \ z) \ y = z \ y" using 20 29 34 by metis have 182: "\x y z . (x \ (- - (y \ z) \ y)) \ z = (x \ y) \ z" using 17 35 123 by metis have 288: "\x y . - x \ - (- x \ y) = top" using 24 38 81 by metis have 315: "\x y . - (- x \ y) \ x = x" using 14 42 288 by metis have 319: "\x y . - x \ - - (- x \ y) = - x" using 9 24 315 by metis have 387: "\x y . - (x \ y) \ - x = - x" using 127 315 by metis have 405: "\x y z . - (x \ (y \ z)) \ - (x \ y) = - (x \ y)" using 13 387 by metis have 419: "\x y . - x \ - - (- x \ y) = - - (- x \ y)" using 315 387 by metis have 1091: "\x y . - (x \ y) \ x = bot" using 9 29 30 65 by metis have 1129: "\x y z . (- (x \ y) \ z) \ x = z \ x" using 17 29 1091 by metis have 1155: "\x y . - - x \ - (- x \ y) = - - x" using 66 103 123 132 by metis have 2097: "\x y . - - (x \ y) \ x = x" using 14 19 1129 by metis have 2124: "\x y . - - (x \ y) \ y = y" using 9 2097 by metis have 2137: "\x y . - x \ - - (x \ y) = top" using 9 288 2097 by metis have 2201: "\x y . - x \ - - (y \ x) = top" using 9 288 2124 by metis have 2343: "\x y . - (- x \ y) \ - - y = top" using 88 2201 by metis have 3022: "\x y . - x \ - (- y \ - x) = top" using 9 132 2343 by metis have 3133: "\x y . - (- x \ - y) \ y = y" using 14 42 3022 by metis have 3134: "\x y . - x \ (- y \ - x) = - y \ - x" using 14 43 3022 by metis have 3961: "\x y . - - (x \ y) \ - - x = - - x" using 14 136 2137 by metis have 9413: "\x y . - - (- x \ y) \ y = - x \ y" using 9 103 182 319 by metis have 12370: "\x y . - x \ - (- - x \ y) = - x" using 132 1155 by metis have 12376: "\x y . - x \ - (x \ y) = - x" using 127 132 1155 by metis have 12383: "\x y . - (x \ y) \ - y = - (x \ y)" using 132 1155 2124 by metis have 12393: "\x y . - - (- x \ - y) = - x \ - y" using 1155 3133 9413 by metis have 12639: "\x y . - x \ - (- y \ x) = - x" using 88 12383 by metis have 24647: "\x y . (- x \ - y) \ - (- x \ - y) = top" using 19 12393 by metis have 28338: "\x y . - (- - (x \ y) \ x) = - (x \ y)" using 123 405 12370 by metis have 28422: "\x y . - (- x \ - y) = - (- y \ - x)" using 13 3134 12393 28338 by metis have 28485: "\x y . - x \ - y = - y \ - x" using 2097 3961 12393 28422 by metis have 52421: "\x y . - (- x \ - (- x \ y)) \ y = y" using 14 144 24647 28485 by metis have 52520: "\x y . - x \ - (- x \ y) = - x \ - y" using 13 12376 12393 12639 28485 52421 by metis have 61156: "\x y . - - (- x \ y) = - x \ - - y" using 419 52520 by metis show ?thesis using 61156 by metis qed text \Theorem 25.1\ subclass subset_boolean_algebra_2 proof show "\x y z. x \ (y \ z) = x \ y \ z" by (simp add: il_associative) show "\x y. x \ y = y \ x" by (simp add: il_commutative) show "\x. x \ x = x" by simp show "\x y. x \ - (y \ - y) = x" using il_bot_unit l12 l6 by auto show "\x y. - (x \ y) = - (- - x \ - - y)" by (metis l15 l4) show "\x y. - x \ - (- x \ y) = - x \ - y" by (smt l11 l15 il_inf_right_dist_sup il_unit_bot l6 l7) qed lemma aa_test: "p = --p \ test p" by (metis ppa_ppd.d_closed) lemma test_aa_increasing: "test p \ p \ --p" by (simp add: ppa_ppd.d_increasing_sub_identity test_sub_identity) lemma "test p \ - - (p \ x) \ p" nitpick [expect=genuine] oops lemma "test p \ --p \ p" nitpick [expect=genuine] oops end class pa_algebra = pa_semiring + minus + assumes pa_minus_def: "-x - -y = -(--x \ -y)" begin subclass subset_boolean_algebra_2_extended proof show "bot = (THE x. \z. x = - (z \ - z))" using l12 l6 by auto thus "top = - (THE x. \z. x = - (z \ - z))" using l2 by blast show "\x y. - x \ - y = - (- - x \ - - y)" by (metis l12 l4) show "\x y. - x - - y = - (- - x \ - y)" by (simp add: pa_minus_def) show "\x y. (x \ y) = (x \ y = y)" by (simp add: il_less_eq) show "\x y. (x < y) = (x \ y = y \ y \ x \ x)" by (simp add: il_less_eq less_le_not_le) qed lemma "\x y. - (x \ - - y) = - (x \ y)" nitpick [expect=genuine] oops end subsection \Antidomain Semirings\ text \Definition 24\ class a_semiring = ppa_semiring + assumes ad3: "-(x \ y) \ -(x \ --y)" begin lemma l16: "- - x \ - (- x \ y)" proof - have 1: "\x y . x \ y \ x \ y = y" by (simp add: il_less_eq) have 3: "\x y z . x \ (y \ z) = (x \ y) \ z" by (simp add: il_associative) have 4: "\x y z . (x \ y) \ z = x \ (y \ z)" using 3 by metis have 5: "\x y . x \ y = y \ x" by (simp add: il_commutative) have 6: "\x . x \ bot = x" by (simp add: il_bot_unit) have 7: "\x . x \ x = x" by simp have 8: "\x y . \(x \ y) \ x \ y = y" using 1 by metis have 9: "\x y . x \ y \ x \ y \ y" using 1 by metis have 10: "\x y z . x \ (y \ z) = (x \ y) \ z" by (simp add: il_inf_associative) have 11: "\x y z . (x \ y) \ z = x \ (y \ z)" using 10 by metis have 12: "\x . top \ x = x" by simp have 13: "\x . x \ top = x" by simp have 14: "\x y z . (x \ y) \ (x \ z) \ x \ (y \ z)" by (simp add: il_sub_inf_right_isotone_var) have 15: "\x y z . (x \ y) \ z = (x \ z) \ (y \ z)" by (simp add: il_inf_right_dist_sup) have 16: "\x y z . (x \ y) \ (z \ y) = (x \ z) \ y" using 15 by metis have 17: "\x . bot \ x = bot" by simp have 18: "\x . - x \ - - x = top" by simp have 19: "\x . - x \ x = bot" by (simp add: a_inf_complement_bot) have 20: "\x y . - (x \ y) \ - (x \ - - y)" by (simp add: ad3) have 22: "\x y z . x \ (y \ z) = y \ (x \ z)" using 4 5 by metis have 25: "\x . bot \ x = x" using 5 6 by metis have 26: "\x y . x \ (x \ y) = x \ y" using 4 7 by metis have 33: "\x y z . (x \ y) \ ((x \ z) \ (x \ (y \ z))) = x \ (y \ z)" using 5 8 14 22 by metis have 47: "\x y . - x \ (- - x \ y) = top \ y" using 4 18 by metis have 48: "\x y . - - x \ (y \ - x) = y \ top" using 4 5 18 by metis have 51: "\x y . - x \ (x \ y) = bot" using 11 17 19 by metis have 52: "- top = bot" using 13 19 by metis have 56: "\x y . (- x \ y) \ x = y \ x" using 16 19 25 by metis have 57: "\x y . (x \ - y) \ y = x \ y" using 5 16 19 25 by metis have 58: "\x y . - (x \ y) \ - (x \ - - y) = - (x \ - - y)" using 8 20 by metis have 60: "\x . - x \ - - - x" using 12 20 by metis have 69: "- bot = top" using 18 25 52 by metis have 74: "\x y . x \ x \ y" using 9 26 by metis have 78: "\x . top \ - x = top" using 5 18 26 by metis have 80: "\x y . x \ y \ x" using 5 74 by metis have 86: "\x y z . x \ y \ x \ (z \ y)" using 22 80 by metis have 95: "\x . - x \ - - - x = - - - x" using 8 60 by metis have 143: "\x y . x \ (x \ - y) = x" using 5 13 26 33 78 by metis have 370: "\x y z . x \ (y \ - z) \ x \ y" using 86 143 by metis have 907: "\x . - x \ - x = - x" using 12 18 57 by metis have 928: "\x y . - x \ (- x \ y) = - x \ y" using 11 907 by metis have 966: "\x y . - (- x \ - - (x \ y)) = top" using 51 58 69 78 by metis have 1535: "\x . - x \ - - - - x = top" using 47 78 95 by metis have 1630: "\x y z . (x \ y) \ - z \ (x \ - z) \ y" using 16 370 by metis have 2422: "\x . - x \ - - - x = - - - x" using 12 57 1535 by metis have 6567: "\x y . - x \ - - (x \ y) = bot" using 12 19 966 by metis have 18123: "\x . - - - x = - x" using 95 143 2422 by metis have 26264: "\x y . - x \ (- y \ - x) \ - - y" using 12 18 1630 by metis have 26279: "\x y . - - (x \ y) \ - - x" using 25 6567 26264 by metis have 26307: "\x y . - - (- x \ y) \ - x" using 928 18123 26279 by metis have 26339: "\x y . - x \ - - (- x \ y) = - x" using 5 8 26307 by metis have 26564: "\x y . - x \ - (- x \ y) = top" using 5 48 78 18123 26339 by metis have 26682: "\x y . - (- x \ y) \ x = x" using 12 56 26564 by metis have 26864: "\x y . - - x \ - (- x \ y)" using 18123 26279 26682 by metis show ?thesis using 26864 by metis qed text \Theorem 25.2\ subclass pa_semiring proof show "\x y. - - x \ - (- x \ y)" by (rule l16) qed lemma l17: "-(x \ y) = -(x \ --y)" by (simp add: ad3 order.antisym l14) lemma a_complement_inf_double_complement: "-(x \ --y) = -(x \ y)" using l17 by auto sublocale a_d: d_semiring_var where d = "\x . --x" proof show "\x y. - - (x \ - - y) \ - - (x \ y)" using l17 by auto show "- - bot = bot" by (simp add: l1 l2) qed lemma "test p \ - - (p \ x) \ p" by (fact a_d.d2) end class a_algebra = a_semiring + minus + assumes a_minus_def: "-x - -y = -(--x \ -y)" begin subclass pa_algebra proof show "\x y. - x - - y = - (- - x \ - y)" by (simp add: a_minus_def) qed text \Theorem 25.4\ subclass subset_boolean_algebra_4_extended proof show "\x y z. x \ (y \ z) = x \ y \ z" by (simp add: il_inf_associative) show "\x y z. (x \ y) \ z = x \ z \ y \ z" by (simp add: il_inf_right_dist_sup) show "\x. - x \ x = bot" by (simp add: a_inf_complement_bot) show "\x. top \ x = x" by simp show "\x y. - (x \ - - y) = - (x \ y)" using l17 by auto show "\x. x \ top = x" by simp show "\x y z. x \ y \ z \ x \ z \ y" by (simp add: il_sub_inf_right_isotone) qed end context subset_boolean_algebra_4_extended begin subclass il_semiring proof show "\x y z. x \ (y \ z) = x \ y \ z" by (simp add: sup_assoc) show "\x y. x \ y = y \ x" by (simp add: sup_commute) show "\x. x \ x = x" by simp show "\x. x \ bot = x" by simp show "\x y z. x \ (y \ z) = x \ y \ z" by (simp add: sba3_inf_associative) show "\x y z. (x \ y) \ z = x \ z \ y \ z" by (simp add: sba3_inf_right_dist_sup) show "\x. top \ x = x" by simp show "\x. x \ top = x" by simp show "\x. bot \ x = bot" by (simp add: inf_left_zero) show "\x y z. x \ y \ z \ x \ z \ y" by (simp add: inf_right_isotone) show "\x y. (x \ y) = (x \ y = y)" by (simp add: le_iff_sup) show "\x y. (x < y) = (x \ y \ \ y \ x)" by (simp add: less_le_not_le) qed subclass a_semiring proof show "\x. - x \ x = bot" by (simp add: sba3_inf_complement_bot) show "\x. - x \ - - x = top" by simp show "\x y. - (x \ y) \ - (x \ - - y)" by (simp add: sba3_complement_inf_double_complement) qed sublocale sba4_a: a_algebra proof show "\x y. - x - - y = - (- - x \ - y)" by (simp add: sub_minus_def) qed end context stone_algebra begin text \Theorem 25.3\ subclass il_semiring proof show "\x y z. x \ (y \ z) = x \ y \ z" by (simp add: sup_assoc) show "\x y. x \ y = y \ x" by (simp add: sup_commute) show "\x. x \ x = x" by simp show "\x. x \ bot = x" by simp show "\x y z. x \ (y \ z) = x \ y \ z" by (simp add: inf.sup_monoid.add_assoc) show "\x y z. (x \ y) \ z = x \ z \ y \ z" by (simp add: inf_sup_distrib2) show "\x. top \ x = x" by simp show "\x. x \ top = x" by simp show "\x. bot \ x = bot" by simp show "\x y z. x \ y \ z \ x \ z \ y" using inf.sup_right_isotone by blast show "\x y. (x \ y) = (x \ y = y)" by (simp add: le_iff_sup) show "\x y. (x < y) = (x \ y \ \ y \ x)" by (simp add: less_le_not_le) qed subclass a_semiring proof show "\x. - x \ x = bot" by simp show "\x. - x \ - - x = top" by simp show "\x y. - (x \ y) \ - (x \ - - y)" by simp qed end end diff --git a/thys/Suppes_Theorem/Suppes_Theorem.thy b/thys/Suppes_Theorem/Suppes_Theorem.thy --- a/thys/Suppes_Theorem/Suppes_Theorem.thy +++ b/thys/Suppes_Theorem/Suppes_Theorem.thy @@ -1,597 +1,597 @@ chapter \ Suppes' Theorem \label{sec:suppes-theorem} \ theory Suppes_Theorem imports Probability_Logic begin no_notation FuncSet.funcset (infixr "\" 60) text \ An elementary completeness theorem for inequalities for probability logic is due to Patrick Suppes @{cite suppesProbabilisticInferenceConcept1966}. \ text \ A consequence of this Suppes' theorem is an elementary form of \<^emph>\collapse\, which asserts that inequalities for probabilities are logically equivalent to the more restricted class of \<^emph>\Dirac measures\ as defined in \S\ref{sec:dirac-measures}. \ section \ Suppes' List Theorem \label{sec:suppes-theorem-for-lists} \ text \ We first establish Suppes' theorem for lists of propositions. This is done by establishing our first completeness theorem using \<^emph>\Dirac measures\. \ text \ First, we use the result from \S\ref{sec:basic-probability-inequality-results} that shows \\ \ \ \ \\ implies \\

\ \ (\\\\. \

\)\. This can be understood as a \<^emph>\soundness\ result. \ text \ To show completeness, assume \\ \ \ \ \ \\. From this obtain a maximally consistent \<^term>\\\ such that \\ \ \ \ \ \\. We then define \<^term>\\ \ = (if (\ \ \) then 1 else 0)\ and show \<^term>\\\ is a \<^emph>\Dirac measure\ such that \\ \ \ (\\\\. \ \)\. \ lemma (in classical_logic) dirac_list_summation_completeness: "(\ \ \ dirac_measures. \ \ \ (\\\\. \ \)) = \ \ \ \ \" proof - { fix \ :: "'a \ real" assume "\ \ dirac_measures" from this interpret probability_logic "(\ \. \ \)" "(\)" "\" "\" unfolding dirac_measures_def by auto assume "\ \ \ \ \" hence "\ \ \ (\\\\. \ \)" using implication_list_summation_inequality by auto } moreover { assume "\ \ \ \ \ \" from this obtain \ where \: "MCS \" "\ \ \" "\ \ \ \" by (meson insert_subset formula_consistent_def formula_maximal_consistency formula_maximally_consistent_extension formula_maximally_consistent_set_def_def set_deduction_base_theory set_deduction_reflection set_deduction_theorem) hence"\ \ \ set \. \ \ \" using arbitrary_disjunction_exclusion_MCS by blast define \ where "\ = (\ \ . if \\\ then (1 :: real) else 0)" from \\ \ \ set \. \ \ \\ have "(\\\\. \ \) = 0" unfolding \_def by (induct \, simp, simp) hence "\ \ \ \ (\\\\. \ \)" unfolding \_def by (simp add: \(2)) hence "\ \ \ dirac_measures. \ (\ \ \ (\\\\. \ \))" unfolding \_def using \(1) MCS_dirac_measure by auto } ultimately show ?thesis by blast qed theorem (in classical_logic) list_summation_completeness: "(\ \

\ probabilities. \

\ \ (\\\\. \

\)) = \ \ \ \ \" (is "?lhs = ?rhs") proof assume ?lhs hence "\ \ \ dirac_measures. \ \ \ (\\\\. \ \)" unfolding dirac_measures_def probabilities_def by blast thus ?rhs using dirac_list_summation_completeness by blast next assume ?rhs show ?lhs proof fix \

:: "'a \ real" assume "\

\ probabilities" from this interpret probability_logic "(\ \. \ \)" "(\)" \ \

unfolding probabilities_def by auto show "\

\ \ (\\\\. \

\)" using \?rhs\ implication_list_summation_inequality by simp qed qed text \ The collapse theorem asserts that to prove an inequalities for all probabilities in probability logic, one only needs to consider the case of functions which take on values of 0 or 1. \ lemma (in classical_logic) suppes_collapse: "(\ \

\ probabilities. \

\ \ (\\\\. \

\)) = (\ \ \ dirac_measures. \ \ \ (\\\\. \ \))" by (simp add: dirac_list_summation_completeness list_summation_completeness) lemma (in classical_logic) probability_member_neg: fixes \

assumes "\

\ probabilities" shows "\

(\ \) = 1 - \

\" proof - from assms interpret probability_logic "(\ \. \ \)" "(\)" \ \

unfolding probabilities_def by auto show ?thesis by (simp add: complementation) qed text \ Suppes' theorem has a philosophical interpretation. It asserts that if \<^term>\\ :\ \\, then our \<^emph>\uncertainty\ in \<^term>\\\ is bounded above by our uncertainty in \<^term>\\\. Here the uncertainty in the proposition \<^term>\\\ is \1 - \

\\. Our uncertainty in \<^term>\\\, on the other hand, is \\\\\. 1 - \

\\. \ theorem (in classical_logic) suppes_list_theorem: "\ :\ \ = (\ \

\ probabilities. (\\\\. 1 - \

\) \ 1 - \

\)" proof - have "\ :\ \ = (\ \

\ probabilities. (\\ \ \<^bold>\ \. \

\) \ \

(\ \))" using list_summation_completeness weak_biconditional_weaken contra_list_curry_uncurry list_deduction_def by blast moreover have "\ \

\ probabilities. (\\ \ (\<^bold>\ \). \

\) = (\\ \ \. \

(\ \))" by (induct \, auto) ultimately show ?thesis using probability_member_neg by (induct \, simp+) qed section \ Suppes' Set Theorem \ text \ Suppes theorem also obtains for \<^emph>\sets\. \ lemma (in classical_logic) dirac_set_summation_completeness: "(\ \ \ dirac_measures. \ \ \ (\\\ set \. \ \)) = \ \ \ \ \" by (metis dirac_list_summation_completeness modus_ponens arbitrary_disjunction_remdups biconditional_left_elimination biconditional_right_elimination hypothetical_syllogism sum.set_conv_list) theorem (in classical_logic) set_summation_completeness: "(\ \ \ probabilities. \ \ \ (\\\ set \. \ \)) = \ \ \ \ \" by (metis dirac_list_summation_completeness dirac_set_summation_completeness list_summation_completeness sum.set_conv_list) lemma (in classical_logic) suppes_set_collapse: "(\ \

\ probabilities. \

\ \ (\\ \ set \. \

\)) = (\ \ \ dirac_measures. \ \ \ (\\ \ set \. \ \))" by (simp add: dirac_set_summation_completeness set_summation_completeness) text \ In our formulation of logic, there is not reason that \\ a = \ b\ while \<^term>\a \ b\. As a consequence the Suppes theorem for sets presented below is different than the one given in \S\ref{sec:suppes-theorem-for-lists}. \ theorem (in classical_logic) suppes_set_theorem: "\ :\ \ = (\ \

\ probabilities. (\\ \ set (\<^bold>\ \). \

\) \ 1 - \

\)" proof - have "\ :\ \ = (\ \

\ probabilities. (\\ \ set (\<^bold>\ \). \

\) \ \

(\ \))" using contra_list_curry_uncurry list_deduction_def set_summation_completeness weak_biconditional_weaken by blast thus ?thesis using probability_member_neg by (induct \, auto) qed section \ Converse Suppes' Theorem \ text \ A formulation of the converse of Suppes' theorem obtains for lists/sets of \<^emph>\logically disjoint\ propositions. \ lemma (in probability_logic) exclusive_sum_list_identity: assumes "\ \ \" shows "\

(\ \) = (\\\\. \

\)" using assms proof (induct \) case Nil then show ?case by (simp add: gaines_weatherson_antithesis) next case (Cons \ \) assume "\ \ (\ # \)" hence "\ \ (\ \ \ \)" "\ \ \" by simp+ hence "\

(\(\ # \)) = \

\ + \

(\ \)" "\

(\ \) = (\\\\. \

\)" using Cons.hyps probability_additivity by auto hence "\

(\(\ # \)) = \

\ + (\\\\. \

\)" by auto thus ?case by simp qed lemma sum_list_monotone: fixes f :: "'a \ real" assumes "\ x. f x \ 0" and "set \ \ set \" and "distinct \" shows "(\\\\. f \) \ (\\\\. f \)" using assms proof - assume "\ x. f x \ 0" have "\\. set \ \ set \ \ distinct \ \ (\\\\. f \) \ (\\\\. f \)" proof (induct \) case Nil then show ?case by simp next case (Cons \ \) { fix \ assume "set \ \ set (\ # \)" and "distinct \" have "(\\\\. f \) \ (\\'\(\ # \). f \')" proof - { assume "\ \ set \" with \set \ \ set (\ # \)\ have "set \ \ set \" by auto hence "(\\\\. f \) \ (\\\\. f \)" using Cons.hyps \distinct \\ by auto moreover have "f \ \ 0" using \\ x. f x \ 0\ by metis ultimately have ?thesis by simp } moreover { assume "\ \ set \" hence "set \ = insert \ (set (removeAll \ \))" by auto with \set \ \ set (\ # \)\ have "set (removeAll \ \) \ set \" by (metis insert_subset list.simps(15) set_removeAll subset_insert_iff) moreover from \distinct \\ have "distinct (removeAll \ \)" by (meson distinct_removeAll) ultimately have "(\\\(removeAll \ \). f \) \ (\\\\. f \)" using Cons.hyps by simp moreover from \\ \ set \\ \distinct \\ have "(\\\\. f \) = f \ + (\\\(removeAll \ \). f \)" using distinct_remove1_removeAll sum_list_map_remove1 by fastforce ultimately have ?thesis using \\ x. f x \ 0\ by simp } ultimately show ?thesis by blast qed } thus ?case by blast qed moreover assume "set \ \ set \" and "distinct \" ultimately show ?thesis by blast qed lemma count_remove_all_sum_list: fixes f :: "'a \ real" shows "real (count_list xs x) * f x + (\x'\(removeAll x xs). f x') = (\x\xs. f x)" by (induct xs, simp, simp, metis combine_common_factor mult_cancel_right1) lemma (in classical_logic) dirac_exclusive_implication_completeness: "(\ \ \ dirac_measures. (\\\\. \ \) \ \ \) = (\ \ \ \ \ \ \ \ \)" proof - { fix \ assume "\ \ dirac_measures" from this interpret probability_logic "(\ \. \ \)" "(\)" "\" "\" unfolding dirac_measures_def by simp assume "\ \ \" "\ \ \ \ \" hence "(\\\\. \ \) \ \ \" using exclusive_sum_list_identity monotonicity by fastforce } moreover { assume "\ \ \ \" hence "(\ \ \ set \. \ \ \ set \. \ \ \ \ \ \ \ (\ \ \)) \ (\ \ \ duplicates \. \ \ \ \)" using exclusive_equivalence set_deduction_base_theory by blast hence "\ (\ \ \ dirac_measures. (\\\\. \ \) \ \ \)" proof (elim disjE) assume "\ \ \ set \. \ \ \ set \. \ \ \ \ \ \ \ (\ \ \)" from this obtain \ and \ where \\_properties: "\ \ set \" "\ \ set \" "\ \ \" "\ \ \ (\ \ \)" by blast from this obtain \ where \: "MCS \" "\ (\ \ \) \ \" by (meson insert_subset formula_consistent_def formula_maximal_consistency formula_maximally_consistent_extension formula_maximally_consistent_set_def_def set_deduction_base_theory set_deduction_reflection set_deduction_theorem) let ?\ = "\ \. if \\\ then (1 :: real) else 0" from \ have "\ \ \" "\ \ \" by (metis formula_maximally_consistent_set_def_implication maximally_consistent_set_def conjunction_def negation_def)+ with \\_properties have "(\\\[\, \]. ?\ \) = 2" "set [\, \] \ set \" "distinct [\, \]" "\\. ?\ \ \ 0" by simp+ hence "(\\\\. ?\ \) \ 2" using sum_list_monotone by metis hence "\ (\\\\. ?\ \) \ ?\ (\)" by auto thus ?thesis using \(1) MCS_dirac_measure by auto next assume "\ \ \ duplicates \. \ \ \ \" from this obtain \ where \: "\ \ duplicates \" "\ \ \ \" using exclusive_equivalence [where \="{}"] set_deduction_base_theory by blast from \ obtain \ where \: "MCS \" "\ \ \ \" by (meson insert_subset formula_consistent_def formula_maximal_consistency formula_maximally_consistent_extension formula_maximally_consistent_set_def_def set_deduction_base_theory set_deduction_reflection set_deduction_theorem) hence "\ \ \" using negation_def by auto let ?\ = "\ \. if \\\ then (1 :: real) else 0" from \ have "count_list \ \ \ 2" using duplicates_alt_def [where xs="\"] by blast hence "real (count_list \ \) * ?\ \ \ 2" using \\ \ \\ by simp moreover { fix \ have "(\\\\. ?\ \) \ 0" by (induct \, simp, simp) } moreover have "(0::real) \ (\a\removeAll \ \. if a \ \ then 1 else 0)" using \\\. 0 \ (\\\\. if \ \ \ then 1 else 0)\ by presburger ultimately have "real (count_list \ \) * ?\ \ + (\ \ \ (removeAll \ \). ?\ \) \ 2" using \2 \ real (count_list \ \) * (if \ \ \ then 1 else 0)\ by linarith hence "(\\\\. ?\ \) \ 2" by (metis count_remove_all_sum_list) hence "\ (\\\\. ?\ \) \ ?\ (\)" by auto thus ?thesis using \(1) MCS_dirac_measure by auto qed } moreover { assume "\ \ \ \ \ \" from this obtain \ \ where \: "MCS \" and \: "\ \ \" and \: "\ \ set \" "\ \ \" by (meson insert_subset formula_consistent_def formula_maximal_consistency formula_maximally_consistent_extension formula_maximally_consistent_set_def_def arbitrary_disjunction_exclusion_MCS set_deduction_base_theory set_deduction_reflection set_deduction_theorem) let ?\ = "\ \. if \\\ then (1 :: real) else 0" from \ have "(\\\\. ?\ \) \ 1" proof (induct \) case Nil then show ?case by simp next case (Cons \' \) obtain f :: "real list \ real" where f: "\rs. f rs \ set rs \ \ 0 \ f rs \ 0 \ sum_list rs" - using sum_list_nonneg by moura + using sum_list_nonneg by metis moreover have "f (map ?\ \) \ set (map ?\ \) \ 0 \ f (map ?\ \)" by fastforce ultimately show ?case by (simp, metis Cons.hyps Cons.prems(1) \(2) set_ConsD) qed hence "\ (\\\\. ?\ \) \ ?\ (\)" using \ by auto hence "\ (\ \ \ dirac_measures. (\\\\. \ \) \ \ \)" using \(1) MCS_dirac_measure by auto } ultimately show ?thesis by blast qed theorem (in classical_logic) exclusive_implication_completeness: "(\ \

\ probabilities. (\\\\. \

\) \ \

\) = (\ \ \ \ \ \ \ \ \)" (is "?lhs = ?rhs") proof assume ?lhs thus ?rhs by (meson dirac_exclusive_implication_completeness dirac_measures_subset subset_eq) next assume ?rhs show ?lhs proof fix \

:: "'a \ real" assume "\

\ probabilities" from this interpret probability_logic "(\ \. \ \)" "(\)" \ \

unfolding probabilities_def by simp show "(\\\\. \

\) \ \

\" using \?rhs\ exclusive_sum_list_identity monotonicity by fastforce qed qed lemma (in classical_logic) dirac_inequality_completeness: "(\ \ \ dirac_measures. \ \ \ \ \) = \ \ \ \" proof - have "\ \ [\]" by (simp add: conjunction_right_elimination negation_def) hence "(\ \ [\] \ \ \ [\] \ \) = \ \ \ \" by (metis arbitrary_disjunction.simps(1) arbitrary_disjunction.simps(2) disjunction_def implication_equivalence negation_def weak_biconditional_weaken) thus ?thesis using dirac_exclusive_implication_completeness [where \="[\]"] by auto qed section \ Implication Inequality Completeness \ text \ The following theorem establishes the converse of \\ \ \ \ \ \

\ \ \

\\, which was proved in \S\ref{sec:prob-logic-alt-def}. \ theorem (in classical_logic) implication_inequality_completeness: "(\ \

\ probabilities. \

\ \ \

\) = \ \ \ \" proof - have "\ \ [\]" by (simp add: conjunction_right_elimination negation_def) hence "(\ \ [\] \ \ \ [\] \ \) = \ \ \ \" by (metis arbitrary_disjunction.simps(1) arbitrary_disjunction.simps(2) disjunction_def implication_equivalence negation_def weak_biconditional_weaken) thus ?thesis using exclusive_implication_completeness [where \="[\]"] by simp qed section \ Characterizing Logical Exclusiveness In Probability Logic \ text \ Finally, we can say that \\

(\ \) = (\\\\. \

\)\ if and only if the propositions in \<^term>\\\ are mutually exclusive (i.e. \\ \ \\). This result also obtains for sets. \ lemma (in classical_logic) dirac_exclusive_list_summation_completeness: "(\ \ \ dirac_measures. \ (\ \) = (\\\\. \ \)) = \ \ \" by (metis antisym_conv dirac_exclusive_implication_completeness dirac_list_summation_completeness trivial_implication) theorem (in classical_logic) exclusive_list_summation_completeness: "(\ \

\ probabilities. \

(\ \) = (\\\\. \

\)) = \ \ \" by (metis antisym_conv exclusive_implication_completeness list_summation_completeness trivial_implication) lemma (in classical_logic) dirac_exclusive_set_summation_completeness: "(\ \ \ dirac_measures. \ (\ \) = (\\ \ set \. \ \)) = \ \ (remdups \)" by (metis (mono_tags) eq_iff dirac_exclusive_implication_completeness dirac_set_summation_completeness trivial_implication set_remdups sum.set_conv_list eq_iff dirac_exclusive_implication_completeness dirac_set_summation_completeness trivial_implication set_remdups sum.set_conv_list antisym_conv) theorem (in classical_logic) exclusive_set_summation_completeness: "(\ \

\ probabilities. \

(\ \) = (\\ \ set \. \

\)) = \ \ (remdups \)" by (metis (mono_tags, opaque_lifting) antisym_conv exclusive_list_summation_completeness exclusive_implication_completeness implication_inequality_completeness set_summation_completeness order.refl sum.set_conv_list) lemma (in probability_logic) exclusive_list_set_inequality: assumes "\ \ \" shows "(\\\\. \

\) = (\\\set \. \

\)" proof - have "distinct (remdups \)" using distinct_remdups by auto hence "duplicates (remdups \) = {}" by (induct "\", simp+) moreover have "set (remdups \) = set \" by (induct "\", simp, simp add: insert_absorb) moreover have "(\\ \ duplicates \. \ \ \) \ (\ \ \ set \. \ \ \ set \. (\ \ \) \ \ \ (\ \ \))" using assms exclusive_elimination1 exclusive_elimination2 set_deduction_base_theory by blast ultimately have "(\\\duplicates (remdups \). \ \ \) \ (\ \ \ set (remdups \). \ \ \ set (remdups \). (\ \ \) \ \ \ (\ \ \))" by auto hence "\ \ (remdups \)" by (meson exclusive_equivalence set_deduction_base_theory) hence "(\\\set \. \

\) = \

(\ \)" by (metis arbitrary_disjunction_remdups biconditional_equivalence exclusive_sum_list_identity sum.set_conv_list) moreover have "(\\\\. \

\) = \

(\ \)" by (simp add: assms exclusive_sum_list_identity) ultimately show ?thesis by metis qed notation FuncSet.funcset (infixr "\" 60) end diff --git a/thys/Virtual_Substitution/ExecutiblePolyProps.thy b/thys/Virtual_Substitution/ExecutiblePolyProps.thy --- a/thys/Virtual_Substitution/ExecutiblePolyProps.thy +++ b/thys/Virtual_Substitution/ExecutiblePolyProps.thy @@ -1,1081 +1,1071 @@ text "Executable Polynomial Properties" theory ExecutiblePolyProps imports Polynomials.MPoly_Type_Univariate MPolyExtension begin text \(Univariate) Polynomial hiding\ lifting_update poly.lifting lifting_forget poly.lifting text \\ subsection "Lemmas with Monomial and Monomials" lemma of_nat_monomial: "of_nat p = monomial p 0" by (auto simp: poly_mapping_eq_iff lookup_of_nat fun_eq_iff lookup_single) lemma of_nat_times_monomial: "of_nat p * monomial c i = monomial (p*c) i" by (auto simp: poly_mapping_eq_iff prod_fun_def fun_eq_iff of_nat_monomial lookup_single mult_single) lemma monomial_adds_nat_iff: "monomial p i adds c \ lookup c i \ p" for p::"nat" apply (auto simp: adds_def lookup_add) by (metis add.left_commute nat_le_iff_add remove_key_sum single_add) lemma update_minus_monomial: "Poly_Mapping.update k i (m - monomial i k) = Poly_Mapping.update k i m" by (auto simp: poly_mapping_eq_iff lookup_update update.rep_eq fun_eq_iff lookup_minus lookup_single) lemma monomials_Var: "monomials (Var x::'a::zero_neq_one mpoly) = {Poly_Mapping.single x 1}" by transfer (auto simp: Var\<^sub>0_def) lemma monomials_Const: "monomials (Const x) = (if x = 0 then {} else {0})" by transfer' (auto simp: Const\<^sub>0_def) lemma coeff_eq_zero_iff: "MPoly_Type.coeff c p = 0 \ p \ monomials c" by transfer (simp add: not_in_keys_iff_lookup_eq_zero) lemma monomials_1[simp]: "monomials 1 = {0}" by transfer auto lemma monomials_and_monoms: shows "(k \ monomials m) = (\ (a::nat). a \ 0 \ (monomials (MPoly_Type.monom k a)) \ monomials m)" proof - show ?thesis using monomials_monom by auto qed lemma mult_monomials_dir_one: shows "monomials (p*q) \ {a+b | a b . a \ monomials p \ b \ monomials q}" using monomials_and_monoms mult_monom by (simp add: keys_mult monomials.rep_eq times_mpoly.rep_eq) lemma monom_eq_zero_iff[simp]: "MPoly_Type.monom a b = 0 \ b = 0" by (metis MPolyExtension.coeff_monom MPolyExtension.monom_zero) lemma update_eq_plus_monomial: "v \ lookup m k \ Poly_Mapping.update k v m = m + monomial (v - lookup m k) k" for v n::nat by transfer auto lemma coeff_monom_mult': "MPoly_Type.coeff ((MPoly_Type.monom m' a) * q) (m'm) = a * MPoly_Type.coeff q (m'm - m')" if *: "m'm = m' + (m'm - m')" by (subst *) (rule More_MPoly_Type.coeff_monom_mult) lemma monomials_zero[simp]: "monomials 0 = {}" by transfer auto lemma in_monomials_iff: "x \ monomials m \ MPoly_Type.coeff m x \ 0" using coeff_eq_zero_iff[of m x] by auto lemma monomials_monom_mult: "monomials (MPoly_Type.monom mon a * q) = (if a = 0 then {} else (+) mon ` monomials q)" for q::"'a::semiring_no_zero_divisors mpoly" apply auto subgoal by transfer' (auto elim!: in_keys_timesE) subgoal by (simp add: in_monomials_iff More_MPoly_Type.coeff_monom_mult) done subsection "Simplification Lemmas for Const 0 and Const 1" lemma add_zero : "P + Const 0 = P" proof - have h:"P + 0 = P" using add_0_right by auto show ?thesis unfolding Const_def using h by (simp add: Const\<^sub>0_zero zero_mpoly.abs_eq) qed (* example *) lemma add_zero_example : "((Var 0)^2 - (Const 1)) + Const 0 = ((Var 0)^2 - (Const 1))" proof - show ?thesis by (simp add : add_zero) qed lemma mult_zero_left : "Const 0 * P = Const 0" proof - have h:"0*P = 0" by simp show ?thesis unfolding Const_def using h by (simp add: Const\<^sub>0_zero zero_mpoly_def) qed lemma mult_zero_right : "P * Const 0 = Const 0" by (metis mult_zero_left mult_zero_right) lemma mult_one_left : "Const 1 * (P :: real mpoly) = P" by (simp add: Const.abs_eq Const\<^sub>0_one one_mpoly_def) lemma mult_one_right : "(P::real mpoly) * Const 1 = P" by (simp add: Const.abs_eq Const\<^sub>0_one one_mpoly_def) subsection "Coefficient Lemmas" lemma coeff_zero[simp]: "MPoly_Type.coeff 0 x = 0" by transfer auto lemma coeff_sum: "MPoly_Type.coeff (sum f S) x = sum (\i. MPoly_Type.coeff (f i) x) S" apply (induction S rule: infinite_finite_induct) apply (auto) by (metis More_MPoly_Type.coeff_add) lemma coeff_mult_Var: "MPoly_Type.coeff (x * Var i ^ p) c = (MPoly_Type.coeff x (c - monomial p i) when lookup c i \ p)" by transfer' (auto simp: Var\<^sub>0_def pprod.monomial_power lookup_times_monomial_right of_nat_times_monomial monomial_adds_nat_iff) lemma lookup_update_self[simp]: "Poly_Mapping.update i (lookup m i) m = m" by (auto simp: lookup_update intro!: poly_mapping_eqI) lemma coeff_Const: "MPoly_Type.coeff (Const p) m = (p when m = 0)" by transfer' (auto simp: Const\<^sub>0_def lookup_single) lemma coeff_Var: "MPoly_Type.coeff (Var p) m = (1 when m = monomial 1 p)" by transfer' (auto simp: Var\<^sub>0_def lookup_single when_def) subsection "Miscellaneous" lemma update_0_0[simp]: "Poly_Mapping.update x 0 0 = 0" by (metis lookup_update_self lookup_zero) lemma mpoly_eq_iff: "p = q \ (\m. MPoly_Type.coeff p m = MPoly_Type.coeff q m)" by transfer (auto simp: poly_mapping_eqI) lemma power_both_sides : assumes Ah : "(A::real) \0" assumes Bh : "(B::real) \0" shows "(A\B) = (A^2\B^2)" using Ah Bh by (meson power2_le_imp_le power_mono) lemma in_list_induct_helper: assumes "set(xs)\X" assumes "P []" assumes "(\x. x\X \ ( \xs. P xs \ P (x # xs)))" shows "P xs" using assms(1) proof(induction xs) case Nil then show ?case using assms by auto next case (Cons a xs) then show ?case using assms(3) by auto qed theorem in_list_induct [case_names Nil Cons]: assumes "P []" assumes "(\x. x\set(xs) \ ( \xs. P xs \ P (x # xs)))" shows "P xs" using in_list_induct_helper[of xs "set(xs)" P] using assms by auto subsubsection "Keys and vars" lemma inKeys_inVars : "a\0 \ x \ keys m \ x \ vars(MPoly_Type.monom m a)" by(simp add: vars_def) lemma notInKeys_notInVars : "x \ keys m \ x \ vars(MPoly_Type.monom m a)" by(simp add: vars_def) lemma lookupNotIn : "x \ keys m \ lookup m x = 0" by (simp add: in_keys_iff) subsection "Degree Lemmas" lemma degree_le_iff: "MPoly_Type.degree p x \ k \ (\m\monomials p. lookup m x \ k)" by transfer simp lemma degree_less_iff: "MPoly_Type.degree p x < k \ (k>0 \ (\m\monomials p. lookup m x < k))" by (transfer fixing: k) (cases "k = 0"; simp) lemma degree_ge_iff: "k > 0 \ MPoly_Type.degree p x \ k \ (\m\monomials p. lookup m x \ k)" using Max_ge_iff by (meson degree_less_iff not_less) lemma degree_greater_iff: "MPoly_Type.degree p x > k \ (\m\monomials p. lookup m x > k)" by transfer' (auto simp: Max_gr_iff) lemma degree_eq_iff: "MPoly_Type.degree p x = k \ (if k = 0 then (\m\monomials p. lookup m x = 0) else (\m\monomials p. lookup m x = k) \ (\m\monomials p. lookup m x \ k))" by (subst eq_iff) (force simp: degree_le_iff degree_ge_iff intro!: antisym) declare poly_mapping.lookup_inject[simp del] lemma lookup_eq_and_update_lemma: "lookup m var = deg \ Poly_Mapping.update var 0 m = x \ m = Poly_Mapping.update var deg x \ lookup x var = 0" unfolding poly_mapping_eq_iff by (force simp: update.rep_eq fun_eq_iff) lemma degree_const : "MPoly_Type.degree (Const (z::real)) (x::nat) = 0" by (simp add: degree_eq_iff monomials_Const) lemma degree_one : "MPoly_Type.degree (Var x :: real mpoly) x = 1" by(simp add: degree_eq_iff monomials_Var) subsection "Lemmas on treating a multivariate polynomial as univariate " lemma coeff_isolate_variable_sparse: "MPoly_Type.coeff (isolate_variable_sparse p var deg) x = (if lookup x var = 0 then MPoly_Type.coeff p (Poly_Mapping.update var deg x) else 0)" apply (transfer fixing: x var deg p) unfolding lookup_sum unfolding lookup_single apply (auto simp: when_def) apply (subst sum.inter_filter[symmetric]) subgoal by simp subgoal by (simp add: lookup_eq_and_update_lemma Collect_conv_if) by (auto intro!: sum.neutral simp add: lookup_update) lemma isovarspar_sum: "isolate_variable_sparse (p+q) var deg = isolate_variable_sparse (p) var deg + isolate_variable_sparse (q) var deg" apply (auto simp add: mpoly_eq_iff coeff_isolate_variable_sparse ) apply (metis More_MPoly_Type.coeff_add coeff_isolate_variable_sparse) by (metis More_MPoly_Type.coeff_add add.comm_neutral coeff_isolate_variable_sparse less_numeral_extra(3)) lemma isolate_zero[simp]: "isolate_variable_sparse 0 var n = 0" by transfer' (auto simp: mpoly_eq_iff) lemma coeff_isolate_variable_sparse_minus_monomial: "MPoly_Type.coeff (isolate_variable_sparse mp var i) (m - monomial i var) = (if lookup m var \ i then MPoly_Type.coeff mp (Poly_Mapping.update var i m) else 0)" by (simp add: coeff_isolate_variable_sparse lookup_minus update_minus_monomial) lemma sum_over_zero: "(mp::real mpoly) = (\i::nat \MPoly_Type.degree mp x. isolate_variable_sparse mp x i * Var x^i)" by (auto simp add: mpoly_eq_iff coeff_sum coeff_mult_Var if_if_eq_conj not_le coeff_isolate_variable_sparse_minus_monomial when_def lookup_update degree_less_iff simp flip: eq_iff intro!: coeff_not_in_monomials) lemma const_lookup_zero : "isolate_variable_sparse (Const p ::real mpoly) (x::nat) (0::nat) = Const p" by (auto simp: mpoly_eq_iff coeff_isolate_variable_sparse coeff_Const when_def) (metis lookup_update_self) lemma const_lookup_suc : "isolate_variable_sparse (Const p :: real mpoly) x (Suc i) = 0" apply(auto simp add: mpoly_eq_iff coeff_isolate_variable_sparse coeff_Const when_def) by (metis lookup_update lookup_zero nat.distinct(1)) lemma isovar_greater_degree : "\i > MPoly_Type.degree p var. isolate_variable_sparse p var i = 0" apply(auto simp add: mpoly_eq_iff coeff_isolate_variable_sparse degree_less_iff) by (metis coeff_not_in_monomials less_irrefl_nat lookup_update) lemma isolate_var_0 : "isolate_variable_sparse (Var x::real mpoly) x 0 = 0" apply(auto simp add: mpoly_eq_iff coeff_isolate_variable_sparse coeff_Var when_def) by (metis gr0I lookup_single_eq lookup_update_self n_not_Suc_n) lemma isolate_var_one : "isolate_variable_sparse (Var x :: real mpoly) x 1 = 1" by (auto simp add: mpoly_eq_iff coeff_isolate_variable_sparse coeff_Var coeff_eq_zero_iff) (smt More_MPoly_Type.coeff_monom One_nat_def add_diff_cancel_left' lookup_eq_and_update_lemma lookup_single_eq lookup_update_self monom_one plus_1_eq_Suc single_diff single_zero update_minus_monomial) lemma isovarspase_monom : assumes notInKeys : "x \ keys m" assumes notZero : "a \ 0" shows "isolate_variable_sparse (MPoly_Type.monom m a) x 0 = (MPoly_Type.monom m a :: real mpoly)" using assms by (auto simp add: mpoly_eq_iff coeff_isolate_variable_sparse coeff_eq_zero_iff in_keys_iff) (metis lookup_update_self) lemma isolate_variable_spase_zero : "lookup m x = 0 \ insertion (nth L) ((MPoly_Type.monom m a)::real mpoly) = 0 \ a \ 0 \ insertion (nth L) (isolate_variable_sparse (MPoly_Type.monom m a) x 0) = 0" by (simp add: isovarspase_monom lookup_eq_zero_in_keys_contradict notInKeys_notInVars) lemma isovarsparNotIn : "x \ vars (p::real mpoly) \ isolate_variable_sparse p x 0 = p" proof(induction p rule: mpoly_induct) case (monom m a) then show ?case apply(cases "a=0") by (simp_all add: isovarspase_monom vars_monom_keys) next case (sum p1 p2 m a) then show ?case by (simp add: monomials.rep_eq vars_add_monom isovarspar_sum) qed lemma degree0isovarspar : assumes deg0 : "MPoly_Type.degree (p::real mpoly) x = 0" shows "isolate_variable_sparse p x 0 = p" proof - have h1 : "p = (\i::nat \MPoly_Type.degree p x. isolate_variable_sparse p x i * Var x ^ i)" using sum_over_zero by auto have h2a : "\f. (\i::nat \0. f i) = f 0" apply(simp add: sum_def) by (metis add.right_neutral comm_monoid_add_class.sum_def finite.emptyI insert_absorb insert_not_empty sum.empty sum.insert) have h2 : "p = isolate_variable_sparse p x 0 * Var x ^ 0" using deg0 h1 h2a by auto show ?thesis using h2 by auto qed subsection "Summation Lemmas" lemma summation_normalized : assumes nonzero : "(B ::real) \0" shows "(\i = 0..<((n::nat)+1). (f i :: real) * B ^ (n - i)) = (\i = 0..<(n+1). (f i) / (B ^ i)) * (B^n)" proof - have h1a : "\x::real. ((\i = 0..<(n+1). (f i) / (B ^ i)) * x = (\i = 0..<(n+1). ((f i) / (B ^ i)) * x))" apply(induction n ) apply(auto) by (simp add: distrib_right) have h1 : "(\i = 0..<(n+1). (f i) / (B ^ i)) * (B^n) = (\i = 0..<(n+1). ((f i) / (B ^ i)) * (B^n))" using h1a by auto have h2 : "(\i = 0..<(n+1). ((f i) / (B ^ i)) * (B^n)) = (\i = 0..<(n+1). (f i) * ((B^n) / (B ^ i)))" by auto have h3 : "(\i = 0..<(n+1). (f i) * ((B^n) / (B ^ i))) = (\i = 0..<(n+1). (f i) * B ^ (n - i))" using add.right_neutral nonzero power_diff by fastforce show ?thesis using h1 h2 h3 by auto qed lemma normalize_summation : assumes nonzero : "(B::real)\0" shows "(\i = 0.. (\i = 0..<(n::nat)+1. (f i::real) / (B ^ i)) = 0" proof - have pow_zero : "\i. B^(i :: nat)\0" using nonzero by(auto) have single_division_zero : "\X. B*(X::real)=0 \ X=0" using nonzero by(auto) have h1: "(\i = 0.. ((\i = 0..i = 0..i = 0..0" shows "(\i = 0..<(n+1). (f i) * B ^ (n - i)) * B ^ (n mod 2) < 0 \ (\i = 0..<((n::nat)+1). (f i::real) / (B ^ i)) < 0" proof - have h1 : "(\i = 0..<(n+1). (f i) * B ^ (n - i)) * B ^ (n mod 2) < 0 \ (\i = 0..<(n+1). (f i) / (B ^ i)) * (B^n) * B ^ (n mod 2) < 0" using summation_normalized nonzero by auto have h2a : "n mod 2 = 0 \ n mod 2 = 1" by auto have h2b : "n mod 2 = 1 \ odd n" by auto have h2c : "(B^n) * B ^ (n mod 2) > 0" using h2a h2b apply auto using nonzero apply presburger by (metis even_Suc mult.commute nonzero power_Suc zero_less_power_eq) have h2 : "\x. ((x * (B^n) * B ^ (n mod 2) < 0) = (x<0))" using h2c using mult.assoc by (metis mult_less_0_iff not_square_less_zero) show ?thesis using h1 h2 by auto qed subsection "Additional Lemmas with Vars" lemma not_in_isovarspar : "isolate_variable_sparse (p::real mpoly) var x = (q::real mpoly) \ var\(vars q)" apply(simp add: isolate_variable_sparse_def vars_def) proof - assume a1: "MPoly (\m | m \ monomials p \ lookup m var = x. monomial (MPoly_Type.coeff p m) (Poly_Mapping.update var 0 m)) = q" { fix pp :: "nat \\<^sub>0 nat" have "isolate_variable_sparse p var x = q" using a1 isolate_variable_sparse.abs_eq by blast then have "var \ keys pp \ pp \ keys (mapping_of q)" by (metis (no_types) coeff_def coeff_isolate_variable_sparse in_keys_iff) } then show "\p\keys (mapping_of q). var \ keys p" by meson qed lemma not_in_add : "var\(vars (p::real mpoly)) \ var\(vars (q::real mpoly)) \ var\(vars (p+q))" by (meson UnE in_mono vars_add) lemma not_in_mult : "var\(vars (p::real mpoly)) \ var\(vars (q::real mpoly)) \ var\(vars (p*q))" by (meson UnE in_mono vars_mult) lemma not_in_neg : "var\(vars(p::real mpoly)) \ var\(vars(-p))" proof - have h: "var \ (vars (-1::real mpoly))" by (metis add_diff_cancel_right' add_neg_numeral_special(8) isolate_var_one isolate_zero isovarsparNotIn isovarspar_sum not_in_isovarspar) show ?thesis using not_in_mult using h by fastforce qed lemma not_in_sub : "var\(vars (p::real mpoly)) \ var\(vars (q::real mpoly)) \ var\(vars (p-q))" using not_in_add not_in_neg by fastforce lemma not_in_pow : "var\(vars(p::real mpoly)) \ var\(vars(p^i))" proof (induction i) case 0 then show ?case using isolate_var_one not_in_isovarspar by (metis power_0) next case (Suc i) then show ?case using not_in_mult by simp qed lemma not_in_sum_var: "(\i. var\(vars(f(i)::real mpoly))) \ var\(vars(\i\{0..<(n::nat)}.f(i)))" proof (induction n) case 0 then show ?case using isolate_zero not_in_isovarspar by fastforce next case (Suc n) have h1: "(sum f {0.. vars (f n)" by (simp add: Suc.prems) then show ?case using h1 vars_add by (simp add: Suc.IH Suc.prems not_in_add) qed lemma not_in_sum : "(\i. var\(vars(f(i)::real mpoly))) \ \(n::nat). var\(vars(\i\{0..x\keys (mapping_of p). var \ keys x \ (\k f. (k \ keys f) = (lookup f k = 0)) \ lookup (mapping_of p) a = 0 \ (\aa. (if aa < length L then L[var := y] ! aa else 0) ^ lookup a aa) = (\aa. (if aa < length L then L[var := x] ! aa else 0) ^ lookup a aa)" apply(cases "lookup (mapping_of p) a = 0") apply auto apply(rule Prod_any.cong) apply auto using lookupNotIn nth_list_update_neq power_0 by smt lemma not_contains_insertion : assumes notIn : "var \ vars (p:: real mpoly)" assumes exists : "insertion (nth_default 0 (list_update L var x)) p = val" shows "insertion (nth_default 0 (list_update L var y)) p = val" using notIn exists apply(simp add: insertion_def insertion_aux_def insertion_fun_def) unfolding vars_def nth_default_def using not_in_keys_iff_lookup_eq_zero apply auto apply(rule Sum_any.cong) apply simp using not_contains_insertion_helper[of p var _ L y x] proof - fix a :: "nat \\<^sub>0 nat" assume a1: "\x\keys (mapping_of p). var \ keys x" assume "\k f. ((k::'a) \ keys f) = (lookup f k = 0)" then show "lookup (mapping_of p) a = 0 \ (\n. (if n < length L then L[var := y] ! n else 0) ^ lookup a n) = (\n. (if n < length L then L[var := x] ! n else 0) ^ lookup a n)" using a1 \\a. \\x\keys (mapping_of p). var \ keys x; \k f. (k \ keys f) = (lookup f k = 0)\ \ lookup (mapping_of p) a = 0 \ (\aa. (if aa < length L then L[var := y] ! aa else 0) ^ lookup a aa) = (\aa. (if aa < length L then L[var := x] ! aa else 0) ^ lookup a aa)\ by blast qed subsection "Insertion Lemmas" lemma insertion_sum_var : "((insertion f (\i\{0..<(n::nat)}.g(i))) = (\i\{0..(n::nat). ((insertion f (\i\{0..i\{0..(n::nat). ((insertion f (\i\n. g(i))) = (\i\n. insertion f (g i)))" by (metis (no_types, lifting) fun_sum_commute insertion_add insertion_zero sum.cong) lemma insertion_pow : "insertion f (p^i) = (insertion f p)^i" proof (induction i) case 0 then show ?case by auto next case (Suc n) then show ?case by (simp add: insertion_mult) qed lemma insertion_neg : "insertion f (-p) = -insertion f p" by (metis add.inverse_inverse insertionNegative) lemma insertion_var : "length L > var \ insertion (nth_default 0 (list_update L var x)) (Var var) = x" by (auto simp: monomials_Var coeff_Var insertion_code nth_default_def) lemma insertion_var_zero : "insertion (nth_default 0 (x#xs)) (Var 0) = x" using insertion_var by fastforce lemma notIn_insertion_sub : "x\vars(p::real mpoly) \ x\vars(q::real mpoly) \ insertion f (p-q) = insertion f p - insertion f q" by (metis ab_group_add_class.ab_diff_conv_add_uminus insertion_add insertion_neg) lemma insertion_sub : "insertion f (A-B :: real mpoly) = insertion f A - insertion f B" using insertion_neg insertion_add by (metis uminus_add_conv_diff) lemma insertion_four : "insertion ((nth_default 0) xs) 4 = 4" by (metis (no_types, lifting) insertion_add insertion_one numeral_plus_numeral one_add_one semiring_norm(2) semiring_norm(6)) lemma insertion_add_ind_basecase: "insertion (nth (list_update L var x)) ((\i::nat \ 0. isolate_variable_sparse p var i * (Var var)^i)) = (\i = 0..<(0+1). insertion (nth (list_update L var x)) (isolate_variable_sparse p var i * (Var var)^i))" proof - have h1: "((\i::nat \ 0. isolate_variable_sparse p var i * (Var var)^i)) = (isolate_variable_sparse p var 0 * (Var var)^0)" by auto show ?thesis using h1 by auto qed lemma insertion_add_ind: "insertion (nth_default 0 (list_update L var x)) ((\i::nat \ d. isolate_variable_sparse p var i * (Var var)^i)) = (\i = 0..<(d+1). insertion (nth_default 0 (list_update L var x)) (isolate_variable_sparse p var i * (Var var)^i))" proof (induction d) case 0 then show ?case using insertion_add_ind_basecase nth_default_def by auto next case (Suc n) then show ?case using insertion_add apply auto by (simp add: insertion_add) qed lemma sum_over_degree_insertion : assumes lLength : "length L > var" assumes deg : "MPoly_Type.degree (p::real mpoly) var = d" shows "(\i = 0..<(d+1). insertion (nth_default 0 (list_update L var x)) (isolate_variable_sparse p var i) * (x^i)) = insertion (nth_default 0 (list_update L var x)) p" proof - have h1: "(p::real mpoly) = (\i::nat \(MPoly_Type.degree p var). isolate_variable_sparse p var i * (Var var)^i)" using sum_over_zero by auto have h2: "insertion (nth_default 0 (list_update L var x)) p = insertion (nth_default 0 (list_update L var x)) ((\i::nat \ d. isolate_variable_sparse p var i * (Var var)^i))" using h1 deg by auto have h3: "insertion (nth_default 0 (list_update L var x)) p = (\i = 0..<(d+1). insertion (nth_default 0 (list_update L var x)) (isolate_variable_sparse p var i * (Var var)^i))" using h2 insertion_add_ind nth_default_def by simp show ?thesis using h3 insertion_var insertion_pow by (metis (no_types, lifting) insertion_mult lLength sum.cong) qed lemma insertion_isovarspars_free : "insertion (nth_default 0 (list_update L var x)) (isolate_variable_sparse (p::real mpoly) var (i::nat)) =insertion (nth_default 0 (list_update L var y)) (isolate_variable_sparse (p::real mpoly) var (i::nat))" proof - have h : "var \ vars(isolate_variable_sparse (p::real mpoly) var (i::nat))" by (simp add: not_in_isovarspar) then show ?thesis using not_contains_insertion by blast qed lemma insertion_zero : "insertion f (Const 0 ::real mpoly) = 0" by (metis add_cancel_right_right add_zero insertion_zero) lemma insertion_one : "insertion f (Const 1 ::real mpoly) = 1" by (metis insertion_one mult.right_neutral mult_one_left) lemma insertion_const : "insertion f (Const c::real mpoly) = (c::real)" by (auto simp: monomials_Const coeff_Const insertion_code) subsection "Putting Things Together" subsubsection "More Degree Lemmas" lemma degree_add_leq : assumes h1 : "MPoly_Type.degree a var \ x" assumes h2 : "MPoly_Type.degree b var \ x" shows "MPoly_Type.degree (a+b) var \ x" using degree_eq_iff coeff_add coeff_not_in_monomials by (smt (z3) More_MPoly_Type.coeff_add add.left_neutral coeff_eq_zero_iff degree_le_iff h1 h2) lemma degree_add_less : assumes h1 : "MPoly_Type.degree a var < x" assumes h2 : "MPoly_Type.degree b var < x" shows "MPoly_Type.degree (a+b) var < x" -proof - - obtain pp :: "nat \ nat \ 'a mpoly \ nat \\<^sub>0 nat" where - "\x0 x1 x2. (\v3. v3 \ monomials x2 \ \ lookup v3 x1 < x0) = (pp x0 x1 x2 \ monomials x2 \ \ lookup (pp x0 x1 x2) x1 < x0)" - by moura - then have f1: "\m n na. (\ MPoly_Type.degree m n < na \ 0 < na \ (\p. p \ monomials m \ lookup p n < na)) \ (MPoly_Type.degree m n < na \ \ 0 < na \ pp na n m \ monomials m \ \ lookup (pp na n m) n < na)" - by (metis (no_types) degree_less_iff) - then have "0 < x \ (\p. p \ monomials a \ lookup p var < x)" - using assms(1) by presburger - then show ?thesis - using f1 by (metis MPolyExtension.coeff_add add.left_neutral assms(2) coeff_eq_zero_iff) -qed + by (metis degree_add_leq h1 h2 linorder_not_le nat_less_le) lemma degree_sum : "(\i\{0..n::nat}. MPoly_Type.degree (f i :: real mpoly) var \ x) \ (MPoly_Type.degree (\x\{0..n}. f x) var) \ x" proof(induction n) case 0 then show ?case by auto next case (Suc n) then show ?case by(simp add: degree_add_leq) qed lemma degree_sum_less : "(\i\{0..n::nat}. MPoly_Type.degree (f i :: real mpoly) var < x) \ (MPoly_Type.degree (\x\{0..n}. f x) var) < x" proof(induction n) case 0 then show ?case by auto next case (Suc n) then show ?case by(simp add: degree_add_less) qed lemma varNotIn_degree : assumes "var \ vars p" shows "MPoly_Type.degree p var = 0" proof- have "\m\monomials p. lookup m var = 0" using assms unfolding vars_def keys_def using monomials.rep_eq by fastforce then show ?thesis using degree_less_iff by blast qed lemma degree_mult_leq : assumes pnonzero: "(p::real mpoly)\0" assumes qnonzero: "(q::real mpoly)\0" shows "MPoly_Type.degree ((p::real mpoly) * (q::real mpoly)) var \ (MPoly_Type.degree p var) + (MPoly_Type.degree q var)" proof(cases "MPoly_Type.degree (p*q) var =0") case True then show ?thesis by simp next case False have hp: "\m\monomials p. lookup m var \ MPoly_Type.degree p var" using degree_eq_iff by (metis zero_le) have hq: "\m\monomials q. lookup m var \ MPoly_Type.degree q var" using degree_eq_iff by (metis zero_le) have hpq: "\m\{a+b | a b . a \ monomials p \ b \ monomials q}. lookup m var \ (MPoly_Type.degree p var) + (MPoly_Type.degree q var)" by (smt add_le_mono hp hq mem_Collect_eq plus_poly_mapping.rep_eq) have h1: "(\m\monomials (p*q). lookup m var \ (MPoly_Type.degree p var) + (MPoly_Type.degree q var))" using mult_monomials_dir_one hpq by blast then show ?thesis using h1 degree_eq_iff False by (simp add: degree_le_iff) qed lemma degree_exists_monom: assumes "p\0" shows "\m\monomials p. lookup m var = MPoly_Type.degree p var" proof(cases "MPoly_Type.degree p var =0") case True have h1: "\m\monomials p. lookup m var = 0" unfolding monomials_def by (metis True assms(1) aux degree_eq_iff in_keys_iff mapping_of_inject monomials.rep_eq monomials_def zero_mpoly.rep_eq) then show ?thesis using h1 using True by simp next case False then show ?thesis using degree_eq_iff assms(1) apply(auto) by (metis degree_eq_iff dual_order.strict_iff_order) qed lemma degree_not_exists_monom: assumes "p\0" shows "\ (\ m\monomials p. lookup m var > MPoly_Type.degree p var)" proof - show ?thesis using degree_less_iff by blast qed lemma degree_monom: "MPoly_Type.degree (MPoly_Type.monom x y) v = (if y = 0 then 0 else lookup x v)" by (auto simp: degree_eq_iff) lemma degree_plus_disjoint: "MPoly_Type.degree (p + MPoly_Type.monom m c) v = max (MPoly_Type.degree p v) (MPoly_Type.degree (MPoly_Type.monom m c) v)" if "m \ monomials p" for p::"real mpoly" using that apply (subst degree_eq_iff) apply (auto simp: monomials_add_disjoint) apply (auto simp: degree_eq_iff degree_monom) apply (metis MPoly_Type.degree_zero degree_exists_monom less_numeral_extra(3)) using degree_le_iff apply blast using degree_eq_iff apply (metis max_def neq0_conv) apply (metis degree_eq_iff max.coboundedI1 neq0_conv) apply (metis MPoly_Type.degree_zero degree_exists_monom max_def zero_le) using degree_le_iff max.cobounded1 by blast subsubsection "More isolate\\_variable\\_sparse lemmas" lemma isolate_variable_sparse_ne_zeroD: "isolate_variable_sparse mp v x \ 0 \ x \ MPoly_Type.degree mp v" using isovar_greater_degree leI by blast context includes poly.lifting begin lift_definition mpoly_to_nested_poly::"'a::comm_monoid_add mpoly \ nat \ 'a mpoly Polynomial.poly" is "\(mp::'a mpoly) (v::nat) (p::nat). isolate_variable_sparse mp v p" \ \note \<^const>\extract_var\ nests the other way around\ unfolding MOST_iff_cofinite proof - fix mp::"'a mpoly" and v::nat have "{p. isolate_variable_sparse mp v p \ 0} \ {0..MPoly_Type.degree mp v}" (is "?s \ _") by (auto dest!: isolate_variable_sparse_ne_zeroD) also have "finite \" by simp finally (finite_subset) show "finite ?s" . qed lemma degree_eq_0_mpoly_to_nested_polyI: "mpoly_to_nested_poly mp v = 0 \ MPoly_Type.degree mp v = 0" apply transfer' apply (simp add: degree_eq_iff) apply transfer' apply (auto simp: fun_eq_iff) proof - fix mpa :: "'a mpoly" and va :: nat and m :: "nat \\<^sub>0 nat" assume a1: "\x. (\m | m \ monomials mpa \ lookup m va = x. monomial (MPoly_Type.coeff mpa m) (Poly_Mapping.update va 0 m)) = 0" assume a2: "m \ monomials mpa" have f3: "\m p. lookup (mapping_of m) p \ (0::'a) \ p \ monomials m" by (metis (full_types) coeff_def coeff_eq_zero_iff) have f4: "\n. mapping_of (isolate_variable_sparse mpa va n) = 0" using a1 by (simp add: isolate_variable_sparse.rep_eq) have "\p n. lookup 0 (p::nat \\<^sub>0 nat) = (0::'a) \ (0::nat) = n" by simp then show "lookup m va = 0" using f4 f3 a2 by (metis coeff_def coeff_isolate_variable_sparse lookup_eq_and_update_lemma) qed lemma coeff_eq_zero_mpoly_to_nested_polyD: "mpoly_to_nested_poly mp v = 0 \ MPoly_Type.coeff mp 0 = 0" apply transfer' apply transfer' by (metis (no_types) coeff_def coeff_isolate_variable_sparse isolate_variable_sparse.rep_eq lookup_zero update_0_0) lemma mpoly_to_nested_poly_eq_zero_iff[simp]: "mpoly_to_nested_poly mp v = 0 \ mp = 0" apply (auto simp: coeff_eq_zero_mpoly_to_nested_polyD degree_eq_0_mpoly_to_nested_polyI) proof - show "mpoly_to_nested_poly mp v = 0 \ mp = 0" apply (frule degree_eq_0_mpoly_to_nested_polyI) apply (frule coeff_eq_zero_mpoly_to_nested_polyD) apply (transfer' fixing: mp v) apply (transfer' fixing: mp v) apply (auto simp: fun_eq_iff mpoly_eq_iff intro!: sum.neutral) proof - fix m :: "nat \\<^sub>0 nat" assume a1: "\x. (\m | m \ monomials mp \ lookup m v = x. monomial (MPoly_Type.coeff mp m) (Poly_Mapping.update v 0 m)) = 0" assume a2: "MPoly_Type.degree mp v = 0" have "\n. isolate_variable_sparse mp v n = 0" using a1 by (simp add: isolate_variable_sparse.abs_eq zero_mpoly.abs_eq) then have f3: "\p. MPoly_Type.coeff mp p = MPoly_Type.coeff 0 p \ lookup p v \ 0" by (metis (no_types) coeff_isolate_variable_sparse lookup_update_self) then show "MPoly_Type.coeff mp m = 0" using a2 coeff_zero by (metis coeff_not_in_monomials degree_eq_iff) qed show "mp = 0 \ mpoly_to_nested_poly 0 v = 0" subgoal apply transfer' apply transfer' by (auto simp: fun_eq_iff intro!: sum.neutral) done qed lemma isolate_variable_sparse_degree_eq_zero_iff: "isolate_variable_sparse p v (MPoly_Type.degree p v) = 0 \ p = 0" apply (transfer') apply auto proof - fix pa :: "'a mpoly" and va :: nat assume "(\m | m \ monomials pa \ lookup m va = MPoly_Type.degree pa va. monomial (MPoly_Type.coeff pa m) (Poly_Mapping.update va 0 m)) = 0" then have "mapping_of (isolate_variable_sparse pa va (MPoly_Type.degree pa va)) = 0" by (simp add: isolate_variable_sparse.rep_eq) then show "pa = 0" by (metis (no_types) coeff_def coeff_eq_zero_iff coeff_isolate_variable_sparse degree_exists_monom lookup_eq_and_update_lemma lookup_zero) qed lemma degree_eq_univariate_degree: "MPoly_Type.degree p v = (if p = 0 then 0 else Polynomial.degree (mpoly_to_nested_poly p v))" apply auto apply (rule antisym) subgoal apply (rule Polynomial.le_degree) apply auto apply transfer' by (simp add: isolate_variable_sparse_degree_eq_zero_iff) subgoal apply (rule Polynomial.degree_le) apply (auto simp: elim!: degree_eq_zeroE) apply transfer' by (simp add: isovar_greater_degree) done lemma compute_mpoly_to_nested_poly[code]: "coeffs (mpoly_to_nested_poly mp v) = (if mp = 0 then [] else map (isolate_variable_sparse mp v) [0.. lookup m v \ i then 0 else MPoly_Type.monom (Poly_Mapping.update v 0 m) a)" proof - have *: "{x. x = m \ lookup x v = i} = (if lookup m v = i then {m} else {})" by auto show ?thesis by (transfer' fixing: m a v i) (simp add:*) qed lemma isolate_variable_sparse_monom_mult: "isolate_variable_sparse (MPoly_Type.monom m a * q) v n = (if n \ lookup m v then MPoly_Type.monom (Poly_Mapping.update v 0 m) a * isolate_variable_sparse q v (n - lookup m v) else 0)" for q::"'a::semiring_no_zero_divisors mpoly" apply (auto simp: MPoly_Type.mult_monom) subgoal apply transfer' subgoal for mon v i a q apply (auto simp add: monomials_monom_mult sum_distrib_left) apply (rule sum.reindex_bij_witness_not_neutral[where j="\a. a - mon" and i="\a. mon + a" and S'="{}" and T'="{}" ]) apply (auto simp: lookup_add) apply (auto simp: poly_mapping_eq_iff fun_eq_iff single.rep_eq Poly_Mapping.mult_single) apply (auto simp: when_def More_MPoly_Type.coeff_monom_mult) by (auto simp: lookup_update lookup_add split: if_splits) done subgoal apply transfer' apply (auto intro!: sum.neutral simp: monomials_monom_mult ) apply (rule poly_mapping_eqI) apply (auto simp: lookup_single when_def) by (simp add: lookup_add) done lemma isolate_variable_sparse_mult: "isolate_variable_sparse (p * q) v n = (\i\n. isolate_variable_sparse p v i * isolate_variable_sparse q v (n - i))" for p q::"'a::semiring_no_zero_divisors mpoly" proof (induction p rule: mpoly_induct) case (monom m a) then show ?case by (cases "a = 0") (auto simp add: mpoly_eq_iff coeff_sum coeff_mult if_conn if_distrib if_distribR isolate_variable_sparse_monom isolate_variable_sparse_monom_mult cong: if_cong) next case (sum p1 p2 m a) then show ?case by (simp add: distrib_right isovarspar_sum sum.distrib) qed subsubsection "More Miscellaneous" lemma var_not_in_Const : "var\vars (Const x :: real mpoly)" unfolding vars_def keys_def by (smt UN_iff coeff_def coeff_isolate_variable_sparse const_lookup_zero keys_def lookup_eq_zero_in_keys_contradict) lemma mpoly_to_nested_poly_mult: "mpoly_to_nested_poly (p * q) v = mpoly_to_nested_poly p v * mpoly_to_nested_poly q v" for p q::"'a::{comm_semiring_0, semiring_no_zero_divisors} mpoly" by (auto simp: poly_eq_iff coeff_mult mpoly_to_nested_poly.rep_eq isolate_variable_sparse_mult) lemma get_if_const_insertion : assumes "get_if_const (p::real mpoly) = Some r" shows "insertion f p = r" proof- have "Set.is_empty (vars p)" apply(cases "Set.is_empty (vars p)") apply(simp) using assms by(simp) then have "(MPoly_Type.coeff p 0) = r" using assms by simp then show ?thesis by (metis Set.is_empty_def \Set.is_empty (vars p)\ empty_iff insertion_irrelevant_vars insertion_trivial) qed subsubsection "Yet more Degree Lemmas" lemma degree_mult: fixes p q::"'a::{comm_semiring_0, ring_1_no_zero_divisors} mpoly" assumes "p \ 0" assumes "q \ 0" shows "MPoly_Type.degree (p * q) v = MPoly_Type.degree p v + MPoly_Type.degree q v" using assms by (auto simp add: degree_eq_univariate_degree mpoly_to_nested_poly_mult Polynomial.degree_mult_eq) lemma degree_eq: assumes "(p::real mpoly) = (q:: real mpoly)" shows "MPoly_Type.degree p x = MPoly_Type.degree q x" by (simp add: assms) lemma degree_var_i : "MPoly_Type.degree (((Var x)^i:: real mpoly)) x = i" proof (induct i) case 0 then show ?case using degree_const by simp next case (Suc i) have multh: "(Var x)^(Suc i) = ((Var x)^i::real mpoly) * (Var x:: real mpoly)" using power_Suc2 by blast have deg0h: "MPoly_Type.degree 0 x = 0" by simp have deg1h: "MPoly_Type.degree (Var x::real mpoly) x = 1" using degree_one by blast have nonzeroh1: "(Var x:: real mpoly) \ 0" using degree_eq deg0h deg1h by auto have nonzeroh2: "((Var x)^i:: real mpoly) \ 0" using degree_eq deg0h Suc.hyps by (metis one_neq_zero power_0) have sumh: "(MPoly_Type.degree (((Var x)^i:: real mpoly) * (Var x:: real mpoly)) x) = (MPoly_Type.degree ((Var x)^i::real mpoly) x) + (MPoly_Type.degree (Var x::real mpoly) x)" using degree_mult[where p = "(Var x)^i::real mpoly", where q = "Var x::real mpoly"] nonzeroh1 nonzeroh2 by blast then show ?case using sumh deg1h by (metis Suc.hyps Suc_eq_plus1 multh) qed lemma degree_less_sum: assumes "MPoly_Type.degree (p::real mpoly) var = n" assumes "MPoly_Type.degree (q::real mpoly) var = m" assumes "m < n" shows "MPoly_Type.degree (p + q) var = n" proof - have h1: "n > 0" using assms(3) neq0_conv by blast have h2: "(\m\monomials p. lookup m var = n) \ (\m\monomials p. lookup m var \ n)" using degree_eq_iff assms(1) by (metis degree_ge_iff h1 order_refl) have h3: "(\m\monomials (p + q). lookup m var = n)" using h2 by (metis MPolyExtension.coeff_add add.right_neutral assms(2) assms(3) coeff_eq_zero_iff degree_not_exists_monom) have h4: "(\m\monomials (p + q). lookup m var \ n)" using h2 assms(3) assms(2) using degree_add_leq degree_le_iff dual_order.strict_iff_order by blast show ?thesis using degree_eq_iff h3 h4 by (metis assms(3) gr_implies_not0) qed lemma degree_less_sum': assumes "MPoly_Type.degree (p::real mpoly) var = n" assumes "MPoly_Type.degree (q::real mpoly) var = m" assumes "n < m" shows "MPoly_Type.degree (p + q) var = m" using degree_less_sum[OF assms(2) assms(1) assms(3)] by (simp add: add.commute) (* Result on the degree of the derivative *) lemma nonzero_const_is_nonzero: assumes "(k::real) \ 0" shows "((Const k)::real mpoly) \ 0" by (metis MPoly_Type.insertion_zero assms insertion_const) lemma degree_derivative : assumes "MPoly_Type.degree p x > 0" shows "MPoly_Type.degree p x = MPoly_Type.degree (derivative x p) x + 1" proof - define f where "f i = (isolate_variable_sparse p x (i+1) * (Var x)^(i) * (Const (i+1)))" for i define n where "n = MPoly_Type.degree p x-1" have d : "\m\monomials p. lookup m x = n+1" using n_def degree_eq_iff assms by (metis add.commute less_not_refl2 less_one linordered_semidom_class.add_diff_inverse) have h1a : "\i. MPoly_Type.degree (isolate_variable_sparse p x i) x = 0" by (simp add: not_in_isovarspar varNotIn_degree) have h1b : "\i::nat. MPoly_Type.degree ((Var x)^i:: real mpoly) x = i" using degree_var_i by auto have h1c1 : "\i. (Var(x)^(i)::real mpoly)\0" by (metis MPoly_Type.degree_zero h1b power_0 zero_neq_one) fix i have h1c1var: "((Var x)^i:: real mpoly) \ 0" using h1c1 by blast have h1c2 : "((Const ((i::nat) + 1))::real mpoly)\0" using nonzero_const_is_nonzero by auto have h1c3: "(isolate_variable_sparse p x (n + 1)) \ 0" using d by (metis One_nat_def Suc_pred add.commute assms isolate_variable_sparse_degree_eq_zero_iff n_def not_gr_zero not_in_isovarspar plus_1_eq_Suc varNotIn_degree) have h3: "(isolate_variable_sparse p x (i+1) = 0) \ (MPoly_Type.degree (f i) x) = 0" by (simp add: f_def) have degh1: "(MPoly_Type.degree (((Const ((i::nat)+1))::real mpoly)*(Var x)^i) x) = ((MPoly_Type.degree ((Const (i+1))::real mpoly) x) + (MPoly_Type.degree ((Var x)^i:: real mpoly) x))" using h1c2 h1c1var degree_mult[where p="((Const ((i::nat)+1))::real mpoly)", where q="((Var x)^i:: real mpoly)"] by blast then have degh2: "(MPoly_Type.degree (((Const ((i::nat)+1))::real mpoly)*(Var x)^i) x) = i" using degree_var_i degree_const by simp have nonzerohyp: "(((Const ((i::nat)+1))::real mpoly)*(Var x)^i) \ 0" proof (induct i) case 0 then show ?case by (simp add: nonzero_const_is_nonzero) next case (Suc i) then show ?case using degree_eq degh2 by (metis Suc_eq_plus1 h1c1 mult_eq_0_iff nat.simps(3) nonzero_const_is_nonzero of_nat_eq_0_iff) qed have h4a1: "(isolate_variable_sparse p x (i+1) \ 0) \ (MPoly_Type.degree (isolate_variable_sparse p x (i+1) * ((Var x)^(i) * (Const (i+1)))::real mpoly) x = (MPoly_Type.degree (isolate_variable_sparse p x (i+1):: real mpoly) x) + (MPoly_Type.degree (((Const (i+1)) * (Var x)^(i))::real mpoly) x))" using nonzerohyp degree_mult[where p = "(isolate_variable_sparse p x (i+1))::real mpoly", where q = "((Const (i+1)) * (Var x)^(i)):: real mpoly", where v = "x"] by (simp add: mult.commute) have h4: "(isolate_variable_sparse p x (i+1) \ 0) \ (MPoly_Type.degree (f i) x) = i" using h4a1 f_def degh2 h1a by (metis (no_types, lifting) degh1 degree_const h1b mult.commute mult.left_commute of_nat_1 of_nat_add) have h5: "(MPoly_Type.degree (f i) x) \ i" using h3 h4 by auto have h6: "(MPoly_Type.degree (f n) x) = n" using h1c3 h4 by (smt One_nat_def add.right_neutral add_Suc_right degree_const degree_mult divisors_zero f_def h1a h1b h1c1 mult.commute nonzero_const_is_nonzero of_nat_1 of_nat_add of_nat_neq_0) have h7a: "derivative x p = (\i\{0..MPoly_Type.degree p x-1}. isolate_variable_sparse p x (i+1) * (Var x)^i * (Const (i+1)))" using derivative_def by auto have h7b: "(\i\{0..MPoly_Type.degree p x-1}. isolate_variable_sparse p x (i+1) * (Var x)^i * (Const (i+1))) = (\i\{0..MPoly_Type.degree p x-1}. (f i))" using f_def by (metis Suc_eq_plus1 add.commute semiring_1_class.of_nat_simps(2)) have h7c: "derivative x p = (\i\{0..MPoly_Type.degree p x-1}. (f i))" using h7a h7b by auto have h7: "derivative x p = (\i\{0..n}. (f i))" using n_def h7c by blast have h8: "n > 0 \ ((MPoly_Type.degree (\i\{0..(n-1)}. (f i)) x) < n)" proof (induct n) case 0 then show ?case by blast next case (Suc n) then show ?case using h5 degree_less_sum by (smt add_cancel_right_right atLeastAtMost_iff degree_const degree_mult degree_sum_less degree_var_i diff_Suc_1 f_def h1a le_imp_less_Suc mult.commute mult_eq_0_iff) qed have h9a: "n = 0 \ MPoly_Type.degree (\i\{0..n}. (f i)) x = n" using h6 by auto have h9b: "n > 0 \ MPoly_Type.degree (\i\{0..n}. (f i)) x = n" proof - have h9bhyp: "n > 0 \ (MPoly_Type.degree (\i\{0..n}. (f i)) x = MPoly_Type.degree ((\i\{0..(n-1)}. (f i)) + (f n)) x)" by (metis Suc_diff_1 sum.atLeast0_atMost_Suc) have h9bhyp2: "n > 0 \ ((MPoly_Type.degree ((\i\{0..(n-1)}. (f i)) + (f n)) x) = n)" using h6 h8 degree_less_sum by (simp add: add.commute) then show ?thesis using h9bhyp2 h9bhyp by linarith qed have h9: "MPoly_Type.degree (\i\{0..n}. (f i)) x = n" using h9a h9b by blast have h10: "MPoly_Type.degree (derivative x p) x = n" using h9 h7 by simp show ?thesis using h10 n_def using assms by linarith qed lemma express_poly : assumes h : "MPoly_Type.degree (p::real mpoly) var = 1 \ MPoly_Type.degree p var = 2" shows "p = (isolate_variable_sparse p var 2)*(Var var)^2 +(isolate_variable_sparse p var 1)*(Var var) +(isolate_variable_sparse p var 0)" proof- have h1a: "MPoly_Type.degree p var = 1 \ p = isolate_variable_sparse p var 0 + isolate_variable_sparse p var 1 * Var var" using sum_over_zero[where mp="p",where x="var"] by auto have h1b: "MPoly_Type.degree p var = 1 \ isolate_variable_sparse p var 2 = 0" using isovar_greater_degree by (simp add: isovar_greater_degree) have h1: "MPoly_Type.degree p var = 1 \ p = isolate_variable_sparse p var 0 + isolate_variable_sparse p var 1 * Var var + isolate_variable_sparse p var 2 * (Var var)^2" using h1a h1b by auto have h2a: "MPoly_Type.degree p var = 2 \ p = (\i::nat \ 2. isolate_variable_sparse p var i * Var var^i)" using sum_over_zero[where mp="p", where x="var"] by auto have h2b: "(\i::nat \ 2. isolate_variable_sparse p var i * Var var^i) = (\i::nat \ 1. isolate_variable_sparse p var i * Var var^i) + isolate_variable_sparse p var 2 * (Var var)^2" apply auto by (simp add: numerals(2)) have h2: "MPoly_Type.degree p var = 2 \ p = isolate_variable_sparse p var 0 + isolate_variable_sparse p var 1 * Var var + isolate_variable_sparse p var 2 * (Var var)^2" using h2a h2b by auto have h3: "isolate_variable_sparse p var 0 + isolate_variable_sparse p var 1 * Var var + isolate_variable_sparse p var 2 * (Var var)^2 = isolate_variable_sparse p var 2 * (Var var)^2 + isolate_variable_sparse p var 1 * Var var + isolate_variable_sparse p var 0" by (simp add: add.commute) show ?thesis using h h1 h2 h3 by presburger qed lemma degree_isovarspar : "MPoly_Type.degree (isolate_variable_sparse (p::real mpoly) x i) x = 0" using not_in_isovarspar varNotIn_degree by blast end diff --git a/thys/Virtual_Substitution/NegInfinityUni.thy b/thys/Virtual_Substitution/NegInfinityUni.thy --- a/thys/Virtual_Substitution/NegInfinityUni.thy +++ b/thys/Virtual_Substitution/NegInfinityUni.thy @@ -1,632 +1,578 @@ theory NegInfinityUni imports UniAtoms NegInfinity QE begin fun allZero' :: "real * real * real \ atomUni fmUni" where "allZero' (a,b,c) = AndUni(AndUni(AtomUni(EqUni(0,0,a))) (AtomUni(EqUni(0,0,b)))) (AtomUni(EqUni(0,0,c)))" lemma convert_allZero : assumes "convert_poly var p (xs'@x#xs) = Some p'" assumes "length xs' = var" shows "eval (allZero p var) (xs'@x#xs) = evalUni (allZero' p') x" proof(cases p') case (fields a b c) then show ?thesis proof(cases "MPoly_Type.degree p var = 0") case True then show ?thesis using assms fields by(simp add:eval_list_conj isovar_greater_degree) next case False then have nonzero : "MPoly_Type.degree p var \ 0" by auto then show ?thesis proof(cases "MPoly_Type.degree p var = 1") case True then show ?thesis using assms fields apply(simp add:eval_list_conj isovar_greater_degree) by (metis) next case False then have degree2 : "MPoly_Type.degree p var = 2" using degree_convert_eq[OF assms(1)] nonzero by auto then show ?thesis using assms apply(simp add:eval_list_conj isovar_greater_degree) using insertion_isovarspars_free list_update_code(2) apply auto by (metis One_nat_def Suc_1 less_2_cases less_Suc_eq numeral_3_eq_3) qed qed qed fun alternateNegInfinity' :: "real * real * real \ atomUni fmUni" where "alternateNegInfinity' (a,b,c) = OrUni(AtomUni(LessUni(0,0,a)))( AndUni(AtomUni(EqUni(0,0,a))) ( OrUni(AtomUni(LessUni(0,0,-b)))( AndUni(AtomUni(EqUni(0,0,b)))( AtomUni(LessUni(0,0,c)) )) )) " lemma convert_alternateNegInfinity : assumes "convert_poly var p (xs'@x#xs) = Some X" assumes "length xs' = var" shows "eval (alternateNegInfinity p var) (xs'@x#xs) = evalUni (alternateNegInfinity' X) x" proof(cases X) case (fields a b c) then show ?thesis proof(cases "MPoly_Type.degree p var = 0") case True then show ?thesis using assms apply (simp add: isovar_greater_degree) apply auto apply (metis aEval.simps(2) eval.simps(1) eval_and eval_false eval_or mult_one_left) by (metis aEval.simps(2) eval.simps(1) eval_or mult_one_left) next case False then have nonzero : "MPoly_Type.degree p var \ 0" by auto then show ?thesis proof(cases "MPoly_Type.degree p var = 1") case True have letbind: "eval (let a_n = isolate_variable_sparse p var (Suc 0) in or (fm.Atom (Less (Const (- 1) * a_n))) (and (fm.Atom (Eq a_n)) (let a_n = isolate_variable_sparse p var 0 in or (fm.Atom (Less (Const 1 * a_n))) (and (fm.Atom (Eq a_n)) FalseF)))) (xs'@x#xs) = eval (or (fm.Atom (Less (Const (- 1) * (isolate_variable_sparse p var (Suc 0))))) (and (fm.Atom (Eq (isolate_variable_sparse p var (Suc 0)))) (or (fm.Atom (Less (Const 1 * (isolate_variable_sparse p var 0)))) (and (fm.Atom (Eq (isolate_variable_sparse p var 0))) FalseF)))) (xs'@x#xs)" by meson show ?thesis using assms True unfolding fields by (simp add: isovar_greater_degree letbind eval_or eval_and insertion_mult insertion_const) next case False then have degree2 : "MPoly_Type.degree p var = 2" using degree_convert_eq[OF assms(1)] nonzero by auto have "[0..<3] = [0,1,2]" by (simp add: upt_rec) then have unfold : " (foldl (\F i. let a_n = isolate_variable_sparse p var i in or (fm.Atom (Less ((if i mod 2 = 0 then Const 1 else Const (- 1)) * a_n))) (and (fm.Atom (Eq a_n)) F)) FalseF [0..<3]) = (let a_n = isolate_variable_sparse p var 2 in or (fm.Atom (Less ((Const 1) * a_n))) (and (fm.Atom (Eq a_n)) (let a_n = isolate_variable_sparse p var (Suc 0) in or (fm.Atom (Less (Const (- 1) * a_n))) (and (fm.Atom (Eq a_n)) (let a_n = isolate_variable_sparse p var 0 in or (fm.Atom (Less (Const 1 * a_n))) (and (fm.Atom (Eq a_n)) FalseF))))))" by auto have letbind : "eval (foldl (\F i. let a_n = isolate_variable_sparse p var i in or (fm.Atom (Less ((if i mod 2 = 0 then Const 1 else Const (- 1)) * a_n))) (and (fm.Atom (Eq a_n)) F)) FalseF [0..<3]) (xs'@x#xs) = eval (or (fm.Atom (Less ( Const 1 * (isolate_variable_sparse p var 2)))) (and (fm.Atom (Eq (isolate_variable_sparse p var 2))) (or (fm.Atom (Less (Const (- 1) * (isolate_variable_sparse p var (Suc 0))))) (and (fm.Atom (Eq (isolate_variable_sparse p var (Suc 0)))) (or (fm.Atom (Less (Const 1 * (isolate_variable_sparse p var 0)))) (and (fm.Atom (Eq (isolate_variable_sparse p var 0))) FalseF)))))) (xs'@x#xs)" apply (simp add:unfold) by metis show ?thesis using degree2 assms unfolding fields by (simp add: isovar_greater_degree eval_or eval_and letbind insertion_mult insertion_const) qed qed qed fun substNegInfinityUni :: "atomUni \ atomUni fmUni" where "substNegInfinityUni (EqUni p) = allZero' p " | "substNegInfinityUni (LessUni p) = alternateNegInfinity' p"| "substNegInfinityUni (LeqUni p) = OrUni (alternateNegInfinity' p) (allZero' p)"| "substNegInfinityUni (NeqUni p) = negUni (allZero' p)" lemma convert_substNegInfinity : assumes "convert_atom var A (xs'@x#xs) = Some(A')" assumes "length xs' = var" shows "eval (substNegInfinity var A) (xs'@x#xs) = evalUni (substNegInfinityUni A') x" using assms proof(cases A) case (Less p) have "\X. convert_poly var p (xs' @ x # xs) = Some X" using assms Less apply(cases "MPoly_Type.degree p var < 3") by (simp_all) then obtain X where X_def: "convert_poly var p (xs' @ x # xs) = Some X" by auto then have A' : "A' = LessUni X" using assms Less apply(cases "MPoly_Type.degree p var < 3") by (simp_all) show ?thesis unfolding Less substNegInfinity.simps unfolding convert_alternateNegInfinity[OF X_def assms(2)] A' apply(cases X) by simp next case (Eq p) then show ?thesis using assms convert_allZero by auto next case (Leq p) define p' where "p' = (case convert_poly var p (xs'@x#xs) of Some p' \ p')" have A'_simp : "A' = LeqUni p'" using assms Leq using p'_def by auto have allZ : "eval (allZero p var) (xs'@x#xs) = evalUni (allZero' p') x" using convert_allZero using Leq assms p'_def by auto have altNeg : "eval (alternateNegInfinity p var) (xs'@x#xs) = evalUni (alternateNegInfinity' p') x" using convert_alternateNegInfinity using Leq assms p'_def by auto show ?thesis unfolding Leq substNegInfinity.simps eval_Or A'_simp substNegInfinityUni.simps evalUni.simps using allZ altNeg by auto next case (Neq p) then show ?thesis using assms convert_allZero convert_neg by auto qed lemma change_eval_eq: fixes x y:: "real" assumes "((aEvalUni (EqUni(a,b,c)) x \ aEvalUni (EqUni(a,b,c)) y) \ x < y)" shows "(\w. x \ w \ w \ y \ a*w^2 + b*w + c = 0)" using assms by auto lemma change_eval_lt: fixes x y:: "real" assumes "((aEvalUni (LessUni (a,b,c)) x \ aEvalUni (LessUni (a,b,c)) y) \ x < y)" shows "(\w. x \ w \ w \ y \ a*w^2 + b*w + c = 0)" proof - let ?p = "[:c, b, a:]" have "sign ?p x \ sign ?p y" using assms unfolding sign_def apply (simp add: distrib_left mult.commute mult.left_commute power2_eq_square) by linarith then have "(\w \ (root_list ?p). x \ w \ w \ y)" using changes_sign assms by auto then obtain w where w_prop: "w \ (root_list ?p) \ x \ w \ w \ y" by auto then have "a*w^2 + b*w + c = 0" unfolding root_list_def using add.commute distrib_right mult.assoc mult.commute power2_eq_square using quadratic_poly_eval by force then show ?thesis using w_prop by auto qed lemma no_change_eval_lt: fixes x y:: "real" assumes "x < y" assumes "\(\w. x \ w \ w \ y \ a*w^2 + b*w + c = 0)" shows "((aEvalUni (LessUni (a,b,c)) x = aEvalUni (LessUni (a,b,c)) y))" using change_eval_lt using assms(1) assms(2) by blast lemma change_eval_neq: fixes x y:: "real" assumes "((aEvalUni (NeqUni (a,b,c)) x \ aEvalUni (NeqUni (a,b,c)) y) \ x < y)" shows "(\w. x \ w \ w \ y \ a*w^2 + b*w + c = 0)" using assms by auto lemma change_eval_leq: fixes x y:: "real" assumes "((aEvalUni (LeqUni (a,b,c)) x \ aEvalUni (LeqUni (a,b,c)) y) \ x < y)" shows "(\w. x \ w \ w \ y \ a*w^2 + b*w + c = 0)" proof - let ?p = "[:c, b, a:]" have "sign ?p x \ sign ?p y" using assms unfolding sign_def apply (simp add: distrib_left mult.commute mult.left_commute power2_eq_square) by linarith then have "(\w \ (root_list ?p). x \ w \ w \ y)" using changes_sign assms by auto then obtain w where w_prop: "w \ (root_list ?p) \ x \ w \ w \ y" by auto then have "a*w^2 + b*w + c = 0" unfolding root_list_def using add.commute distrib_right mult.assoc mult.commute power2_eq_square using quadratic_poly_eval by force then show ?thesis using w_prop by auto qed lemma change_eval: fixes x y:: "real" fixes At:: "atomUni" assumes xlt: "x < y" assumes noteq: "((aEvalUni At) x \ aEvalUni (At) y)" assumes "getPoly At = (a, b, c)" shows "(\w. x \ w \ w \ y \ a*w^2 + b*w + c = 0)" proof - have four_types: "At = (EqUni (a,b,c)) \ At = (LessUni (a,b,c)) \ At = (LeqUni (a,b,c)) \ At = (NeqUni (a,b,c))" using getPoly.elims assms(3) by force have eq_h: "At = (EqUni (a,b,c)) \ (\w. x \ w \ w \ y \ a*w^2 + b*w + c = 0)" using assms(1) assms(2) change_eval_eq by blast have less_h: "At = (LessUni(a,b,c)) \ (\w. x \ w \ w \ y \ a*w^2 + b*w + c = 0)" using change_eval_lt assms by blast have leq_h: "At = (LeqUni(a,b,c)) \ (\w. x \ w \ w \ y \ a*w^2 + b*w + c = 0)" using change_eval_leq assms by blast have neq_h: "At = (NeqUni(a,b,c)) \ (\w. x \ w \ w \ y \ a*w^2 + b*w + c = 0)" using change_eval_neq assms by blast show ?thesis using four_types eq_h less_h leq_h neq_h by blast qed lemma no_change_eval: fixes x y:: "real" assumes "getPoly At = (a, b, c)" assumes "x < y" assumes "\(\w. x \ w \ w \ y \ a*w^2 + b*w + c = 0)" shows "((aEvalUni At) x = aEvalUni (At) y \ x < y)" using change_eval using assms(1) assms(2) assms(3) by blast lemma same_eval'' : assumes "getPoly At = (a, b, c)" assumes nonz: "a \ 0 \ b \ 0 \ c \ 0" shows "\x. \yy. poly ?p y = a*y^2 + b*y + c" by (simp add: distrib_left power2_eq_square) have "?p \ 0" using nonz by auto then have "finite {y. poly ?p y = 0}" using poly_roots_finite by blast then have "finite {y. c + (a * y\<^sup>2 + b * y) = 0} \ \y. y * (b + y * a) = a * y\<^sup>2 + b * y \ finite {y. a * y\<^sup>2 + b * y + c = 0}" proof - assume a1: "finite {y. c + (a * y\<^sup>2 + b * y) = 0}" have "\x0. c + (a * x0\<^sup>2 + b * x0) = a * x0\<^sup>2 + b * x0 + c" by simp then show ?thesis using a1 by presburger qed then have finset: "finite {y. a*y^2 + b*y + c = 0}" using poly_eval by (metis \finite {y. poly [:c, b, a:] y = 0}\ poly_roots_set_same) then have "\x. (\y. a*y^2 + b*y + c = 0 \ x < y)" proof - let ?l = "sorted_list_of_set {y. a*y^2 + b*y + c = 0}" have "\x. x < ?l ! 0" using infzeros nonz by blast then obtain x where x_prop: "x < ?l! 0" by auto then have "\ y. List.member ?l y \ x < y" proof clarsimp fix y assume "List.member ?l y" then have "\n. n < length ?l \ ?l ! n = y" by (meson in_set_conv_nth in_set_member) then obtain n where n_prop: "n < length ?l \ ?l ! n = y" by auto have h: "\n < length ?l. ?l ! n \ ?l !0" using sorted_iff_nth_mono using sorted_sorted_list_of_set by blast then have h: "y \ ?l ! 0" using n_prop by auto then show "x < y" using x_prop by auto qed then show ?thesis using finset set_sorted_list_of_set in_set_member by (metis (mono_tags, lifting) mem_Collect_eq) qed then obtain x where x_prop: "(\y. a*y^2 + b*y + c = 0 \ x < y)" by auto then have same_as: "\yx0. (x0 < x) = (\ 0 \ x0 + - 1 * x)" - by auto - have f2: "(0 \ - 1 * x + v0_0) = (x + - 1 * v0_0 \ 0)" - by auto - have f3: "v0_0 + - 1 * x = - 1 * x + v0_0" - by auto - have f4: "\x0 x1 x2 x3. (x3::real) * x0\<^sup>2 + x2 * x0 + x1 = x1 + x3 * x0\<^sup>2 + x2 * x0" - by auto - have f5: "\x3 x4 x5. (aEvalUni x3 x5 \ aEvalUni x3 x4) = ((\ aEvalUni x3 x5) = aEvalUni x3 x4)" - by fastforce - have f6: "\x0 x1 x2 x3 x4 x5. (x5 < x4 \ (\ aEvalUni x3 x5) = aEvalUni x3 x4 \ getPoly x3 = (x2, x1, x0) \ (\v6\x5. v6 \ x4 \ x0 + x2 * v6\<^sup>2 + x1 * v6 = 0)) = ((\ x5 < x4 \ (\ aEvalUni x3 x5) \ aEvalUni x3 x4 \ getPoly x3 \ (x2, x1, x0)) \ (\v6\x5. v6 \ x4 \ x0 + x2 * v6\<^sup>2 + x1 * v6 = 0))" - by fastforce - have f7: "\x0 x5. ((x0::real) \ x5) = (x0 + - 1 * x5 \ 0)" - by auto - have f8: "\x0 x6. ((x6::real) \ x0) = (0 \ x0 + - 1 * x6)" - by auto - have "\x4 x5. ((x5::real) < x4) = (\ x4 + - 1 * x5 \ 0)" - by auto - then have "(\r ra a rb rc rd. r < ra \ aEvalUni a r \ aEvalUni a ra \ getPoly a = (rb, rc, rd) \ (\re\r. re \ ra \ rb * re\<^sup>2 + rc * re + rd = 0)) = (\r ra a rb rc rd. (ra + - 1 * r \ 0 \ (\ aEvalUni a r) \ aEvalUni a ra \ getPoly a \ (rb, rc, rd)) \ (\re. 0 \ re + - 1 * r \ re + - 1 * ra \ 0 \ rd + rb * re\<^sup>2 + rc * re = 0))" - using f8 f7 f6 f5 f4 by presburger - then have f9: "\r ra a rb rc rd. (ra + - 1 * r \ 0 \ (\ aEvalUni a r) \ aEvalUni a ra \ getPoly a \ (rb, rc, rd)) \ (\re. 0 \ re + - 1 * r \ re + - 1 * ra \ 0 \ rd + rb * re\<^sup>2 + rc * re = 0)" - by (meson change_eval) - obtain rr :: "real \ real \ real \ real \ real \ real" where - "\x0 x1 x2 x4 x5. (\v6. 0 \ v6 + - 1 * x5 \ v6 + - 1 * x4 \ 0 \ x0 + x2 * v6\<^sup>2 + x1 * v6 = 0) = (0 \ rr x0 x1 x2 x4 x5 + - 1 * x5 \ rr x0 x1 x2 x4 x5 + - 1 * x4 \ 0 \ x0 + x2 * (rr x0 x1 x2 x4 x5)\<^sup>2 + x1 * rr x0 x1 x2 x4 x5 = 0)" - by moura - then have f10: "\r ra a rb rc rd. ra + - 1 * r \ 0 \ aEvalUni a r = aEvalUni a ra \ getPoly a \ (rb, rc, rd) \ r + - 1 * rr rd rc rb ra r \ 0 \ 0 \ ra + - 1 * rr rd rc rb ra r \ rd + rb * (rr rd rc rb ra r)\<^sup>2 + rc * rr rd rc rb ra r = 0" - using f9 by simp - have f11: "(rr c b a x v0_0 + - 1 * x \ 0) = (0 \ x + - 1 * rr c b a x v0_0)" - by force - have "\x0. (x < x0) = (\ x0 + - 1 * x \ 0)" - by auto - then have f12: "\r. c + a * r\<^sup>2 + b * r \ 0 \ \ r + - 1 * x \ 0" - using x_prop by auto - obtain rra :: real where - "(\v0. \ 0 \ v0 + - 1 * x \ aEvalUni At v0 \ aEvalUni At x) = (\ 0 \ rra + - 1 * x \ aEvalUni At rra \ aEvalUni At x)" - by moura - then show ?thesis - using f12 f11 f10 f3 f2 f1 - proof - - have f1: "\x0. (x0 < x) = (\ 0 \ x0 + - 1 * x)" - by auto - have f2: "(0 \ v0_0a + - 1 * x) = (x + - 1 * v0_0a \ 0)" - by auto - have f3: "(rr c b a x v0_0a + - 1 * x \ 0) = (0 \ x + - 1 * rr c b a x v0_0a)" - by auto - obtain rrb :: real where - "(\v0. \ 0 \ v0 + - 1 * x \ aEvalUni At v0 \ aEvalUni At x) = (\ 0 \ rrb + - 1 * x \ aEvalUni At rrb \ aEvalUni At x)" - by blast - then show ?thesis - using f3 f2 f1 assms(1) f10 f12 - by smt - qed - qed + using no_change_eval change_eval assms by (smt (verit, ccfv_SIG)) then show ?thesis by auto qed lemma inequality_case : "(\(x::real). \(y::real)2 + (b::real) * y + (c::real) < 0) = (a < 0 \ a = 0 \ (0 < b \ b = 0 \ c < 0))" proof- let ?At = "(LessUni (a, b, c))" have firsth : "\x. (\y2 + b * y + c < 0 \ a\0)" proof - fix x assume lt: "\y2 + b * y + c < 0" have "\w. \y < w. y^2 > (-b/a)*y - c/a" using ysq_dom_y_plus_coeff[where b = "-b/a", where c = "-c/a"] by auto then have gtdomhelp: "a > 0 \ \w. \y < w. a*y^2 > a*((-b/a)*y - c/a)" by auto have "\y. (a > 0 \ a*((-b/a)*y - c/a) = -b*y - c)" by (simp add: right_diff_distrib') then have gtdom: "a > 0 \ \w.\y < w. a*y^2 > -b*y - c" using gtdomhelp by simp then have " a > 0 \ False" proof - assume "a > 0" then have "\w.\y < w. a*y^2 > -b*y - c" using gtdom by auto then obtain w where w_prop: "\y < w. a*y^2 + b*y + c > 0" by fastforce let ?mn = "min w x - 1" have gtz: "a*?mn^2 + b*?mn + c > 0" using w_prop by auto have ltz: "a*?mn^2 + b*?mn + c < 0" using lt by auto then show "False" using gtz ltz by auto qed then show "a \ 0" by fastforce qed have bleq0 : "\x. (\y b\0)" proof - fix x assume lt: "\y \w.\y < w. b*y > - c" by (metis mult.commute neg_less_divide_eq) then have "b < 0 \ False" proof - assume "b < 0" then have "\w.\y < w. b*y > - c" using gtdom by auto then obtain w where w_prop: "\y < w .b*y + c > 0" by fastforce let ?mn = "min w x - 1" have gtz: "b*?mn + c > 0" using w_prop by auto have ltz: "b*?mn + c < 0" using lt by auto then show "False" using gtz ltz by auto qed then show "b \ 0" by fastforce qed have secondh: "\x. (\y2 + b * y + c < 0 \ \ a < 0 \ \ 0 < b \ b = 0)" using firsth bleq0 by (metis add.commute add.right_neutral less_eq_real_def mult_zero_class.mult_zero_left) have thirdh : "\x. \y2 + b * y + c < 0 \ \ a < 0 \ \ 0 < b \ c < 0" using firsth secondh add.commute add.right_neutral infzeros mult_zero_class.mult_zero_left not_numeral_le_zero order_refl by (metis less_eq_real_def) have fourthh : "a < 0 \ \x. \y2 + b * y + c < 0" proof - assume aleq: "a < 0" have "\(w::real). \(y::real). (y < w \ y^2 > (-b/a)*y + (-c/a))" using ysq_dom_y_plus_coeff[where b = "-b/a", where c = "-c/a"] by blast then have hyp:"\(w::real). \(y::real). (y < w \ a*y^2 \ a*(-b/a)*y + a*(-c/a))" by (metis (no_types, opaque_lifting) \a < 0\ distrib_left less_eq_real_def linorder_not_le mult.assoc mult_less_cancel_left) have "\y. a*(-b/a)*y + a*(-c/a) = -b*y -c" using \a < 0\ by auto then have "\(w::real). \(y::real). (y < w \ a*y^2 \ -b*y - c)" using hyp by auto then have "\(w::real). \(y::real). (y < w \ a*y^2 + b*y + c \ 0)" by (metis add.commute add_uminus_conv_diff le_diff_eq mult_minus_left real_add_le_0_iff) then obtain w where w_prop: "\(y::real). (y < w \ a*y^2 + b*y + c \ 0)" by auto have "\x. \y < x. aEvalUni ?At x = aEvalUni ?At y" using same_eval'' proof - have f1: "\x0 x1. ((x0::real) < x1) = (\ 0 \ x0 + - 1 * x1)" by linarith have "a \ 0" using \a < 0\ by force then obtain rr :: "atomUni \ real" where "\r. 0 \ r + - 1 * rr (LessUni (a, b, c)) \ aEvalUni (LessUni (a, b, c)) r = aEvalUni (LessUni (a, b, c)) (rr (LessUni (a, b, c)))" using f1 by (metis getPoly.simps(4) same_eval'') then show ?thesis using f1 by meson qed then obtain x where x_prop: "\y < x. aEvalUni ?At x = aEvalUni ?At y" by auto let ?mn = "min x w - 1" have "\y < ?mn. aEvalUni ?At y = True \ aEvalUni ?At y = False" using x_prop by auto have "\ y < ?mn. aEvalUni ?At y = False \ a*y^2 + b*y + c \ 0" by auto then have "\y. \y2 + b * y + c \ 0 \ y < min x w - 1 \ \ a * y\<^sup>2 + b * y + c < 0 \ a * y\<^sup>2 + b * y + c = 0" proof - fix y :: real assume a1: "y < min x w - 1" assume a2: "\ a * y\<^sup>2 + b * y + c < 0" assume a3: "\y2 + b * y + c \ 0" have "y < w" using a1 by linarith then show "a * y\<^sup>2 + b * y + c = 0" using a3 a2 less_eq_real_def by blast qed then have "\ y < ?mn. aEvalUni ?At y = False \ a*y^2 + b*y + c = 0" using w_prop by auto then have "\ y < ?mn. aEvalUni ?At y = False \ False" using infzeros aleq by (metis power_zero_numeral zero_less_power2) then have "\ y < ?mn. aEvalUni ?At y = True" proof - { fix rr :: real have "\r ra. (ra::real) < r \ \ ra < r + - 1" by linarith then have "\ rr < min x w - 1 \ aEvalUni (LessUni (a, b, c)) rr" by (metis (no_types) \\y False\ ab_group_add_class.ab_diff_conv_add_uminus less_eq_real_def min_less_iff_disj not_le x_prop) } then show ?thesis by blast qed then show ?thesis by auto qed have fifthh : "b > 0 \ \x. \y 0" show "\x. \y(x::real). \(y::real)2 + (b::real) * y + (c::real) > 0) = (a > 0 \ a = 0 \ (0 > b \ b = 0 \ c > 0))" proof - have s1: "\y. - 1 * a * y\<^sup>2 + - 1 * b * y + - 1 * c < 0 \ a * y\<^sup>2 + b * y + c > 0" by auto have s2: "(- 1 * a < 0 \ - 1 * a = 0 \ (0 < - 1 * b \ - 1 * b = 0 \ - 1 * c < 0)) \ (a > 0 \ a = 0 \ (0 > b \ b = 0 \ c > 0)) " by auto have "(\x. \y2 + - 1 * b * y + - 1 * c < 0) = (- 1 * a < 0 \ - 1 * a = 0 \ (0 < - 1 * b \ - 1 * b = 0 \ - 1 * c < 0))" using inequality_case[where a = "-1*a", where b = "-1*b", where c= "-1*c"] by auto then show ?thesis using s1 s2 by auto qed lemma infinity_evalUni_LessUni : "(\x. \yx. \yx. \yx. \yx. \y2 + b * y + c < 0) \ (\x. \y2 + b * y + c = 0)) \ (\x. \y2 + b * y + c \ 0)" using less_eq_real_def by auto have h2: "(\x. \y2 + b * y + c \ 0) \ ((\x. \y2 + b * y + c < 0) \ (\x. \y2 + b * y + c = 0))" proof - assume a1: "(\x. \y2 + b * y + c \ 0)" have "\(\x. \y2 + b * y + c = 0) \ (\x. \y2 + b * y + c < 0) " proof - assume a2: "\(\x. \y2 + b * y + c = 0)" then have "a \ 0 \ b \ 0 \ c \ 0" by auto then have "(a < 0 \ a = 0 \ (0 < b \ b = 0 \ c < 0))" proof - have x1: "a > 0 \ False" proof - assume "a > 0" then have "(\(x::real). \(y::real)2 + (b::real) * y + (c::real) > 0)" using inequality_case_geq by auto then show ?thesis using a1 by (meson a2 linorder_not_le min_less_iff_conj) qed then have x2: "a = 0 \ 0 > b \ False" proof - assume "a = 0 \ 0 > b" then have "(\(x::real). \(y::real)2 + (b::real) * y + (c::real) > 0)" using inequality_case_geq by blast then show ?thesis using a1 by (meson a2 linorder_not_le min_less_iff_conj) qed then have x3: "a = 0 \ b = 0 \ c > 0 \ False " using a1 a2 by auto show ?thesis using x1 x2 x3 by (meson \a \ 0 \ b \ 0 \ c \ 0\ linorder_neqE_linordered_idom) qed then show "(\x. \y2 + b * y + c < 0)" using inequality_case by auto qed then show ?thesis by auto qed have "(\x. \y2 + b * y + c \ 0) = (\x. \y2 + b * y + c < 0) \ (\x. \y2 + b * y + c = 0)" using h1 h2 by auto then show "(\x. \y2 + b * y + c \ 0) = (a < 0 \ a = 0 \ (0 < b \ b = 0 \ c < 0) \ a = 0 \ b = 0 \ c = 0)" using inequality_case[of a b c] infzeros[of _ a b c] by auto qed qed text "This is the vertical translation for substNegInfinityUni where we represent the virtual substution of negative infinity in the univariate case" lemma infinity_evalUni : shows "(\x. \y real \ int" where "sign p x \ (if poly p x = 0 then 0 else (if poly p x > 0 then 1 else -1))" definition sign_num:: "real \ int" where "sign_num x \ (if x = 0 then 0 else (if x > 0 then 1 else -1))" definition root_list:: "real Polynomial.poly \ real set" where "root_list p \ ({(x::real). poly p x = 0}::real set)" definition root_set:: "(real \ real \ real) set \ real set" where "root_set les \ ({(x::real). (\ (a, b, c) \ les. a*x^2 + b*x + c = 0)}::real set)" definition sorted_root_list_set:: "(real \ real \ real) set \ real list" where "sorted_root_list_set p \ sorted_list_of_set (root_set p)" definition nonzero_root_set:: "(real \ real \ real) set \ real set" where "nonzero_root_set les \ ({(x::real). (\ (a, b, c) \ les. (a \ 0 \ b \ 0) \ a*x^2 + b*x + c = 0)}::real set)" definition sorted_nonzero_root_list_set:: "(real \ real \ real) set \ real list" where "sorted_nonzero_root_list_set p \ sorted_list_of_set (nonzero_root_set p)" (* Important property of sorted lists *) lemma sorted_list_prop: fixes l::"real list" fixes x::"real" assumes sorted: "sorted l" assumes lengt: "length l > 0" assumes xgt: "x > l ! 0" assumes xlt: "x \ l ! (length l - 1)" shows "\n. (n+1) < (length l) \ x \ l !n \ x \ l ! (n + 1)" proof - have "\(\n. (n+1) < (length l) \ x \ l !n \ x \ l ! (n + 1)) \ False" proof clarsimp fix n assume alln: "\n. l ! n \ x \ Suc n < length l \ \ x \ l ! Suc n" have "\k. (k < length l \ x > l ! k)" proof clarsimp fix k show "k < length l \ l ! k < x" proof (induct k) case 0 then show ?case using xgt by auto next case (Suc k) then show ?case using alln using less_eq_real_def by auto qed qed then show "False" using xlt diff_Suc_less lengt not_less by (metis One_nat_def) qed then show ?thesis by auto qed subsection "Quadratic polynomial properties" lemma quadratic_poly_eval: fixes a b c::"real" fixes x::"real" shows "poly [:c, b, a:] x = a*x^2 + b*x + c" proof - have "x * (b + x * a) = a * x\<^sup>2 + b * x" by (metis add.commute distrib_right mult.assoc mult.commute power2_eq_square) then show ?thesis by auto qed lemma poly_roots_set_same: fixes a b c:: "real" shows "{(x::real). a * x\<^sup>2 + b * x + c = 0} = {x. poly [:c, b, a:] x = 0}" proof - have "\x. a*x^2 + b*x + c = poly [:c, b, a:] x" proof clarsimp fix x show "a * x\<^sup>2 + b * x = x * (b + x * a)" using quadratic_poly_eval[of c b a x] by auto qed then show ?thesis by auto qed lemma root_set_finite: assumes fin: "finite les" assumes nex: "\(\ (a, b, c) \ les. a = 0 \ b = 0 \ c = 0 )" shows "finite (root_set les)" proof - have "\(a, b, c) \ les. finite {(x::real). a*x^2 + b*x + c = 0}" proof clarsimp fix a b c assume "(a, b, c) \ les" then have "[:c, b, a:] \ 0" using nex by auto then have "finite {x. poly [:c, b, a:] x = 0}" using poly_roots_finite[where p = "[:c, b, a:]"] by auto then show "finite {x. a * x\<^sup>2 + b * x + c = 0}" using poly_roots_set_same by auto qed then show ?thesis using fin unfolding root_set_def by auto qed lemma nonzero_root_set_finite: assumes fin: "finite les" shows "finite (nonzero_root_set les)" proof - have "\(a, b, c) \ les. (a \ 0 \ b \ 0) \ finite {(x::real). a*x^2 + b*x + c = 0}" proof clarsimp fix a b c assume ins: "(a, b, c) \ les" assume "a = 0 \ b \ 0" then have "[:c, b, a:] \ 0" using ins by auto then have "finite {x. poly [:c, b, a:] x = 0}" using poly_roots_finite[where p = "[:c, b, a:]"] by auto then show "finite {x. a * x\<^sup>2 + b * x + c = 0}" using poly_roots_set_same by auto qed then show ?thesis using fin unfolding nonzero_root_set_def by auto qed lemma discriminant_lemma: fixes a b c r::"real" assumes aneq: "a \ 0" assumes beq: "b = 2 * a * r" assumes root: " a * r^2 - 2 * a * r*r + c = 0" shows "\x. a * x\<^sup>2 + b * x + c = 0 \ x = -r" proof - have "c = a*r^2" using root by (simp add: power2_eq_square) then have same: "b^2 - 4*a*c = (2 * a * r)^2 - 4*a*(a*r^2)" using beq by blast have "(2 * a * r)^2 = 4*a^2*r^2" by (simp add: mult.commute power2_eq_square) then have "(2 * a * r)^2 - 4*a*(a*(r)^2) = 0" using power2_eq_square by auto then have "b^2 - 4*a*c = 0" using same by simp then have "\x. a * x\<^sup>2 + b * x + c = 0 \ x = -b / (2 * a)" using discriminant_zero aneq unfolding discrim_def by auto then show ?thesis using beq by (simp add: aneq) qed (* Show a polynomial only changes sign when it passes through a root *) lemma changes_sign: fixes p:: "real Polynomial.poly" shows "\x::real. \ y::real. ((sign p x \ sign p y \ x < y) \ (\c \ (root_list p). x \ c \ c \ y))" proof clarsimp fix x y assume "sign p x \ sign p y" assume "x < y" then show "\c\root_list p. x \ c \ c \ y" using poly_IVT_pos[of x y p] poly_IVT_neg[of x y p] by (metis (mono_tags) \sign p x \ sign p y\ less_eq_real_def linorder_neqE_linordered_idom mem_Collect_eq root_list_def sign_def) qed (* Show a polynomial only changes sign if it passes through a root *) lemma changes_sign_var: fixes a b c x y:: "real" shows "((sign_num (a*x^2 + b*x + c) \ sign_num (a*y^2 + b*y + c) \ x < y) \ (\q. (a*q^2 + b*q + c = 0 \ x \ q \ q \ y)))" proof clarsimp assume sn: "sign_num (a * x\<^sup>2 + b * x + c) \ sign_num (a * y\<^sup>2 + b * y + c)" assume xy: " x < y" let ?p = "[:c, b, a:]" have cs: "((sign ?p x \ sign ?p y \ x < y) \ (\c \ (root_list ?p). x \ c \ c \ y))" using changes_sign[of ?p] by auto have "(sign ?p x \ sign ?p y \ x < y)" using sn xy unfolding sign_def sign_num_def using quadratic_poly_eval by presburger then have "(\c \ (root_list ?p). x \ c \ c \ y)" using cs by auto then obtain q where "q \ root_list ?p \ x \ q \ q \ y" by auto then have "a*q^2 + b*q + c = 0 \ x \ q \ q \ y" unfolding root_list_def using quadratic_poly_eval[of c b a q] by auto then show "\q. a * q\<^sup>2 + b * q + c = 0 \ x \ q \ q \ y" by auto qed subsection "Continuity Properties" lemma continuity_lem_eq0: fixes p::"real" shows "r < p \ \x\{r <..p}. a * x\<^sup>2 + b * x + c = 0 \ (a = 0 \ b = 0 \ c = 0)" proof - assume r_lt: "r < p" assume inf_zer: "\x\{r <..p}. a * x\<^sup>2 + b * x + c = 0" have nf: "\finite {r..(a = 0 \ b = 0 \ c = 0) \ False" proof - assume "\(a = 0 \ b = 0 \ c = 0)" then have "[:c, b, a:] \ 0" by auto then have fin: "finite {x. poly [:c, b, a:] x = 0}" using poly_roots_finite[where p = "[:c, b, a:]"] by auto have "{x. a*x^2 + b*x + c = 0} = {x. poly [:c, b, a:] x = 0}" using quadratic_poly_eval by auto then have finset: "finite {x. a*x^2 + b*x + c = 0}" using fin by auto have "{r <..p} \ {x. a*x^2 + b*x + c = 0}" using inf_zer by blast then show "False" using finset nf using finite_subset by (metis (no_types, lifting) infinite_Ioc_iff r_lt) qed then show "(a = 0 \ b = 0 \ c = 0)" by auto qed lemma continuity_lem_lt0: fixes r:: "real" fixes a b c:: "real" shows "poly [:c, b, a:] r < 0 \ \y'> r. \x\{r<..y'}. poly [:c, b, a:] x < 0" proof - let ?f = "poly [:c,b,a:]" assume r_ltz: "poly [:c, b, a:] r < 0" then have "[:c, b, a:] \ 0" by auto then have "finite {x. poly [:c, b, a:] x = 0}" using poly_roots_finite[where p = "[:c, b, a:]"] by auto then have fin: "finite {x. x > r \ poly [:c, b, a:] x = 0}" using finite_Collect_conjI by blast let ?l = "sorted_list_of_set {x. x > r \ poly [:c, b, a:] x = 0}" show ?thesis proof (cases "length ?l = 0") case True then have no_zer: "\(\x>r. poly [:c, b, a:] x = 0)" using sorted_list_of_set_eq_Nil_iff fin by auto then have "\y. y > r \ y < (r + 1) \ poly [:c, b, a:] y < 0 " proof - fix y assume "y > r \ y < r + 1" then show "poly [:c, b, a:] y < 0" using r_ltz no_zer poly_IVT_pos[where a = "r", where p = "[:c, b, a:]", where b = "y"] by (meson linorder_neqE_linordered_idom) qed then show ?thesis by (metis greaterThanAtMost_iff less_add_one less_eq_real_def linorder_not_le no_zer poly_IVT_pos r_ltz) next case False then have len_nonz: "length (sorted_list_of_set {x. r < x \ poly [:c, b, a:] x = 0}) \ 0" by blast then have "\n \ {x. x > r \ poly [:c, b, a:] x = 0}. (nth_default 0 ?l 0) \ n" using fin set_sorted_list_of_set sorted_sorted_list_of_set using in_set_conv_nth leI not_less0 sorted_nth_mono by (smt not_less_iff_gr_or_eq nth_default_def) then have no_zer: "\(\x>r. (x < (nth_default 0 ?l 0) \ poly [:c, b, a:] x = 0))" using sorted_sorted_list_of_set by auto then have fa: "\y. y > r \ y < (nth_default 0 ?l 0) \ poly [:c, b, a:] y < 0 " proof - fix y assume "y > r \ y < (nth_default 0 ?l 0)" then show "poly [:c, b, a:] y < 0" using r_ltz no_zer poly_IVT_pos[where a = "r", where p = "[:c, b, a:]", where b = "y"] by (meson less_imp_le less_le_trans linorder_neqE_linordered_idom) qed have "nth_default 0 ?l 0 > r" using fin set_sorted_list_of_set using len_nonz length_0_conv length_greater_0_conv mem_Collect_eq nth_mem by (metis (no_types, lifting) nth_default_def) then have "\(y'::real). r < y' \ y' < (nth_default 0 ?l 0)" using dense by blast then obtain y' where y_prop:"r < y' \y' < (nth_default 0 ?l 0)" by auto then have "\x\{r<..y'}. poly [:c, b, a:] x < 0" using fa by auto then show ?thesis using y_prop by blast qed qed lemma continuity_lem_gt0: fixes r:: "real" fixes a b c:: "real" shows "poly [:c, b, a:] r > 0 \ \y'> r. \x\{r<..y'}. poly [:c, b, a:] x > 0" proof - assume r_gtz: "poly [:c, b, a:] r > 0 " let ?p = "[:-c, -b, -a:]" have revpoly: "\x. -1*(poly [:c, b, a:] x) = poly [:-c, -b, -a:] x" by (metis (no_types, opaque_lifting) Polynomial.poly_minus minus_pCons mult_minus1 neg_equal_0_iff_equal) then have "poly ?p r < 0" using r_gtz by (metis mult_minus1 neg_less_0_iff_less) then have "\y'> r. \x\{r<..y'}. poly ?p x < 0" using continuity_lem_lt0 by blast then obtain y' where y_prop: "y' > r \ (\x\{r<..y'}. poly ?p x < 0)" by auto then have "\x\{r<..y'}. poly [:c, b, a:] x > 0 " using revpoly using neg_less_0_iff_less by fastforce then show ?thesis using y_prop by blast qed lemma continuity_lem_lt0_expanded: fixes r:: "real" fixes a b c:: "real" shows "a*r^2 + b*r + c < 0 \ \y'> r. \x\{r<..y'}. a*x^2 + b*x + c < 0" using quadratic_poly_eval continuity_lem_lt0 by (simp add: add.commute) lemma continuity_lem_gt0_expanded: fixes r:: "real" fixes a b c:: "real" fixes k::"real" assumes kgt: "k > r" shows "a*r^2 + b*r + c > 0 \ \x\{r<..k}. a*x^2 + b*x + c > 0" proof - assume "a*r^2 + b*r + c > 0" then have "\y'> r. \x\{r<..y'}. poly [:c, b, a:] x > 0" using continuity_lem_gt0 quadratic_poly_eval[of c b a r] by auto then obtain y' where y_prop: "y' > r \ (\x\{r<..y'}. poly [:c, b, a:] x > 0)" by auto then have "\q. q > r \ q < min k y'" using kgt dense by (metis min_less_iff_conj) then obtain q where q_prop: "q > r \q < min k y'" by auto then have "a*q^2 + b*q + c > 0" using y_prop quadratic_poly_eval[of c b a q] by (metis greaterThanAtMost_iff less_eq_real_def min_less_iff_conj) then show ?thesis using q_prop by auto qed subsection "Negative Infinity (Limit) Properties" lemma ysq_dom_y: fixes b:: "real" fixes c:: "real" shows "\(w::real). \(y:: real). (y < w \ y^2 > b*y)" proof - have c1: "b \ 0 ==> \(w::real). \(y:: real). (y < w \ y^2 > b*y)" proof - assume "b \ 0" then have p1: "\(y:: real). (y < -1 \ y*b \ 0)" by (simp add: mult_nonneg_nonpos2) have p2: "\(y:: real). (y < -1 \ y^2 > 0)" by auto then have h1: "\(y:: real). (y < -1 \ y^2 > b*y)" using p1 p2 by (metis less_eq_real_def less_le_trans mult.commute) then show ?thesis by auto qed have c2: "b < 0 \ b > -1 ==> \(w::real). \(y:: real). (y < w \ y^2 > b*y)" proof - assume "b < 0 \ b > -1 " then have h1: "\(y:: real). (y < -1 \ y^2 > b*y)" by (simp add: power2_eq_square) then show ?thesis by auto qed have c3: "b \ -1 ==> \(w::real). \(y:: real). (y < w \ y^2 > b*y)" proof - assume "b \ -1 " then have h1: "\(y:: real). (y < b \ y^2 > b*y)" by (metis le_minus_one_simps(3) less_irrefl less_le_trans mult.commute mult_less_cancel_left power2_eq_square) then show ?thesis by auto qed then show ?thesis using c1 c2 c3 by (metis less_trans linorder_not_less) qed lemma ysq_dom_y_plus_coeff: fixes b:: "real" fixes c:: "real" shows "\(w::real). \(y::real). (y < w \ y^2 > b*y + c)" proof - have "\(w::real). \(y:: real). (y < w \ y^2 > b*y)" using ysq_dom_y by auto then obtain w where w_prop: "\(y:: real). (y < w \ y^2 > b*y)" by auto have "c \ 0 \ \(y::real). (y < w \ y^2 > b*y + c)" using w_prop by auto then have c1: "c \ 0 \ \(w::real). \(y::real). (y < w \ y^2 > b*y + c)" by auto have "\(w::real). \(y:: real). (y < w \ y^2 > (b-c)*y)" using ysq_dom_y by auto then obtain k where k_prop: "\(y:: real). (y < k \ y^2 > (b-c)*y)" by auto let ?mn = "min k (-1)" have "(c> 0 \ (\ y < -1. -c*y > c))" proof - assume cgt: " c> 0" show "\(y::real) < -1. -c*y > c" proof clarsimp fix y::"real" assume "y < -1" then have "-y > 1" by auto then have "c < c*(-y)" using cgt by (metis \1 < - y\ mult.right_neutral mult_less_cancel_left_pos) then show " c < - (c * y) " by auto qed qed then have "(c> 0 \ (\ y < ?mn. (b-c)*y > b*y + c))" by (simp add: left_diff_distrib) then have c2: "c > 0 \ \(y::real). (y < ?mn \ y^2 > b*y + c)" using k_prop by force then have c2: "c > 0 \ \(w::real). \(y::real). (y < w \ y^2 > b*y + c)" by blast show ?thesis using c1 c2 by fastforce qed lemma ysq_dom_y_plus_coeff_2: fixes b:: "real" fixes c:: "real" shows "\(w::real). \(y::real). (y > w \ y^2 > b*y + c)" proof - have "\(w::real). \(y::real). (y < w \ y^2 > -b*y + c)" using ysq_dom_y_plus_coeff[where b = "-b", where c = "c"] by auto then obtain w where w_prop: "\(y::real). (y < w \ y^2 > -b*y + c)" by auto let ?mn = "min w (-1)" have "\(y::real). (y < ?mn \ y^2 > -b*y + c)" using w_prop by auto then have "\(y::real). (y > (-1*?mn) \ y^2 > b*y + c)" by (metis (no_types, opaque_lifting) minus_less_iff minus_mult_commute mult_1 power2_eq_iff) then show ?thesis by auto qed lemma neg_lc_dom_quad: fixes a:: "real" fixes b:: "real" fixes c:: "real" assumes alt: "a < 0" shows "\(w::real). \(y::real). (y > w \ a*y^2 + b*y + c < 0)" proof - have "\(w::real). \(y::real). (y > w \ y^2 > (-b/a)*y + (-c/a))" using ysq_dom_y_plus_coeff_2[where b = "-b/a", where c = "-c/a"] by auto then have keyh: "\(w::real). \(y::real). (y > w \ a*y^2 < a*((-b/a)*y + (-c/a)))" using alt by auto have simp1: "\y. a*((-b/a)*y + (-c/a)) = a*(-b/a)*y + a*(-c/a)" using diff_divide_eq_iff by fastforce have simp2: "\y. a*(-b/a)*y + a*(-c/a) = -b*y + a*(-c/a)" using assms by auto have simp3: "\y. -b*y + a*(-c/a) = -b*y - c" using assms by auto then have "\y. a*((-b/a)*y + (-c/a)) = -b*y - c" using simp1 simp2 simp3 by auto then have keyh2: "\(w::real). \(y::real). (y > w \ a*y^2 < -b*y-c)" using keyh by auto have "\y. a*y^2 < -b*y-c \ a*y^2 + b*y + c < 0" by auto then show ?thesis using keyh2 by auto qed lemma pos_lc_dom_quad: fixes a:: "real" fixes b:: "real" fixes c:: "real" assumes alt: "a > 0" shows "\(w::real). \(y::real). (y > w \ a*y^2 + b*y + c > 0)" proof - have "-a < 0" using alt by simp then have "\(w::real). \(y::real). (y > w \ -a*y^2 - b*y - c < 0)" using neg_lc_dom_quad[where a = "-a", where b = "-b", where c = "-c"] by auto then obtain w where w_prop: "\(y::real). (y > w \ -a*y^2 - b*y - c < 0)" by auto then have "\(y::real). (y > w \ a*y^2 + b*y + c > 0)" by auto then show ?thesis by auto qed (* lemma interval_infinite: fixes r p::"real" assumes "r < p" shows "infinite {r<..(d, e, f)\set les. \y'> q. \x\{q<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\y'>q. (\(d, e, f)\set les. \x\{q<..y'}. d * x\<^sup>2 + e * x + f < 0))" proof (induct les) case Nil then show ?case using gt_ex by auto next case (Cons z les) have "\a\set les. case a of (d, e, f) \ \y'>q. \x\{q<..y'}. d * x\<^sup>2 + e * x + f < 0" using Cons.prems by auto then have " \y'>q. \a\set les. case a of (d, e, f) \ \x\{q<..y'}. d * x\<^sup>2 + e * x + f < 0" using Cons.hyps by auto then obtain y1 where y1_prop : "y1>q \ (\a\set les. case a of (d, e, f) \ \x\{q<..y1}. d * x\<^sup>2 + e * x + f < 0)" by auto have "case z of (d, e, f) \ \y'>q. \x\{q<..y'}. d * x\<^sup>2 + e * x + f < 0" using Cons.prems by auto then obtain y2 where y2_prop: "y2>q \ (case z of (d, e, f) \ (\x\{q<..y2}. d * x\<^sup>2 + e * x + f < 0))" by auto let ?y = "min y1 y2" have "?y > q \ (\a\set (z#les). case a of (d, e, f) \ \x\{q<..?y}. d * x\<^sup>2 + e * x + f < 0)" using y1_prop y2_prop by force then show ?case by blast qed lemma have_inbetween_point_les: fixes r::"real" assumes "(\(d, e, f)\set les. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f < 0)" shows "(\x. (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" proof - have "(\(d, e, f)\set les. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\y'>r. (\(d, e, f)\set les. \x\{r<..y'}. d * x\<^sup>2 + e * x + f < 0))" using les_qe_inf_helper assms by auto then have "(\y'>r. (\(d, e, f)\set les. \x\{r<..y'}. d * x\<^sup>2 + e * x + f < 0))" using assms by blast then obtain y where y_prop: "y > r \ (\(d, e, f)\set les. \x\{r<..y}. d * x\<^sup>2 + e * x + f < 0)" by auto have "\q. q > r \q < y" using y_prop dense by auto then obtain q where q_prop: "q > r \ q < y" by auto then have "(\(d, e, f)\set les. d*q^2 + e*q + f < 0)" using y_prop by auto then show ?thesis by auto qed lemma one_root_a_gt0: fixes a b c r:: "real" shows "\y'. b = 2 * a * r \ \ a < 0 \ a * r^2 - 2 * a *r*r + c = 0 \ - r < y' \ \x\{-r<..y'}. \ a * x\<^sup>2 + 2 * a * r*x + c < 0" proof - fix y' assume beq: "b = 2 * a * r" assume aprop: " \ a < 0" assume root: " a * r\<^sup>2 - 2 * a *r*r + c = 0" assume rootlt: "- r < y'" show " \x\{- r<..y'}. \ a * x\<^sup>2 + 2 * a* r*x+ c < 0" proof - have h: "a = 0 \ (b = 0 \ c = 0)" using beq root by auto then have aeq: "a = 0 \ \x\{- r<..y'}. \ a * x\<^sup>2 + 2 * a*r*x + c < 0" using rootlt by (metis add.left_neutral continuity_lem_eq0 less_numeral_extra(3) mult_zero_left mult_zero_right) then have alt: "a > 0 \ \x\{- r<..y'}. \ a * x\<^sup>2 + 2 * a *r*x + c < 0" proof - assume agt: "a > 0" then have "\(w::real). \(y::real). (y > w \ a*y^2 + b*y + c > 0)" using pos_lc_dom_quad by auto then obtain w where w_prop: "\y::real. (y > w \ a*y^2 + b*y + c > 0)" by auto have isroot: "a*(-r)^2 + b*(-r) + c = 0" using root beq by auto then have wgteq: "w \ -(r)" proof - have "w < -r \ False" using w_prop isroot by auto then show ?thesis using not_less by blast qed then have w1: "w + 1 > -r" by auto have w2: "a*(w + 1)^2 + b*(w+1) + c > 0" using w_prop by auto have rootiff: "\x. a * x\<^sup>2 + b * x + c = 0 \ x = -r" using discriminant_lemma[where a = "a", where b = "b", where c= "c", where r = "r"] isroot agt beq by auto have allgt: "\x > -r. a*x^2 + b*x + c > 0" proof clarsimp fix x assume "x > -r" have xgtw: "x > w + 1 \ a*x^2 + b*x + c > 0 " using w1 w2 rootiff poly_IVT_neg[where a = "w+1", where b = "x", where p = "[:c,b,a:]"] quadratic_poly_eval by (metis less_eq_real_def linorder_not_less) have xltw: "x < w + 1 \ a*x^2 + b*x + c > 0" using w1 w2 rootiff poly_IVT_pos[where a= "x", where b = "w + 1", where p = "[:c,b,a:]"] quadratic_poly_eval less_eq_real_def linorder_not_less by (metis \- r < x\) then show "a*x^2 + b*x + c > 0" using w2 xgtw xltw by fastforce qed have "\z. z > -r \ z < y'" using rootlt dense[where x = "-r", where y = "y'"] by auto then obtain z where z_prop: " z > -r \ z < y'" by auto then have "a*z^2 + b*z + c > 0" using allgt by auto then show ?thesis using z_prop using beq greaterThanAtMost_iff by force qed then show ?thesis using aeq alt aprop by linarith qed qed lemma leq_qe_inf_helper: fixes q:: "real" shows"(\(d, e, f)\set leq. \y'> q. \x\{q<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\y'>q. (\(d, e, f)\set leq. \x\{q<..y'}. d * x\<^sup>2 + e * x + f \ 0))" proof (induct leq) case Nil then show ?case using gt_ex by auto next case (Cons z leq) have "\a\set leq. case a of (d, e, f) \ \y'>q. \x\{q<..y'}. d * x\<^sup>2 + e * x + f \ 0" using Cons.prems by auto then have " \y'>q. \a\set leq. case a of (d, e, f) \ \x\{q<..y'}. d * x\<^sup>2 + e * x + f \ 0" using Cons.hyps by auto then obtain y1 where y1_prop : "y1>q \ (\a\set leq. case a of (d, e, f) \ \x\{q<..y1}. d * x\<^sup>2 + e * x + f \ 0)" by auto have "case z of (d, e, f) \ \y'>q. \x\{q<..y'}. d * x\<^sup>2 + e * x + f \ 0" using Cons.prems by auto then obtain y2 where y2_prop: "y2>q \ (case z of (d, e, f) \ (\x\{q<..y2}. d * x\<^sup>2 + e * x + f \ 0))" by auto let ?y = "min y1 y2" have "?y > q \ (\a\set (z#leq). case a of (d, e, f) \ \x\{q<..?y}. d * x\<^sup>2 + e * x + f \ 0)" using y1_prop y2_prop by force then show ?case by blast qed lemma neq_qe_inf_helper: fixes q:: "real" shows"(\(d, e, f)\set neq. \y'> q. \x\{q<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\y'>q. (\(d, e, f)\set neq. \x\{q<..y'}. d * x\<^sup>2 + e * x + f \ 0))" proof (induct neq) case Nil then show ?case using gt_ex by auto next case (Cons z neq) have "\a\set neq. case a of (d, e, f) \ \y'>q. \x\{q<..y'}. d * x\<^sup>2 + e * x + f \ 0" using Cons.prems by auto then have " \y'>q. \a\set neq. case a of (d, e, f) \ \x\{q<..y'}. d * x\<^sup>2 + e * x + f \ 0" using Cons.hyps by auto then obtain y1 where y1_prop : "y1>q \ (\a\set neq. case a of (d, e, f) \ \x\{q<..y1}. d * x\<^sup>2 + e * x + f \ 0)" by auto have "case z of (d, e, f) \ \y'>q. \x\{q<..y'}. d * x\<^sup>2 + e * x + f \ 0" using Cons.prems by auto then obtain y2 where y2_prop: "y2>q \ (case z of (d, e, f) \ (\x\{q<..y2}. d * x\<^sup>2 + e * x + f \ 0))" by auto let ?y = "min y1 y2" have "?y > q \ (\a\set (z#neq). case a of (d, e, f) \ \x\{q<..?y}. d * x\<^sup>2 + e * x + f \ 0)" using y1_prop y2_prop by force then show ?case by blast qed subsection "Some Casework" lemma quadratic_shape1a: fixes a b c x y::"real" assumes agt: "a > 0" assumes xyroots: "x < y \ a*x^2 + b*x + c = 0 \ a*y^2 + b*y + c = 0" shows "\z. (z > x \ z < y \ a*z^2 + b*z + c < 0)" proof clarsimp fix z assume zgt: "z > x" assume zlt: "z < y" have frac_gtz: "(1/(2*a)) > 0" using agt by simp have xy_prop:"(x = (-b + sqrt(b^2 - 4*a*c))/(2*a) \ y = (-b - sqrt(b^2 - 4*a*c))/(2*a)) \ (y = (-b + sqrt(b^2 - 4*a*c))/(2*a) \ x = (-b - sqrt(b^2 - 4*a*c))/(2*a))" using xyroots agt discriminant_iff unfolding discrim_def by auto have "b^2 - 4*a*c \ 0" using xyroots discriminant_iff using assms(1) discrim_def by auto then have pos_discrim: "b^2 - 4*a*c > 0" using xyroots discriminant_zero using \0 \ b\<^sup>2 - 4 * a * c\ assms(1) discrim_def less_eq_real_def linorder_not_less by metis then have sqrt_gt: "sqrt(b^2 - 4*a*c) > 0" using real_sqrt_gt_0_iff by blast then have "(- b - sqrt(b^2 - 4*a*c)) < (- b + sqrt(b^2 - 4*a*c))" by auto then have "(- b - sqrt(b^2 - 4*a*c))*(1/(2*a)) < (- b + sqrt(b^2 - 4*a*c))*(1/(2*a)) " using frac_gtz by (simp add: divide_strict_right_mono) then have "(- b - sqrt(b^2 - 4*a*c))/(2*a) < (- b + sqrt(b^2 - 4*a*c))/(2*a)" by auto then have xandy: "x = (- b - sqrt(b^2 - 4*a*c))/(2*a) \ y = (- b + sqrt(b^2 - 4*a*c))/(2*a)" using xy_prop xyroots by auto let ?mdpt = "-b/(2*a)" have xlt: "x < ?mdpt" using xandy sqrt_gt using frac_gtz divide_minus_left divide_strict_right_mono sqrt_gt by (smt (verit) agt) have ylt: "?mdpt < y" using xandy sqrt_gt frac_gtz by (smt (verit, del_insts) divide_strict_right_mono zero_less_divide_1_iff) have mdpt_val: "a*?mdpt^2 + b*?mdpt + c < 0" proof - have firsteq: "a*?mdpt^2 + b*?mdpt + c = (a*b^2)/(4*a^2) - (b^2)/(2*a) + c" by (simp add: power2_eq_square) have h1: "(a*b^2)/(4*a^2) = (b^2)/(4*a)" by (simp add: power2_eq_square) have h2: "(b^2)/(2*a) = (2*b^2)/(4*a)" by linarith have h3: "c = (4*a*c)/(4*a)" using agt by auto have "a*?mdpt^2 + b*?mdpt + c = (b^2)/(4*a) - (2*b^2)/(4*a) + (4*a*c)/(4*a) " using firsteq h1 h2 h3 by linarith then have "a*?mdpt^2 + b*?mdpt + c = (b^2 - 2*b^2 + 4*a*c)/(4*a)" by (simp add: diff_divide_distrib) then have eq2: "a*?mdpt^2 + b*?mdpt + c = (4*a*c - b^2)/(4*a)" by simp have h: "4*a*c - b^2 < 0" using pos_discrim by auto have "1/(4*a) > 0" using agt by auto then have "(4*a*c - b^2)*(1/(4*a)) < 0" using h using mult_less_0_iff by blast then show ?thesis using eq2 by auto qed have nex: "\ (\k> x. k < y \ poly [:c, b, a:] k = 0)" using discriminant_iff agt by (metis discrim_def less_irrefl quadratic_poly_eval xandy) have nor2: "\ (\x>z. x < - b / (2 * a) \ poly [:c, b, a:] x = 0)" using nex xlt ylt zgt zlt by auto have nor: "\ (\x>- b / (2 * a). x < z \ poly [:c, b, a:] x = 0)" using nex xlt ylt zgt zlt discriminant_iff agt by auto then have mdpt_lt: "?mdpt < z \ a*z^2 + b*z + c < 0 " using mdpt_val zgt zlt xlt ylt poly_IVT_pos[where p = "[:c, b, a:]", where a= "?mdpt", where b = "z"] quadratic_poly_eval[of c b a] by (metis \\ (\k>x. k < y \ poly [:c, b, a:] k = 0)\ linorder_neqE_linordered_idom) have mdpt_gt: "?mdpt > z \ a*z^2 + b*z + c < 0 " using zgt zlt mdpt_val xlt ylt nor2 poly_IVT_neg[where p = "[:c, b, a:]", where a = "z", where b = "?mdpt"] quadratic_poly_eval[of c b a] by (metis linorder_neqE_linordered_idom nex) then show "a*z^2 + b*z + c < 0" using mdpt_lt mdpt_gt mdpt_val by fastforce qed lemma quadratic_shape1b: fixes a b c x y::"real" assumes agt: "a > 0" assumes xy_roots: "x < y \ a*x^2 + b*x + c = 0 \ a*y^2 + b*y + c = 0" shows "\z. (z > y \ a*z^2 + b*z + c > 0)" proof - fix z assume z_gt :"z > y" have nogt: "\(\w. w > y \ a*w^2 + b*w + c = 0)" using xy_roots discriminant_iff by (metis agt less_eq_real_def linorder_not_less) have "\(w::real). \(y::real). (y > w \ a*y^2 + b*y + c > 0)" using agt pos_lc_dom_quad by auto then have "\k > y. a*k^2 + b*k + c > 0" by (metis add.commute agt less_add_same_cancel1 linorder_neqE_linordered_idom pos_add_strict) then obtain k where k_prop: "k > y \ a*k^2 + b*k + c > 0" by auto have kgt: "k > z \ a*z^2 + b*z + c > 0" proof - assume kgt: "k > z" then have zneq: "a*z^2 + b*z + c = 0 \ False" using nogt using z_gt by blast have znlt: "a*z^2 + b*z + c < 0 \ False" using kgt k_prop quadratic_poly_eval[of c b a] z_gt nogt poly_IVT_pos[where a= "z", where b = "k", where p = "[:c, b, a:]"] by (metis less_eq_real_def less_le_trans) then show "a*z^2 + b*z + c > 0" using zneq znlt using linorder_neqE_linordered_idom by blast qed have klt: "k < z \ a*z^2 + b*z + c > 0" proof - assume klt: "k < z" then have zneq: "a*z^2 + b*z + c = 0 \ False" using nogt using z_gt by blast have znlt: "a*z^2 + b*z + c < 0 \ False" using klt k_prop quadratic_poly_eval[of c b a] z_gt nogt poly_IVT_neg[where a= "k", where b = "z", where p = "[:c, b, a:]"] by (metis add.commute add_less_cancel_left add_mono_thms_linordered_field(3) less_eq_real_def) then show "a*z^2 + b*z + c > 0" using zneq znlt using linorder_neqE_linordered_idom by blast qed then show "a*z^2 + b*z + c > 0" using k_prop kgt klt by fastforce qed lemma quadratic_shape2a: fixes a b c x y::"real" assumes "a < 0" assumes "x < y \ a*x^2 + b*x + c = 0 \ a*y^2 + b*y + c = 0" shows "\z. (z > x \ z < y \ a*z^2 + b*z + c > 0)" using quadratic_shape1a[where a= "-a", where b = "-b", where c = "-c", where x = "x", where y = "y"] using assms(1) assms(2) by fastforce lemma quadratic_shape2b: fixes a b c x y::"real" assumes "a < 0" assumes "x < y \ a*x^2 + b*x + c = 0 \ a*y^2 + b*y + c = 0" shows "\z. (z > y \ a*z^2 + b*z + c < 0)" using quadratic_shape1b[where a= "-a", where b = "-b", where c = "-c", where x = "x", where y = "y"] using assms(1) assms(2) by force lemma case_d1: fixes a b c r::"real" shows "b < 2 * a * r \ a * r^2 - b*r + c = 0 \ \y'>- r. \x\{-r<..y'}. a * x\<^sup>2 + b * x + c < 0" proof - assume b_lt: "b < 2*a*r" assume root: "a*r^2 - b*r + c = 0" then have "c = b*r - a*r^2" by auto have aeq: "a = 0 \ \y'>- r. \x\{-r<..y'}. a * x\<^sup>2 + b * x + c < 0" proof - assume azer: "a = 0" then have bltz: "b < 0" using b_lt by auto then have "c = b*r" using azer root by auto then have eval: "\x. a*x^2 + b*x + c = b*(x + r)" using azer by (simp add: distrib_left) have "\x > -r. b*(x + r) < 0" proof clarsimp fix x assume xgt: "- r < x" then have "x + r > 0" by linarith then show "b * (x + r) < 0" using bltz using mult_less_0_iff by blast qed then show ?thesis using eval using less_add_same_cancel1 zero_less_one by (metis greaterThanAtMost_iff) qed have aneq: "a \ 0 \\y'>- r. \x\{-r<..y'}. a * x\<^sup>2 + b * x + c < 0" proof - assume aneq: "(a::real) \ 0" have "b^2 - 4*a*c < 0 \ a * r\<^sup>2 + b * r + c \ 0" using root discriminant_negative[of a b c r] unfolding discrim_def using aneq by auto then have " a * r\<^sup>2 + b * r + c \ 0 \ a * r\<^sup>2 - b * r + c = 0 \ b\<^sup>2 < 4 * a * c \ False" proof - assume a1: "a * r\<^sup>2 - b * r + c = 0" assume a2: "b\<^sup>2 < 4 * a * c" have f3: "(0 \ - 1 * (4 * a * c) + (- 1 * b)\<^sup>2) = (4 * a * c + - 1 * (- 1 * b)\<^sup>2 \ 0)" by simp have f4: "(- 1 * b)\<^sup>2 + - 1 * (4 * a * c) = - 1 * (4 * a * c) + (- 1 * b)\<^sup>2" by auto have f5: "c + a * r\<^sup>2 + - 1 * b * r = a * r\<^sup>2 + c + - 1 * b * r" by auto have f6: "\x0 x1 x2 x3. (x3::real) * x0\<^sup>2 + x2 * x0 + x1 = x1 + x3 * x0\<^sup>2 + x2 * x0" by simp have f7: "\x1 x2 x3. (discrim x3 x2 x1 < 0) = (\ 0 \ discrim x3 x2 x1)" by auto have f8: "\r ra rb. discrim r ra rb = ra\<^sup>2 + - 1 * (4 * r * rb)" using discrim_def by auto have "\ 4 * a * c + - 1 * (- 1 * b)\<^sup>2 \ 0" using a2 by simp then have "a * r\<^sup>2 + c + - 1 * b * r \ 0" using f8 f7 f6 f5 f4 f3 by (metis (no_types) aneq discriminant_negative) then show False using a1 by linarith qed then have "\(b^2 - 4*a*c < 0)" using root using \b\<^sup>2 - 4 * a * c < 0 \ a * r\<^sup>2 + b * r + c \ 0\ by auto then have discrim: "b\<^sup>2 \ 4 * a * c " by auto then have req: "r = (b + sqrt(b^2 - 4*a*c))/(2*a) \ r = (b - sqrt(b^2 - 4*a*c))/(2*a)" using aneq root discriminant_iff[where a="a", where b ="-b", where c="c", where x="r"] unfolding discrim_def by auto then have "r = (b - sqrt(b^2 - 4*a*c))/(2*a) \ b > 2*a*r" proof - assume req: "r = (b - sqrt(b^2 - 4*a*c))/(2*a)" then have h1: "2*a*r = 2*a*((b - sqrt(b^2 - 4*a*c))/(2*a))" by auto then have h2: "2*a*((b - sqrt(b^2 - 4*a*c))/(2*a)) = b - sqrt(b^2 - 4*a*c)" using aneq by auto have h3: "sqrt(b^2 - 4*a*c) \ 0" using discrim by auto then have "b - sqrt(b^2 - 4*a*c) < b" using b_lt h1 h2 by linarith then show ?thesis using req h2 by auto qed then have req: "r = (b + sqrt(b^2 - 4*a*c))/(2*a)" using req b_lt by auto then have discrim2: "b^2 - 4 *a*c > 0" using aneq b_lt by auto then have "\x y. x \ y \ a * x\<^sup>2 + b * x + c = 0 \ a * y\<^sup>2 + b * y + c = 0" using aneq discriminant_pos_ex[of a b c] unfolding discrim_def by auto then obtain x y where xy_prop: "x < y \ a * x\<^sup>2 + b * x + c = 0 \ a * y\<^sup>2 + b * y + c = 0" by (meson linorder_neqE_linordered_idom) then have "(x = (-b + sqrt(b^2 - 4*a*c))/(2*a) \ y = (-b - sqrt(b^2 - 4*a*c))/(2*a)) \ (y = (-b + sqrt(b^2 - 4*a*c))/(2*a) \ x = (-b - sqrt(b^2 - 4*a*c))/(2*a))" using aneq discriminant_iff unfolding discrim_def by auto then have xy_prop2: "(x = (-b + sqrt(b^2 - 4*a*c))/(2*a) \ y = -r) \ (y = (-b + sqrt(b^2 - 4*a*c))/(2*a) \ x = -r)" using req by (simp add: \x = (- b + sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a) \ y = (- b - sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a) \ y = (- b + sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a) \ x = (- b - sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a)\ minus_divide_left) (* When a < 0, -r is the bigger root *) have alt: "a < 0 \ \k > -r. a * k^2 + b * k + c < 0" proof clarsimp fix k assume alt: " a < 0" assume "- r < k" have alt2: " (1/(2*a)::real) < 0" using alt by simp have "(-b - sqrt(b^2 - 4*a*c)) < (-b + sqrt(b^2 - 4*a*c))" using discrim2 by auto then have "(-b - sqrt(b^2 - 4*a*c))* (1/(2*a)::real) > (-b + sqrt(b^2 - 4*a*c))* (1/(2*a)::real)" using alt2 using mult_less_cancel_left_neg by fastforce then have rgtroot: "-r > (-b + sqrt(b^2 - 4*a*c))/(2*a)" using req \x = (- b + sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a) \ y = (- b - sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a) \ y = (- b + sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a) \ x = (- b - sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a)\ xy_prop2 by auto then have "(y = -r \ x = (-b + sqrt(b^2 - 4*a*c))/(2*a))" using xy_prop xy_prop2 by auto then show "a * k^2 + b * k + c < 0" using xy_prop \- r < k\ alt quadratic_shape2b xy_prop by blast qed (* When a > 0, -r is the smaller root *) have agt: "a > 0 \ \y'>- r. \x\{-r<..y'}. a * x\<^sup>2 + b * x + c < 0" proof - assume agt: "a> 0" have alt2: " (1/(2*a)::real) > 0" using agt by simp have "(-b - sqrt(b^2 - 4*a*c)) < (-b + sqrt(b^2 - 4*a*c))" using discrim2 by auto then have "(-b - sqrt(b^2 - 4*a*c))* (1/(2*a)::real) < (-b + sqrt(b^2 - 4*a*c))* (1/(2*a)::real)" using alt2 proof - have f1: "- b - sqrt (b\<^sup>2 - c * (4 * a)) < - b + sqrt (b\<^sup>2 - c * (4 * a))" by (metis \- b - sqrt (b\<^sup>2 - 4 * a * c) < - b + sqrt (b\<^sup>2 - 4 * a * c)\ mult.commute) have "0 < a * 2" using \0 < 1 / (2 * a)\ by auto then show ?thesis using f1 by (simp add: divide_strict_right_mono mult.commute) qed then have rlltroot: "-r < (-b + sqrt(b^2 - 4*a*c))/(2*a)" using req \x = (- b + sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a) \ y = (- b - sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a) \ y = (- b + sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a) \ x = (- b - sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a)\ xy_prop2 by auto then have "(x = -r \ y = (-b + sqrt(b^2 - 4*a*c))/(2*a))" using xy_prop xy_prop2 by auto have "\k. x < k \ k < y" using xy_prop dense by auto then obtain k where k_prop: "x < k \ k < y" by auto then have "\x\{-r<..k}. a * x\<^sup>2 + b * x + c < 0" using agt quadratic_shape1a[where a= "a", where b = "b", where c= "c", where x = "x", where y = "y"] using \x = - r \ y = (- b + sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a)\ greaterThanAtMost_iff xy_prop by auto then show "\y'>- r. \x\{-r<..y'}. a * x\<^sup>2 + b * x + c < 0" using k_prop using \x = - r \ y = (- b + sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a)\ by blast qed show ?thesis using alt agt by (metis aneq greaterThanAtMost_iff less_add_same_cancel1 linorder_neqE_linordered_idom zero_less_one) qed show "\y'>- r. \x\{-r<..y'}. a * x\<^sup>2 + b * x + c < 0" using aeq aneq by blast qed lemma case_d4: fixes a b c r::"real" shows "\y'. b \ 2 * a * r \ \ b < 2 * a * r \ a *r^2 - b * r + c = 0 \ -r < y' \ \x\{-r<..y'}. \ a * x\<^sup>2 + b * x + c < 0" proof - fix y' assume bneq: "b \ 2 * a * r" assume bnotless: "\ b < 2 * a * r" assume root: "a *r^2 - b * r + c = 0" assume y_prop: "-r < y'" have b_gt: "b > 2*a*r" using bneq bnotless by auto have aeq: "a = 0 \ \y'>- r. \x\{-r<..y'}. a * x\<^sup>2 + b * x + c > 0" proof - assume azer: "a = 0" then have bgt: "b > 0" using b_gt by auto then have "c = b*r" using azer root by auto then have eval: "\x. a*x^2 + b*x + c = b*(x + r)" using azer by (simp add: distrib_left) have "\x > -r. b*(x + r) > 0" proof clarsimp fix x assume xgt: "- r < x" then have "x + r > 0" by linarith then show "b * (x + r) > 0" using bgt by auto qed then show ?thesis using eval using less_add_same_cancel1 zero_less_one by (metis greaterThanAtMost_iff) qed have aneq: "a \ 0 \\y'>- r. \x\{-r<..y'}. a * x\<^sup>2 + b * x + c > 0" proof - assume aneq: "a\0" { assume a1: "a * r\<^sup>2 - b * r + c = 0" assume a2: "b\<^sup>2 < 4 * a * c" have f3: "(0 \ - 1 * (4 * a * c) + (- 1 * b)\<^sup>2) = (4 * a * c + - 1 * (- 1 * b)\<^sup>2 \ 0)" by simp have f4: "(- 1 * b)\<^sup>2 + - 1 * (4 * a * c) = - 1 * (4 * a * c) + (- 1 * b)\<^sup>2" by auto have f5: "c + a * r\<^sup>2 + - 1 * b * r = a * r\<^sup>2 + c + - 1 * b * r" by auto have f6: "\x0 x1 x2 x3. (x3::real) * x0\<^sup>2 + x2 * x0 + x1 = x1 + x3 * x0\<^sup>2 + x2 * x0" by simp have f7: "\x1 x2 x3. (discrim x3 x2 x1 < 0) = (\ 0 \ discrim x3 x2 x1)" by auto have f8: "\r ra rb. discrim r ra rb = ra\<^sup>2 + - 1 * (4 * r * rb)" using discrim_def by auto have "\ 4 * a * c + - 1 * (- 1 * b)\<^sup>2 \ 0" using a2 by simp then have "a * r\<^sup>2 + c + - 1 * b * r \ 0" using f8 f7 f6 f5 f4 f3 by (metis (no_types) aneq discriminant_negative) then have False using a1 by linarith } note * = this have "b^2 - 4*a*c < 0 \ a * r\<^sup>2 + b * r + c \ 0" using root discriminant_negative[of a b c r] unfolding discrim_def using aneq by auto then have "\(b^2 - 4*a*c < 0)" using root * by auto then have discrim: "b\<^sup>2 \ 4 * a * c " by auto then have req: "r = (b + sqrt(b^2 - 4*a*c))/(2*a) \ r = (b - sqrt(b^2 - 4*a*c))/(2*a)" using aneq root discriminant_iff[where a="a", where b ="-b", where c="c", where x="r"] unfolding discrim_def by auto then have "r = (b + sqrt(b^2 - 4*a*c))/(2*a) \ b < 2*a*r" proof - assume req: "r = (b + sqrt(b^2 - 4*a*c))/(2*a)" then have h1: "2*a*r = 2*a*((b + sqrt(b^2 - 4*a*c))/(2*a))" by auto then have h2: "2*a*((b + sqrt(b^2 - 4*a*c))/(2*a)) = b + sqrt(b^2 - 4*a*c)" using aneq by auto have h3: "sqrt(b^2 - 4*a*c) \ 0" using discrim by auto then have "b + sqrt(b^2 - 4*a*c) > b" using b_gt h1 h2 by linarith then show ?thesis using req h2 by auto qed then have req: "r = (b - sqrt(b^2 - 4*a*c))/(2*a)" using req b_gt using aneq discrim by auto then have discrim2: "b^2 - 4 *a*c > 0" using aneq b_gt by auto then have "\x y. x \ y \ a * x\<^sup>2 + b * x + c = 0 \ a * y\<^sup>2 + b * y + c = 0" using aneq discriminant_pos_ex[of a b c] unfolding discrim_def by auto then obtain x y where xy_prop: "x < y \ a * x\<^sup>2 + b * x + c = 0 \ a * y\<^sup>2 + b * y + c = 0" by (meson linorder_neqE_linordered_idom) then have "(x = (-b + sqrt(b^2 - 4*a*c))/(2*a) \ y = (-b - sqrt(b^2 - 4*a*c))/(2*a)) \ (y = (-b + sqrt(b^2 - 4*a*c))/(2*a) \ x = (-b - sqrt(b^2 - 4*a*c))/(2*a))" using aneq discriminant_iff unfolding discrim_def by auto then have xy_prop2: "(x = (-b - sqrt(b^2 - 4*a*c))/(2*a) \ y = -r) \ (y = (-b - sqrt(b^2 - 4*a*c))/(2*a) \ x = -r)" using req divide_inverse minus_diff_eq mult.commute mult_minus_right by (smt (verit, ccfv_SIG) uminus_add_conv_diff) (* When a > 0, -r is the greater root *) have agt: "a > 0 \ \k > -r. a * k^2 + b * k + c > 0" proof clarsimp fix k assume agt: " a > 0" assume "- r < k" have agt2: " (1/(2*a)::real) > 0" using agt by simp have "(-b - sqrt(b^2 - 4*a*c)) < (-b + sqrt(b^2 - 4*a*c))" using discrim2 by auto then have "(-b - sqrt(b^2 - 4*a*c))* (1/(2*a)::real) < (-b + sqrt(b^2 - 4*a*c))* (1/(2*a)::real)" using agt2 by (simp add: divide_strict_right_mono) then have rgtroot: "-r > (-b - sqrt(b^2 - 4*a*c))/(2*a)" using req \x = (- b + sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a) \ y = (- b - sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a) \ y = (- b + sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a) \ x = (- b - sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a)\ xy_prop2 by auto then have "(x = (-b - sqrt(b^2 - 4*a*c))/(2*a)) \ y = -r" using xy_prop xy_prop2 by auto then show "a * k^2 + b * k + c > 0" using \- r < k\ xy_prop agt quadratic_shape1b[where a= "a", where b ="b", where c="c", where x = "x", where y = "-r", where z = "k"] by blast qed (* When a < 0, -r is the smaller root *) have agt2: "a < 0 \ \y'>- r. \x\{-r<..y'}. a * x\<^sup>2 + b * x + c > 0" proof - assume alt: "a<0" have alt2: " (1/(2*a)::real) < 0" using alt by simp have "(-b - sqrt(b^2 - 4*a*c)) < (-b + sqrt(b^2 - 4*a*c))" using discrim2 by auto then have "(-b - sqrt(b^2 - 4*a*c))* (1/(2*a)::real) > (-b + sqrt(b^2 - 4*a*c))* (1/(2*a)::real)" using alt2 using mult_less_cancel_left_neg by fastforce then have rlltroot: "-r < (-b - sqrt(b^2 - 4*a*c))/(2*a)" using req using \x = (- b + sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a) \ y = (- b - sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a) \ y = (- b + sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a) \ x = (- b - sqrt (b\<^sup>2 - 4 * a * c)) / (2 * a)\ xy_prop2 by auto then have h: "(x = -r \ y = (-b - sqrt(b^2 - 4*a*c))/(2*a))" using xy_prop xy_prop2 by auto have "\k. x < k \ k < y" using xy_prop dense by auto then obtain k where k_prop: "x < k \ k < y" by auto then have "\x\{-r<..k}. a * x\<^sup>2 + b * x + c > 0" using alt quadratic_shape2a[where a= "a", where b = "b", where c= "c", where x = "x", where y = "y"] xy_prop h greaterThanAtMost_iff by auto then show "\y'>- r. \x\{-r<..y'}. a * x\<^sup>2 + b * x + c > 0" using k_prop using h by blast qed show ?thesis using aneq agt agt2 by (meson greaterThanAtMost_iff linorder_neqE_linordered_idom y_prop) qed show "\x\{-r<..y'}. \ a * x\<^sup>2 + b * x + c < 0" using aneq aeq by (metis greaterThanAtMost_iff less_eq_real_def linorder_not_less y_prop) qed lemma one_root_a_lt0: fixes a b c r y'::"real" assumes alt: "a < 0" assumes beq: "b = 2 * a * r" assumes root: " a * r^2 - 2*a*r*r + c = 0" shows "\y'>- r. \x\{- r<..y'}. a * x\<^sup>2 + 2*a*r*x + c < 0" proof - have root_iff: "\x. a * x\<^sup>2 + b * x + c = 0 \ x = -r" using alt root discriminant_lemma[where r = "r"] beq by auto have "a < 0 \ (\y. \x > y. a*x^2 + b*x + c < 0)" using neg_lc_dom_quad by auto then obtain k where k_prop: "\x > k. a*x^2 + b*x + c < 0" using alt by auto let ?mx = "max (k+1) (-r + 1)" have "a*?mx^2 + b*?mx + c < 0" using k_prop by auto then have "\y > -r. a*y^2 + b*y + c < 0" by force then obtain z where z_prop: "z > -r \ a*z^2 + b*z + c < 0" by auto have poly_eval_prop: "\(x::real). poly [:c, b, a :] x = a*x^2 + b*x + c" using quadratic_poly_eval by auto then have nozer: "\(\x. (x > -r \ poly [:c, b, a :] x = 0))" using root_iff by (simp add: add.commute) have poly_z: "poly [:c, b, a:] z < 0" using z_prop poly_eval_prop by auto have "\y > -r. a*y^2 + b*y + c < 0" proof clarsimp fix y assume ygt: "- r < y" have h1: "y = z \ a * y\<^sup>2 + b * y + c < 0" using z_prop by auto have h2: "y < z \ a * y\<^sup>2 + b * y + c < 0" proof - assume ylt: "y < z" have notz: "a*y^2 + b*y + c \ 0" using ygt nozer poly_eval_prop by auto have h1: "a *y^2 + b*y + c > 0 \ poly [:c, b, a:] y > 0" using poly_eval_prop by auto have ivtprop: "poly [:c, b, a:] y > 0 \ (\x. y < x \ x < z \ poly [:c, b, a:] x = 0)" using ylt poly_z poly_IVT_neg[where a = "y", where b = "z", where p = "[:c, b, a:]"] by auto then have "a*y^2 + b*y + c > 0 \ False" using h1 ivtprop ygt nozer by auto then show "a*y^2 + b*y + c < 0" using notz using linorder_neqE_linordered_idom by blast qed have h3: "y > z \ a * y\<^sup>2 + b * y + c < 0" proof - assume ygtz: "y > z" have notz: "a*y^2 + b*y + c \ 0" using ygt nozer poly_eval_prop by auto have h1: "a *y^2 + b*y + c > 0 \ poly [:c, b, a:] y > 0" using poly_eval_prop by auto have ivtprop: "poly [:c, b, a:] y > 0 \ (\x. z < x \ x < y \ poly [:c, b, a:] x = 0)" using ygtz poly_z using poly_IVT_pos by blast then have "a*y^2 + b*y + c > 0 \ False" using h1 ivtprop z_prop nozer by auto then show "a*y^2 + b*y + c < 0" using notz using linorder_neqE_linordered_idom by blast qed show "a * y\<^sup>2 + b * y + c < 0" using h1 h2 h3 using linorder_neqE_linordered_idom by blast qed then show ?thesis using \\y>- r. a * y\<^sup>2 + b * y + c < 0\ beq by auto qed lemma one_root_a_lt0_var: fixes a b c r y'::"real" assumes alt: "a < 0" assumes beq: "b = 2 * a * r" assumes root: " a * r^2 - 2*a*r*r + c = 0" shows "\y'>- r. \x\{- r<..y'}. a * x\<^sup>2 + 2*a*r*x + c \ 0" proof - have h1: "\y'>- r. \x\{- r<..y'}. a * x\<^sup>2 + 2 * a * r * x + c < 0 \ \y'>-r. \x\{- r<..y'}. a * x\<^sup>2 + 2 * a *r * x + c \ 0" using less_eq_real_def by blast then show ?thesis using one_root_a_lt0[of a b r] assms by auto qed subsection "More Continuity Properties" lemma continuity_lem_gt0_expanded_var: fixes r:: "real" fixes a b c:: "real" fixes k::"real" assumes kgt: "k > r" shows "a*r^2 + b*r + c > 0 \ \x\{r<..k}. a*x^2 + b*x + c \ 0" proof - assume a: "a*r^2 + b*r + c > 0 " have h: "\x\{r<..k}. a*x^2 + b*x + c > 0 \ \x\{r<..k}. a*x^2 + b*x + c \ 0" using less_eq_real_def by blast have "\x\{r<..k}. a*x^2 + b*x + c > 0" using a continuity_lem_gt0_expanded[of r k a b c] assms by auto then show "\x\{r<..k}. a*x^2 + b*x + c \ 0" using h by auto qed lemma continuity_lem_lt0_expanded_var: fixes r:: "real" fixes a b c:: "real" shows "a*r^2 + b*r + c < 0 \ \y'> r. \x\{r<..y'}. a*x^2 + b*x + c \ 0" proof - assume "a*r^2 + b*r + c < 0 " then have " \y'> r. \x\{r<..y'}. a*x^2 + b*x + c < 0" using continuity_lem_lt0_expanded by auto then show " \y'> r. \x\{r<..y'}. a*x^2 + b*x + c \ 0" using less_eq_real_def by auto qed lemma nonzcoeffs: fixes a b c r::"real" shows "a\0 \ b\0 \ c\0 \ \y'>r. \x\{r<..y'}. a * x\<^sup>2 + b * x + c \ 0 " proof - assume "a\0 \ b\0 \ c\0" then have fin: "finite {x. a*x^2 + b*x + c = 0}" by (metis pCons_eq_0_iff poly_roots_finite poly_roots_set_same) (* then have fin2: "finite {x. a*x^2 + b*x + c = 0 \ x > r}" using finite_Collect_conjI by blast *) let ?s = "{x. a*x^2 + b*x + c = 0}" have imp: "(\q \ ?s. q > r) \ (\q \ ?s. (q > r \ (\x \ ?s. x > r \ x \ q)))" proof - assume asm: "(\q \ ?s. q > r)" then have none: "{q. q \ ?s \ q > r} \ {}" by blast have fin2: "finite {q. q \ ?s \ q > r}" using fin by simp have "\k. k = Min {q. q \ ?s \ q > r}" using fin2 none by blast then obtain k where k_prop: "k = Min {q. q \ ?s \ q > r}" by auto then have kp1: "k \ ?s \ k > r" using Min_in fin2 none by blast then have kp2: "\x \ ?s. x > r \ x \ k" using Min_le fin2 using k_prop by blast show "(\q \ ?s. (q > r \ (\x \ ?s. x > r \ x \ q)))" using kp1 kp2 by blast qed have h2: "(\q \ ?s. q > r) \ \y'>r. \x\{r<..y'}. a * x\<^sup>2 + b * x + c \ 0" proof - assume "(\q \ ?s. q > r)" then obtain q where q_prop: "q \ ?s \ (q > r \ (\x \ ?s. x > r \ x \ q))" using imp by blast then have "\w. w > r \ w < q" using dense by blast then obtain w where w_prop: "w > r \ w < q" by auto then have "\(\x\{r<..w}. x \ ?s)" using w_prop q_prop by auto then have "\x\{r<..w}. a * x\<^sup>2 + b * x + c \ 0" by blast then show "\y'>r. \x\{r<..y'}. a * x\<^sup>2 + b * x + c \ 0" using w_prop by blast qed have h1: "\(\q \ ?s. q > r) \ \y'>r. \x\{r<..y'}. a * x\<^sup>2 + b * x + c \ 0" proof - assume "\(\q \ ?s. q > r)" then have "\x\{r<..(r+1)}. a * x\<^sup>2 + b * x + c \ 0" using greaterThanAtMost_iff by blast then show ?thesis using less_add_same_cancel1 less_numeral_extra(1) by blast qed then show "\y'>r. \x\{r<..y'}. a * x\<^sup>2 + b * x + c \ 0" using h2 by blast qed (* Show if there are infinitely many values of x where a*x^2 + b*x + c is 0, then the a*x^2 + b*x + c is the zero polynomial *) lemma infzeros : fixes y:: "real" assumes "\x::real < (y::real). a * x\<^sup>2 + b * x + c = 0" shows "a = 0 \ b=0 \ c=0" proof - let ?A = "{(x::real). x < y}" have "\ (n::nat) f. ?A = f ` {i. i < n} \ inj_on f {i. i < n} \ False" proof clarsimp fix n:: "nat" fix f assume xlt: "{x. x < y} = f ` {i. i < n}" assume injh: "inj_on f {i. i < n}" have "?A \ {}" by (simp add: linordered_field_no_lb) then have ngtz: "n > 0" using xlt injh using gr_implies_not_zero by auto have cardisn: "card ?A = n" using xlt injh by (simp add: card_image) have "\k::nat. ((y - (k::nat) - 1) \ ?A)" by auto then have subs: "{k. \(x::nat). k = y - x - 1 \ 0 \ x \ x \ n} \ ?A" by auto have seteq: "(\x. y - real x - 1) ` {0..n} ={k. \(x::nat). k = y - x - 1 \ 0 \ x \ x \ n}" by auto have injf: "inj_on (\x. y - real x - 1) {0..n}" unfolding inj_on_def by auto have "card {k. \(x::nat). k = y - x - 1 \ 0 \ x \ x \ n} = n + 1" using injf seteq card_atMost inj_on_iff_eq_card[where A = "{0..n}", where f = "\x. y - x - 1"] by auto then have if_fin: "finite ?A \ card ?A \ n + 1" using subs card_mono by (metis (lifting) card_mono) then have if_inf: "infinite ?A \ card ?A = 0" by (meson card.infinite) then show "False" using if_fin if_inf cardisn ngtz by auto qed then have nfin: "\ finite {(x::real). x < y}" using finite_imp_nat_seg_image_inj_on by blast have "{(x::real). x < y} \ {x. a*x^2 + b*x + c = 0}" using assms by auto then have nfin2: "\ finite {x. a*x^2 + b*x + c = 0}" using nfin finite_subset by blast { fix x assume "a * x\<^sup>2 + b * x + c = 0" then have f1: "a * (x * x) + x * b + c = 0" by (simp add: Groups.mult_ac(2) power2_eq_square) have f2: "\r. c + (r + (c + - c)) = r + c" by simp have f3: "\r ra rb. (rb::real) * ra + ra * r = (rb + r) * ra" by (metis Groups.mult_ac(2) Rings.ring_distribs(2)) have "\r. r + (c + - c) = r" by simp then have "c + x * (b + x * a) = 0" using f3 f2 f1 by (metis Groups.add_ac(3) Groups.mult_ac(1) Groups.mult_ac(2)) } hence "{x. a*x^2 + b*x + c = 0} \ {x. poly [:c, b, a:] x = 0}" by auto then have " \ finite {x. poly [:c, b, a:] x = 0}" using nfin2 using finite_subset by blast then have "[:c, b, a:] = 0" using poly_roots_finite[where p = "[:c, b, a:]"] by auto then show ?thesis by auto qed lemma have_inbetween_point_leq: fixes r::"real" assumes "(\((d::real), (e::real), (f::real))\set leq. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0)" shows "(\x. (\(a, b, c)\set leq. a * x\<^sup>2 + b * x + c \ 0))" proof - have "(\(d, e, f)\set leq. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\y'>r. (\(d, e, f)\set leq. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using leq_qe_inf_helper assms by auto then have "(\y'>r. (\(d, e, f)\set leq. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using assms by blast then obtain y where y_prop: "y > r \ (\(d, e, f)\set leq. \x\{r<..y}. d * x\<^sup>2 + e * x + f \ 0)" by auto have "\q. q > r \q < y" using y_prop dense by auto then obtain q where q_prop: "q > r \ q < y" by auto then have "(\(d, e, f)\set leq. d*q^2 + e*q + f \ 0)" using y_prop by auto then show ?thesis by auto qed lemma have_inbetween_point_neq: fixes r::"real" assumes "(\((d::real), (e::real), (f::real))\set neq. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0)" shows "(\x. (\(a, b, c)\set neq. a * x\<^sup>2 + b * x + c \ 0))" proof - have "(\(d, e, f)\set neq. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\y'>r. (\(d, e, f)\set neq. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using neq_qe_inf_helper assms by auto then have "(\y'>r. (\(d, e, f)\set neq. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using assms by blast then obtain y where y_prop: "y > r \ (\(d, e, f)\set neq. \x\{r<..y}. d * x\<^sup>2 + e * x + f \ 0)" by auto have "\q. q > r \q < y" using y_prop dense by auto then obtain q where q_prop: "q > r \ q < y" by auto then have "(\(d, e, f)\set neq. d*q^2 + e*q + f \ 0)" using y_prop by auto then show ?thesis by auto qed subsection "Setting up and Helper Lemmas" subsubsection "The les\\_qe lemma" lemma les_qe_forward : shows "(((\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \ (\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ ((\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0))))) \ ((\x. (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)))" proof - assume big_asm: "(((\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \ (\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ ((\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)))))" then have big_or: "(\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \ (\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0)) \ (\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)) \ (\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)) " by auto have h1_helper: "(\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \ (\y.\x(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" proof - show "(\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \ (\y.\x(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" proof (induct les) case Nil then show ?case by auto next case (Cons q les) have ind: " \a\set (q # les). case a of (a, ba, c) \ \x. \y2 + ba * y + c < 0" using Cons.prems by auto then have "case q of (a, ba, c) \ \x. \y2 + ba * y + c < 0 " by simp then obtain y2 where y2_prop: "case q of (a, ba, c) \ (\y2 + ba * y + c < 0)" by auto have "\a\set les. case a of (a, ba, c) \ \x. \y2 + ba * y + c < 0" using ind by simp then have " \y. \xa\set les. case a of (a, ba, c) \ a * x\<^sup>2 + ba * x + c < 0" using Cons.hyps by blast then obtain y1 where y1_prop: "\xa\set les. case a of (a, ba, c) \ a * x^2 + ba * x + c < 0" by blast let ?y = "min y1 y2" have "\x < ?y. (\a\set (q #les). case a of (a, ba, c) \ a * x^2 + ba * x + c < 0)" using y1_prop y2_prop by fastforce then show ?case by blast qed qed then have h1: "(\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \(\x. (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" by (smt (z3) infzeros less_eq_real_def not_numeral_le_zero) (* apply (auto) by (metis (lifting) infzeros zero_neq_numeral) *) have h2: " (\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0)) \ (\x. (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" proof - assume asm: "(\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0))" then obtain a' b' c' where abc_prop: "(a', b', c') \set les \ a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0)" by auto then show "(\x. (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" using have_inbetween_point_les by auto qed have h3: " (\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)) \ ((\x. (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)))" proof - assume asm: "\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)" then obtain a' b' c' where abc_prop: "(a', b', c')\set les \ a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)" by auto then show "(\x. (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" using have_inbetween_point_les by auto qed have h4: "(\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)) \ (\x. (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" proof - assume asm: "(\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)) " then obtain a' b' c' where abc_prop: "(a', b', c')\set les \ a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)" by auto then have "(\x. (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" using have_inbetween_point_les by auto then show ?thesis using asm by auto qed show ?thesis using big_or h1 h2 h3 h4 by blast qed (*sample points, some starter proofs below in comments *) lemma les_qe_backward : shows "(\x. (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)) \ ((\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \ (\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ ((\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0))))" proof - assume havex: "(\x. (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" then obtain x where x_prop: "\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0" by auto have h: "(\ (\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \ \ (\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0)) \ \ (\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)) \ \ (\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0))) \ False" proof - assume big: "(\ (\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \ \ (\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0)) \ \ (\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)) \ \ (\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)))" have notneginf: "\ (\(a, b, c)\set les. \x. \y2 + b * y + c < 0)" using big by auto have notlinroot: "\ (\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0))" using big by auto have notquadroot1: " \ (\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0))" using big by auto have notquadroot2:" \ (\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0))" using big by auto have nok: "\ (\k. \(a, b, c)\set les. a*k^2 + b*k + c = 0 \ (\(d, e, f)\set les. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f < 0))" proof - have "(\k. \(a, b, c)\set les. a*k^2 + b*k + c = 0 \ (\(d, e, f)\set les. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f < 0)) \ False" proof - assume "(\k. \(a, b, c)\set les. a*k^2 + b*k + c = 0 \ (\(d, e, f)\set les. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f < 0))" then obtain k a b c where k_prop: "(a, b, c) \ set les \ a*k^2 + b*k + c = 0 \ (\(d, e, f)\set les. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f < 0)" by auto have azer: "a = 0 \ False" proof - assume azer: "a = 0" then have "b = 0 \ c = 0" using k_prop by auto then have bnonz: "b\ 0" using azer x_prop k_prop by auto then have "k = -c/b" using k_prop azer by (smt (verit, best) mult_eq_0_iff nonzero_mult_div_cancel_left) then have " (\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0))" using k_prop azer bnonz by auto then show "False" using notlinroot by auto qed have anonz: "a \ 0 \ False" proof - assume anonz: "a \ 0 " let ?r1 = "(- b - sqrt (b^2 - 4 * a * c)) / (2 * a)" let ?r2 = "(- b + sqrt (b^2 - 4 * a * c)) / (2 * a)" have discr: "4 * a * c \ b^2" using anonz k_prop discriminant_negative[of a b c] unfolding discrim_def by fastforce then have "k = ?r1 \ k = ?r2" using k_prop discriminant_nonneg[of a b c] unfolding discrim_def using anonz by auto then have "(\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)) \ (\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0))" using discr anonz notquadroot1 notquadroot2 k_prop by auto then show "False" using notquadroot1 notquadroot2 by auto qed show "False" using azer anonz by auto qed then show ?thesis by auto qed have finset: "finite (set les)" by blast have h1: "(\(a, b, c)\set les. a = 0 \ b = 0 \ c = 0) \ False" using x_prop by fastforce then have h2: "\(\(a, b, c)\set les. a = 0 \ b = 0 \ c = 0) \ False" proof - assume nozer: "\(\(a, b, c)\set les. a = 0 \ b = 0 \ c = 0)" then have same_set: "root_set (set les) = set (sorted_root_list_set (set les))" using root_set_finite finset set_sorted_list_of_set by (simp add: nozer root_set_finite sorted_root_list_set_def) have xnotin: "x \ root_set (set les)" unfolding root_set_def using x_prop by auto let ?srl = "sorted_root_list_set (set les)" have notinlist: "\ List.member ?srl x" using xnotin same_set by (simp add: in_set_member) then have notmem: "\n < (length ?srl). x \ nth_default 0 ?srl n" using nth_mem same_set xnotin nth_default_def by metis show ?thesis proof (induct ?srl) case Nil then have "(\(a, b, c)\set les. \x. \y2 + b * y + c < 0)" proof clarsimp fix a b c assume noroots: "[] = sorted_root_list_set (set les)" assume inset: "(a, b, c) \ set les" have "{} = root_set (set les)" using noroots same_set by auto then have nozero: "\(\x. a*x^2 + b*x + c = 0)" using inset unfolding root_set_def by auto have "\y2 + b * y + c < 0" proof clarsimp fix y assume "y < x" then have "sign_num (a*x^2 + b*x + c) = sign_num (a*y^2 + b*y + c)" using nozero by (metis changes_sign_var) then show "a * y\<^sup>2 + b * y + c < 0" unfolding sign_num_def using x_prop inset by (smt split_conv) qed then show "\x. \y2 + b * y + c < 0" by auto qed then show ?case using notneginf by auto next case (Cons w xa) (* Need to argue that x isn't greater than the largest element of ?srl *) (* that if srl has length \ 2, x isn't in between any of the roots of ?srl*) (* and that x isn't less than the lowest root in ?srl *) then have lengthsrl: "length ?srl > 0" by auto have neginf: "x < nth_default 0 ?srl 0 \ False" proof - assume xlt: "x < nth_default 0 ?srl 0" have all: "(\(a, b, c)\set les. \y2 + b * y + c < 0)" proof clarsimp fix a b c y assume inset: "(a, b, c) \ set les" assume "y < x" have xl: "a*x^2 + b*x + c < 0" using x_prop inset by auto have "\(\q. q < nth_default 0 ?srl 0 \ a*q^2 + b*q + c = 0)" proof - have "(\q. q < nth_default 0 ?srl 0 \ a*q^2 + b*q + c = 0) \ False" proof - assume "\q. q < nth_default 0 ?srl 0 \ a*q^2 + b*q + c = 0" then obtain q where q_prop: "q < nth_default 0 ?srl 0 \a*q^2 + b*q + c = 0" by auto then have " q \ root_set (set les)" unfolding root_set_def using inset by auto then have "List.member ?srl q" using same_set by (simp add: in_set_member) then have "q \ nth_default 0 ?srl 0" using sorted_sorted_list_of_set[where A = "root_set (set les)"] unfolding sorted_root_list_set_def by (metis \q \ root_set (set les)\ in_set_conv_nth le_less_linear lengthsrl not_less0 nth_default_nth same_set sorted_nth_mono sorted_root_list_set_def) then show "False" using q_prop by auto qed then show ?thesis by auto qed then have "\(\q. q < x \ a*q^2 + b*q + c = 0)" using xlt by auto then show " a * y\<^sup>2 + b * y + c < 0" using xl changes_sign_var[where a = "a", where b = "b", where c = "c", where x = "y", where y = "x"] unfolding sign_num_def using \y < x\ less_eq_real_def zero_neq_numeral by fastforce qed have "(\(a, b, c)\set les. \x. \y2 + b * y + c < 0)" proof clarsimp fix a b c assume "(a, b, c)\set les" then show "\x. \y2 + b * y + c < 0" using all by blast qed then show "False" using notneginf by auto qed have "x > nth_default 0 ?srl (length ?srl - 1) \ (\k. \(a, b, c)\set les. a*k^2 + b*k + c = 0 \ (\(d, e, f)\set les. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f < 0))" proof - assume xgt: "x > nth_default 0 ?srl (length ?srl - 1)" let ?lg = "nth_default 0 ?srl (length ?srl - 1)" have "List.member ?srl ?lg" by (metis diff_less in_set_member lengthsrl nth_default_def nth_mem zero_less_one) then have "?lg \ root_set (set les) " using same_set in_set_member[of ?lg ?srl] by auto then have exabc: "\(a, b, c)\set les. a*?lg^2 + b*?lg + c = 0" unfolding root_set_def by auto have "(\(d, e, f)\set les. \q\{?lg<..x}. d * q^2 + e * q + f < 0)" proof clarsimp fix d e f q assume inset: "(d, e, f) \ set les" assume qgt: "(nth_default 0) (sorted_root_list_set (set les)) (length (sorted_root_list_set (set les)) - Suc 0) < q" assume qlt: "q \ x" have nor: "\(\r. d * r^2 + e * r + f = 0 \ r > ?lg)" proof - have "(\r. d * r^2 + e * r + f = 0 \ r > ?lg) \ False " proof - assume "\r. d * r^2 + e * r + f = 0 \ r > ?lg" then obtain r where r_prop: "d*r^2 + e*r + f = 0 \ r > ?lg" by auto then have "r \ root_set (set les)" using inset unfolding root_set_def by auto then have "List.member ?srl r" using same_set in_set_member by (simp add: in_set_member) then have " r \ ?lg" using sorted_sorted_list_of_set nth_default_def by (metis One_nat_def Suc_pred \r \ root_set (set les)\ in_set_conv_nth lengthsrl lessI less_Suc_eq_le same_set sorted_nth_mono sorted_root_list_set_def) then show "False" using r_prop by auto qed then show ?thesis by auto qed then have xltz_helper: "\(\r. r \ q \ d * r^2 + e * r + f = 0)" using qgt by auto then have xltz: "d*x^2 + e*x + f < 0" using inset x_prop by auto show "d * q\<^sup>2 + e * q + f < 0" using qlt qgt nor changes_sign_var[of d _ e f _] xltz xltz_helper unfolding sign_num_def apply (auto) by smt qed then have " (\(d, e, f)\set les. \y'>?lg. \x\{?lg<..y'}. d * x\<^sup>2 + e * x + f < 0)" using xgt by auto then have "(\(a, b, c)\set les. a*?lg^2 + b*?lg + c = 0 \ (\(d, e, f)\set les. \y'>?lg. \x\{?lg<..y'}. d * x\<^sup>2 + e * x + f < 0))" using exabc by auto then show "(\k. \(a, b, c)\set les. a*k^2 + b*k + c = 0 \ (\(d, e, f)\set les. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f < 0))" by auto qed then have posinf: "x > nth_default 0 ?srl (length ?srl - 1) \ False" using nok by auto have "(\n. (n+1) < (length ?srl) \ x > (nth_default 0 ?srl) n \ x < (nth_default 0 ?srl (n + 1))) \ (\k. \(a, b, c)\set les. a*k^2 + b*k + c = 0 \ (\(d, e, f)\set les. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f < 0))" proof - assume "\n. (n+1) < (length ?srl) \ x > nth_default 0 ?srl n \ x < nth_default 0 ?srl (n + 1)" then obtain n where n_prop: "(n+1) < (length ?srl) \ x > nth_default 0 ?srl n \ x < nth_default 0 ?srl (n + 1)" by auto let ?elt = "nth_default 0 ?srl n" let ?elt2 = "nth_default 0 ?srl (n + 1)" have "List.member ?srl ?elt" using n_prop nth_default_def by (metis add_lessD1 in_set_member nth_mem) then have "?elt \ root_set (set les) " using same_set in_set_member[of ?elt ?srl] by auto then have exabc: "\(a, b, c)\set les. a*?elt^2 + b*?elt + c = 0" unfolding root_set_def by auto then obtain a b c where "(a, b, c)\set les \ a*?elt^2 + b*?elt + c = 0" by auto have xltel2: "x < ?elt2" using n_prop by auto have xgtel: "x > ?elt " using n_prop by auto have "(\(d, e, f)\set les. \q\{?elt<..x}. d * q^2 + e * q + f < 0)" proof clarsimp fix d e f q assume inset: "(d, e, f) \ set les" assume qgt: "nth_default 0 (sorted_root_list_set (set les)) n < q" assume qlt: "q \ x" have nor: "\(\r. d * r^2 + e * r + f = 0 \ r > ?elt \r < ?elt2)" proof - have "(\r. d * r^2 + e * r + f = 0 \ r > ?elt \ r < ?elt2) \ False " proof - assume "\r. d * r^2 + e * r + f = 0 \ r > ?elt \ r < ?elt2" then obtain r where r_prop: "d*r^2 + e*r + f = 0 \ r > ?elt \ r < ?elt2" by auto then have "r \ root_set (set les)" using inset unfolding root_set_def by auto then have "List.member ?srl r" using same_set in_set_member by (simp add: in_set_member) then have "\i < (length ?srl). r = nth_default 0 ?srl i" by (metis \r \ root_set (set les)\ in_set_conv_nth same_set nth_default_def) then obtain i where i_prop: "i < (length ?srl) \ r = nth_default 0 ?srl i" by auto have "r > ?elt" using r_prop by auto then have igt: " i > n" using i_prop sorted_sorted_list_of_set by (smt add_lessD1 leI n_prop nth_default_def sorted_nth_mono sorted_root_list_set_def) have "r < ?elt2" using r_prop by auto then have ilt: " i < n + 1" using i_prop sorted_sorted_list_of_set by (smt leI n_prop nth_default_def sorted_nth_mono sorted_root_list_set_def) then show "False" using igt ilt by auto qed then show ?thesis by auto qed then have nor: "\(\r. d * r^2 + e * r + f = 0 \ r > ?elt \r \ x)" using xltel2 xgtel by auto then have xltz: "d*x^2 + e*x + f < 0" using inset x_prop by auto show "d * q\<^sup>2 + e * q + f < 0" using qlt qgt nor changes_sign_var[of d _ e f _] xltz unfolding sign_num_def by smt qed then have " (\(d, e, f)\set les. \y'>?elt. \x\{?elt<..y'}. d * x\<^sup>2 + e * x + f < 0)" using xgtel xltel2 by auto then have "(\(a, b, c)\set les. a*?elt^2 + b*?elt + c = 0 \ (\(d, e, f)\set les. \y'>?elt. \x\{?elt<..y'}. d * x\<^sup>2 + e * x + f < 0))" using exabc by auto then show "(\k. \(a, b, c)\set les. a*k^2 + b*k + c = 0 \ (\(d, e, f)\set les. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f < 0))" by auto qed then have inbetw: "(\n. (n+1) < (length ?srl) \ x > nth_default 0 ?srl n \ x < nth_default 0 ?srl (n + 1)) \ False" using nok by auto have lenzer: "length xa = 0 \ False" proof - assume "length xa = 0" have xis: "x > w \ x < w" using notmem Cons.hyps by (smt list.set_intros(1) same_set xnotin) have xgt: "x > w \ False" proof - assume xgt: "x > w" show "False" using posinf Cons.hyps by (metis One_nat_def Suc_eq_plus1 \length xa = 0\ cancel_comm_monoid_add_class.diff_cancel list.size(4) nth_default_Cons_0 xgt) qed have xlt: "x < w \ False" proof - assume xlt: "x < w" show "False" using neginf Cons.hyps by (metis nth_default_Cons_0 xlt) qed show "False" using xis xgt xlt by auto qed have lengt: "length xa > 0 \ False" proof - assume "length xa > 0" have "x \ nth_default 0 ?srl 0" using neginf by fastforce then have xgtf: "x > nth_default 0 ?srl 0" using notmem using Cons.hyps(2) by fastforce have "x \ nth_default 0 ?srl (length ?srl - 1)" using posinf by fastforce then have "(\n. (n+1) < (length ?srl) \ x \ nth_default 0 ?srl n \ x \ nth_default 0 ?srl (n + 1))" using lengthsrl xgtf notmem sorted_list_prop[where l = ?srl, where x = "x"] by (metis add_lessD1 diff_less nth_default_nth sorted_root_list_set_def sorted_sorted_list_of_set zero_less_one) then obtain n where n_prop: "(n+1) < (length ?srl) \ x \ nth_default 0 ?srl n \ x \ nth_default 0 ?srl (n + 1)" by auto then have "x > nth_default 0 ?srl n \ x < nth_default 0 ?srl (n+1)" using notmem by (metis Suc_eq_plus1 Suc_lessD less_eq_real_def) then have "(\n. (n+1) < (length ?srl) \ x > nth_default 0 ?srl n \ x < nth_default 0 ?srl (n + 1))" using n_prop by blast then show "False" using inbetw by auto qed then show ?case using lenzer lengt by auto qed qed show "False" using h1 h2 by auto qed then have equiv_false: "\((\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \ (\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0)) \ (\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)) \ (\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0))) \ False" by linarith have "\((\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \ (\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ ((\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)))) \ False" proof - assume "\((\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \ (\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ ((\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0))))" then have "\((\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \ (\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0)) \ (\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)) \ (\(a', b', c')\set les. a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)))" by auto then show ?thesis using equiv_false by auto qed then show ?thesis by blast qed lemma les_qe : shows "(\x. (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)) = ((\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \ (\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ ((\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0))))" proof - have first: "(\x. (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)) \ ((\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \ (\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ ((\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0))))" using les_qe_backward by auto have second: "((\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \ (\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ ((\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0)))) \ (\x. (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)) " using les_qe_forward by auto have "(\x. (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)) \ ((\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \ (\(a', b', c')\set les. a' = 0 \ b' \ 0 \ (\(d, e, f)\set les. \y'>- (c' / b'). \x\{- (c' / b')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ a' \ 0 \ 4 * a' * c' \ b'\<^sup>2 \ ((\(d, e, f)\set les. \y'>(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a'). \x\{(sqrt (b'\<^sup>2 - 4 * a' * c') - b') / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set les. \y'>(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' - sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0))))" using first second by meson then show ?thesis by blast qed subsubsection "equiv\\_lemma" lemma equiv_lemma: assumes big_asm: "(\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0)) \ (\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))) \ (\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0)) \ ((\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" shows "((\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))) \ (\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" proof - let ?t = " ((\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))) \ (\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" have h1: "(\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0)) \ ?t" by auto have h2: "(\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))) \ ?t" by auto have h3: "(\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0)) \ ?t" by auto show ?thesis using big_asm h1 h2 h3 by presburger qed subsubsection "The eq\\_qe lemma" lemma eq_qe_forwards: shows "(\x. (\(a, b, c)\set eq. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)) \ ((\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))) \ (\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" proof - let ?big_or = "(\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0)) \ (\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))) \ (\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0)) \ ((\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" assume asm: "(\x. (\(a, b, c)\set eq. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)) " then obtain x where x_prop: "(\(a, b, c)\set eq. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)" by auto have "\ (\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0)) \ \ (\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))) \ \ (\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0)) \ \ ((\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)) \ False" proof - assume big_conj: "\ (\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0)) \ \ (\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))) \ \ (\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0)) \ \ ((\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" have not_lin: "\(\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0))" using big_conj by auto have not_quad1: "\(\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0)))" using big_conj by auto have not_quad2: "\(\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))" using big_conj by auto have not_zer: "\((\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" using big_conj by auto then have not_zer1: "\(\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)" by auto have "(\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)" using asm by auto then have "\(\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0)" using not_zer1 by auto then have "\ (d, e, f)\set eq. d \ 0 \ e \ 0 \ f \ 0 " by auto then obtain d e f where def_prop: "(d, e, f) \ set eq \ (d \ 0 \ e \ 0 \ f \ 0)" by auto then have eval_at_x: "d*x^2 + e*x + f = 0" using x_prop by auto have dnonz: "d \ 0 \ False" proof - assume dneq: "d \ 0" then have discr: "-(e^2) + 4 *d *f \ 0" using discriminant_negative[of d e f x] eval_at_x unfolding discrim_def by linarith let ?r1 = "(- e + - 1 * sqrt (e^2 - 4 *d *f)) / (2 * d)" let ?r2 = "(- e + 1 * sqrt (e^2 - 4 *d *f)) / (2 * d)" have xis: "x = ?r1 \ x = ?r2" using dneq discr discriminant_nonneg[of d e f x] eval_at_x unfolding discrim_def by auto have xr1: "x = ?r1 \ False" using not_quad2 x_prop discr def_prop dneq by auto have xr2: "x = ?r2 \ False" using not_quad1 x_prop discr def_prop dneq by auto show "False" using xr1 xr2 xis by auto qed then have dz: "d = 0" by auto have enonz: "e \ 0 \ False" proof - assume enonz: "e\ 0" then have "x = -f/e" using dz eval_at_x by (metis add.commute minus_add_cancel mult.commute mult_zero_class.mult_zero_left nonzero_eq_divide_eq) then show "False" using not_lin x_prop enonz dz def_prop by auto qed then have ez: "e = 0" by auto have fnonz: "f \ 0 \ False" using ez dz eval_at_x by auto show "False" using def_prop dnonz enonz fnonz by auto qed then have h: "\(?big_or) \ False" by auto then show ?thesis using equiv_lemma by presburger qed lemma eq_qe_backwards: "((\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))) \ (\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)) \ (\x. ((\(a, b, c)\set eq. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))) " proof - assume "((\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))) \ (\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" then have bigor: "(\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0)) \ (\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))) \ (\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0)) \ ((\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" by auto have h1: "(\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0)) \ (\(x::real). (\(a, b, c)\set eq. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" proof - assume "(\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0))" then obtain a' b' c' where abc_prop: "(a', b', c')\set eq \ (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0)" by auto let ?x = "(-c' /b')::real" have "(\(d, e, f)\set eq. d * ?x\<^sup>2 + e * ?x + f = 0) \ (\(d, e, f)\set les. d * ?x^2 + e * ?x + f < 0)" using abc_prop by auto then show ?thesis using abc_prop by blast qed have h2: " (\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))) \ (\x. (\(a, b, c)\set eq. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" proof - assume "(\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0)))" then obtain a' b' c' where abc_prop: "(a', b', c')\set eq \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))" by auto let ?x = "((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')::real)" have anonz: "a' \ 0" using abc_prop by auto then have "\(q::real). q = ?x" by auto then obtain q where q_prop: "q = ?x" by auto have "(\(d, e, f)\set eq. d * (?x)\<^sup>2 + e * (?x) + f = 0) \ (\(d, e, f)\set les. d * (?x)\<^sup>2 + e * (?x) + f < 0)" using abc_prop by auto then have "(\(d, e, f)\set eq. d * q\<^sup>2 + e * q + f = 0) \ (\(d, e, f)\set les. d * q\<^sup>2 + e * q + f < 0)" using q_prop by auto then show ?thesis by auto qed have h3: "(\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0)) \ (\x. (\(a, b, c)\set eq. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" proof - assume "(\(a', b', c')\set eq. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))" then obtain a' b' c' where abc_prop: "a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (a', b', c')\set eq \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0)" by auto let ?x = "(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')" have anonz: "a' \ 0" using abc_prop by auto then have "\(q::real). q = ?x" by auto then obtain q where q_prop: "q = ?x" by auto have "(\(d, e, f)\set eq. d * (?x)\<^sup>2 + e * (?x) + f = 0) \ (\(d, e, f)\set les. d * (?x)\<^sup>2 + e * (?x) + f < 0)" using abc_prop by auto then have "(\(d, e, f)\set eq. d * q\<^sup>2 + e * q + f = 0) \ (\(d, e, f)\set les. d * q\<^sup>2 + e * q + f < 0)" using q_prop by auto then show ?thesis by auto qed have h4: "((\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)) \ (\x. (\(a, b, c)\set eq. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" proof - assume asm: "((\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" then have allzer: "(\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0)" by auto have "(\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)" using asm by auto then obtain x where x_prop: " \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0" by auto then have "\(d, e, f)\set eq. d*x^2 + e*x + f = 0" using allzer by auto then show ?thesis using x_prop by auto qed show ?thesis using bigor h1 h2 h3 h4 by blast qed lemma eq_qe : "(\x. ((\(a, b, c)\set eq. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))) = ((\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))) \ (\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" proof - have h1: "(\x. (\(a, b, c)\set eq. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)) \ ((\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))) \ (\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" using eq_qe_forwards by auto have h2: "((\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))) \ (\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)) \ (\x. (\(a, b, c)\set eq. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" using eq_qe_backwards by auto have h3: "(\x. (\(a, b, c)\set eq. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0)) \ ((\(a', b', c')\set eq. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set les. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set eq. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set les. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0))) \ (\(d, e, f)\set eq. d = 0 \ e = 0 \ f = 0) \ (\x. \(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" using h1 h2 by smt then show ?thesis by (auto) qed subsubsection "The qe\\_forwards lemma" lemma qe_forwards_helper_gen: fixes r:: "real" assumes f8: "\(\((a'::real), (b'::real), (c'::real))\set c. ((a'\ 0 \ b'\ 0) \ a'*r^2 + b'*r + c' = 0) \ ((\(d, e, f)\set a. d * r\<^sup>2 + e * r + f = 0) \ (\(d, e, f)\set b. d * r^2 + e * r + f < 0) \ (\(d, e, f)\set c. d * r^2 + e * r + f \ 0) \ (\(d, e, f)\set d. d * r^2 + e * r + f \ 0)))" assumes alleqset: "\x. (\(d, e, f)\set a. d * x^2 + e * x + f = 0)" assumes f5: "\(\(a', b', c')\set b. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0))" assumes f6: "\ (\(a', b', c')\set b. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" assumes f7: "\ (\(a', b', c')\set b. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))" assumes f10: "\(\(a', b', c')\set d. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0))" assumes f11: "\(\(a', b', c')\set d. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" assumes f12: "\(\(a', b', c')\set d. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))" shows "\(\(a', b', c')\set c. ((a'\ 0 \ b'\ 0) \ a'*r^2 + b'*r + c' = 0) \ (\(d, e, f)\set a. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0))" proof - have "(\(a', b', c')\set c. ((a'\ 0 \ b'\ 0) \ a'*r^2 + b'*r + c' = 0) \ (\(d, e, f)\set a. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0)) \ False" proof - assume "(\(a', b', c')\set c. ((a'\ 0 \ b'\ 0) \ a'*r^2 + b'*r + c' = 0) \ (\(d, e, f)\set a. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0))" then obtain a' b' c' where abc_prop: "(a', b', c')\set c \ ((a'\ 0 \ b'\ 0) \ a'*r^2 + b'*r + c' = 0) \ (\(d, e, f)\set a. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0)" by auto have h1: "(\(d, e, f)\set a. d * r^2 + e * r + f = 0)" using alleqset by blast have c_prop: "(\(d, e, f)\set c. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0)" using abc_prop by auto have h2: "(\(d, e, f)\set c. d *r^2 + e * r + f \ 0)" proof - have c1: "\ (d, e, f) \ set c. d * (r)\<^sup>2 + e * (r) + f > 0 \ False" proof - assume "\ (d, e, f) \ set c. d * (r)\<^sup>2 + e * (r) + f > 0" then obtain d e f where def_prop: "(d, e, f) \ set c \ d * (r)\<^sup>2 + e * r + f > 0" by auto have "\y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0" using def_prop c_prop by auto then obtain y' where y_prop: " y' >r \ (\x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0)" by auto have "\x\{r<..y'}. d*x^2 + e*x + f > 0" using def_prop continuity_lem_gt0_expanded[of "r" y' d e f] using y_prop by linarith then show "False" using y_prop by auto qed then show ?thesis by fastforce qed have b_prop: "(\(d, e, f)\set b. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f < 0)" using abc_prop by auto have h3: "(\(d, e, f)\set b. d * r\<^sup>2 + e * r + f < 0)" proof - have c1: "\ (d, e, f) \ set b. d * r\<^sup>2 + e * r + f > 0 \ False" proof - assume "\ (d, e, f) \ set b. d * r\<^sup>2 + e * r + f > 0" then obtain d e f where def_prop: "(d, e, f) \ set b \ d * r\<^sup>2 + e * r + f > 0" by auto then have "\y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f < 0" using b_prop by auto then obtain y' where y_prop: " y' >r \ (\x\{r<..y'}. d * x\<^sup>2 + e * x + f < 0)" by auto then have "\k. k > r \ k < y' \ d * k^2 + e * k + f < 0" using dense by (meson dense greaterThanAtMost_iff less_eq_real_def) then obtain k where k_prop: "k > r \ k < y' \ d * k^2 + e * k + f < 0" by auto then have "\(\x>r. x < y' \ d * x\<^sup>2 + e * x + f = 0)" using y_prop by force then show "False" using k_prop def_prop y_prop poly_IVT_neg[of "r" "k" "[:f, e, d:]"] poly_IVT_pos[of "-c'/b'" "k" "[:f, e, d:]"] by (smt quadratic_poly_eval) qed have c2: "\ (d, e, f) \ set b. d * r\<^sup>2 + e * r + f = 0 \ False" proof - assume "\ (d, e, f) \ set b. d * r\<^sup>2 + e * r + f = 0" then obtain d' e f where def_prop: "(d', e, f) \ set b \ d' * r\<^sup>2 + e * r + f = 0" by auto then have same: "(d' = 0 \ e \ 0) \ (-f/e = r)" proof - assume asm: "(d' = 0 \ e \ 0)" then have " e * r + f = 0" using def_prop by auto then show "-f/e = r" using asm by (metis (no_types) add.commute diff_0 divide_minus_left minus_add_cancel nonzero_mult_div_cancel_left uminus_add_conv_diff) qed let ?r = "-f/e" have "(d' = 0 \ e \ 0) \ ((d', e, f) \ set b \ ((\(d, e, f)\set a. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using same def_prop abc_prop by auto then have "(d' = 0 \ e \ 0) \ (\(a', b', c')\set b. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0))" by auto then have f1: "(d' = 0 \ e \ 0) \ False" using f5 by auto have f2: "(d' = 0 \ e = 0 \ f = 0) \ False" proof - assume "(d' = 0 \ e = 0 \ f = 0)" then have allzer: "\x. d'*x^2 + e*x + f = 0" by auto have "\y'>r. \x\{r<..y'}. d' * x\<^sup>2 + e * x + f < 0" using b_prop def_prop by auto then obtain y' where y_prop: " y' >r \ (\x\{r<..y'}. d' * x\<^sup>2 + e * x + f < 0)" by auto then have "\k. k > r \ k < y' \ d' * k^2 + e * k + f < 0" using dense by (meson dense greaterThanAtMost_iff less_eq_real_def) then show "False" using allzer by auto qed have f3: "d' \ 0 \ False" proof - assume dnonz: "d' \ 0" have discr: " - e\<^sup>2 + 4 * d' * f \ 0" using def_prop discriminant_negative[of d' e f] unfolding discrim_def using def_prop by fastforce then have two_cases: "r = (- e + - 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d') \ r = (- e + 1 * sqrt (e\<^sup>2 - 4 * d' * f)) / (2 * d')" using def_prop dnonz discriminant_nonneg[of d' e f] unfolding discrim_def by fastforce have some_props: "((d', e, f) \ set b \ d' \ 0 \ - e\<^sup>2 + 4 * d' * f \ 0)" using dnonz def_prop discr by auto let ?r1 = "(- e + - 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d')" let ?r2 = "(- e + 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d')" have cf1: "r = (- e + - 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d') \ False" proof - assume "r = (- e + - 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d')" then have "(d', e, f) \ set b \ d' \ 0 \ - e\<^sup>2 + 4 * d' * f \ 0 \ ((\(d, e, f)\set a. \y'>?r1. \x\{?r1<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?r1. \x\{?r1<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?r1. \x\{?r1<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?r1. \x\{?r1<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using abc_prop some_props by auto then show "False" using f7 by auto qed have cf2: "r = (- e + 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d') \ False" proof - assume "r = (- e + 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d')" then have "(d', e, f) \ set b \ d' \ 0 \ - e\<^sup>2 + 4 * d' * f \ 0 \ ((\(d, e, f)\set a. \y'>?r2. \x\{?r2<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?r2. \x\{?r2<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?r2. \x\{?r2<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?r2. \x\{?r2<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using abc_prop some_props by auto then show "False" using f6 by auto qed then show "False" using two_cases cf1 cf2 by auto qed (* discriminant_nonnegative *) have eo: "(d' \ 0) \ (d' = 0 \ e \ 0) \ (d' = 0 \ e = 0 \ f = 0)" using def_prop by auto then show "False" using f1 f2 f3 by auto qed show ?thesis using c1 c2 by fastforce qed have d_prop: "(\(d, e, f)\set d. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0)" using abc_prop by auto have h4: "(\(d, e, f)\set d. d * r\<^sup>2 + e * r + f \ 0)" proof - have "(\(d, e, f)\set d. d * r\<^sup>2 + e * r + f = 0) \ False" proof - assume "\ (d, e, f) \ set d. d * r\<^sup>2 + e * r + f = 0" then obtain d' e f where def_prop: "(d', e, f) \ set d \ d' * r\<^sup>2 + e * r + f = 0" by auto then have same: "(d' = 0 \ e \ 0) \ (-f/e = r)" proof - assume asm: "(d' = 0 \ e \ 0)" then have " e * r + f = 0" using def_prop by auto then show "-f/e = r" using asm by (metis (no_types) add.commute diff_0 divide_minus_left minus_add_cancel nonzero_mult_div_cancel_left uminus_add_conv_diff) qed let ?r = "-f/e" have "(d' = 0 \ e \ 0) \ ((d', e, f) \ set d \ ((\(d, e, f)\set a. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using same def_prop abc_prop by auto then have "(d' = 0 \ e \ 0) \ (\(a', b', c')\set d. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'> -c'/b'. \x\{ -c'/b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'> -c'/b'. \x\{ -c'/b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'> -c'/b'. \x\{ -c'/b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'> -c'/b'. \x\{ -c'/b'<..y'}. d * x\<^sup>2 + e * x + f \ 0))" by auto then have f1: "(d' = 0 \ e \ 0) \ False" using f10 by auto have f2: "(d' = 0 \ e = 0 \ f = 0) \ False" proof - assume "(d' = 0 \ e = 0 \ f = 0)" then have allzer: "\x. d'*x^2 + e*x + f = 0" by auto have "\y'> r. \x\{ r<..y'}. d' * x\<^sup>2 + e * x + f \ 0" using d_prop def_prop by auto then obtain y' where y_prop: " y' >r \ (\x\{r<..y'}. d' * x\<^sup>2 + e * x + f \ 0)" by auto then have "\k. k > r \ k < y' \ d' * k^2 + e * k + f \ 0" using dense by (meson dense greaterThanAtMost_iff less_eq_real_def) then show "False" using allzer by auto qed have f3: "d' \ 0 \ False" proof - assume dnonz: "d' \ 0" have discr: " - e\<^sup>2 + 4 * d' * f \ 0" using def_prop discriminant_negative[of d' e f] unfolding discrim_def by fastforce then have two_cases: "r = (- e + - 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d') \ r = (- e + 1 * sqrt (e\<^sup>2 - 4 * d' * f)) / (2 * d')" using def_prop dnonz discriminant_nonneg[of d' e f] unfolding discrim_def by fastforce have some_props: "((d', e, f) \ set d \ d' \ 0 \ - e\<^sup>2 + 4 * d' * f \ 0)" using dnonz def_prop discr by auto let ?r1 = "(- e + - 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d')" let ?r2 = "(- e + 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d')" have cf1: "r = (- e + - 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d') \ False" proof - assume "r = (- e + - 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d')" then have "(d', e, f) \ set d \ d' \ 0 \ - e\<^sup>2 + 4 * d' * f \ 0 \ ((\(d, e, f)\set a. \y'>?r1. \x\{?r1<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?r1. \x\{?r1<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?r1. \x\{?r1<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?r1. \x\{?r1<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using abc_prop some_props by auto then show "False" using f12 by auto qed have cf2: "r = (- e + 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d') \ False" proof - assume "r = (- e + 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d')" then have "(d', e, f) \ set d \ d' \ 0 \ - e\<^sup>2 + 4 * d' * f \ 0 \ ((\(d, e, f)\set a. \y'>?r2. \x\{?r2<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?r2. \x\{?r2<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?r2. \x\{?r2<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?r2. \x\{?r2<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using abc_prop some_props by auto then show "False" using f11 by auto qed then show "False" using two_cases cf1 cf2 by auto qed (* discriminant_nonnegative *) have eo: "(d' \ 0) \ (d' = 0 \ e \ 0) \ (d' = 0 \ e = 0 \ f = 0)" using def_prop by auto then show "False" using f1 f2 f3 by auto qed then show ?thesis by auto qed have "(\(a', b', c')\set c. ((a'\ 0 \ b'\ 0) \ a'*r^2 + b'*r + c' = 0) \ (\(d, e, f)\set a. d * r\<^sup>2 + e * r + f = 0) \ (\(d, e, f)\set b. d * r\<^sup>2 + e * r + f < 0) \ (\(d, e, f)\set c. d * r\<^sup>2 + e * r + f \ 0) \ (\(d, e, f)\set d. d * r\<^sup>2 + e * r + f \ 0))" using h1 h2 h3 h4 abc_prop by auto then show "False" using f8 by auto qed then show ?thesis by auto qed lemma qe_forwards_helper_lin: assumes alleqset: "\x. (\(d, e, f)\set a. d * x^2 + e * x + f = 0)" assumes f5: "\(\(a', b', c')\set b. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0))" assumes f6: "\ (\(a', b', c')\set b. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" assumes f7: "\ (\(a', b', c')\set b. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))" assumes f8: "\(\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0))" assumes f10: "\(\(a', b', c')\set d. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0))" assumes f11: "\(\(a', b', c')\set d. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" assumes f12: "\(\(a', b', c')\set d. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))" shows "\(\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0))" proof - have "(\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0)) \ False" proof - assume "(\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0))" then obtain a' b' c' where abc_prop: "(a', b', c')\set c \ (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0)" by auto then have bnonz: "b'\0" by auto have h1: "(\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0)" using bnonz alleqset by blast have c_prop: "(\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0)" using abc_prop by auto have h2: "(\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0)" proof - have c1: "\ (d, e, f) \ set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f > 0 \ False" proof - assume "\ (d, e, f) \ set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f > 0" then obtain d e f where def_prop: "(d, e, f) \ set c \ d * (- c' / b')\<^sup>2 + e * (- c' / b') + f > 0" by auto have "\y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0" using def_prop c_prop by auto then obtain y' where y_prop: " y' >- c' / b' \ (\x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0)" by auto have "\x\{(-c'/b')<..y'}. d*x^2 + e*x + f > 0" using def_prop continuity_lem_gt0_expanded[of "(-c'/b')" y' d e f] using y_prop by linarith then show "False" using y_prop by auto qed then show ?thesis by fastforce qed have b_prop: "(\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0)" using abc_prop by auto have h3: "(\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0)" proof - have c1: "\ (d, e, f) \ set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f > 0 \ False" proof - assume "\ (d, e, f) \ set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f > 0" then obtain d e f where def_prop: "(d, e, f) \ set b \ d * (- c' / b')\<^sup>2 + e * (- c' / b') + f > 0" by auto then have "\y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0" using b_prop by auto then obtain y' where y_prop: " y' >- c' / b' \ (\x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0)" by auto then have "\k. k > -c'/b' \ k < y' \ d * k^2 + e * k + f < 0" using dense by (meson dense greaterThanAtMost_iff less_eq_real_def) then obtain k where k_prop: "k > -c'/b' \ k < y' \ d * k^2 + e * k + f < 0" by auto then have "\(\x>(-c'/b'). x < y' \ d * x\<^sup>2 + e * x + f = 0)" using y_prop by force then show "False" using k_prop def_prop y_prop poly_IVT_neg[of "-c'/b'" "k" "[:f, e, d:]"] poly_IVT_pos[of "-c'/b'" "k" "[:f, e, d:]"] by (smt quadratic_poly_eval) qed have c2: "\ (d, e, f) \ set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0 \ False" proof - assume "\ (d, e, f) \ set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0" then obtain d' e f where def_prop: "(d', e, f) \ set b \ d' * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0" by auto then have same: "(d' = 0 \ e \ 0) \ (-f/e = -c'/b')" proof - assume asm: "(d' = 0 \ e \ 0)" then have " e * (- c' / b') + f = 0" using def_prop by auto then show "-f/e = -c'/b'" using asm by auto qed let ?r = "-f/e" have "(d' = 0 \ e \ 0) \ ((d', e, f) \ set b \ ((\(d, e, f)\set a. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using same def_prop abc_prop by auto then have "(d' = 0 \ e \ 0) \ (\(a', b', c')\set b. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0))" by auto then have f1: "(d' = 0 \ e \ 0) \ False" using f5 by auto have f2: "(d' = 0 \ e = 0 \ f = 0) \ False" proof - assume "(d' = 0 \ e = 0 \ f = 0)" then have allzer: "\x. d'*x^2 + e*x + f = 0" by auto have "\y'>- c' / b'. \x\{- c' / b'<..y'}. d' * x\<^sup>2 + e * x + f < 0" using b_prop def_prop by auto then obtain y' where y_prop: " y' >- c' / b' \ (\x\{- c' / b'<..y'}. d' * x\<^sup>2 + e * x + f < 0)" by auto then have "\k. k > -c'/b' \ k < y' \ d' * k^2 + e * k + f < 0" using dense by (meson dense greaterThanAtMost_iff less_eq_real_def) then show "False" using allzer by auto qed have f3: "d' \ 0 \ False" proof - assume dnonz: "d' \ 0" have discr: " - e\<^sup>2 + 4 * d' * f \ 0" using def_prop discriminant_negative[of d' e f] unfolding discrim_def by fastforce then have two_cases: "(- c' / b') = (- e + - 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d') \ (- c' / b') = (- e + 1 * sqrt (e\<^sup>2 - 4 * d' * f)) / (2 * d')" using def_prop dnonz discriminant_nonneg[of d' e f] unfolding discrim_def by fastforce have some_props: "((d', e, f) \ set b \ d' \ 0 \ - e\<^sup>2 + 4 * d' * f \ 0)" using dnonz def_prop discr by auto let ?r1 = "(- e + - 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d')" let ?r2 = "(- e + 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d')" have cf1: "(- c' / b') = (- e + - 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d') \ False" proof - assume "(- c' / b') = (- e + - 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d')" then have "(d', e, f) \ set b \ d' \ 0 \ - e\<^sup>2 + 4 * d' * f \ 0 \ ((\(d, e, f)\set a. \y'>?r1. \x\{?r1<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?r1. \x\{?r1<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?r1. \x\{?r1<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?r1. \x\{?r1<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using abc_prop some_props by auto then show "False" using f7 by auto qed have cf2: "(- c' / b') = (- e + 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d') \ False" proof - assume "(- c' / b') = (- e + 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d')" then have "(d', e, f) \ set b \ d' \ 0 \ - e\<^sup>2 + 4 * d' * f \ 0 \ ((\(d, e, f)\set a. \y'>?r2. \x\{?r2<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?r2. \x\{?r2<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?r2. \x\{?r2<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?r2. \x\{?r2<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using abc_prop some_props by auto then show "False" using f6 by auto qed then show "False" using two_cases cf1 cf2 by auto qed (* discriminant_nonnegative *) have eo: "(d' \ 0) \ (d' = 0 \ e \ 0) \ (d' = 0 \ e = 0 \ f = 0)" using def_prop by auto then show "False" using f1 f2 f3 by auto qed show ?thesis using c1 c2 by fastforce qed have d_prop: "(\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0)" using abc_prop by auto have h4: "(\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0)" proof - have "(\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ False" (* begin *) proof - assume "\ (d, e, f) \ set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0" then obtain d' e f where def_prop: "(d', e, f) \ set d \ d' * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0" by auto then have same: "(d' = 0 \ e \ 0) \ (-f/e = -c'/b')" proof - assume asm: "(d' = 0 \ e \ 0)" then have " e * (- c' / b') + f = 0" using def_prop by auto then show "-f/e = -c'/b'" using asm by auto qed let ?r = "-f/e" have "(d' = 0 \ e \ 0) \ ((d', e, f) \ set d \ ((\(d, e, f)\set a. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using same def_prop abc_prop by auto then have "(d' = 0 \ e \ 0) \ (\(a', b', c')\set d. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0))" by auto then have f1: "(d' = 0 \ e \ 0) \ False" using f10 by auto have f2: "(d' = 0 \ e = 0 \ f = 0) \ False" proof - assume "(d' = 0 \ e = 0 \ f = 0)" then have allzer: "\x. d'*x^2 + e*x + f = 0" by auto have "\y'>- c' / b'. \x\{- c' / b'<..y'}. d' * x\<^sup>2 + e * x + f \ 0" using d_prop def_prop by auto then obtain y' where y_prop: " y' >- c' / b' \ (\x\{- c' / b'<..y'}. d' * x\<^sup>2 + e * x + f \ 0)" by auto then have "\k. k > -c'/b' \ k < y' \ d' * k^2 + e * k + f \ 0" using dense by (meson dense greaterThanAtMost_iff less_eq_real_def) then show "False" using allzer by auto qed have f3: "d' \ 0 \ False" proof - assume dnonz: "d' \ 0" have discr: " - e\<^sup>2 + 4 * d' * f \ 0" using def_prop discriminant_negative[of d' e f] unfolding discrim_def by fastforce then have two_cases: "(- c' / b') = (- e + - 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d') \ (- c' / b') = (- e + 1 * sqrt (e\<^sup>2 - 4 * d' * f)) / (2 * d')" using def_prop dnonz discriminant_nonneg[of d' e f] unfolding discrim_def by fastforce have some_props: "((d', e, f) \ set d \ d' \ 0 \ - e\<^sup>2 + 4 * d' * f \ 0)" using dnonz def_prop discr by auto let ?r1 = "(- e + - 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d')" let ?r2 = "(- e + 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d')" have cf1: "(- c' / b') = (- e + - 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d') \ False" proof - assume "(- c' / b') = (- e + - 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d')" then have "(d', e, f) \ set d \ d' \ 0 \ - e\<^sup>2 + 4 * d' * f \ 0 \ ((\(d, e, f)\set a. \y'>?r1. \x\{?r1<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?r1. \x\{?r1<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?r1. \x\{?r1<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?r1. \x\{?r1<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using abc_prop some_props by auto then show "False" using f12 by auto qed have cf2: "(- c' / b') = (- e + 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d') \ False" proof - assume "(- c' / b') = (- e + 1 * sqrt (e^2 - 4 * d' * f)) / (2 * d')" then have "(d', e, f) \ set d \ d' \ 0 \ - e\<^sup>2 + 4 * d' * f \ 0 \ ((\(d, e, f)\set a. \y'>?r2. \x\{?r2<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?r2. \x\{?r2<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?r2. \x\{?r2<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?r2. \x\{?r2<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using abc_prop some_props by auto then show "False" using f11 by auto qed then show "False" using two_cases cf1 cf2 by auto qed (* discriminant_nonnegative *) have eo: "(d' \ 0) \ (d' = 0 \ e \ 0) \ (d' = 0 \ e = 0 \ f = 0)" using def_prop by auto then show "False" using f1 f2 f3 by auto qed then show ?thesis by auto qed have "(\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0))" using h1 h2 h3 h4 bnonz abc_prop by auto then show "False" using f8 by auto qed then show ?thesis by auto qed lemma qe_forwards_helper: assumes alleqset: "\x. (\(d, e, f)\set a. d * x^2 + e * x + f = 0)" assumes f1: "\((\(a, b, c)\set a. a = 0 \ b = 0 \ c = 0) \ (\(a, b, c)\set b. \x. \y2 + b * y + c < 0) \ (\(a, b, c)\set c. \x. \y2 + b * y + c \ 0) \ (\(a, b, c)\set d. \x. \y2 + b * y + c \ 0))" assumes f5: "\(\(a', b', c')\set b. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0))" assumes f6: "\ (\(a', b', c')\set b. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" assumes f7: "\ (\(a', b', c')\set b. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))" assumes f8: "\(\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0))" assumes f13: "\(\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)))" assumes f9: "\(\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))" assumes f10: "\(\(a', b', c')\set d. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0))" assumes f11: "\(\(a', b', c')\set d. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" assumes f12: "\(\(a', b', c')\set d. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))" shows "\(\x. (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0))" proof - have nor: "\r. \(\(a', b', c')\set c. ((a'\ 0 \ b'\ 0) \ a'*r^2 + b'*r + c' = 0) \ ((\(d, e, f)\set a. d * r\<^sup>2 + e * r + f = 0) \ (\(d, e, f)\set b. d * r^2 + e * r + f < 0) \ (\(d, e, f)\set c. d * r^2 + e * r + f \ 0) \ (\(d, e, f)\set d. d * r^2 + e * r + f \ 0)))" proof clarsimp fix r t u v assume inset: "(t, u, v) \ set c" assume eo: "t = 0 \ u \ 0 " assume zero_eq: "t*r^2 + u*r + v = 0" assume ah: "\x\set a. case x of (d, e, f) \ d * r\<^sup>2 + e * r + f = 0" assume bh: "\x\set b. case x of (d, e, f) \ d * r\<^sup>2 + e * r + f < 0" assume ch: "\x\set c. case x of (d, e, f) \ d * r\<^sup>2 + e * r + f \ 0" assume dh: "\x\set d. case x of (d, e, f) \ d * r\<^sup>2 + e * r + f \ 0" have two_cases: "t \ 0 \ (t = 0 \ u \ 0)" using eo by auto have c1: "t \ 0 \ False" proof - assume tnonz: "t \ 0" then have discr_prop: "- u\<^sup>2 + 4 * t * v \ 0 " using discriminant_negative[of t u v] zero_eq unfolding discrim_def by force then have ris: "r = ((-u + - 1 * sqrt (u^2 - 4 * t * v)) / (2 * t)) \ r = ((-u + 1 * sqrt (u^2 - 4 * t * v)) / (2 * t)) " using tnonz discriminant_nonneg[of t u v] zero_eq unfolding discrim_def by auto let ?r1 = "((-u + - 1 * sqrt (u^2 - 4 * t * v)) / (2 * t))" let ?r2 = "((-u + 1 * sqrt (u^2 - 4 * t * v)) / (2 * t))" have ris1: "r = ?r1 \ False" proof - assume "r = ?r1" then have "(\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))" using inset ah bh ch dh discr_prop tnonz by auto then show ?thesis using f9 by auto qed have ris2: "r = ?r2 \ False" proof - assume "r = ?r2" then have "(\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)))" using inset ah bh ch dh discr_prop tnonz by auto then show ?thesis using f13 by auto qed show "False" using ris ris1 ris2 by auto qed have c2: "(t = 0 \ u \ 0) \ False" proof - assume asm: "t = 0 \ u \ 0" then have "r = -v/u" using zero_eq add.right_neutral nonzero_mult_div_cancel_left by (metis add.commute divide_divide_eq_right divide_eq_0_iff neg_eq_iff_add_eq_0) then have "(\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0))" using asm inset ah bh ch dh by auto then show "False" using f8 by auto qed then show "False" using two_cases c1 c2 by auto qed have keyh: "\r. \(\(a', b', c')\set c. ((a'\ 0 \ b'\ 0) \ a'*r^2 + b'*r + c' = 0) \ (\(d, e, f)\set a. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0))" proof - fix r have h8: "\(\(a', b', c')\set c. ((a'\ 0 \ b'\ 0) \ a'*r^2 + b'*r + c' = 0) \ ((\(d, e, f)\set a. d * r\<^sup>2 + e * r + f = 0) \ (\(d, e, f)\set b. d * r^2 + e * r + f < 0) \ (\(d, e, f)\set c. d * r^2 + e * r + f \ 0) \ (\(d, e, f)\set d. d * r^2 + e * r + f \ 0)))" using nor by auto show "\(\(a', b', c')\set c. ((a'\ 0 \ b'\ 0) \ a'*r^2 + b'*r + c' = 0) \ (\(d, e, f)\set a. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>r. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using qe_forwards_helper_gen[of c r a b d] alleqset f5 f6 f7 h8 f10 f11 f12 by auto qed have f8a: "\(\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using qe_forwards_helper_lin[of a b c d] alleqset f5 f6 f7 f8 f10 f11 f12 by blast have f13a: "\ (\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" proof - have "(\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))) \ False" proof - assume "(\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" then obtain a' b' c' where abc_prop: "(a', b', c')\set c \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))" by auto let ?r = "(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')" have somek: "\k. k = ?r" by auto then obtain k where k_prop: "k = ?r" by auto have "(a'\ 0 \ b'\ 0) \ (a'*?r^2 + b'*?r + c' = 0)" using abc_prop discriminant_nonneg[of a' b' c'] unfolding discrim_def apply (auto) by (metis (mono_tags, lifting) times_divide_eq_right) then have "(\(a', b', c')\set c. ((a'\ 0 \ b'\ 0) \ a'*?r^2 + b'*?r + c' = 0) \ (\(d, e, f)\set a. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using abc_prop by auto then have "(\(a', b', c')\set c. ((a'\ 0 \ b'\ 0) \ a'*k^2 + b'*k + c' = 0) \ (\(d, e, f)\set a. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using k_prop by auto then have "\k. (\(a', b', c')\set c. ((a'\ 0 \ b'\ 0) \ a'*k^2 + b'*k + c' = 0) \ (\(d, e, f)\set a. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f \ 0))" by auto then show "False" using keyh by auto qed then show ?thesis by auto qed have f9a: "\ (\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))" proof - have "(\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)) \ False" proof - assume "(\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))" then obtain a' b' c' where abc_prop: "(a', b', c')\set c \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)" by auto let ?r = "(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')" have somek: "\k. k = ?r" by auto then obtain k where k_prop: "k = ?r" by auto have "(a'\ 0 \ b'\ 0) \ (a'*?r^2 + b'*?r + c' = 0)" using abc_prop discriminant_nonneg[of a' b' c'] unfolding discrim_def apply (auto) by (metis (mono_tags, lifting) times_divide_eq_right) then have "(\(a', b', c')\set c. ((a'\ 0 \ b'\ 0) \ a'*?r^2 + b'*?r + c' = 0) \ (\(d, e, f)\set a. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?r. \x\{?r<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using abc_prop by auto then have "(\(a', b', c')\set c. ((a'\ 0 \ b'\ 0) \ a'*k^2 + b'*k + c' = 0) \ (\(d, e, f)\set a. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using k_prop by auto then have "\k. (\(a', b', c')\set c. ((a'\ 0 \ b'\ 0) \ a'*k^2 + b'*k + c' = 0) \ (\(d, e, f)\set a. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f \ 0))" by auto then show "False" using keyh by auto qed then show ?thesis by auto qed (* We need to show that the point is in one of these ranges *) have "(\x. (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0)) \ False" proof - assume "(\x. (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0))" then obtain x where x_prop: "(\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0)" by auto (* Need this sorted_nonzero_root_list_set in case some of the tuples from set c are (0, 0, 0) *) let ?srl = "sorted_nonzero_root_list_set (((set b) \ (set c))\ (set d))" have alleqsetvar: "\(t, u, v) \ set a. t = 0 \ u = 0 \ v = 0" proof clarsimp fix t u v assume "(t, u, v) \ set a" then have "\x. t*x^2 + u*x + v = 0" using alleqset by auto then have "\x\{0<..1}. t * x\<^sup>2 + u * x + v = 0" by auto then show "t = 0 \ u = 0 \ v = 0" using continuity_lem_eq0[of 0 1 t u v] by auto qed (* Should violate f1 *) have lenzero: "length ?srl = 0 \ False" proof - assume lenzero: "length ?srl = 0" have ina: "(\(a, b, c)\set a. a = 0 \ b = 0 \ c = 0)" using alleqsetvar by auto have inb: "(\(a, b, c)\set b. \y. a * y\<^sup>2 + b * y + c < 0)" proof clarsimp fix t u v y assume insetb: "(t, u, v) \ set b" then have "t * x\<^sup>2 + u * x + v < 0" using x_prop by auto then have tuv_prop: "t \ 0 \ u \ 0 \ v \ 0" by auto then have tuzer: "(t = 0 \ u = 0) \ \(\q. t * q\<^sup>2 + u * q + v = 0)" by simp then have tunonz: "(t \ 0 \ u \ 0) \ \(\q. t * q\<^sup>2 + u * q + v = 0)" proof - assume tuv_asm: "t \ 0 \ u \ 0" have "\q. t * q\<^sup>2 + u * q + v = 0 \ False" proof - assume "\ q. t * q\<^sup>2 + u * q + v = 0" then obtain q where "t * q\<^sup>2 + u * q + v = 0" by auto then have qin: "q \ {x. \(a, b, c)\set b \ set c \ set d. (a \ 0 \ b \ 0) \ a * x\<^sup>2 + b * x + c = 0}" using insetb tuv_asm tuv_prop by auto have "set ?srl = nonzero_root_set (set b \ set c \ set d)" unfolding sorted_nonzero_root_list_set_def using set_sorted_list_of_set[of "nonzero_root_set (set b \ set c \ set d)"] nonzero_root_set_finite[of "(set b \ set c \ set d)"] by auto then have "q \ set ?srl" using qin unfolding nonzero_root_set_def by auto then have "List.member ?srl q" using in_set_member[of q ?srl] by auto then show "False" using lenzero by (simp add: member_rec(2)) qed then show ?thesis by auto qed have nozer: "\(\q. t * q\<^sup>2 + u * q + v = 0)" using tuzer tunonz by blast have samesn: "sign_num (t*x^2 + u*x + v) = sign_num (t*y^2 + u*y + v)" proof - have "x < y \ sign_num (t*x^2 + u*x + v) = sign_num (t*y^2 + u*y + v)" using changes_sign_var[of t x u v y] nozer by auto have "y < x \ sign_num (t*x^2 + u*x + v) = sign_num (t*y^2 + u*y + v)" using changes_sign_var[of t y u v x] nozer proof - assume "y < x" then show ?thesis using \\q. t * q\<^sup>2 + u * q + v = 0\ \sign_num (t * y\<^sup>2 + u * y + v) \ sign_num (t * x\<^sup>2 + u * x + v) \ y < x \ \q. t * q\<^sup>2 + u * q + v = 0 \ y \ q \ q \ x\ by presburger qed show ?thesis using changes_sign_var using \x < y \ sign_num (t * x\<^sup>2 + u * x + v) = sign_num (t * y\<^sup>2 + u * y + v)\ \y < x \ sign_num (t * x\<^sup>2 + u * x + v) = sign_num (t * y\<^sup>2 + u * y + v)\ by fastforce qed (* changes_sign_var *) have "sign_num (t*x^2 + u*x + v) = -1" using insetb unfolding sign_num_def using x_prop by auto then have "sign_num (t*y^2 + u*y + v) = -1" using samesn by auto then show "t * y\<^sup>2 + u * y + v < 0" unfolding sign_num_def by smt qed have inc: "(\(a, b, c)\set c. \y. a * y\<^sup>2 + b * y + c \ 0)" proof clarsimp fix t u v y assume insetc: "(t, u, v) \ set c" then have "t * x\<^sup>2 + u * x + v \ 0" using x_prop by auto then have tuzer: "t = 0 \ u = 0 \ t*y^2 + u*y + v \ 0 " proof - assume tandu: "t = 0 \ u = 0" then have "v \ 0" using insetc x_prop by auto then show "t*y^2 + u*y + v \ 0" using tandu by auto qed have tunonz: "t \ 0 \ u \ 0 \ t*y^2 + u*y + v \ 0" proof - assume tuv_asm: "t \ 0 \ u \ 0" have insetcvar: "t*x^2 + u*x + v < 0" proof - have "t*x^2 + u*x + v = 0 \ False" proof - assume zer: "t*x^2 + u*x + v = 0" then have xin: "x \ {x. \(a, b, c)\set b \ set c \ set d. (a \ 0 \ b \ 0) \ a * x\<^sup>2 + b * x + c = 0}" using insetc tuv_asm by auto have "set ?srl = nonzero_root_set (set b \ set c \ set d)" unfolding sorted_nonzero_root_list_set_def using set_sorted_list_of_set[of "nonzero_root_set (set b \ set c \ set d)"] nonzero_root_set_finite[of "(set b \ set c \ set d)"] by auto then have "x \ set ?srl" using xin unfolding nonzero_root_set_def by auto then have "List.member ?srl x" using in_set_member[of x ?srl] by auto then show "False" using lenzero by (simp add: member_rec(2)) qed then show ?thesis using \t * x\<^sup>2 + u * x + v \ 0\ by fastforce qed then have tunonz: "\(\q. t * q\<^sup>2 + u * q + v = 0)" proof - have "\q. t * q\<^sup>2 + u * q + v = 0 \ False" proof - assume "\ q. t * q\<^sup>2 + u * q + v = 0" then obtain q where "t * q\<^sup>2 + u * q + v = 0" by auto then have qin: "q \ {x. \(a, b, c)\set b \ set c \ set d. (a \ 0 \ b \ 0) \ a * x\<^sup>2 + b * x + c = 0}" using insetc tuv_asm by auto have "set ?srl = nonzero_root_set (set b \ set c \ set d)" unfolding sorted_nonzero_root_list_set_def using set_sorted_list_of_set[of "nonzero_root_set (set b \ set c \ set d)"] nonzero_root_set_finite[of "(set b \ set c \ set d)"] by auto then have "q \ set ?srl" using qin unfolding nonzero_root_set_def by auto then have "List.member ?srl q" using in_set_member[of q ?srl] by auto then show "False" using lenzero by (simp add: member_rec(2)) qed then show ?thesis by auto qed have nozer: "\(\q. t * q\<^sup>2 + u * q + v = 0)" using tuzer tunonz by blast have samesn: "sign_num (t*x^2 + u*x + v) = sign_num (t*y^2 + u*y + v)" proof - have "x < y \ sign_num (t*x^2 + u*x + v) = sign_num (t*y^2 + u*y + v)" using changes_sign_var[of t x u v y] nozer by auto have "y < x \ sign_num (t*x^2 + u*x + v) = sign_num (t*y^2 + u*y + v)" using changes_sign_var[of t y u v x] nozer proof - assume "y < x" then show ?thesis using \\q. t * q\<^sup>2 + u * q + v = 0\ \sign_num (t * y\<^sup>2 + u * y + v) \ sign_num (t * x\<^sup>2 + u * x + v) \ y < x \ \q. t * q\<^sup>2 + u * q + v = 0 \ y \ q \ q \ x\ by presburger qed show ?thesis using changes_sign_var using \x < y \ sign_num (t * x\<^sup>2 + u * x + v) = sign_num (t * y\<^sup>2 + u * y + v)\ \y < x \ sign_num (t * x\<^sup>2 + u * x + v) = sign_num (t * y\<^sup>2 + u * y + v)\ by fastforce qed (* changes_sign_var *) have "sign_num (t*x^2 + u*x + v) = -1" using insetcvar unfolding sign_num_def using x_prop by auto then have "sign_num (t*y^2 + u*y + v) = -1" using samesn by auto then show "t * y\<^sup>2 + u * y + v \ 0" unfolding sign_num_def by smt qed then show "t * y\<^sup>2 + u * y + v \ 0" using tuzer tunonz by blast qed have ind: "(\(a, b, c)\set d. \y. a * y\<^sup>2 + b * y + c \ 0)" proof clarsimp fix t u v y assume insetd: "(t, u, v) \ set d" assume falseasm: "t * y\<^sup>2 + u * y + v = 0" then have snz: "sign_num (t*y^2 + u*y + v) = 0" unfolding sign_num_def by auto have "t * x\<^sup>2 + u * x + v \ 0" using insetd x_prop by auto then have tuv_prop: "t \ 0 \ u \ 0 \ v \ 0" by auto then have tuzer: "(t = 0 \ u = 0) \ \(\q. t * q\<^sup>2 + u * q + v = 0)" by simp then have tunonz: "(t \ 0 \ u \ 0) \ \(\q. t * q\<^sup>2 + u * q + v = 0)" proof - assume tuv_asm: "t \ 0 \ u \ 0" have "\q. t * q\<^sup>2 + u * q + v = 0 \ False" proof - assume "\ q. t * q\<^sup>2 + u * q + v = 0" then obtain q where "t * q\<^sup>2 + u * q + v = 0" by auto then have qin: "q \ {x. \(a, b, c)\set b \ set c \ set d. (a \ 0 \ b \ 0) \ a * x\<^sup>2 + b * x + c = 0}" using insetd tuv_asm tuv_prop by auto have "set ?srl = nonzero_root_set (set b \ set c \ set d)" unfolding sorted_nonzero_root_list_set_def using set_sorted_list_of_set[of "nonzero_root_set (set b \ set c \ set d)"] nonzero_root_set_finite[of "(set b \ set c \ set d)"] by auto then have "q \ set ?srl" using qin unfolding nonzero_root_set_def by auto then have "List.member ?srl q" using in_set_member[of q ?srl] by auto then show "False" using lenzero by (simp add: member_rec(2)) qed then show ?thesis by auto qed have nozer: "\(\q. t * q\<^sup>2 + u * q + v = 0)" using tuzer tunonz by blast have samesn: "sign_num (t*x^2 + u*x + v) = sign_num (t*y^2 + u*y + v)" proof - have "x < y \ sign_num (t*x^2 + u*x + v) = sign_num (t*y^2 + u*y + v)" using changes_sign_var[of t x u v y] nozer by auto have "y < x \ sign_num (t*x^2 + u*x + v) = sign_num (t*y^2 + u*y + v)" using changes_sign_var[of t y u v x] nozer proof - assume "y < x" then show ?thesis using \\q. t * q\<^sup>2 + u * q + v = 0\ \sign_num (t * y\<^sup>2 + u * y + v) \ sign_num (t * x\<^sup>2 + u * x + v) \ y < x \ \q. t * q\<^sup>2 + u * q + v = 0 \ y \ q \ q \ x\ by presburger qed show ?thesis using changes_sign_var using \x < y \ sign_num (t * x\<^sup>2 + u * x + v) = sign_num (t * y\<^sup>2 + u * y + v)\ \y < x \ sign_num (t * x\<^sup>2 + u * x + v) = sign_num (t * y\<^sup>2 + u * y + v)\ by fastforce qed (* changes_sign_var *) have "sign_num (t*x^2 + u*x + v) = -1 \ sign_num (t*x^2 + u*x + v) = 1 " using insetd unfolding sign_num_def using x_prop by auto then have "sign_num (t*y^2 + u*y + v) = -1 \ sign_num (t*y^2 + u*y + v) = 1" using samesn by auto then show "False" using snz by auto qed (* Show all the polynomials never change sign *) have "((\(a, b, c)\set a. a = 0 \ b = 0 \ c = 0) \ (\(a, b, c)\set b. \y. a * y\<^sup>2 + b * y + c < 0) \ (\(a, b, c)\set c. \y. a * y\<^sup>2 + b * y + c \ 0) \ (\(a, b, c)\set d. \y. a * y\<^sup>2 + b * y + c \ 0))" using ina inb inc ind by auto then show "False" using f1 by auto qed have cases_mem: "(List.member ?srl x) \ False" proof - assume "(List.member ?srl x)" then have "x \ {x. \(a, b, c)\set b \ set c \ set d. (a \ 0 \ b \ 0) \ a * x\<^sup>2 + b * x + c = 0}" using set_sorted_list_of_set nonzero_root_set_finite in_set_member by (metis List.finite_set finite_Un nonzero_root_set_def sorted_nonzero_root_list_set_def) then have "\ (a, b, c) \ (((set b) \ (set c))\ (set d)) . (a \ 0 \ b \ 0) \ a*x^2 + b*x + c = 0" by blast then obtain t u v where def_prop: "(t, u, v) \ (((set b) \ (set c))\ (set d)) \ (t \ 0 \ u \ 0) \ t*x^2 + u*x + v = 0" by auto have notinb: "(t, u, v) \ (set b)" proof - have "(t, u, v) \ (set b ) \ False" proof - assume "(t, u, v) \ (set b)" then have "t*x^2 + u*x + v < 0" using x_prop by blast then show "False" using def_prop by simp qed then show ?thesis by auto qed have notind: "(t, u, v) \ (set d)" proof - have "(t, u, v) \ (set d) \ False" proof - assume "(t, u, v) \ (set d)" then have "t*x^2 + u*x + v \ 0" using x_prop by blast then show "False" using def_prop by simp qed then show ?thesis by auto qed then have inset: "(t, u, v) \ (set c)" using def_prop notinb notind by blast have case1: "t \ 0 \ False" proof - assume tnonz: "t \ 0" then have r1or2:"x = (- u + - 1 * sqrt (u\<^sup>2 - 4 * t * v)) / (2 * t) \ x = (- u + 1 * sqrt (u^2 - 4 * t * v)) / (2 * t) " using def_prop discriminant_negative[of t u v] discriminant_nonneg[of t u v] apply (auto) using notinb apply (force) apply (simp add: discrim_def discriminant_iff) using notind by force have discrh: "-1*u^2 + 4 * t * v \ 0" using tnonz discriminant_negative[of t u v] unfolding discrim_def using def_prop by force have r1: "x = (- u + - 1 * sqrt (u\<^sup>2 - 4 * t * v)) / (2 * t) \ False" proof - assume xis: "x = (- u + - 1 * sqrt (u\<^sup>2 - 4 * t * v)) / (2 * t)" have " t \ 0 \ - 1*u^2 + 4 * t * v \ 0 \ (\(d, e, f)\set a. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. d * x\<^sup>2 + e * x + f \ 0)" using tnonz alleqset discrh x_prop by auto then have "(\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))" using xis inset by auto then show "False" using f9 by auto qed have r2: "x = (- u + 1 * sqrt (u\<^sup>2 - 4 * t * v)) / (2 * t) \ False" proof - assume xis: "x = (- u + 1 * sqrt (u\<^sup>2 - 4 * t * v)) / (2 * t)" have " t \ 0 \ - 1*u^2 + 4 * t * v \ 0 \ (\(d, e, f)\set a. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. d * x\<^sup>2 + e * x + f \ 0)" using tnonz alleqset discrh x_prop by auto then have "(\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))" using xis inset by auto then show "False" using f13 by auto qed then show "False" using r1or2 r1 r2 by auto qed have case2: "(t = 0 \ u \ 0) \ False" proof - assume asm: "t = 0 \ u \ 0" then have xis: "x = - v / u" using def_prop notinb add.commute diff_0 divide_non_zero minus_add_cancel uminus_add_conv_diff by (metis mult_zero_left) have "((t = 0 \ u \ 0) \ (\(d, e, f)\set a. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. d * x^2 + e * x + f < 0) \ (\(d, e, f)\set c. d * x^2 + e * x + f \ 0) \ (\(d, e, f)\set d. d * x^2 + e * x + f \ 0))" using asm x_prop alleqset by auto then have "(\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0))" using xis inset by auto then show "False" using f8 by auto qed show "False" using def_prop case1 case2 by auto qed have lengt0: "length ?srl \ 1 \ False" proof- assume asm: "length ?srl \ 1" (* should violate f1 *) have cases_lt: "x < ?srl ! 0 \ False" proof - assume xlt: "x < ?srl ! 0" have samesign: "\ (a, b, c) \ (set b \ set c \ set d). (\y < x. sign_num (a * y\<^sup>2 + b * y + c) = sign_num (a*x^2 + b*x + c))" proof clarsimp fix t u v y assume insetunion: "(t, u, v) \ set b \ (t, u, v) \ set c \ (t, u, v) \ set d" assume ylt: "y < x" have tuzer: "t = 0 \ u = 0 \ sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" unfolding sign_num_def by auto have tunonzer: "t \ 0 \ u \ 0 \ sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" proof - assume tuv_asm: "t\ 0 \ u \ 0" have "\(\q. q < ?srl ! 0 \ t * q\<^sup>2 + u * q + v = 0)" proof clarsimp fix q assume qlt: "q < sorted_nonzero_root_list_set (set b \ set c \ set d) ! 0" assume "t * q\<^sup>2 + u * q + v = 0" then have qin: "q \ {x. \(a, b, c)\set b \ set c \ set d. (a \ 0 \ b \ 0) \ a * x\<^sup>2 + b * x + c = 0}" using insetunion tuv_asm by auto have "set ?srl = nonzero_root_set (set b \ set c \ set d)" unfolding sorted_nonzero_root_list_set_def using set_sorted_list_of_set[of "nonzero_root_set (set b \ set c \ set d)"] nonzero_root_set_finite[of "(set b \ set c \ set d)"] by auto then have "q \ set ?srl" using qin unfolding nonzero_root_set_def by auto then have lm: "List.member ?srl q" using in_set_member[of q ?srl] by auto then have " List.member (sorted_list_of_set (nonzero_root_set (set b \ set c \ set d))) q \ q < sorted_list_of_set (nonzero_root_set (set b \ set c \ set d)) ! 0 \ (\x xs. (x \ set xs) = (\i (\x xs. (x \ set xs) = List.member xs x) \ (\y x. \ y \ x \ x < y) \ (\xs. sorted xs = (\i j. i \ j \ j < length xs \ xs ! i \ xs ! j)) \ (\p. sorted_nonzero_root_list_set p \ sorted_list_of_set (nonzero_root_set p)) \ False" proof - assume a1: "List.member (sorted_list_of_set (nonzero_root_set (set b \ set c \ set d))) q" assume a2: "q < sorted_list_of_set (nonzero_root_set (set b \ set c \ set d)) ! 0" have f3: "List.member (sorted_list_of_set {r. \p. p \ set b \ set c \ set d \ (case p of (ra, rb, rc) \ (ra \ 0 \ rb \ 0) \ ra * r\<^sup>2 + rb * r + rc = 0)}) q" using a1 by (metis nonzero_root_set_def) have f4: "q < sorted_list_of_set {r. \p. p \ set b \ set c \ set d \ (case p of (ra, rb, rc) \ (ra \ 0 \ rb \ 0) \ ra * r\<^sup>2 + rb * r + rc = 0)} ! 0" using a2 by (metis nonzero_root_set_def) have f5: "q \ set (sorted_list_of_set {r. \p. p \ set b \ set c \ set d \ (case p of (ra, rb, rc) \ (ra \ 0 \ rb \ 0) \ ra * r\<^sup>2 + rb * r + rc = 0)})" using f3 by (meson in_set_member) have "\rs r. \n. ((r::real) \ set rs \ n < length rs) \ (r \ set rs \ rs ! n = r)" by (metis in_set_conv_nth) then obtain nn :: "real list \ real \ nat" where f6: "\r rs. (r \ set rs \ nn rs r < length rs) \ (r \ set rs \ rs ! nn rs r = r)" by moura then have "sorted_list_of_set {r. \p. p \ set b \ set c \ set d \ (case p of (ra, rb, rc) \ (ra \ 0 \ rb \ 0) \ ra * r\<^sup>2 + rb * r + rc = 0)} ! nn (sorted_list_of_set {r. \p. p \ set b \ set c \ set d \ (case p of (ra, rb, rc) \ (ra \ 0 \ rb \ 0) \ ra * r\<^sup>2 + rb * r + rc = 0)}) q = q" using f5 by blast then have "\n. \ sorted (sorted_list_of_set {r. \p. p \ set b \ set c \ set d \ (case p of (ra, rb, rc) \ (ra \ 0 \ rb \ 0) \ ra * r\<^sup>2 + rb * r + rc = 0)}) \ \ n \ nn (sorted_list_of_set {r. \p. p \ set b \ set c \ set d \ (case p of (ra, rb, rc) \ (ra \ 0 \ rb \ 0) \ ra * r\<^sup>2 + rb * r + rc = 0)}) q \ \ nn (sorted_list_of_set {r. \p. p \ set b \ set c \ set d \ (case p of (ra, rb, rc) \ (ra \ 0 \ rb \ 0) \ ra * r\<^sup>2 + rb * r + rc = 0)}) q < length (sorted_list_of_set {r. \p. p \ set b \ set c \ set d \ (case p of (ra, rb, rc) \ (ra \ 0 \ rb \ 0) \ ra * r\<^sup>2 + rb * r + rc = 0)}) \ sorted_list_of_set {r. \p. p \ set b \ set c \ set d \ (case p of (ra, rb, rc) \ (ra \ 0 \ rb \ 0) \ ra * r\<^sup>2 + rb * r + rc = 0)} ! n \ q" - using not_less not_less0 sorted_iff_nth_mono - by (metis (no_types, lifting)) + using not_less0 sorted_iff_nth_mono by (metis (lifting)) then show ?thesis using f6 f5 f4 by (meson le0 not_less sorted_sorted_list_of_set) qed then show "False" using lm qlt in_set_conv_nth in_set_member not_le_imp_less not_less0 sorted_iff_nth_mono sorted_nonzero_root_list_set_def sorted_sorted_list_of_set by auto qed then have "\(\q. q \ x \ t * q\<^sup>2 + u * q + v = 0)" using xlt by auto then show " sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" using ylt changes_sign_var[of t y u v x] by blast qed then show " sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" using tuzer by blast qed have bseth: "(\(a, b, c)\set b. \y2 + b * y + c < 0)" proof clarsimp fix t u v y assume insetb: "(t, u, v) \ set b" assume yltx: "y < x" have "(t, u, v) \ (set b \ set c \ set d)" using insetb by auto then have samesn: "sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" using samesign insetb yltx by blast have "sign_num (t*x^2 + u*x + v) = -1" using x_prop insetb unfolding sign_num_def by auto then show "t * y\<^sup>2 + u * y + v < 0" using samesn unfolding sign_num_def by (metis add.right_inverse add.right_neutral linorder_neqE_linordered_idom one_add_one zero_neq_numeral) qed have bset: " (\(a, b, c)\set b. \x. \y2 + b * y + c < 0)" proof clarsimp fix t u v assume inset: "(t, u, v) \ set b" then have " \y2 + u * y + v < 0 " using bseth by auto then show "\x. \y2 + u * y + v < 0" by auto qed have cseth: "(\(a, b, c)\set c. \y2 + b * y + c \ 0)" proof clarsimp fix t u v y assume insetc: "(t, u, v) \ set c" assume yltx: "y < x" have "(t, u, v) \ (set b \ set c \ set d)" using insetc by auto then have samesn: "sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" using samesign insetc yltx by blast have "sign_num (t*x^2 + u*x + v) = -1 \ sign_num (t*x^2 + u*x + v) = 0" using x_prop insetc unfolding sign_num_def by auto then show "t * y\<^sup>2 + u * y + v \ 0" using samesn unfolding sign_num_def using zero_neq_one by fastforce qed have cset: " (\(a, b, c)\set c. \x. \y2 + b * y + c \ 0)" proof clarsimp fix t u v assume inset: "(t, u, v) \ set c" then have " \y2 + u * y + v \ 0 " using cseth by auto then show "\x. \y2 + u * y + v \0" by auto qed have dseth: "(\(a, b, c)\set d. \y2 + b * y + c \ 0)" proof clarsimp fix t u v y assume insetd: "(t, u, v) \ set d" assume yltx: "y < x" assume contrad: "t * y\<^sup>2 + u * y + v = 0" have "(t, u, v) \ (set b \ set c \ set d)" using insetd by auto then have samesn: "sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" using samesign insetd yltx by blast have "sign_num (t*x^2 + u*x + v) = -1 \ sign_num (t*x^2 + u*x + v) = 1" using x_prop insetd unfolding sign_num_def by auto then have "t * y\<^sup>2 + u * y + v \ 0" using samesn unfolding sign_num_def by auto then show "False" using contrad by auto qed have dset: " (\(a, b, c)\set d. \x. \y2 + b * y + c \ 0)" proof clarsimp fix t u v assume inset: "(t, u, v) \ set d" then have " \y2 + u * y + v \ 0 " using dseth by auto then show "\x. \y2 + u * y + v \ 0" by auto qed have "(\(a, b, c)\set a. a = 0 \ b = 0 \ c = 0)" using alleqsetvar by auto then have "((\(a, b, c)\set a. a = 0 \ b = 0 \ c = 0) \ (\(a, b, c)\set b. \x. \y2 + b * y + c < 0) \ (\(a, b, c)\set c. \x. \y2 + b * y + c \ 0) \ (\(a, b, c)\set d. \x. \y2 + b * y + c \ 0))" using bset cset dset by auto then show "False" using f1 by auto qed (* should violate one of the infinitesmials *) have cases_gt: " x > ?srl ! (length ?srl - 1) \ False" proof - assume xgt: "x > ?srl ! (length ?srl - 1)" let ?bgrt = "?srl ! (length ?srl - 1)" have samesign: "\ (a, b, c) \ (set b \ set c \ set d). (\y > ?bgrt. sign_num (a * y\<^sup>2 + b * y + c) = sign_num (a*x^2 + b*x + c))" proof clarsimp fix t u v y assume insetunion: "(t, u, v) \ set b \ (t, u, v) \ set c \ (t, u, v) \ set d" assume ygt: "sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - Suc 0) < y" have tuzer: "t = 0 \ u = 0 \ sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" unfolding sign_num_def by auto have tunonzer: "t \ 0 \ u \ 0 \ sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" proof - assume tuv_asm: "t\ 0 \ u \ 0" have "\(\q. q > ?srl ! (length ?srl - 1) \ t * q\<^sup>2 + u * q + v = 0)" proof clarsimp fix q assume qgt: "sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - Suc 0) < q" assume "t * q\<^sup>2 + u * q + v = 0" then have qin: "q \ {x. \(a, b, c)\set b \ set c \ set d. (a \ 0 \ b \ 0) \ a * x\<^sup>2 + b * x + c = 0}" using insetunion tuv_asm by auto have "set ?srl = nonzero_root_set (set b \ set c \ set d)" unfolding sorted_nonzero_root_list_set_def using set_sorted_list_of_set[of "nonzero_root_set (set b \ set c \ set d)"] nonzero_root_set_finite[of "(set b \ set c \ set d)"] by auto then have "q \ set ?srl" using qin unfolding nonzero_root_set_def by auto then have "List.member ?srl q" using in_set_member[of q ?srl] by auto then show "False" using qgt in_set_conv_nth in_set_member not_le_imp_less not_less0 sorted_iff_nth_mono sorted_nonzero_root_list_set_def sorted_sorted_list_of_set by (smt (z3) Suc_diff_Suc Suc_n_not_le_n \q \ set (sorted_nonzero_root_list_set (set b \ set c \ set d))\ in_set_conv_nth length_0_conv length_greater_0_conv length_sorted_list_of_set lenzero less_Suc_eq_le minus_nat.diff_0 not_le sorted_nth_mono sorted_sorted_list_of_set) qed then have nor: "\(\q. q > ?bgrt \ t * q\<^sup>2 + u * q + v = 0)" using xgt by auto have c1: " x > y \ sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" using nor changes_sign_var[of t y u v x] xgt ygt by fastforce then have c2: " y > x \ sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" using nor changes_sign_var[of t x u v y] xgt ygt by force then have c3: " x = y \ sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" unfolding sign_num_def by auto then show "sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" using c1 c2 c3 by linarith qed then show " sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" using tuzer by blast qed have "(\(a, b, c)\set a. a = 0 \ b = 0 \ c = 0)" using alleqsetvar by auto have " ?bgrt \ set ?srl" using set_sorted_list_of_set nonzero_root_set_finite in_set_member using asm by auto then have "?bgrt \ nonzero_root_set (set b \ set c \ set d )" unfolding sorted_nonzero_root_list_set_def using set_sorted_list_of_set nonzero_root_set_finite by auto then have "\t u v. (t, u, v) \ set b \ set c \ set d \(t \ 0 \ u \ 0) \ (t * ?bgrt\<^sup>2 + u * ?bgrt + v = 0)" unfolding nonzero_root_set_def by auto then obtain t u v where tuvprop1: "(t, u, v) \ set b \ set c \ set d \(t \ 0 \ u \ 0) \ (t * ?bgrt\<^sup>2 + u * ?bgrt + v = 0)" by auto then have tuvprop: "((t, u, v) \ set b \ (t \ 0 \ u \ 0) \ (t * ?bgrt\<^sup>2 + u * ?bgrt + v = 0)) \ ((t, u, v) \ set c \ (t \ 0 \ u \ 0) \ (t * ?bgrt\<^sup>2 + u * ?bgrt + v = 0)) \ ((t, u, v) \ set d \ (t \ 0 \ u \ 0) \ (t * ?bgrt\<^sup>2 + u * ?bgrt + v = 0)) " by auto have tnonz: "t\ 0 \ (-1*u^2 + 4 * t * v \ 0 \ (?bgrt = (- u + 1 * sqrt (u^2 - 4 * t * v)) / (2 * t) \ ?bgrt = (- u + -1 * sqrt (u^2 - 4 * t * v)) / (2 * t)))" proof - assume "t\ 0" have "-1*u^2 + 4 * t * v \ 0 " using tuvprop1 discriminant_negative[of t u v] unfolding discrim_def using \t \ 0\ by force then show ?thesis using tuvprop discriminant_nonneg[of t u v] unfolding discrim_def using \t \ 0\ by auto qed have unonz: "(t = 0 \ u \ 0) \ ?bgrt = - v / u" proof - assume "(t = 0 \ u \ 0)" then have "u*?bgrt + v = 0" using tuvprop1 by simp then show "?bgrt = - v / u" by (simp add: \t = 0 \ u \ 0\ eq_minus_divide_eq mult.commute) qed have allpropb: "(\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0)" proof clarsimp fix t1 u1 v1 y1 x1 assume ins: "(t1, u1, v1) \ set b" assume x1gt: " sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - Suc 0) < x1" assume "x1 \ y1" have xsn: "sign_num (t1 * x^2 + u1 * x + v1 ) = -1" using ins x_prop unfolding sign_num_def by auto have "sign_num (t1 * x1\<^sup>2 + u1 * x1 + v1 ) = sign_num (t1 * x^2 + u1 * x + v1 ) " using ins x1gt samesign apply (auto) by blast then show "t1 * x1\<^sup>2 + u1 * x1 + v1 < 0" using xsn unfolding sign_num_def by (metis add.right_inverse add.right_neutral linorder_neqE_linordered_idom one_add_one zero_neq_numeral) qed have allpropbvar: "(\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0)" proof clarsimp fix t1 u1 v1 assume "(t1, u1, v1) \ set b" then have "\x\{?bgrt<..(?bgrt + 1)}. t1 * x\<^sup>2 + u1 * x + v1 < 0" using allpropb by force then show "\y'>sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - Suc 0). \x\{sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - Suc 0)<..y'}. t1 * x\<^sup>2 + u1 * x + v1 < 0" using less_add_one by (metis One_nat_def) qed have allpropc: "(\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)" proof clarsimp fix t1 u1 v1 y1 x1 assume ins: "(t1, u1, v1) \ set c" assume x1gt: " sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - Suc 0) < x1" assume "x1 \ y1" have xsn: "sign_num (t1 * x^2 + u1 * x + v1 ) = -1 \ sign_num (t1 * x^2 + u1 * x + v1 ) = 0" using ins x_prop unfolding sign_num_def by auto have "sign_num (t1 * x1\<^sup>2 + u1 * x1 + v1 ) = sign_num (t1 * x^2 + u1 * x + v1 ) " using ins x1gt samesign One_nat_def proof - have "case (t1, u1, v1) of (r, ra, rb) \ \raa>sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - 1). sign_num (r * raa\<^sup>2 + ra * raa + rb) = sign_num (r * x\<^sup>2 + ra * x + rb)" by (smt (z3) Un_iff ins samesign) then show ?thesis by (simp add: x1gt) qed then show "t1 * x1\<^sup>2 + u1 * x1 + v1 \ 0" using xsn unfolding sign_num_def by (metis equal_neg_zero less_numeral_extra(3) linorder_not_less zero_neq_one) qed have allpropcvar: "(\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)" proof clarsimp fix t1 u1 v1 assume "(t1, u1, v1) \ set c" then have "\x\{?bgrt<..(?bgrt + 1)}. t1 * x\<^sup>2 + u1 * x + v1 \ 0" using allpropc by force then show "\y'>sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - Suc 0). \x\{sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - Suc 0)<..y'}. t1 * x\<^sup>2 + u1 * x + v1 \ 0" using less_add_one One_nat_def by metis qed have allpropd: "(\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)" proof clarsimp fix t1 u1 v1 y1 x1 assume ins: "(t1, u1, v1) \ set d" assume contrad:"t1 * x1\<^sup>2 + u1 * x1 + v1 = 0" assume x1gt: " sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - Suc 0) < x1" assume "x1 \ y1" have xsn: "sign_num (t1 * x^2 + u1 * x + v1 ) = -1 \ sign_num (t1 * x^2 + u1 * x + v1 ) = 1" using ins x_prop unfolding sign_num_def by auto have "sign_num (t1 * x1\<^sup>2 + u1 * x1 + v1 ) = sign_num (t1 * x^2 + u1 * x + v1 ) " using ins x1gt samesign apply (auto) by blast then have "t1 * x1\<^sup>2 + u1 * x1 + v1 \ 0" using xsn unfolding sign_num_def by auto then show "False" using contrad by auto qed have allpropdvar: "(\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)" proof clarsimp fix t1 u1 v1 assume "(t1, u1, v1) \ set d" then have "\x\{?bgrt<..(?bgrt + 1)}. t1 * x\<^sup>2 + u1 * x + v1 \ 0" using allpropd by force then show "\y'>sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - Suc 0). \x\{sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - Suc 0)<..y'}. t1 * x\<^sup>2 + u1 * x + v1 \ 0" using less_add_one by force qed have "\x. (\(d, e, f)\set a. d * x\<^sup>2 + e * x + f = 0)" using alleqsetvar by auto then have ast: "(\(d, e, f)\set a. \x\{?bgrt<..(?bgrt + 1)}. d * x\<^sup>2 + e * x + f = 0)" by auto have allpropavar: "(\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0)" proof clarsimp fix t1 u1 v1 assume "(t1, u1, v1) \ set a" then have "\x\{?bgrt<..(?bgrt + 1)}. t1 * x\<^sup>2 + u1 * x + v1 = 0 " using ast by auto then show "\y'>sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - Suc 0). \x\{sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - Suc 0)<..y'}. t1 * x\<^sup>2 + u1 * x + v1 = 0" using less_add_one One_nat_def by metis qed have quadsetb: "((t, u, v) \ set b \ t\ 0) \ False" proof - assume asm: "(t, u, v) \ set b \ t\ 0" have bgrt1: "(?bgrt = (- u + 1 * sqrt (u^2 - 4 * t * v)) / (2 * t)) \ False " proof - assume bgrtis: "?bgrt = (- u + 1 * sqrt (u^2 - 4 * t * v)) / (2 * t)" have discrim_prop: "-1*u^2 + 4 * t * v \ 0" using asm tnonz using \sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - 1) = (- u + 1 * sqrt (u\<^sup>2 - 4 * t * v)) / (2 * t)\ by auto have "((t, u, v)\set b \ t \ 0 \ - 1*u^2 + 4 * t * v \ 0 \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm discrim_prop allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using f6 bgrtis by auto qed have bgrt2: "(?bgrt = (- u + -1 * sqrt (u^2 - 4 * t * v)) / (2 * t)) \ False " proof - assume bgrtis: "?bgrt = (- u + -1 * sqrt (u^2 - 4 * t * v)) / (2 * t)" have discrim_prop: "-1*u^2 + 4 * t * v \ 0" using asm tnonz using \sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - 1) = (- u + -1 * sqrt (u\<^sup>2 - 4 * t * v)) / (2 * t)\ by auto have "((t, u, v)\set b \ t \ 0 \ - 1*u^2 + 4 * t * v \ 0 \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm discrim_prop allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using f7 bgrtis by auto qed show "False" using tnonz bgrt1 bgrt2 asm by auto qed have linsetb: "((t, u, v) \ set b \ (t = 0 \ u \ 0)) \ False" proof - assume asm: "(t, u, v) \ set b \ (t = 0 \ u \ 0)" then have bgrtis: "?bgrt = (- v / u)" using unonz by blast have "((t, u, v)\set b \ (t = 0 \ u \ 0) \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using bgrtis f5 by auto qed have insetb: "((t, u, v) \ set b \ (t \ 0 \ u \ 0) \ (t * ?bgrt\<^sup>2 + u * ?bgrt + v = 0)) \ False" using quadsetb linsetb by auto have quadsetc: "(t, u, v) \ set c \ t\ 0 \ False" proof - assume asm: "(t, u, v) \ set c \ t\ 0" have bgrt1: "(?bgrt = (- u + 1 * sqrt (u^2 - 4 * t * v)) / (2 * t)) \ False " proof - assume bgrtis: "?bgrt = (- u + 1 * sqrt (u^2 - 4 * t * v)) / (2 * t)" have discrim_prop: "-1*u^2 + 4 * t * v \ 0" using asm tnonz using \sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - 1) = (- u + 1 * sqrt (u\<^sup>2 - 4 * t * v)) / (2 * t)\ by auto have "((t, u, v)\set c \ t \ 0 \ - 1*u^2 + 4 * t * v \ 0 \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm discrim_prop allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using f13a bgrtis by auto qed have bgrt2: "(?bgrt = (- u + -1 * sqrt (u^2 - 4 * t * v)) / (2 * t)) \ False " proof - assume bgrtis: "?bgrt = (- u + -1 * sqrt (u^2 - 4 * t * v)) / (2 * t)" have discrim_prop: "-1*u^2 + 4 * t * v \ 0" using asm tnonz using \sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - 1) = (- u + -1 * sqrt (u\<^sup>2 - 4 * t * v)) / (2 * t)\ by auto have "((t, u, v)\set c \ t \ 0 \ - 1*u^2 + 4 * t * v \ 0 \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm discrim_prop allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using f9a bgrtis by auto qed show "False" using tnonz bgrt1 bgrt2 asm by auto qed have linsetc: "(t, u, v) \ set c \ (t = 0 \ u \ 0) \ False" proof - assume asm: "(t, u, v) \ set c \ (t = 0 \ u \ 0)" then have bgrtis: "?bgrt = (- v / u)" using unonz by blast have "((t, u, v)\set c \ (t = 0 \ u \ 0) \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using bgrtis f8a by auto qed have insetc: "((t, u, v) \ set c \ (t \ 0 \ u \ 0) \ (t * ?bgrt\<^sup>2 + u * ?bgrt + v = 0)) \ False" using quadsetc linsetc by auto have quadsetd: "(t, u, v) \ set d \ t\ 0 \ False" proof - assume asm: "(t, u, v) \ set d \ t\ 0" have bgrt1: "(?bgrt = (- u + 1 * sqrt (u^2 - 4 * t * v)) / (2 * t)) \ False " proof - assume bgrtis: "?bgrt = (- u + 1 * sqrt (u^2 - 4 * t * v)) / (2 * t)" have discrim_prop: "-1*u^2 + 4 * t * v \ 0" using asm tnonz using \sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - 1) = (- u + 1 * sqrt (u\<^sup>2 - 4 * t * v)) / (2 * t)\ by auto have "((t, u, v)\set d \ t \ 0 \ - 1*u^2 + 4 * t * v \ 0 \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm discrim_prop allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using f11 bgrtis by auto qed have bgrt2: "(?bgrt = (- u + -1 * sqrt (u^2 - 4 * t * v)) / (2 * t)) \ False " proof - assume bgrtis: "?bgrt = (- u + -1 * sqrt (u^2 - 4 * t * v)) / (2 * t)" have discrim_prop: "-1*u^2 + 4 * t * v \ 0" using asm tnonz using \sorted_nonzero_root_list_set (set b \ set c \ set d) ! (length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - 1) = (- u + -1 * sqrt (u\<^sup>2 - 4 * t * v)) / (2 * t)\ by auto have "((t, u, v)\set d \ t \ 0 \ - 1*u^2 + 4 * t * v \ 0 \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm discrim_prop allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using f12 bgrtis by auto qed show "False" using tnonz bgrt1 bgrt2 asm by auto qed have linsetd: "(t, u, v) \ set d \ (t = 0 \ u \ 0) \ False" proof - assume asm: "(t, u, v) \ set d \ (t = 0 \ u \ 0)" then have bgrtis: "?bgrt = (- v / u)" using unonz by blast have "((t, u, v)\set d \ (t = 0 \ u \ 0) \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using bgrtis f10 by auto qed have insetd: "((t, u, v) \ set d \ (t \ 0 \ u \ 0) \ (t * ?bgrt\<^sup>2 + u * ?bgrt + v = 0)) \ False" using quadsetd linsetd by auto then show "False" using insetb insetc insetd tuvprop by auto qed have len1: "length ?srl = 1 \ False" proof - assume len1: "length ?srl = 1" have cases: "(List.member ?srl x) \ x < ?srl ! 0 \ x > ?srl ! 0" using in_set_member lenzero nth_mem by fastforce then show "False" using len1 cases_mem cases_lt cases_gt by auto qed have lengtone: "length ?srl > 1 \ False" proof - assume lengt1: "length ?srl > 1" have cases: "(List.member ?srl x) \ x < ?srl ! 0 \ x > ?srl ! (length ?srl -1) \ (\k \ (length ?srl - 2). (?srl ! k < x \ x x > ?srl ! (length ?srl -1) \ (x \ ?srl ! 0 \ x \ ?srl ! (length ?srl -1))" by auto have ifo: "(x \ ?srl ! 0 \ x \ ?srl ! (length ?srl -1)) \ ((List.member ?srl x) \ (\k \ (length ?srl - 2). ?srl ! k < x \ x ?srl ! 0 \ x \ ?srl ! (length ?srl -1)" then have "\(List.member ?srl x) \ (\k \ (length ?srl - 2). ?srl ! k < x \ x (List.member ?srl x)" have "\(\k \ (length ?srl - 2). ?srl ! k < x \ x False" proof clarsimp assume "\k. sorted_nonzero_root_list_set (set b \ set c \ set d) ! k < x \ k \ length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - 2 \ \ x < sorted_nonzero_root_list_set (set b \ set c \ set d) ! Suc k" then have allk: "(\k \ length ?srl - 2. ?srl ! k < x \ \ x < ?srl ! Suc k)" by auto have basec: "x \ ?srl ! 0" using xinbtw by auto have "\k \ length ?srl - 2. ?srl ! k < x" proof clarsimp fix k assume klteq: "k \ length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - 2" show "sorted_nonzero_root_list_set (set b \ set c \ set d) ! k < x" using nonmem klteq basec proof (induct k) case 0 then show ?case using in_set_member lenzero nth_mem by fastforce next case (Suc k) then show ?case by (smt Suc_leD Suc_le_lessD \\k. sorted_nonzero_root_list_set (set b \ set c \ set d) ! k < x \ k \ length (sorted_nonzero_root_list_set (set b \ set c \ set d)) - 2 \ \ x < sorted_nonzero_root_list_set (set b \ set c \ set d) ! Suc k\ diff_less in_set_member length_0_conv length_greater_0_conv lenzero less_trans_Suc nth_mem pos2) qed qed then have "x \ ?srl ! (length ?srl -1)" using allk by (metis One_nat_def Suc_diff_Suc lengt1 less_eq_real_def less_or_eq_imp_le one_add_one plus_1_eq_Suc xinbtw) then have "x > ?srl ! (length ?srl - 1)" using nonmem by (metis One_nat_def Suc_le_D asm diff_Suc_Suc diff_zero in_set_member lessI less_eq_real_def nth_mem) then show "False" using xinbtw by auto qed then show "(\k \ (length ?srl - 2). ?srl ! k < x \ x (\k \ (length ?srl - 2). ?srl ! k < x \ x k \ (length ?srl - 2). ?srl ! k < x \ x False" proof - assume "(\k \ (length ?srl - 2). ?srl ! k < x \ x (length ?srl - 2) \ ?srl ! k < x \ x (a, b, c) \ (set b \ set c \ set d). (\y. (?srl ! k < y \ y sign_num (a * y\<^sup>2 + b * y + c) = sign_num (a*x^2 + b*x + c))" proof clarsimp fix t u v y assume insetunion: "(t, u, v) \ set b \ (t, u, v) \ set c \ (t, u, v) \ set d" assume ygt: " sorted_nonzero_root_list_set (set b \ set c \ set d) ! k < y" assume ylt: "y < sorted_nonzero_root_list_set (set b \ set c \ set d) ! Suc k" have tuzer: "t = 0 \ u = 0 \ sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" unfolding sign_num_def by auto have tunonzer: "t \ 0 \ u \ 0 \ sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" proof - assume tuv_asm: "t\ 0 \ u \ 0" have nor: "\(\q. q > ?srl ! k \ q < ?srl ! (k + 1) \ t * q\<^sup>2 + u * q + v = 0)" proof clarsimp fix q assume qlt: "q < sorted_nonzero_root_list_set (set b \ set c \ set d) ! Suc k" assume qgt: "sorted_nonzero_root_list_set (set b \ set c \ set d) ! k < q" assume "t * q\<^sup>2 + u * q + v = 0" then have qin: "q \ {x. \(a, b, c)\set b \ set c \ set d. (a \ 0 \ b \ 0) \ a * x\<^sup>2 + b * x + c = 0}" using insetunion tuv_asm by auto have "set ?srl = nonzero_root_set (set b \ set c \ set d)" unfolding sorted_nonzero_root_list_set_def using set_sorted_list_of_set[of "nonzero_root_set (set b \ set c \ set d)"] nonzero_root_set_finite[of "(set b \ set c \ set d)"] by auto then have "q \ set ?srl" using qin unfolding nonzero_root_set_def by auto then have "List.member ?srl q" using in_set_member[of q ?srl] by auto then have "\n < length ?srl. q = ?srl ! n" by (metis \q \ set (sorted_nonzero_root_list_set (set b \ set c \ set d))\ in_set_conv_nth) then obtain n where nprop: "n < length ?srl \ q = ?srl ! n" by auto then have ngtk: "n > k" proof - have sortedh: "sorted ?srl" by (simp add: sorted_nonzero_root_list_set_def) then have nlteq: "n \ k \ ?srl ! n \ ?srl ! k" using nprop k_prop sorted_iff_nth_mono using sorted_nth_mono by (metis Suc_1 \q \ set (sorted_nonzero_root_list_set (set b \ set c \ set d))\ diff_less length_pos_if_in_set sup.absorb_iff2 sup.strict_boundedE zero_less_Suc) have "?srl ! n > ?srl ! k" using nprop qgt by auto then show ?thesis using nlteq by linarith qed then have nltkp1: "n < k+1" proof - have sortedh: "sorted ?srl" by (simp add: sorted_nonzero_root_list_set_def) then have ngteq: "k+1 \ n \ ?srl ! (k+1) \ ?srl ! n" using nprop k_prop sorted_iff_nth_mono by auto have "?srl ! n < ?srl ! (k + 1)" using nprop qlt by auto then show ?thesis using ngteq by linarith qed then show "False" using ngtk nltkp1 by auto qed have c1: " x > y \ sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" using nor changes_sign_var[of t y u v x] k_prop ygt ylt by fastforce then have c2: " y > x \ sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" using nor changes_sign_var[of t x u v y] k_prop ygt ylt by force then have c3: " x = y \ sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" unfolding sign_num_def by auto then show "sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" using c1 c2 c3 by linarith qed then show " sign_num (t * y\<^sup>2 + u * y + v) = sign_num (t * x\<^sup>2 + u * x + v)" using tuzer by blast qed let ?bgrt = "?srl ! k" have "(\(a, b, c)\set a. a = 0 \ b = 0 \ c = 0)" using alleqsetvar by auto have " ?bgrt \ set ?srl" using set_sorted_list_of_set nonzero_root_set_finite in_set_member k_prop asm by (smt diff_Suc_less le_eq_less_or_eq less_le_trans nth_mem one_add_one plus_1_eq_Suc zero_less_one) then have "?bgrt \ nonzero_root_set (set b \ set c \ set d )" unfolding sorted_nonzero_root_list_set_def using set_sorted_list_of_set nonzero_root_set_finite by auto then have "\t u v. (t, u, v) \ set b \ set c \ set d \(t \ 0 \ u \ 0) \ (t * ?bgrt\<^sup>2 + u * ?bgrt + v = 0)" unfolding nonzero_root_set_def by auto then obtain t u v where tuvprop1: "(t, u, v) \ set b \ set c \ set d \(t \ 0 \ u \ 0) \ (t * ?bgrt\<^sup>2 + u * ?bgrt + v = 0)" by auto then have tuvprop: "((t, u, v) \ set b \ (t \ 0 \ u \ 0) \ (t * ?bgrt\<^sup>2 + u * ?bgrt + v = 0)) \ ((t, u, v) \ set c \ (t \ 0 \ u \ 0) \ (t * ?bgrt\<^sup>2 + u * ?bgrt + v = 0)) \ ((t, u, v) \ set d \ (t \ 0 \ u \ 0) \ (t * ?bgrt\<^sup>2 + u * ?bgrt + v = 0)) " by auto have tnonz: "t\ 0 \ (-1*u^2 + 4 * t * v \ 0 \ (?bgrt = (- u + 1 * sqrt (u^2 - 4 * t * v)) / (2 * t) \ ?bgrt = (- u + -1 * sqrt (u^2 - 4 * t * v)) / (2 * t)))" proof - assume "t\ 0" have "-1*u^2 + 4 * t * v \ 0 " using tuvprop1 discriminant_negative[of t u v] unfolding discrim_def using \t \ 0\ by force then show ?thesis using tuvprop discriminant_nonneg[of t u v] unfolding discrim_def using \t \ 0\ by auto qed have unonz: "(t = 0 \ u \ 0) \ ?bgrt = - v / u" proof - assume "(t = 0 \ u \ 0)" then have "u*?bgrt + v = 0" using tuvprop1 by simp then show "?bgrt = - v / u" by (simp add: \t = 0 \ u \ 0\ eq_minus_divide_eq mult.commute) qed have "\y'. y' > x \ y' < ?srl ! (k+1)" using k_prop dense by blast then obtain y1 where y1_prop: "y1 > x \ y1 < ?srl ! (k+1)" by auto then have y1inbtw: "y1 > ?srl ! k \ y1 < ?srl ! (k+1)" using k_prop by auto have allpropb: "(\(d, e, f)\set b. \x\{?bgrt<..y1}. d * x\<^sup>2 + e * x + f < 0)" proof clarsimp fix t1 u1 v1 x1 assume ins: "(t1, u1, v1) \ set b" assume x1gt: "sorted_nonzero_root_list_set (set b \ set c \ set d) ! k < x1" assume x1lt: "x1 \ y1" have x1inbtw: "x1 > ?srl ! k \ x1 < ?srl ! (k+1)" using x1gt x1lt y1inbtw by (smt One_nat_def cases_gt k_prop) have xsn: "sign_num (t1 * x^2 + u1 * x + v1 ) = -1" using ins x_prop unfolding sign_num_def by auto have "sign_num (t1 * x1\<^sup>2 + u1 * x1 + v1 ) = sign_num (t1 * x^2 + u1 * x + v1 ) " using ins x1inbtw samesign by blast then show "t1 * x1\<^sup>2 + u1 * x1 + v1 < 0" using xsn unfolding sign_num_def by (metis add.right_inverse add.right_neutral linorder_neqE_linordered_idom one_add_one zero_neq_numeral) qed have allpropbvar: "(\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0)" proof clarsimp fix t1 u1 v1 assume "(t1, u1, v1) \ set b" then have "\x\{?bgrt<..y1}. t1 * x\<^sup>2 + u1 * x + v1 < 0" using allpropb by force then show " \y'>sorted_nonzero_root_list_set (set b \ set c \ set d) ! k. \x\{sorted_nonzero_root_list_set (set b \ set c \ set d) ! k<..y'}. t1 * x\<^sup>2 + u1 * x + v1 < 0" using y1inbtw by blast qed have allpropc: "(\(d, e, f)\set c. \x\{?bgrt<..y1}. d * x\<^sup>2 + e * x + f \ 0)" proof clarsimp fix t1 u1 v1 x1 assume ins: "(t1, u1, v1) \ set c" assume x1gt: " sorted_nonzero_root_list_set (set b \ set c \ set d) ! k < x1" assume x1lt: "x1 \ y1" have x1inbtw: "x1 > ?srl ! k \ x1 < ?srl ! (k+1)" using x1gt x1lt y1inbtw by (smt One_nat_def cases_gt k_prop) have xsn: "sign_num (t1 * x^2 + u1 * x + v1 ) = -1 \ sign_num (t1 * x^2 + u1 * x + v1 ) = 0" using ins x_prop unfolding sign_num_def by auto have "sign_num (t1 * x1\<^sup>2 + u1 * x1 + v1 ) = sign_num (t1 * x^2 + u1 * x + v1 ) " using ins x1inbtw samesign by blast then show "t1 * x1\<^sup>2 + u1 * x1 + v1 \ 0" using xsn unfolding sign_num_def by (smt (verit, del_insts)) qed have allpropcvar: "(\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)" proof clarsimp fix t1 u1 v1 assume "(t1, u1, v1) \ set c" then have "\x\{?bgrt<..y1}. t1 * x\<^sup>2 + u1 * x + v1 \ 0" using allpropc by force then show " \y'>sorted_nonzero_root_list_set (set b \ set c \ set d) ! k. \x\{sorted_nonzero_root_list_set (set b \ set c \ set d) ! k<..y'}. t1 * x\<^sup>2 + u1 * x + v1 \ 0" using y1inbtw by blast qed have allpropd: "(\(d, e, f)\set d. \x\{?bgrt<..y1}. d * x\<^sup>2 + e * x + f \ 0)" proof clarsimp fix t1 u1 v1 x1 assume ins: "(t1, u1, v1) \ set d" assume contrad:"t1 * x1\<^sup>2 + u1 * x1 + v1 = 0" assume x1gt: " sorted_nonzero_root_list_set (set b \ set c \ set d) ! k < x1" assume x1lt: "x1 \ y1" have x1inbtw: "x1 > ?srl ! k \ x1 < ?srl ! (k+1)" using x1gt x1lt y1inbtw by (smt One_nat_def cases_gt k_prop) have xsn: "sign_num (t1 * x^2 + u1 * x + v1 ) = -1 \ sign_num (t1 * x^2 + u1 * x + v1 ) = 1" using ins x_prop unfolding sign_num_def by auto have "sign_num (t1 * x1\<^sup>2 + u1 * x1 + v1 ) = sign_num (t1 * x^2 + u1 * x + v1 ) " using ins x1inbtw samesign by blast then have "t1 * x1\<^sup>2 + u1 * x1 + v1 \ 0" using xsn unfolding sign_num_def by auto then show "False" using contrad by auto qed have allpropdvar: "(\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)" proof clarsimp fix t1 u1 v1 assume "(t1, u1, v1) \ set d" then have "\x\{?bgrt<..y1}. t1 * x\<^sup>2 + u1 * x + v1 \ 0" using allpropd by force then show " \y'>sorted_nonzero_root_list_set (set b \ set c \ set d) ! k. \x\{sorted_nonzero_root_list_set (set b \ set c \ set d) ! k<..y'}. t1 * x\<^sup>2 + u1 * x + v1 \ 0" using y1inbtw by blast qed have "\x. (\(d, e, f)\set a. d * x\<^sup>2 + e * x + f = 0)" using alleqsetvar by auto then have ast: "(\(d, e, f)\set a. \x\{?bgrt<..(?bgrt + 1)}. d * x\<^sup>2 + e * x + f = 0)" by auto have allpropavar: "(\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0)" proof clarsimp fix t1 u1 v1 assume "(t1, u1, v1) \ set a" then have "\x\{?bgrt<..(?bgrt + 1)}. t1 * x\<^sup>2 + u1 * x + v1 = 0 " using ast by auto then show "\y'>sorted_nonzero_root_list_set (set b \ set c \ set d) ! k. \x\{sorted_nonzero_root_list_set (set b \ set c \ set d) ! k<..y'}. t1 * x\<^sup>2 + u1 * x + v1 = 0" using less_add_one by blast qed have quadsetb: "((t, u, v) \ set b \ t\ 0) \ False" proof - assume asm: "(t, u, v) \ set b \ t\ 0" have bgrt1: "(?bgrt = (- u + 1 * sqrt (u^2 - 4 * t * v)) / (2 * t)) \ False " proof - assume bgrtis: "?bgrt = (- u + 1 * sqrt (u^2 - 4 * t * v)) / (2 * t)" have discrim_prop: "-1*u^2 + 4 * t * v \ 0" using asm tnonz by blast have "((t, u, v)\set b \ t \ 0 \ - 1*u^2 + 4 * t * v \ 0 \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm discrim_prop allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using f6 bgrtis by auto qed have bgrt2: "(?bgrt = (- u + -1 * sqrt (u^2 - 4 * t * v)) / (2 * t)) \ False " proof - assume bgrtis: "?bgrt = (- u + -1 * sqrt (u^2 - 4 * t * v)) / (2 * t)" have discrim_prop: "-1*u^2 + 4 * t * v \ 0" using asm tnonz by blast have "((t, u, v)\set b \ t \ 0 \ - 1*u^2 + 4 * t * v \ 0 \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm discrim_prop allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using f7 bgrtis by auto qed show "False" using tnonz bgrt1 bgrt2 asm by auto qed have linsetb: "((t, u, v) \ set b \ (t = 0 \ u \ 0)) \ False" proof - assume asm: "(t, u, v) \ set b \ (t = 0 \ u \ 0)" then have bgrtis: "?bgrt = (- v / u)" using unonz by blast have "((t, u, v)\set b \ (t = 0 \ u \ 0) \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using bgrtis f5 by auto qed have insetb: "((t, u, v) \ set b \ (t \ 0 \ u \ 0) \ (t * ?bgrt\<^sup>2 + u * ?bgrt + v = 0)) \ False" using quadsetb linsetb by auto have quadsetc: "(t, u, v) \ set c \ t\ 0 \ False" proof - assume asm: "(t, u, v) \ set c \ t\ 0" have bgrt1: "(?bgrt = (- u + 1 * sqrt (u^2 - 4 * t * v)) / (2 * t)) \ False " proof - assume bgrtis: "?bgrt = (- u + 1 * sqrt (u^2 - 4 * t * v)) / (2 * t)" have discrim_prop: "-1*u^2 + 4 * t * v \ 0" using asm tnonz by blast have "((t, u, v)\set c \ t \ 0 \ - 1*u^2 + 4 * t * v \ 0 \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm discrim_prop allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using f13a bgrtis by auto qed have bgrt2: "(?bgrt = (- u + -1 * sqrt (u^2 - 4 * t * v)) / (2 * t)) \ False " proof - assume bgrtis: "?bgrt = (- u + -1 * sqrt (u^2 - 4 * t * v)) / (2 * t)" have discrim_prop: "-1*u^2 + 4 * t * v \ 0" using asm tnonz by blast have "((t, u, v)\set c \ t \ 0 \ - 1*u^2 + 4 * t * v \ 0 \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm discrim_prop allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using f9a bgrtis by auto qed show "False" using tnonz bgrt1 bgrt2 asm by auto qed have linsetc: "(t, u, v) \ set c \ (t = 0 \ u \ 0) \ False" proof - assume asm: "(t, u, v) \ set c \ (t = 0 \ u \ 0)" then have bgrtis: "?bgrt = (- v / u)" using unonz by blast have "((t, u, v)\set c \ (t = 0 \ u \ 0) \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using bgrtis f8a by auto qed have insetc: "((t, u, v) \ set c \ (t \ 0 \ u \ 0) \ (t * ?bgrt\<^sup>2 + u * ?bgrt + v = 0)) \ False" using quadsetc linsetc by auto have quadsetd: "(t, u, v) \ set d \ t\ 0 \ False" proof - assume asm: "(t, u, v) \ set d \ t\ 0" have bgrt1: "(?bgrt = (- u + 1 * sqrt (u^2 - 4 * t * v)) / (2 * t)) \ False " proof - assume bgrtis: "?bgrt = (- u + 1 * sqrt (u^2 - 4 * t * v)) / (2 * t)" have discrim_prop: "-1*u^2 + 4 * t * v \ 0" using asm tnonz by blast have "((t, u, v)\set d \ t \ 0 \ - 1*u^2 + 4 * t * v \ 0 \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm discrim_prop allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using f11 bgrtis by auto qed have bgrt2: "(?bgrt = (- u + -1 * sqrt (u^2 - 4 * t * v)) / (2 * t)) \ False " proof - assume bgrtis: "?bgrt = (- u + -1 * sqrt (u^2 - 4 * t * v)) / (2 * t)" have discrim_prop: "-1*u^2 + 4 * t * v \ 0" using asm tnonz by blast have "((t, u, v)\set d \ t \ 0 \ - 1*u^2 + 4 * t * v \ 0 \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm discrim_prop allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using f12 bgrtis by auto qed show "False" using tnonz bgrt1 bgrt2 asm by auto qed have linsetd: "(t, u, v) \ set d \ (t = 0 \ u \ 0) \ False" proof - assume asm: "(t, u, v) \ set d \ (t = 0 \ u \ 0)" then have bgrtis: "?bgrt = (- v / u)" using unonz by blast have "((t, u, v)\set d \ (t = 0 \ u \ 0) \ ((\(d, e, f)\set a. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>?bgrt. \x\{?bgrt<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" using asm allpropavar allpropbvar allpropcvar allpropdvar by linarith then show "False" using bgrtis f10 by auto qed have insetd: "((t, u, v) \ set d \ (t \ 0 \ u \ 0) \ (t * ?bgrt\<^sup>2 + u * ?bgrt + v = 0)) \ False" using quadsetd linsetd by auto then show "False" using insetb insetc insetd tuvprop by auto qed show "False" using cases cases_btw cases_mem cases_lt cases_gt by auto qed show "False" using asm len1 lengtone by linarith qed show "False" using lenzero lengt0 by linarith qed then show ?thesis by blast qed lemma qe_forwards: assumes "(\x. (\(a, b, c)\set a. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0))" shows "((\(a, b, c)\set a. a = 0 \ b = 0 \ c = 0) \ (\(a, b, c)\set b. \x. \y2 + b * y + c < 0) \ (\(a, b, c)\set c. \x. \y2 + b * y + c \ 0) \ (\(a, b, c)\set d. \x. \y2 + b * y + c \ 0) \ (\(a', b', c')\set a. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))) \ (\(a', b', c')\set b. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))) \ (\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))) \ (\(a', b', c')\set d. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))))" (* using eq_qe_1 les_qe_1 *) proof - let ?e2 = "(((\(a, b, c)\set a. a = 0 \ b = 0 \ c = 0) \ (\(a, b, c)\set b. \x. \y2 + b * y + c < 0) \ (\(a, b, c)\set c. \x. \y2 + b * y + c \ 0) \ (\(a, b, c)\set d. \x. \y2 + b * y + c \ 0) \ (\(a', b', c')\set a. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))) \ (\(a', b', c')\set b. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))) \ (\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))) \ (\(a', b', c')\set d. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))))" let ?f10orf11orf12 = "(\(a', b', c')\set d. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" let ?f8orf9 = "(\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)))" let ?f5orf6orf7 = "(\(a', b', c')\set b. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" let ?f2orf3orf4 = "(\(a', b', c')\set a. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)))" let ?e1 = "(\x. (\(a, b, c)\set a. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0))" let ?f1 = "((\(a, b, c)\set a. a = 0 \ b = 0 \ c = 0) \ (\(a, b, c)\set b. \x. \y2 + b * y + c < 0) \ (\(a, b, c)\set c. \x. \y2 + b * y + c \ 0) \ (\(a, b, c)\set d. \x. \y2 + b * y + c \ 0))" let ?f2 = "(\(a', b', c')\set a. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0))" let ?f3 = "(\(a', b', c')\set a. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))" let ?f4 = "(\(a', b', c')\set a. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)) " let ?f5 = "(\(a', b', c')\set b. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0))" let ?f6 = "(\(a', b', c')\set b. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" let ?f7 = "(\(a', b', c')\set b. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))" let ?f8 = "(\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0))" let ?f13 = "(\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)))" let ?f9 = "(\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))" let ?f10 = "(\(a', b', c')\set d. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0))" let ?f11 = "(\(a', b', c')\set d. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" let ?f12 = "(\(a', b', c')\set d. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))" have h1a: "(?f1 \ ?f2orf3orf4 \ ?f5orf6orf7 \ ?f8orf9 \ ?f10orf11orf12) \ ?e2" by auto have h2: "(?f2 \ ?f3 \ ?f4) \ ?f2orf3orf4" by auto then have h1b: "(?f1 \ ?f2 \ ?f3 \ ?f4 \ ?f5orf6orf7 \ ?f8orf9 \ ?f10orf11orf12) \ ?e2" using h1a by auto have h3: "(?f5 \ ?f6 \ ?f7) \ ?f5orf6orf7" by auto then have h1c: "(?f1 \ ?f2 \ ?f3 \ ?f4 \ ?f5 \ ?f6 \ ?f7 \ ?f8orf9 \ ?f10orf11orf12) \ ?e2" using h1b by smt have h4: "(?f8 \ ?f9 \ ?f13) \ ?f8orf9" by auto then have h1d: "(?f1 \ ?f2 \ ?f3 \ ?f4 \ ?f5 \ ?f6 \ ?f7 \ ?f8 \ ?f9 \ ?f13 \ ?f10orf11orf12) \ ?e2" using h1c by smt have h5: "(?f10 \ ?f11 \ ?f12) \ ?f10orf11orf12" by auto then have bigor: "(?f1 \ ?f2 \ ?f3 \ ?f4 \ ?f5 \ ?f6 \ ?f7 \ ?f8 \ ?f13 \ ?f9 \ ?f10 \ ?f11 \ ?f12) \ ?e2 " using h1d by smt then have bigor_var: "\?e2 \ \(?f1 \ ?f2 \ ?f3 \ ?f4 \ ?f5 \ ?f6 \ ?f7 \ ?f8 \ ?f13 \ ?f9 \ ?f10 \ ?f11 \ ?f12) " using contrapos_nn by smt have not_eq: "\(?f1 \ ?f2 \ ?f3 \ ?f4 \ ?f5 \ ?f6 \ ?f7 \ ?f8 \ ?f13 \ ?f9 \ ?f10 \ ?f11 \ ?f12) =(\?f1 \ \?f2 \ \?f3 \ \?f4 \ \?f5 \ \?f6 \ \?f7 \ \?f8 \ \?f13 \ \?f9 \ \?f10 \ \?f11 \ \?f12) " by linarith obtain x where x_prop: "(\(a, b, c)\set a. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0)" using assms by auto have "(\?f1 \ \?f2 \ \?f3 \ \?f4 \ \?f5 \ \?f6 \ \?f7 \ \?f8 \ \?f13 \ \?f9 \ \?f10 \ \?f11 \ \?f12) \ False" proof - assume big_not: " \ ((\(a, b, c)\set a. a = 0 \ b = 0 \ c = 0) \ (\(a, b, c)\set b. \x. \y2 + b * y + c < 0) \ (\(a, b, c)\set c. \x. \y2 + b * y + c \ 0) \ (\(a, b, c)\set d. \x. \y2 + b * y + c \ 0)) \ \ (\(a', b', c')\set a. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0)) \ \ (\(a', b', c')\set a. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)) \ \ (\(a', b', c')\set a. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)) \ \ (\(a', b', c')\set b. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0)) \ \ (\(a', b', c')\set b. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)) \ \ (\(a', b', c')\set b. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)) \ \ (\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0)) \ \ (\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)) \ \ (\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)) \ \ (\(a', b', c')\set d. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0)) \ \ (\(a', b', c')\set d. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)) \ \ (\(a', b', c')\set d. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))" have c1: "(\ (d, e, f) \ set a. d \ 0 \ - e\<^sup>2 + 4 * d * f \ 0) \ False" proof - assume "(\ (d, e, f) \ set a. d \ 0 \ - e\<^sup>2 + 4 * d * f \ 0)" then obtain a' b' c' where abc_prop: "(a', b', c') \ set a \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0" by auto then have "a'*x^2 + b'*x + c' = 0" using x_prop by auto then have xis: "x = (- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a') \ x = (- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a') " using abc_prop discriminant_nonneg[of a' b' c'] unfolding discrim_def by auto then have "((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)) \ ((\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))" using x_prop by auto then have "(\(a', b', c')\set a. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)) \ (\(a', b', c')\set a. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))" using abc_prop xis by auto then show "False" using big_not by auto qed have c2: "(\ (d, e, f) \ set a. d = 0 \ e \ 0) \ False" proof - assume "(\ (d, e, f) \ set a. d = 0 \ e \ 0)" then obtain a' b' c' where abc_prop: "(a', b', c') \ set a \ a' = 0 \ b' \ 0" by auto then have "a'*x^2 + b'*x + c' = 0" using x_prop by auto then have "b'*x + c' = 0" using abc_prop by auto then have xis: "x = - c' / b'" using abc_prop by (smt divide_non_zero) then have "(\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0)" using x_prop by auto then have "(\(a', b', c')\set a. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0))" using abc_prop xis by auto then show "False" using big_not by auto qed have c3: "(\ (d, e, f) \ set a. d = 0 \ e = 0 \ f = 0) \ False" proof - assume "(\ (d, e, f) \ set a. d = 0 \ e = 0 \ f = 0)" then have equalset: "\x. (\(d, e, f)\set a. d * x^2 + e * x + f = 0)" using case_prodE by auto have "\?f5 \ \?f6 \ \?f7 \ \?f8 \ \?f13 \ \?f9 \ \?f10 \ \?f11 \ \?f12" using big_not by auto then have "\(\x. (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0))" using equalset big_not qe_forwards_helper[of a b c d] by auto then show "False" using x_prop by auto qed have eo: "(\ (d, e, f) \ set a. d \ 0 \ - e\<^sup>2 + 4 * d * f \ 0) \ (\ (d, e, f) \ set a. d = 0 \ e \ 0) \ (\ (d, e, f) \ set a. d = 0 \ e = 0 \ f = 0)" proof - have "(\ (d, e, f) \ set a. (d \ 0 \ - e\<^sup>2 + 4 * d * f \ 0))" proof clarsimp fix d e f assume in_set: " (d, e, f) \ set a" assume dnonz: "d \ 0" have "d*x^2 + e*x + f = 0" using in_set x_prop by auto then show " 4 * d * f \ e\<^sup>2" using dnonz discriminant_negative[of d e f] unfolding discrim_def by fastforce qed then have discrim_prop: "\(\ (d, e, f) \ set a. d \ 0 \ - e\<^sup>2 + 4 * d * f \ 0) \ \(\ (d, e, f) \ set a. d \ 0)" by auto have "\(\ (d, e, f) \ set a. d \ 0) \ \(\ (d, e, f) \ set a. d = 0 \ e \ 0) \ (\ (d, e, f) \ set a. d = 0 \ e = 0 \ f = 0)" proof - assume ne: "\(\ (d, e, f) \ set a. d \ 0) \ \(\ (d, e, f) \ set a. d = 0 \ e \ 0)" show "(\ (d, e, f) \ set a. d = 0 \ e = 0 \ f = 0)" proof clarsimp fix d e f assume in_set: "(d, e, f) \set a" then have xzer: "d*x^2 + e*x + f = 0" using x_prop by auto have dzer: "d = 0" using ne in_set by auto have ezer: "e = 0" using ne in_set by auto show "d = 0 \ e = 0 \ f = 0" using xzer dzer ezer by auto qed qed then show ?thesis using discrim_prop by auto qed show "False" using c1 c2 c3 eo by auto qed then have " \?e2 \ False" using bigor_var not_eq by presburger (* Takes a second *) then have " \?e2 \ False" using impI[of "\?e2" "False"] by blast then show ?thesis by auto qed subsubsection "Some Cases and Misc" lemma quadratic_linear : assumes "b\0" assumes "a \ 0" assumes "4 * a * ba \ aa\<^sup>2" assumes "b * (sqrt (aa\<^sup>2 - 4 * a * ba) - aa) / (2 * a) + c = 0" assumes "\x\set eq. case x of (d, e, f) \ d * ((sqrt (aa\<^sup>2 - 4 * a * ba) - aa) / (2 * a))\<^sup>2 + e * (sqrt (aa\<^sup>2 - 4 * a * ba) - aa) / (2 * a) + f = 0" assumes "(aaa, aaaa, baa) \ set eq" shows "aaa * (c / b)\<^sup>2 - aaaa * c / b + baa = 0" proof- have h: "-(c/b) = (sqrt (aa\<^sup>2 - 4 * a * ba) - aa) / (2 * a)" using assms by (smt divide_minus_left nonzero_mult_div_cancel_left times_divide_eq_right) have h1 : "\x\set eq. case x of (d, e, f) \ d * (c / b)\<^sup>2 + e * - (c / b) + f = 0" using assms(5) unfolding h[symmetric] Fields.division_ring_class.times_divide_eq_right[symmetric] Power.ring_1_class.power2_minus . show ?thesis using bspec[OF h1 assms(6)] by simp qed lemma quadratic_linear1: assumes "b\0" assumes "a \ 0" assumes "4 * a * ba \ aa\<^sup>2" assumes "(b::real) * (sqrt ((aa::real)\<^sup>2 - 4 * (a::real) * (ba::real)) - (aa::real)) / (2 * a) + (c::real) = 0" assumes " (\x\set (les::(real*real*real)list). case x of (d, e, f) \ d * ((sqrt (aa\<^sup>2 - 4 * a * ba) - aa) / (2 * a))\<^sup>2 + e * (sqrt (aa\<^sup>2 - 4 * a * ba) - aa) / (2 * a) + f < 0)" assumes "(aaa, aaaa, baa) \ set les" shows "aaa * (c / b)\<^sup>2 - aaaa * c / b + baa < 0" proof- have h: "-(c/b) = (sqrt (aa\<^sup>2 - 4 * a * ba) - aa) / (2 * a)" using assms by (smt divide_minus_left nonzero_mult_div_cancel_left times_divide_eq_right) have h1 : "\x\set les. case x of (d, e, f) \ d * (c / b)\<^sup>2 + e * - (c / b) + f < 0" using assms(5) unfolding h[symmetric] Fields.division_ring_class.times_divide_eq_right[symmetric] Power.ring_1_class.power2_minus . show ?thesis using bspec[OF h1 assms(6)] by simp qed lemma quadratic_linear2 : assumes "b\0" assumes "a \ 0" assumes "4 * a * ba \ aa\<^sup>2" assumes "b * (- aa -sqrt (aa\<^sup>2 - 4 * a * ba)) / (2 * a) + c = 0" assumes "\x\set eq. case x of (d, e, f) \ d * ((- aa -sqrt (aa\<^sup>2 - 4 * a * ba)) / (2 * a))\<^sup>2 + e * (- aa -sqrt (aa\<^sup>2 - 4 * a * ba)) / (2 * a) + f = 0" assumes "(aaa, aaaa, baa) \ set eq" shows "aaa * (c / b)\<^sup>2 - aaaa * c / b + baa = 0" proof- have h: "-((c::real)/(b::real)) = (- (aa::real) -sqrt (aa\<^sup>2 - 4 * (a::real) * (ba::real))) / (2 * a)" using assms by (smt divide_minus_left nonzero_mult_div_cancel_left times_divide_eq_right) have h1 : "\x\set eq. case x of (d, e, f) \ d * (c / b)\<^sup>2 + e * - (c / b) + f = 0" using assms(5) unfolding h[symmetric] Fields.division_ring_class.times_divide_eq_right[symmetric] Power.ring_1_class.power2_minus . show ?thesis using bspec[OF h1 assms(6)] by simp qed lemma quadratic_linear3: assumes "b\0" assumes "a \ 0" assumes "4 * a * ba \ aa\<^sup>2" assumes "(b::real) * (- (aa::real)- sqrt ((aa::real)\<^sup>2 - 4 * (a::real) * (ba::real)) ) / (2 * a) + (c::real) = 0" assumes "(\x\set (les::(real*real*real)list). case x of (d, e, f) \ d * ((- aa - sqrt (aa\<^sup>2 - 4 * a * ba)) / (2 * a))\<^sup>2 + e * (- aa - sqrt (aa\<^sup>2 - 4 * a * ba)) / (2 * a) + f < 0)" assumes "(aaa, aaaa, baa) \ set les" shows "aaa * (c / b)\<^sup>2 - aaaa * c / b + baa < 0" proof- have h: "-((c::real)/(b::real)) = (- (aa::real) -sqrt (aa\<^sup>2 - 4 * (a::real) * (ba::real))) / (2 * a)" using assms by (smt divide_minus_left nonzero_mult_div_cancel_left times_divide_eq_right) have h1 : "\x\set les. case x of (d, e, f) \ d * (c / b)\<^sup>2 + e * - (c / b) + f < 0" using assms(5) unfolding h[symmetric] Fields.division_ring_class.times_divide_eq_right[symmetric] Power.ring_1_class.power2_minus . show ?thesis using bspec[OF h1 assms(6)] by simp qed lemma h1b_helper_les: "(\((a::real), (b::real), (c::real))\set les. \x. \y2 + b * y + c < 0) \ (\y.\x(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" proof - show "(\(a, b, c)\set les. \x. \y2 + b * y + c < 0) \ (\y.\x(a, b, c)\set les. a * x\<^sup>2 + b * x + c < 0))" proof (induct les) case Nil then show ?case by auto next case (Cons q les) have ind: " \a\set (q # les). case a of (a, ba, c) \ \x. \y2 + ba * y + c < 0" using Cons.prems by auto then have "case q of (a, ba, c) \ \x. \y2 + ba * y + c < 0 " by simp then obtain y2 where y2_prop: "case q of (a, ba, c) \ (\y2 + ba * y + c < 0)" by auto have "\a\set les. case a of (a, ba, c) \ \x. \y2 + ba * y + c < 0" using ind by simp then have " \y. \xa\set les. case a of (a, ba, c) \ a * x\<^sup>2 + ba * x + c < 0" using Cons.hyps by blast then obtain y1 where y1_prop: "\xa\set les. case a of (a, ba, c) \ a * x^2 + ba * x + c < 0" by blast let ?y = "min y1 y2" have "\x < ?y. (\a\set (q #les). case a of (a, ba, c) \ a * x^2 + ba * x + c < 0)" using y1_prop y2_prop by fastforce then show ?case by blast qed qed lemma h1b_helper_leq: "(\((a::real), (b::real), (c::real))\set leq. \x. \y2 + b * y + c \ 0) \ (\y.\x(a, b, c)\set leq. a * x\<^sup>2 + b * x + c \ 0))" proof - show "(\(a, b, c)\set leq. \x. \y2 + b * y + c \ 0) \ (\y.\x(a, b, c)\set leq. a * x\<^sup>2 + b * x + c \ 0))" proof (induct leq) case Nil then show ?case by auto next case (Cons q leq) have ind: " \a\set (q # leq). case a of (a, ba, c) \ \x. \y2 + ba * y + c \ 0" using Cons.prems by auto then have "case q of (a, ba, c) \ \x. \y2 + ba * y + c \ 0 " by simp then obtain y2 where y2_prop: "case q of (a, ba, c) \ (\y2 + ba * y + c \ 0)" by auto have "\a\set leq. case a of (a, ba, c) \ \x. \y2 + ba * y + c \ 0" using ind by simp then have " \y. \xa\set leq. case a of (a, ba, c) \ a * x\<^sup>2 + ba * x + c \ 0" using Cons.hyps by blast then obtain y1 where y1_prop: "\xa\set leq. case a of (a, ba, c) \ a * x^2 + ba * x + c \ 0" by blast let ?y = "min y1 y2" have "\x < ?y. (\a\set (q #leq). case a of (a, ba, c) \ a * x^2 + ba * x + c \ 0)" using y1_prop y2_prop by fastforce then show ?case by blast qed qed lemma h1b_helper_neq: "(\((a::real), (b::real), (c::real))\set neq. \x. \y2 + b * y + c \ 0) \ (\y.\x(a, b, c)\set neq. a * x\<^sup>2 + b * x + c \ 0))" proof - show "(\(a, b, c)\set neq. \x. \y2 + b * y + c \ 0) \ (\y.\x(a, b, c)\set neq. a * x\<^sup>2 + b * x + c \ 0))" proof (induct neq) case Nil then show ?case by auto next case (Cons q neq) have ind: " \a\set (q # neq). case a of (a, ba, c) \ \x. \y2 + ba * y + c \ 0" using Cons.prems by auto then have "case q of (a, ba, c) \ \x. \y2 + ba * y + c \ 0 " by simp then obtain y2 where y2_prop: "case q of (a, ba, c) \ (\y2 + ba * y + c \ 0)" by auto have "\a\set neq. case a of (a, ba, c) \ \x. \y2 + ba * y + c \ 0" using ind by simp then have " \y. \xa\set neq. case a of (a, ba, c) \ a * x\<^sup>2 + ba * x + c \ 0" using Cons.hyps by blast then obtain y1 where y1_prop: "\xa\set neq. case a of (a, ba, c) \ a * x^2 + ba * x + c \ 0" by blast let ?y = "min y1 y2" have "\x < ?y. (\a\set (q #neq). case a of (a, ba, c) \ a * x^2 + ba * x + c \ 0)" using y1_prop y2_prop by fastforce then show ?case by blast qed qed lemma min_lem: fixes r::"real" assumes a1: "(\y'>r. (\((d::real), (e::real), (f::real))\set b. \x\{r<..y'}. d * x\<^sup>2 + e * x + f < 0))" assumes a2: "(\y'>r. (\((d::real), (e::real), (f::real))\set c. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0))" assumes a3: "(\y'>r. (\((d::real), (e::real), (f::real))\set d. \x\{r<..y'}. d * x\<^sup>2 + e * x + f \ 0))" shows "(\x. (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0))" proof - obtain y1 where y1_prop: "y1 > r \ (\(d, e, f)\set b. \x\{r<..y1}. d * x\<^sup>2 + e * x + f < 0)" using a1 by auto obtain y2 where y2_prop: "y2 > r \ (\(d, e, f)\set c. \x\{r<..y2}. d * x\<^sup>2 + e * x + f \ 0)" using a2 by auto obtain y3 where y3_prop: "y3 > r \ (\(d, e, f)\set d. \x\{r<..y3}. d * x\<^sup>2 + e * x + f \ 0)" using a3 by auto let ?y = "(min (min y1 y2) y3)" have "?y > r" using y1_prop y2_prop y3_prop by auto then have "\x. x > r \ x < ?y" using dense[of r ?y] by auto then obtain x where x_prop: "x > r \ x < ?y" by auto have bp: "(\(a, b, c)\set b. a *x\<^sup>2 + b * x + c < 0)" using x_prop y1_prop by auto have cp: "(\(a, b, c)\set c. a * x^2 + b * x + c \ 0)" using x_prop y2_prop by auto have dp: "(\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0)" using x_prop y3_prop by auto then have "(\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0)" using bp cp dp by auto then show ?thesis by auto qed lemma qe_infinitesimals_helper: fixes k::"real" assumes asm: "(\(d, e, f)\set a. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f \ 0)" shows "(\x. (\(a, b, c)\set a. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0))" proof - have "\(d, e, f)\set a. d = 0 \ e = 0 \ f = 0" proof clarsimp fix d e f assume "(d, e, f) \ set a" then have "\y'>k. \x\{k<..y'}. d * x\<^sup>2 + e * x + f = 0" using asm by auto then obtain y' where y_prop: "y'>k \ (\x\{k<..y'}. d * x\<^sup>2 + e * x + f = 0)" by auto then show "d = 0 \ e = 0 \ f = 0" using continuity_lem_eq0[of "k" "y'" d e f] by auto qed then have eqprop: "\x. (\(a, b, c)\set a. a * x\<^sup>2 + b * x + c = 0) " by auto have lesprop: "(\y'>k. (\(d, e, f)\set b. \x\{k<..y'}. d * x\<^sup>2 + e * x + f < 0))" using les_qe_inf_helper[of b "k"] asm by blast have leqprop: "(\y'>k. (\(d, e, f)\set c. \x\{(k)<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using leq_qe_inf_helper[of c "k"] asm by blast have neqprop: "(\y'>(k). (\(d, e, f)\set d. \x\{(k)<..y'}. d * x\<^sup>2 + e * x + f \ 0))" using neq_qe_inf_helper[of d "k"] asm by blast then have "(\x. (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0)) " using lesprop leqprop neqprop min_lem[of "k" b c d] by auto then show ?thesis using eqprop by auto qed subsubsection "The qe\\_backwards lemma" lemma qe_backwards: assumes "(((\(a, b, c)\set a. a = 0 \ b = 0 \ c = 0) \ (\(a, b, c)\set b. \x. \y2 + b * y + c < 0) \ (\(a, b, c)\set c. \x. \y2 + b * y + c \ 0) \ (\(a, b, c)\set d. \x. \y2 + b * y + c \ 0) \ (\(a', b', c')\set a. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))) \ (\(a', b', c')\set b. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))) \ (\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))) \ (\(a', b', c')\set d. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))))" shows " (\x. (\(a, b, c)\set a. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0))" proof - let ?e2 = "(((\(a, b, c)\set a. a = 0 \ b = 0 \ c = 0) \ (\(a, b, c)\set b. \x. \y2 + b * y + c < 0) \ (\(a, b, c)\set c. \x. \y2 + b * y + c \ 0) \ (\(a, b, c)\set d. \x. \y2 + b * y + c \ 0) \ (\(a', b', c')\set a. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))) \ (\(a', b', c')\set b. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))) \ (\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))) \ (\(a', b', c')\set d. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))))" let ?f10orf11orf12 = "(\(a', b', c')\set d. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" let ?f8orf9 = "(\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)))" let ?f5orf6orf7 = "(\(a', b', c')\set b. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" let ?f2orf3orf4 = "(\(a', b', c')\set a. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)))" let ?e1 = "(\x. (\(a, b, c)\set a. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0))" let ?f1 = "((\(a, b, c)\set a. a = 0 \ b = 0 \ c = 0) \ (\(a, b, c)\set b. \x. \y2 + b * y + c < 0) \ (\(a, b, c)\set c. \x. \y2 + b * y + c \ 0) \ (\(a, b, c)\set d. \x. \y2 + b * y + c \ 0))" let ?f2 = "(\(a', b', c')\set a. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0))" let ?f3 = "(\(a', b', c')\set a. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))" let ?f4 = "(\(a', b', c')\set a. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)) " let ?f5 = "(\(a', b', c')\set b. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0))" let ?f6 = "(\(a', b', c')\set b. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" let ?f7 = "(\(a', b', c')\set b. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))" let ?f8 = "(\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0))" let ?f13 = "(\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)))" let ?f9 = "(\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))" let ?f10 = "(\(a', b', c')\set d. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0))" let ?f11 = "(\(a', b', c')\set d. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" let ?f12 = "(\(a', b', c')\set d. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))" have h1a: "?e2 \ (?f1 \ ?f2orf3orf4 \ ?f5orf6orf7 \ ?f8orf9 \ ?f10orf11orf12)" by auto have h2: "?f2orf3orf4 \ (?f2 \ ?f3 \ ?f4)" by auto then have h1b: "?e2 \ (?f1 \ ?f2 \ ?f3 \ ?f4 \ ?f5orf6orf7 \ ?f8orf9 \ ?f10orf11orf12) " using h1a by auto have h3: "?f5orf6orf7 \ (?f5 \ ?f6 \ ?f7)" by auto then have h1c: "?e2 \ (?f1 \ ?f2 \ ?f3 \ ?f4 \ ?f5 \ ?f6 \ ?f7 \ ?f8orf9 \ ?f10orf11orf12) " using h1b by smt have h4: "?f8orf9 \ (?f8 \ ?f9 \ ?f13)" by auto then have h1d: "?e2 \ (?f1 \ ?f2 \ ?f3 \ ?f4 \ ?f5 \ ?f6 \ ?f7 \ ?f8 \ ?f9 \ ?f13 \ ?f10orf11orf12) " using h1c by smt have h5: "?f10orf11orf12 \ (?f10 \ ?f11 \ ?f12)" by auto then have bigor: "?e2 \ (?f1 \ ?f2 \ ?f3 \ ?f4 \ ?f5 \ ?f6 \ ?f7 \ ?f8 \ ?f13 \ ?f9 \ ?f10 \ ?f11 \ ?f12) " using h1d by smt have "?f1 \ ?e1" proof - assume asm: "(\(a, b, c)\set a. a = 0 \ b = 0 \ c = 0) \ (\(a, b, c)\set b. \x. \y2 + b * y + c < 0) \ (\(a, b, c)\set c. \x. \y2 + b * y + c \ 0) \ (\(a, b, c)\set d. \x. \y2 + b * y + c \ 0)" then have eqprop: "\x. \(a, b, c)\set a. a * x\<^sup>2 + b * x + c = 0" by auto have "\y. \x(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0" using asm h1b_helper_les by auto then obtain y1 where y1_prop: "\x(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0" by auto have "\y. \x(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0" using asm h1b_helper_leq by auto then obtain y2 where y2_prop: "\x(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0" by auto have "\y. \x(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0" using asm h1b_helper_neq by auto then obtain y3 where y3_prop: "\x(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0" by auto let ?y = "(min (min y1 y2) y3) - 1" have y_prop: "?y < y1 \ ?y < y2 \ ?y < y3" by auto have ap: "(\(a, b, c)\set a. a * ?y\<^sup>2 + b * ?y + c = 0)" using eqprop by auto have bp: "(\(a, b, c)\set b. a * ?y\<^sup>2 + b * ?y + c < 0)" using y_prop y1_prop by auto have cp: "(\(a, b, c)\set c. a * ?y\<^sup>2 + b * ?y + c \ 0)" using y_prop y2_prop by auto have dp: "(\(a, b, c)\set d. a * ?y\<^sup>2 + b * ?y + c \ 0)" using y_prop y3_prop by auto then have "(\(a, b, c)\set a. a * ?y\<^sup>2 + b * ?y + c = 0) \ (\(a, b, c)\set b. a * ?y\<^sup>2 + b * ?y + c < 0) \ (\(a, b, c)\set c. a * ?y\<^sup>2 + b * ?y + c \ 0) \ (\(a, b, c)\set d. a * ?y\<^sup>2 + b * ?y + c \ 0)" using ap bp cp dp by auto then show ?thesis by auto qed then have h1: "?f1 \ ?e1" by auto have "?f2 \ ?e1" proof - assume " \(a', b', c')\set a. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0)" then obtain a' b' c' where abc_prop: "(a', b', c')\set a \ (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0)" by auto then have "\(x::real). x = -c'/b'" by auto then obtain x where x_prop: "x = - c' / b'" by auto then have "(\xa\set a. case xa of (a, b, c) \ a * x\<^sup>2 + b * x + c = 0) \ (\xa\set b. case xa of (a, b, c) \ a * x\<^sup>2 + b * x + c < 0) \ (\xa\set c. case xa of (a, b, c) \ a * x\<^sup>2 + b * x + c \ 0) \ (\xa\set d. case xa of (a, b, c) \ a * x\<^sup>2 + b * x + c \ 0)" using abc_prop by auto then show ?thesis by auto qed then have h2: "?f2 \ ?e1" by auto have "?f3 \ ?e1" proof - assume "\(a', b', c')\set a. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)" then obtain a' b' c' where abc_prop: "(a', b', c')\set a \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)" by auto then have "\(x::real). x = (- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')" by auto then obtain x where x_prop: " x = (- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')" by auto then have "(\(a, b, c)\set a. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0)" using abc_prop by auto then show ?thesis by auto qed then have h3: "?f3 \ ?e1" by auto have "?f4 \ ?e1" proof - assume " \(a', b', c')\set a. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)" then obtain a' b' c' where abc_prop: "(a', b', c')\set a \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)" by auto then have "\(x::real). x = (- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')" by auto then obtain x where x_prop: " x = (- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')" by auto then have "(\(a, b, c)\set a. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0)" using abc_prop by auto then show ?thesis by auto qed then have "?f4 \ ?e1" by auto have "?f5 \ ?e1" proof - assume asm: "\(a', b', c')\set b. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0)" then obtain a' b' c' where abc_prop: "(a', b', c')\set b \ (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0)" by auto then show ?thesis using qe_infinitesimals_helper[of a "- c' / b'" b c d] by auto qed then have h5: "?f5 \ ?e1" by auto have "?f6 \ ?e1" proof - assume "\(a', b', c')\set b. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)" then obtain a' b' c' where abc_prop: "(a', b', c')\set b \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)" by auto then show ?thesis using qe_infinitesimals_helper[of a "(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')" b c d] by auto qed then have h6: "?f6 \ ?e1" by auto have "?f7 \ ?e1" proof - assume "\(a', b', c')\set b. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)" then obtain a' b' c' where abc_prop: "(a', b', c')\set b \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)" by auto then show ?thesis using qe_infinitesimals_helper[of a "(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')" b c d] by auto qed then have h7: "?f7 \ ?e1" by auto have "?f8 \ ?e1" proof - assume "\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0)" then obtain a' b' c' where abc_prop: "(a', b', c')\set c \ (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0)" by auto then have "\(x::real). x = (- c' / b')" by auto then obtain x where x_prop: " x = - c' / b'" by auto then have "(\(a, b, c)\set a. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0)" using abc_prop by auto then show ?thesis by auto qed then have h8: "?f8 \ ?e1" by auto have "?f9 \ ?e1" proof - assume "\(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)" then obtain a' b' c' where abc_prop: "(a', b', c')\set c \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)" by auto then have "\(x::real). x = (- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')" by auto then obtain x where x_prop: " x = (- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')" by auto then have "(\(a, b, c)\set a. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0)" using abc_prop by auto then show ?thesis by auto qed then have h9: "?f9 \ ?e1" by auto have "?f10 \ ?e1" proof - assume asm: "\(a', b', c')\set d. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0)" then obtain a' b' c' where abc_prop: "(a', b', c')\set d \ (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0)" by auto then show ?thesis using qe_infinitesimals_helper[of a "- c' / b'" b c d] by auto qed then have h10: "?f10 \ ?e1" by auto have "?f11 \ ?e1" proof - assume "\(a', b', c')\set d. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)" then obtain a' b' c' where abc_prop: "(a', b', c')\set d \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)" by auto then show ?thesis using qe_infinitesimals_helper[of a "(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')" b c d] by auto qed then have h11: "?f11 \ ?e1" by auto have "?f12 \ ?e1" proof - assume "\(a', b', c')\set d. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)" then obtain a' b' c' where abc_prop: "(a', b', c')\set d \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)" by auto then show ?thesis using qe_infinitesimals_helper[of a "(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')" b c d] by auto qed then have h12: "?f12 \ ?e1" by auto have "?f13 \ ?e1" proof - assume " \(a', b', c')\set c. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)" then obtain a' b' c' where abc_prop: "(a', b', c')\set c \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)" by auto then have "\(x::real). x = (- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')" by auto then obtain x where x_prop: " x = (- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')" by auto then have "(\(a, b, c)\set a. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0)" using abc_prop by auto then show ?thesis by auto qed then have h13: "?f13 \ ?e1" by auto show ?thesis using bigor h1 h2 h3 h4 h5 h6 h7 h8 h9 h10 h11 h12 h13 using assms by (smt \\(a', b', c')\set a. a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ \x. (\(a, b, c)\set a. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0)\) (* by force *) qed subsection "General QE lemmas" lemma qe: "(\x. (\(a, b, c)\set a. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0)) = ((\(a, b, c)\set a. a = 0 \ b = 0 \ c = 0) \ (\(a, b, c)\set b. \x. \y2 + b * y + c < 0) \ (\(a, b, c)\set c. \x. \y2 + b * y + c \ 0) \ (\(a, b, c)\set d. \x. \y2 + b * y + c \ 0) \ (\(a', b', c')\set a. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))) \ (\(a', b', c')\set b. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))) \ (\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))) \ (\(a', b', c')\set d. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))))" proof - let ?e1 = "((\(a, b, c)\set a. a = 0 \ b = 0 \ c = 0) \ (\(a, b, c)\set b. \x. \y2 + b * y + c < 0) \ (\(a, b, c)\set c. \x. \y2 + b * y + c \ 0) \ (\(a, b, c)\set d. \x. \y2 + b * y + c \ 0) \ (\(a', b', c')\set a. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))) \ (\(a', b', c')\set b. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))) \ (\(a', b', c')\set c. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\(d, e, f)\set b. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\(d, e, f)\set c. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\(d, e, f)\set d. d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set a. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\(d, e, f)\set b. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\(d, e, f)\set c. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\(d, e, f)\set d. d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0))) \ (\(a', b', c')\set d. (a' = 0 \ b' \ 0) \ (\(d, e, f)\set a. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set a. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set a. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set b. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set c. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set d. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0))))" let ?e2 = "(\x. (\(a, b, c)\set a. a * x\<^sup>2 + b * x + c = 0) \ (\(a, b, c)\set b. a * x\<^sup>2 + b * x + c < 0) \ (\(a, b, c)\set c. a * x\<^sup>2 + b * x + c \ 0) \ (\(a, b, c)\set d. a * x\<^sup>2 + b * x + c \ 0))" have h1: "?e1 \ ?e2" using qe_backwards by auto have h2: "?e2 \ ?e1" using qe_forwards by auto have "(?e2 \ ?e1) \ (?e1 \ ?e2) " using h1 h2 by force then have "?e2 \ ?e1" using iff_conv_conj_imp[of ?e1 ?e2] by presburger then show ?thesis by auto qed fun eq_fun :: "real \ real \ real \ (real*real*real) list \ (real*real*real) list \ (real*real*real) list \ (real*real*real) list \ bool" where "eq_fun a' b' c' eq les leq neq = ((a' = 0 \ b' \ 0) \ (\a\set eq. case a of (d, e, f) \ d * (- c' / b')\<^sup>2 + e * (- c' / b') + f = 0) \ (\a\set les. case a of (d, e, f) \ d * (- c' / b')\<^sup>2 + e * (- c' / b') + f < 0) \ (\a\set leq. case a of (d, e, f) \ d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ (\a\set neq. case a of (d, e, f) \ d * (- c' / b')\<^sup>2 + e * (- c' / b') + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\a\set eq. case a of (d, e, f) \ d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\a\set les. case a of (d, e, f) \ d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\a\set leq. case a of (d, e, f) \ d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\a\set neq. case a of (d, e, f) \ d * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\a\set eq. case a of (d, e, f) \ d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f = 0) \ (\a\set les. case a of (d, e, f) \ d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f < 0) \ (\a\set leq. case a of (d, e, f) \ d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0) \ (\a\set neq. case a of (d, e, f) \ d * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'))\<^sup>2 + e * ((- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')) + f \ 0)))" fun les_fun :: "real \ real \ real \ (real*real*real) list \ (real*real*real) list \ (real*real*real) list \ (real*real*real) list \ bool" where "les_fun a' b' c' eq les leq neq = ((a' = 0 \ b' \ 0) \ (\(d, e, f)\set eq. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set les. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set leq. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set neq. \y'>- c' / b'. \x\{- c' / b'<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ a' \ 0 \ - b'\<^sup>2 + 4 * a' * c' \ 0 \ ((\(d, e, f)\set eq. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set les. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set leq. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set neq. \y'>(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set eq. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f = 0) \ (\(d, e, f)\set les. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f < 0) \ (\(d, e, f)\set leq. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0) \ (\(d, e, f)\set neq. \y'>(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a'). \x\{(- b' + - 1 * sqrt (b'\<^sup>2 - 4 * a' * c')) / (2 * a')<..y'}. d * x\<^sup>2 + e * x + f \ 0)))" lemma general_qe' : (* Direct substitution F(x) *) assumes "F = (\x. (\(a,b,c)\set eq . a*x\<^sup>2+b*x+c=0)\ (\(a,b,c)\set les. a*x\<^sup>2+b*x+c<0)\ (\(a,b,c)\set leq. a*x\<^sup>2+b*x+c\0)\ (\(a,b,c)\set neq. a*x\<^sup>2+b*x+c\0))" (* Substitution of r+\ into F *) assumes "F\ = (\r. (\(a,b,c)\set eq. \y>r.\x\{r<..y}. a*x\<^sup>2+b*x+c=0) \ (\(a,b,c)\set les. \y>r.\x\{r<..y}. a*x\<^sup>2+b*x+c<0) \ (\(a,b,c)\set leq. \y>r.\x\{r<..y}. a*x\<^sup>2+b*x+c\0) \ (\(a,b,c)\set neq. \y>r.\x\{r<..y}. a*x\<^sup>2+b*x+c\0) )" (* Substitution of -\ into F *) assumes "F\<^sub>i\<^sub>n\<^sub>f = ( (\(a,b,c)\set eq. \x. \y2+b*y+c=0) \ (\(a,b,c)\set les. \x. \y2+b*y+c<0) \ (\(a,b,c)\set leq. \x. \y2+b*y+c\0) \ (\(a,b,c)\set neq. \x. \y2+b*y+c\0) )" (* finds linear or quadratic roots of a polynomial *) assumes "roots = (\(a,b,c). if a=0 \ b\0 then {-c/b}::real set else if a\0 \ b\<^sup>2-4*a*c\0 then {(-b+sqrt(b\<^sup>2-4*a*c))/(2*a)}\{(-b-sqrt(b\<^sup>2-4*a*c))/(2*a)} else {})" (* all the root of each atom *) assumes "A = \(roots ` (set eq))" assumes "B = \(roots ` (set les))" assumes "C = \(roots ` (set leq))" assumes "D = \(roots ` (set neq))" (* Quantifier Elimination *) shows "(\x. F(x)) = (F\<^sub>i\<^sub>n\<^sub>f\(\r\A. F r)\(\r\B. F\ r)\(\r\C. F r)\(\r\D. F\ r))" proof- { fix X have "(\(a, b, c)\set X. eq_fun a b c eq les leq neq) = (\x\F ` \(roots ` (set X)). x)" proof(induction X) case Nil then show ?case by auto next case (Cons p X) have h1: "(\x\F ` \ (roots ` set (p # X)). x) = ((\x\F ` roots p. x) \ (\x\F ` \ (roots ` set X). x))" by auto have h2 :"(case p of (a,b,c) \ eq_fun a b c eq les leq neq) = (\x\F ` roots p. x)" apply(cases p) unfolding assms apply simp by linarith show ?case unfolding h1 Cons[symmetric] using h2 by auto qed } then have eq : "\X. (\(a, b, c)\set X. eq_fun a b c eq les leq neq) = (\x\F ` \ (roots ` set X). x)" by auto { fix X have "(\(a, b, c)\set X. les_fun a b c eq les leq neq) = (\x\F\ ` \(roots ` (set X)). x)" proof(induction X) case Nil then show ?case by auto next case (Cons p X) have h1: "(\x\F\ ` \ (roots ` set (p # X)). x) = ((\x\F\ ` roots p. x) \ (\x\F\ ` \ (roots ` set X). x))" by auto have h2 :"(case p of (a,b,c) \ les_fun a b c eq les leq neq) = (\x\F\ ` roots p. x)" apply(cases p) unfolding assms apply simp by linarith show ?case unfolding h1 Cons[symmetric] using h2 by auto qed } then have les : "\X. (\(a, b, c)\set X. les_fun a b c eq les leq neq) = (\x\F\ ` \ (roots ` set X). x)" by auto have inf : "(\(a, b, c)\set eq. a = 0 \ b = 0 \ c = 0) = (\x\set eq. case x of (a, b, c) \ \x. \y2 + b * y + c = 0)" proof(induction eq) case Nil then show ?case by auto next case (Cons p eq) then show ?case proof(cases p) case (fields a b c) show ?thesis unfolding fields using infzeros[of _ a b c] Cons by auto qed qed show ?thesis using qe[of "eq" "les" "leq" "neq"] using eq[of eq] eq[of leq] les[of les] les[of neq] unfolding inf assms by auto qed lemma general_qe'' : (* Direct substitution F(x) *) assumes "S = {(=), (<), (\), (\)}" assumes "finite (X (=))" assumes "finite (X (<))" assumes "finite (X (\))" assumes "finite (X (\))" assumes "F = (\x. \rel\S. \(a,b,c)\(X rel). rel (a*x\<^sup>2+b*x+c) 0)" (* Substitution of r+\ into F *) assumes "F\ = (\r. \rel\S. \(a,b,c)\(X rel). \y>r.\x\{r<..y}. rel (a*x\<^sup>2+b*x+c) 0)" (* Substitution of -\ into F *) assumes "F\<^sub>i\<^sub>n\<^sub>f = (\rel\S. \(a,b,c)\(X rel). \x. \y2+b*y+c) 0)" (* finds linear or quadratic roots of a polynomial *) assumes "roots = (\(a,b,c). if a=0 \ b\0 then {-c/b}::real set else if a\0 \ b\<^sup>2-4*a*c\0 then {(-b+sqrt(b\<^sup>2-4*a*c))/(2*a)}\{(-b-sqrt(b\<^sup>2-4*a*c))/(2*a)} else {})" (* all the root of each atom *) assumes "A = \(roots ` ((X (=))))" assumes "B = \(roots ` ((X (<))))" assumes "C = \(roots ` ((X (\))))" assumes "D = \(roots ` ((X (\))))" (* Quantifier Elimination *) shows "(\x. F(x)) = (F\<^sub>i\<^sub>n\<^sub>f\(\r\A. F r)\(\r\B. F\ r)\(\r\C. F r)\(\r\D. F\ r))" proof- define less where "less = (\(a::real,b::real,c::real).\(a',b',c'). a (a=a'\ (b(b=b'\cx.\y. x=y \ less x y)" have linorder: "class.linorder leq less" unfolding class.linorder_def class.order_def class.preorder_def class.order_axioms_def class.linorder_axioms_def less_def leq_def by auto show ?thesis using assms(6-8) unfolding assms(1) apply simp using general_qe'[OF _ _ _ assms(9), of F "List.linorder.sorted_list_of_set leq (X (=))" "List.linorder.sorted_list_of_set leq (X (<))" "List.linorder.sorted_list_of_set leq (X (\))" "List.linorder.sorted_list_of_set leq (X (\))" F\ F\<^sub>i\<^sub>n\<^sub>f A B C D] unfolding List.linorder.set_sorted_list_of_set[OF linorder assms(2)] List.linorder.set_sorted_list_of_set[OF linorder assms(3)] List.linorder.set_sorted_list_of_set[OF linorder assms(4)] List.linorder.set_sorted_list_of_set[OF linorder assms(5)] using assms(10-13)by auto qed theorem general_qe : (* finite sets of atoms involving = < \ and \*) assumes "R = {(=), (<), (\), (\)}" assumes "\rel\R. finite (Atoms rel)" (* Direct substitution F(x) *) assumes "F = (\x. \rel\R. \(a,b,c)\(Atoms rel). rel (a*x\<^sup>2+b*x+c) 0)" (* Substitution of r+\ into F *) assumes "F\ = (\r. \rel\R. \(a,b,c)\(Atoms rel). \y>r.\x\{r<..y}. rel (a*x\<^sup>2+b*x+c) 0)" (* Substitution of -\ into F *) assumes "F\<^sub>i\<^sub>n\<^sub>f = (\rel\R. \(a,b,c)\(Atoms rel). \x. \y2+b*y+c) 0)" (* finds linear or quadratic roots of a polynomial *) assumes "roots = (\(a,b,c). if a=0 \ b\0 then {-c/b} else if a\0 \ b\<^sup>2-4*a*c\0 then {(-b+sqrt(b\<^sup>2-4*a*c))/(2*a)}\{(-b-sqrt(b\<^sup>2-4*a*c))/(2*a)} else {})" (* Quantifier Elimination *) shows "(\x. F(x)) = (F\<^sub>i\<^sub>n\<^sub>f \ (\r\\(roots ` (Atoms (=) \ Atoms (\))). F r) \ (\r\\(roots ` (Atoms (<) \ Atoms (\))). F\ r))" using general_qe''[OF assms(1) _ _ _ _ assms(3-6) refl refl refl refl] using assms(2) unfolding assms(1) by auto end \ No newline at end of file diff --git a/thys/Virtual_Substitution/QuadraticCase.thy b/thys/Virtual_Substitution/QuadraticCase.thy --- a/thys/Virtual_Substitution/QuadraticCase.thy +++ b/thys/Virtual_Substitution/QuadraticCase.thy @@ -1,969 +1,947 @@ subsection "Quadratic Case" theory QuadraticCase imports VSAlgos begin (*-------------------------------------------------------------------------------------------------------------*) lemma quad_part_1_eq : assumes lLength : "length L > var" assumes hdeg : "MPoly_Type.degree (p::real mpoly) var = (deg::nat)" assumes nonzero : "D \ 0" assumes ha : "\x. insertion (nth_default 0 (list_update L var x)) a = (A::real)" assumes hb : "\x. insertion (nth_default 0 (list_update L var x)) b = (B::real)" assumes hd : "\x. insertion (nth_default 0 (list_update L var x)) d = (D::real)" shows "aEval (Eq p) (list_update L var ((A+B*C)/D)) = aEval (Eq(quadratic_part_1 var a b d (Eq p))) (list_update L var C)" proof - define f where "f i = insertion (nth_default 0 (list_update L var C)) (isolate_variable_sparse p var i) * ((A + B * C) ^ i)" for i have h1 : "\i. (insertion (nth_default 0 (list_update L var C)) (isolate_variable_sparse p var i)) = (insertion (nth_default 0 (list_update L var ((A+B*C)/D))) (isolate_variable_sparse p var i))" by(simp add: insertion_isovarspars_free) have h2 : "((\i = 0..i = 0..i = 0..i = 0..i = 0..i = 0..i = 0.. var" assumes hdeg : "MPoly_Type.degree (p::real mpoly) var = (deg::nat)" assumes nonzero : "D \ 0" assumes ha : "\x. insertion (nth_default 0 (list_update L var x)) a = (A::real)" assumes hb : "\x. insertion (nth_default 0 (list_update L var x)) b = (B::real)" assumes hd : "\x. insertion (nth_default 0 (list_update L var x)) d = (D::real)" shows "aEval (Less p) (list_update L var ((A+B*C)/D)) = aEval (Less(quadratic_part_1 var a b d (Less p))) (list_update L var C)" proof - define f where "f i = insertion (nth_default 0 (list_update L var C)) (isolate_variable_sparse p var i) * ((A + B * C) ^ i)" for i have h1a : "((\i = 0..i = 0..i. (insertion (nth_default 0 (list_update L var C)) (isolate_variable_sparse p var i)) = (insertion (nth_default 0 (list_update L var ((A+B*C)/D))) (isolate_variable_sparse p var i))" by(simp add: insertion_isovarspars_free) have "((\i = 0..i = 0..i = 0..i = 0..i = 0.. var" assumes hdeg : "MPoly_Type.degree (p::real mpoly) var = (deg::nat)" assumes nonzero : "D \ 0" assumes ha : "\x. insertion (nth_default 0 (list_update L var x)) a = (A::real)" assumes hb : "\x. insertion (nth_default 0 (list_update L var x)) b = (B::real)" assumes hd : "\x. insertion (nth_default 0 (list_update L var x)) d = (D::real)" shows "aEval (Leq p) (list_update L var ((A+B*C)/D)) = aEval (Leq(quadratic_part_1 var a b d (Leq p))) (list_update L var C)" proof - define f where "f i = insertion (nth_default 0 (list_update L var C)) (isolate_variable_sparse p var i) * ((A + B * C) ^ i)" for i have h1a : "((\i = 0..i = 0..i = 0..i = 0..i = 0.. 0) =((\i = 0.. 0)" using h1a h1b nonzero by auto have h4a : "\i. (insertion (nth_default 0 (list_update L var C)) (isolate_variable_sparse p var i)) = (insertion (nth_default 0 (list_update L var ((A+B*C)/D))) (isolate_variable_sparse p var i))" by(simp add: insertion_isovarspars_free) have "((\i = 0.. 0)= ((\i = 0.. 0)" using h1c f_def by auto also have "...=((\i = 0.. 0)" by auto also have "...=((\i = 0.. 0)" by (metis (no_types, lifting) power_divide sum.cong) also have "...=((\i = 0..0)" using h4a by auto also have "... = (insertion (nth_default 0 (list_update L var ((A+B*C)/D))) p\0)" using sum_over_degree_insertion hdeg lLength by auto finally show ?thesis by(simp add: hdeg lLength insertion_add insertion_mult ha hb hd insertion_sum insertion_pow insertion_var) qed (*------------------------------------------------------------------------------------------------*) lemma quad_part_1_neq : assumes lLength : "length L > var" assumes hdeg : "MPoly_Type.degree (p::real mpoly) var = (deg::nat)" assumes nonzero : "D \ 0" assumes ha : "\x. insertion (nth_default 0 (list_update L var x)) a = (A::real)" assumes hb : "\x. insertion (nth_default 0 (list_update L var x)) b = (B::real)" assumes hd : "\x. insertion (nth_default 0 (list_update L var x)) d = (D::real)" shows "aEval (Neq p) (list_update L var ((A+B*C)/D)) = aEval (Neq(quadratic_part_1 var a b d (Neq p))) (list_update L var C)" proof - have "aEval (Eq(quadratic_part_1 var a b d (Eq p))) (list_update L var C) = aEval (Eq p) (list_update L var ((A+B*C)/D))" using quad_part_1_eq assms by blast then show ?thesis by auto qed (*------------------------------------------------------------------------------------------------*) lemma sqrt_case : assumes detGreater0 : "SQ \ 0" shows "((SQ^(i div 2)) * real (i mod 2) * sqrt SQ + SQ ^ (i div 2) * (1 - real (i mod 2))) = (sqrt SQ) ^ i" proof - have h1 : "i mod 2 = 0 \ (odd i \ (i mod 2 = 1))" by auto have h2 : "i mod 2 = 0 \ ((SQ^(i div 2)) * real (i mod 2) * sqrt SQ + SQ ^ (i div 2) * (1 - real (i mod 2))) = (sqrt SQ) ^ i" using detGreater0 apply auto by (simp add: real_sqrt_power_even) have h3 : "(odd i \ (i mod 2 = 1)) \ ((SQ^(i div 2)) * real (i mod 2) * sqrt SQ + SQ ^ (i div 2) * (1 - real (i mod 2))) = (sqrt SQ) ^ i" using detGreater0 apply auto by (smt One_nat_def add_Suc_right mult.commute nat_arith.rule0 odd_two_times_div_two_succ power.simps(2) power_mult real_sqrt_pow2) show ?thesis using h1 h2 h3 by linarith qed lemma sum_over_sqrt : assumes detGreater0 : "SQ \ 0" shows "(\i\{0..i\{0.. var" assumes detGreater0 : "SQ\0" assumes hdeg : "MPoly_Type.degree (p::real mpoly) var = (deg ::nat)" assumes hsq : "\x. insertion (nth_default 0 (list_update L var x)) sq = (SQ::real)" shows "aEval (Eq p) (list_update L var (sqrt SQ)) = aEval (Eq(quadratic_part_2 var sq p)) (list_update L var (sqrt SQ))" proof - define f where "f i = insertion (nth_default 0 (list_update L var (sqrt SQ))) (isolate_variable_sparse p var i)" for i have h1a : "(\i\{0..i\{0..i\{0..i\{0.. var" assumes detGreater0 : "SQ\0" assumes hdeg : "MPoly_Type.degree (p::real mpoly) var = (deg ::nat)" assumes hsq : "\x. insertion (nth_default 0 (list_update L var x)) sq = (SQ::real)" shows "aEval (Less p) (list_update L var (sqrt SQ)) = aEval (Less(quadratic_part_2 var sq p)) (list_update L var (sqrt SQ))" proof - define f where "f i = insertion (nth_default 0 (list_update L var (sqrt SQ))) (isolate_variable_sparse p var i)" for i have h1a : "(\i\{0..i\{0..i\{0..i\{0.. var" assumes detGreater0 : "SQ\0" assumes hdeg : "MPoly_Type.degree (p::real mpoly) var = (deg ::nat)" assumes hsq : "\x. insertion (nth_default 0 (list_update L var x)) sq = (SQ::real)" shows "aEval (Neq p) (list_update L var (sqrt SQ)) = aEval (Neq(quadratic_part_2 var sq p)) (list_update L var (sqrt SQ))" proof - define f where "f i = insertion (nth_default 0 (list_update L var (sqrt SQ))) (isolate_variable_sparse p var i)" for i have h1a : "(\i\{0..i\{0..i\{0..i\{0.. var" assumes detGreater0 : "SQ\0" assumes hdeg : "MPoly_Type.degree (p::real mpoly) var = (deg ::nat)" assumes hsq : "\x. insertion (nth_default 0 (list_update L var x)) sq = (SQ::real)" shows "aEval (Leq p) (list_update L var (sqrt SQ)) = aEval (Leq(quadratic_part_2 var sq p)) (list_update L var (sqrt SQ))" proof - define f where "f i = insertion (nth_default 0 (list_update L var (sqrt SQ))) (isolate_variable_sparse p var i)" for i have h1a : "(\i\{0..i\{0..i\{0..i\{0..vars(sq::real mpoly)" shows "MPoly_Type.degree (quadratic_part_2 var sq p) var \ 1" proof - define deg where "deg = MPoly_Type.degree (p::real mpoly) var" define f where "f i = isolate_variable_sparse p var i * sq ^ (i div 2) * Const (real (i mod 2)) * Var var" for i define g where "g i = isolate_variable_sparse p var i * sq ^ (i div 2) * Const (1 - real (i mod 2))" for i have h1a : "\i. MPoly_Type.degree (isolate_variable_sparse p var i) var = 0" by (simp add: varNotIn_degree not_in_isovarspar) have h1b : "\i. MPoly_Type.degree (sq ^ (i div 2)) var = 0" by (simp add: sqfree varNotIn_degree not_in_pow) have h1c : "\i. MPoly_Type.degree (Const (real (i mod 2))) var = 0" using degree_const by blast have h1d : "MPoly_Type.degree (Var var :: real mpoly) var = 1" using degree_one by auto have h1 : "\i 1" using f_def degree_mult h1a h1b h1c h1d by (smt ExecutiblePolyProps.degree_one add.right_neutral mult.commute mult_eq_0_iff nat_le_linear not_one_le_zero) have h2a : "\i. MPoly_Type.degree (Const (1 - real (i mod 2))) var = 0" using degree_const by blast have h2 : "\ii 1" using h1 h2 by (simp add: degree_add_leq) show ?thesis using atLeastLessThanSuc_atLeastAtMost degree_sum f_def g_def h3 deg_def by auto qed (*------------------------------------------------------------------------------------------------*) lemma quad_equality_helper : assumes lLength : "length L > var" assumes detGreat0 : "Cv\0" assumes hC : "\x. insertion (nth_default 0 (list_update L var x)) (C::real mpoly) = (Cv::real)" assumes hA : "\x. insertion (nth_default 0 (list_update L var x)) (A::real mpoly) = (Av::real)" assumes hB : "\x. insertion (nth_default 0 (list_update L var x)) (B::real mpoly) = (Bv::real)" shows "aEval (Eq (A + B * Var var)) (list_update L var (sqrt Cv)) = eval (And (Atom(Leq (A*B))) (Atom (Eq (A^2-B^2*C)))) (list_update L var (sqrt Cv))" proof- have h1 : "\x. insertion (nth_default 0 (list_update L var x)) (A^2-(B^2)*C) = Av^2-(Bv^2)*Cv" by(simp add: hA hB hC insertion_add insertion_mult insertion_sub insertion_pow) have h2a : "(Av + Bv * sqrt Cv = 0) = (Av = - Bv * sqrt Cv)" by auto have h2b : "(Av = - Bv * sqrt Cv) \ (Av^2 = (- Bv * sqrt Cv)^2)" by simp have h2c : "(Av^2 = (- Bv * sqrt Cv)^2) = (Av^2 = Bv^2 * (sqrt Cv)^2)" by (simp add: power_mult_distrib) have h2d : "(Av^2 = Bv^2 * (sqrt Cv)^2) = (Av^2 = Bv^2 * Cv)" by (simp add: detGreat0) have h2 : "(Av + Bv * sqrt Cv = 0) \ (Av^2 = Bv^2 * Cv)" using h2a h2b h2c h2d by blast have h3a : "(Av*Bv > 0) \ (Av + Bv * sqrt Cv \ 0)" by (smt detGreat0 mult_nonneg_nonneg real_sqrt_ge_zero zero_less_mult_iff) have h3 : "(Av + Bv * sqrt Cv = 0) \ (Av*Bv\ 0)" using h3a by linarith have h4 : "(Av * Bv \ 0 \ Av\<^sup>2 = Bv\<^sup>2 * Cv) \ (Av + Bv * sqrt Cv = 0)" apply(cases "Av>0") apply (metis detGreat0 h2a h2c h2d mult_minus_left not_le power2_eq_iff real_sqrt_lt_0_iff zero_less_mult_iff) by (smt h2a real_sqrt_abs real_sqrt_mult zero_less_mult_iff) show ?thesis apply(simp add: hA hB h1 insertion_add insertion_mult insertion_var lLength) using h2 h3 h4 by blast qed lemma quadratic_sub_eq : assumes lLength : "length L > var" assumes nonzero : "Dv \ 0" assumes detGreater0 : "Cv \ 0" assumes freeC : "var \ vars c" assumes ha : "\x. insertion (nth_default 0 (list_update L var x)) (a::real mpoly) = (Av :: real)" assumes hb : "\x. insertion (nth_default 0 (list_update L var x)) (b::real mpoly) = (Bv :: real)" assumes hc : "\x. insertion (nth_default 0 (list_update L var x)) (c::real mpoly) = (Cv :: real)" assumes hd : "\x. insertion (nth_default 0 (list_update L var x)) (d::real mpoly) = (Dv :: real)" shows "aEval (Eq p) (list_update L var ((Av+Bv*sqrt(Cv))/Dv)) = eval (quadratic_sub var a b c d (Eq p)) (list_update L var (sqrt Cv))" proof - define p1 where "(p1::real mpoly) = quadratic_part_1 var a b d (Eq p)" define p2 where "(p2::real mpoly) = quadratic_part_2 var c p1" define A where "A = isolate_variable_sparse p2 var 0" define B where "B = isolate_variable_sparse p2 var 1" have h3c : "MPoly_Type.degree p2 var = 0 \ MPoly_Type.degree p2 var = 1" using freeC quad_part_2_deg p2_def by (meson le_neq_implies_less less_one) have h3d : "MPoly_Type.degree p2 var = 0 \ B = 0" by(simp add: B_def isovar_greater_degree) then have h3f : "MPoly_Type.degree p2 var = 0 \ p2 = A + B * Var var" by(simp add: h3d A_def degree0isovarspar) have h3g1 : "MPoly_Type.degree p2 var = 1 \ p2 = (\i\1. isolate_variable_sparse p2 var i * Var var ^ i)" using sum_over_zero by metis have h3g2a : "\f. (\i::nat\1. f i) = f 0 + f 1" by simp have h3g2 : "(\i::nat\1. isolate_variable_sparse p2 var i * Var var ^ i) = isolate_variable_sparse p2 var 0 * Var var ^ 0 + isolate_variable_sparse p2 var 1 * Var var ^ 1" using h3g2a by blast have h3g : "MPoly_Type.degree p2 var = 1 \ p2 = A + B * Var var" apply(simp add: sum_over_zero A_def B_def) using h3g1 h3g2 by (metis (no_types, lifting) One_nat_def mult_cancel_left2 power_0 power_one_right) have h3h : "p2 = A + B * Var var" using h3c h3f h3g by auto have h4a : "\x::real. \y::real. insertion (nth_default 0 (list_update L var y)) A = x" using not_contains_insertion not_in_isovarspar A_def by blast have h4b : "\x::real. \y::real. insertion (nth_default 0 (list_update L var y)) B = x" using not_contains_insertion not_in_isovarspar B_def by blast have "aEval (Eq p) (list_update L var ((Av+Bv*sqrt(Cv))/Dv)) = aEval (Eq p1) (list_update L var (sqrt Cv))" using p1_def quad_part_1_eq nonzero ha hb hd lLength by blast also have h2 : "... = aEval (Eq p2) (list_update L var (sqrt Cv))" using p2_def quad_part_2_eq lLength detGreater0 hc by metis also have "... = aEval (Eq (A + B * Var var)) (list_update L var (sqrt Cv))" using h3h by auto also have "... = eval (And (Atom(Leq (A*B))) (Atom (Eq (A^2-B^2*c)))) (list_update L var (sqrt Cv))" using quad_equality_helper hc detGreater0 h4a h4b lLength by blast also have "... = eval (quadratic_sub var a b c d (Eq p)) (list_update L var (sqrt Cv))" using p2_def A_def B_def p1_def quadratic_sub.simps(1) by metis finally show ?thesis by blast qed (*------------------------------------------------------------------------------------------------*) lemma quadratic_sub_less_helper : assumes lLength : "length L > var" assumes detGreat0 : "Cv\0" assumes hC : "\x. insertion (nth_default 0 (list_update L var x)) (C::real mpoly) = (Cv::real)" assumes hA : "\x. insertion (nth_default 0 (list_update L var x)) (A::real mpoly) = (Av::real)" assumes hB : "\x. insertion (nth_default 0 (list_update L var x)) (B::real mpoly) = (Bv::real)" shows "aEval (Less (A + B * Var var)) (list_update L var (sqrt Cv)) = eval (Or (And (fm.Atom (Less A)) (fm.Atom (Less (B\<^sup>2 * C - A\<^sup>2)))) (And (fm.Atom (Leq B)) (Or (fm.Atom (Less A)) (fm.Atom (Less (A\<^sup>2 - B\<^sup>2 * C)))))) (list_update L var (sqrt Cv)) " proof- have h1 : "\x. insertion (nth_default 0 (list_update L var x)) (A^2-(B^2)*C) = Av^2-(Bv^2)*Cv" by(simp add: hA hB hC insertion_add insertion_mult insertion_sub insertion_pow) have h2 : "\x. insertion (nth_default 0 (list_update L var x)) ((B^2)*C-A^2) = (Bv^2)*Cv-Av^2" by(simp add: hA hB hC insertion_add insertion_mult insertion_sub insertion_pow) have h3 : "Av=0 \ Bv=0 \ (Av + Bv * sqrt Cv < 0) = (Av < 0 \ Bv\<^sup>2 * Cv < Av\<^sup>2 \ Bv \ 0 \ (Av < 0 \ Av\<^sup>2 < Bv\<^sup>2 * Cv))" by simp have h4 : "Av<0 \ Bv\0 \ (Av + Bv * sqrt Cv < 0) = (Av < 0 \ Bv\<^sup>2 * Cv < Av\<^sup>2 \ Bv \ 0 \ (Av < 0 \ Av\<^sup>2 < Bv\<^sup>2 * Cv))" by (metis add.right_neutral add_mono_thms_linordered_field(5) detGreat0 less_eq_real_def mult_less_0_iff mult_zero_class.mult_zero_left mult_zero_class.mult_zero_right real_sqrt_eq_zero_cancel_iff real_sqrt_gt_0_iff) have h5a : "Av\0 \ Bv\0 \ (Av < -Bv * sqrt Cv) \ (Av\<^sup>2 < Bv\<^sup>2 * Cv)" proof - assume a1: "0 \ Av" assume a2: "Av < - Bv * sqrt Cv" assume "Bv \ 0" then have "Av < sqrt (Cv * (Bv * Bv))" using a2 by (simp add: mult.commute real_sqrt_mult) then show ?thesis using a1 by (metis (no_types) mult.commute power2_eq_square real_sqrt_less_iff real_sqrt_mult real_sqrt_pow2_iff) qed have h5b : "Av\0 \ Bv\0 \ (Av\<^sup>2 < Bv\<^sup>2 * Cv) \ (Av < -Bv * sqrt Cv)" using real_less_rsqrt real_sqrt_mult by fastforce have h5 : "Av\0 \ Bv\0 \ (Av + Bv * sqrt Cv < 0) = (Av < 0 \ Bv\<^sup>2 * Cv < Av\<^sup>2 \ Bv \ 0 \ (Av < 0 \ Av\<^sup>2 < Bv\<^sup>2 * Cv))" using h5a h5b by linarith have h6 : "Av\0 \ Bv>0 \ (Av + Bv * sqrt Cv < 0) = (Av < 0 \ Bv\<^sup>2 * Cv < Av\<^sup>2 \ Bv \ 0 \ (Av < 0 \ Av\<^sup>2 < Bv\<^sup>2 * Cv))" by (smt detGreat0 mult_nonneg_nonneg real_sqrt_ge_zero) have h7a : "Av<0 \ Bv>0 \ (Av < -Bv * sqrt Cv) \ (Bv\<^sup>2 * Cv < Av\<^sup>2)" by (smt mult_minus_left real_sqrt_abs real_sqrt_le_mono real_sqrt_mult) have h7b : "Av<0 \ Bv>0 \ (Bv\<^sup>2 * Cv < Av\<^sup>2) \ (Av < -Bv * sqrt Cv)" by (metis abs_of_nonneg abs_real_def add.commute less_eq_real_def mult.assoc mult_minus_left power2_eq_square real_add_less_0_iff real_sqrt_less_iff real_sqrt_mult real_sqrt_mult_self) have h7 : "Av<0 \ Bv>0 \ (Av + Bv * sqrt Cv < 0) = (Av < 0 \ Bv\<^sup>2 * Cv < Av\<^sup>2 \ Bv \ 0 \ (Av < 0 \ Av\<^sup>2 < Bv\<^sup>2 * Cv))" using h7a h7b by linarith show ?thesis apply(simp add: hA hB h1 h2 insertion_add insertion_mult insertion_var lLength) using h3 h4 h5 h6 h7 by smt qed lemma quadratic_sub_less : assumes lLength : "length L > var" assumes nonzero : "Dv \ 0" assumes detGreater0 : "Cv \ 0" assumes freeC : "var \ vars c" assumes ha : "\x. insertion (nth_default 0 (list_update L var x)) (a::real mpoly) = (Av :: real)" assumes hb : "\x. insertion (nth_default 0 (list_update L var x)) (b::real mpoly) = (Bv :: real)" assumes hc : "\x. insertion (nth_default 0 (list_update L var x)) (c::real mpoly) = (Cv :: real)" assumes hd : "\x. insertion (nth_default 0 (list_update L var x)) (d::real mpoly) = (Dv :: real)" shows "aEval (Less p) (list_update L var ((Av+Bv*sqrt(Cv))/Dv)) = eval (quadratic_sub var a b c d (Less p)) (list_update L var (sqrt Cv))" proof - define p1 where "(p1::real mpoly) = quadratic_part_1 var a b d (Less p)" define p2 where "(p2::real mpoly) = quadratic_part_2 var c p1" define A where "A = isolate_variable_sparse p2 var 0" define B where "B = isolate_variable_sparse p2 var 1" have h3b : "MPoly_Type.degree p2 var \ 1" using freeC quad_part_2_deg p2_def by blast then have h3c : "MPoly_Type.degree p2 var = 0 \ MPoly_Type.degree p2 var = 1" by auto have h3d : "MPoly_Type.degree p2 var = 0 \ B = 0" by(simp add: B_def isovar_greater_degree) then have h3f : "MPoly_Type.degree p2 var = 0 \ p2 = A + B * Var var" by(simp add: h3d A_def degree0isovarspar) have h3g1 : "MPoly_Type.degree p2 var = 1 \ p2 = (\i\1. isolate_variable_sparse p2 var i * Var var ^ i)" using sum_over_zero by metis have h3g2a : "\f. (\i::nat\1. f i) = f 0 + f 1" by simp have h3g2 : "(\i::nat\1. isolate_variable_sparse p2 var i * Var var ^ i) = isolate_variable_sparse p2 var 0 * Var var ^ 0 + isolate_variable_sparse p2 var 1 * Var var ^ 1" using h3g2a by blast have h3g : "MPoly_Type.degree p2 var = 1 \ p2 = A + B * Var var" apply(simp add: sum_over_zero A_def B_def) using h3g1 h3g2 by (metis (no_types, lifting) One_nat_def mult_cancel_left2 power_0 power_one_right) have h3h : "p2 = A + B * Var var" using h3c h3f h3g by auto have h4a : "\x::real. \y::real. insertion (nth_default 0(list_update L var y)) A = x" using not_contains_insertion not_in_isovarspar A_def by blast have h4b : "\x::real. \y::real. insertion (nth_default 0(list_update L var y)) B = x" using not_contains_insertion not_in_isovarspar B_def by blast have h1 : "aEval (Less p) (list_update L var ((Av+Bv*sqrt(Cv))/Dv)) = aEval (Less (quadratic_part_1 var a b d (Less p))) (list_update L var (sqrt Cv))" using quad_part_1_less assms by blast also have "... = aEval (Less p1) (list_update L var (sqrt Cv))" using p1_def by auto also have "... = aEval (Less (quadratic_part_2 var c p1)) (list_update L var (sqrt Cv))" using quad_part_2_less assms by blast also have "... = aEval (Less p2) (list_update L var (sqrt Cv))" using p2_def by auto also have "... = aEval (Less (A + B * Var var)) (list_update L var (sqrt Cv))" using h3h by auto also have "... = eval (Or (And (fm.Atom (Less A)) (fm.Atom (Less (B\<^sup>2 * c - A\<^sup>2)))) (And (fm.Atom (Leq B)) (Or (fm.Atom (Less A)) (fm.Atom (Less (A\<^sup>2 - B\<^sup>2 * c)))))) (list_update L var (sqrt Cv))" using quadratic_sub_less_helper hc detGreater0 h4a h4b lLength by blast also have "... = eval (quadratic_sub var a b c d (Less p)) (list_update L var (sqrt Cv))" using p2_def A_def B_def p1_def quadratic_sub.simps(2) by metis finally show ?thesis by blast qed (*------------------------------------------------------------------------------------------------*) lemma quadratic_sub_leq_helper : assumes lLength : "length L > var" assumes detGreat0 : "Cv\0" assumes hC : "\x. insertion (nth_default 0 (list_update L var x)) (C::real mpoly) = (Cv::real)" assumes hA : "\x. insertion (nth_default 0 (list_update L var x)) (A::real mpoly) = (Av::real)" assumes hB : "\x. insertion (nth_default 0 (list_update L var x)) (B::real mpoly) = (Bv::real)" shows "aEval (Leq (A + B * Var var)) (list_update L var (sqrt Cv)) = eval (Or(And(Atom(Leq(A)))(Atom (Leq(B^2*C-A^2))))(And (Atom(Leq B)) (Atom(Leq (A^2-B^2*C))))) (list_update L var (sqrt Cv))" proof- have h1 : "\x. insertion (nth_default 0 (list_update L var x)) (A^2-(B^2)*C) = Av^2-(Bv^2)*Cv" by(simp add: hA hB hC insertion_add insertion_mult insertion_sub insertion_pow) have h2 : "\x. insertion (nth_default 0 (list_update L var x)) ((B^2)*C-A^2) = (Bv^2)*Cv-Av^2" by(simp add: hA hB hC insertion_add insertion_mult insertion_sub insertion_pow) have h3 : "Av=0 \ Bv=0 \ (Av + Bv * sqrt Cv \ 0) = (Av \ 0 \ Bv\<^sup>2 * Cv \ Av\<^sup>2 \ Bv \ 0 \ Av\<^sup>2 \ Bv\<^sup>2 * Cv)" by simp have h4 : "Av<0 \ Bv\0 \ (Av + Bv * sqrt Cv \ 0) = (Av \ 0 \ Bv\<^sup>2 * Cv \ Av\<^sup>2 \ Bv \ 0 \ Av\<^sup>2 \ Bv\<^sup>2 * Cv)" by (smt detGreat0 real_sqrt_ge_zero zero_less_mult_iff) have h5 : "Av=0 \ Bv\0 \ (Av + Bv * sqrt Cv \ 0) = (Av \ 0 \ Bv\<^sup>2 * Cv \ Av\<^sup>2 \ Bv \ 0 \ Av\<^sup>2 \ Bv\<^sup>2 * Cv)" by (smt detGreat0 real_sqrt_ge_zero zero_less_mult_iff) have h6 : "Av\0 \ Bv>0 \ (Av + Bv * sqrt Cv \ 0) = (Av \ 0 \ Bv\<^sup>2 * Cv \ Av\<^sup>2 \ Bv \ 0 \ Av\<^sup>2 \ Bv\<^sup>2 * Cv)" by (smt detGreat0 mult_nonneg_nonneg mult_pos_pos real_sqrt_gt_0_iff real_sqrt_zero zero_le_power2 zero_less_mult_pos zero_less_power2) have h7a : "Av<0 \ Bv>0 \ (Av + Bv * sqrt Cv \ 0) \ Bv\<^sup>2 * Cv \ Av\<^sup>2" by (smt real_sqrt_abs real_sqrt_less_mono real_sqrt_mult) have h7b : "Av<0 \ Bv>0 \ Bv\<^sup>2 * Cv \ Av\<^sup>2 \ (Av + Bv * sqrt Cv \ 0) " by (smt real_sqrt_abs real_sqrt_less_mono real_sqrt_mult) have h7 : "Av<0 \ Bv>0 \ (Av + Bv * sqrt Cv \ 0) = (Av \ 0 \ Bv\<^sup>2 * Cv \ Av\<^sup>2 \ Bv \ 0 \ Av\<^sup>2 \ Bv\<^sup>2 * Cv)" using h7a h7b by linarith have h8c : "(-Bv * sqrt Cv)^2 = Bv\<^sup>2 * Cv" by (simp add: detGreat0 power_mult_distrib) have h8a : "Av>0 \ Bv\0 \ (Av \ -Bv * sqrt Cv) \ Av\<^sup>2 \ Bv\<^sup>2 * Cv" using detGreat0 h8c power_both_sides by smt have h8b : "Av>0 \ Bv\0 \ Av\<^sup>2 \ Bv\<^sup>2 * Cv \ (Av + Bv * sqrt Cv \ 0) " using detGreat0 h8c power_both_sides by (smt mult_minus_left real_sqrt_ge_zero zero_less_mult_iff) have h8 : "Av>0 \ Bv\0 \ (Av + Bv * sqrt Cv \ 0) = (Av \ 0 \ Bv\<^sup>2 * Cv \ Av\<^sup>2 \ Bv \ 0 \ Av\<^sup>2 \ Bv\<^sup>2 * Cv)" using h8a h8b by linarith show ?thesis apply(simp add: hA hB h1 h2 insertion_add insertion_mult insertion_var lLength) using h3 h4 h5 h6 h7 h8 by smt qed lemma quadratic_sub_leq : assumes lLength : "length L > var" assumes nonzero : "Dv \ 0" assumes detGreater0 : "Cv \ 0" assumes freeC : "var \ vars c" assumes ha : "\x. insertion (nth_default 0 (list_update L var x)) (a::real mpoly) = (Av :: real)" assumes hb : "\x. insertion (nth_default 0 (list_update L var x)) (b::real mpoly) = (Bv :: real)" assumes hc : "\x. insertion (nth_default 0 (list_update L var x)) (c::real mpoly) = (Cv :: real)" assumes hd : "\x. insertion (nth_default 0 (list_update L var x)) (d::real mpoly) = (Dv :: real)" shows "aEval (Leq p) (list_update L var ((Av+Bv*sqrt(Cv))/Dv)) = eval (quadratic_sub var a b c d (Leq p)) (list_update L var (sqrt Cv))" proof - define p1 where "(p1::real mpoly) = quadratic_part_1 var a b d (Leq p)" define p2 where "(p2::real mpoly) = quadratic_part_2 var c p1" define A where "A = isolate_variable_sparse p2 var 0" define B where "B = isolate_variable_sparse p2 var 1" have h3b : "MPoly_Type.degree p2 var \ 1" using freeC quad_part_2_deg p2_def lLength by metis then have h3c : "MPoly_Type.degree p2 var = 0 \ MPoly_Type.degree p2 var = 1" by auto have h3d : "MPoly_Type.degree p2 var = 0 \ B = 0" by(simp add: B_def isovar_greater_degree) then have h3f : "MPoly_Type.degree p2 var = 0 \ p2 = A + B * Var var" by(simp add: h3d A_def degree0isovarspar) have h3g1 : "MPoly_Type.degree p2 var = 1 \ p2 = (\i\1. isolate_variable_sparse p2 var i * Var var ^ i)" using sum_over_zero by metis have h3g2a : "\f. (\i::nat\1. f i) = f 0 + f 1" by simp have h3g2 : "(\i::nat\1. isolate_variable_sparse p2 var i * Var var ^ i) = isolate_variable_sparse p2 var 0 * Var var ^ 0 + isolate_variable_sparse p2 var 1 * Var var ^ 1" using h3g2a by blast have h3g : "MPoly_Type.degree p2 var = 1 \ p2 = A + B * Var var" apply(simp add: sum_over_zero A_def B_def) using h3g1 h3g2 by (metis (no_types, lifting) One_nat_def mult_cancel_left2 power_0 power_one_right) have h3h : "p2 = A + B * Var var" using h3c h3f h3g by auto have h4a : "\x::real. \y::real. insertion (nth_default 0 (list_update L var y)) A = x" using not_contains_insertion not_in_isovarspar A_def by blast have h4b : "\x::real. \y::real. insertion (nth_default 0 (list_update L var y)) B = x" using not_contains_insertion not_in_isovarspar B_def by blast have "aEval (Leq p) (list_update L var ((Av+Bv*sqrt(Cv))/Dv)) = aEval (Leq p1) (list_update L var (sqrt Cv))" using quad_part_1_leq nonzero ha hb hd p1_def lLength by metis also have "... = aEval (Leq p2) (list_update L var (sqrt Cv))" using p2_def quad_part_2_leq hc detGreater0 lLength by metis also have "... = aEval (Leq (A + B * Var var)) (list_update L var (sqrt Cv))" using h3h by auto also have h4 : "... = eval (Or (And (Atom(Leq(A))) (Atom (Leq(B^2*c-A^2)))) (And (Atom(Leq B)) (Atom(Leq (A^2-B^2*c))))) (list_update L var (sqrt Cv))" using quadratic_sub_leq_helper hc detGreater0 h4a h4b lLength by blast also have "... = eval (quadratic_sub var a b c d (Leq p)) (list_update L var (sqrt Cv))" using p1_def quadratic_sub.simps(3) p2_def A_def B_def by metis finally show ?thesis by blast qed (*------------------------------------------------------------------------------------------------*) lemma quadratic_sub_neq : assumes lLength : "length L > var" assumes nonzero : "Dv \ 0" assumes detGreater0 : "Cv \ 0" assumes freeC : "var \ vars c" assumes ha : "\x. insertion (nth_default 0 (list_update L var x)) (a::real mpoly) = (Av :: real)" assumes hb : "\x. insertion (nth_default 0 (list_update L var x)) (b::real mpoly) = (Bv :: real)" assumes hc : "\x. insertion (nth_default 0 (list_update L var x)) (c::real mpoly) = (Cv :: real)" assumes hd : "\x. insertion (nth_default 0 (list_update L var x)) (d::real mpoly) = (Dv :: real)" shows "aEval (Neq p) (list_update L var ((Av+Bv*sqrt(Cv))/Dv)) = eval (quadratic_sub var a b c d (Neq p)) (list_update L var (sqrt Cv))" proof - define p1 where "(p1::real mpoly) = quadratic_part_1 var a b d (Neq p)" define p2 where "(p2::real mpoly) = quadratic_part_2 var c p1" define A where "A = isolate_variable_sparse p2 var 0" define B where "B = isolate_variable_sparse p2 var 1" have h3b : "MPoly_Type.degree p2 var \ 1" using freeC quad_part_2_deg p2_def lLength by metis then have h3c : "MPoly_Type.degree p2 var = 0 \ MPoly_Type.degree p2 var = 1" by auto have h3d : "MPoly_Type.degree p2 var = 0 \ B = 0" by(simp add: B_def isovar_greater_degree) then have h3f : "MPoly_Type.degree p2 var = 0 \ p2 = A + B * Var var" by(simp add: h3d A_def degree0isovarspar) have h3g1 : "MPoly_Type.degree p2 var = 1 \ p2 = (\i\1. isolate_variable_sparse p2 var i * Var var ^ i)" using sum_over_zero by metis have h3g2a : "\f. (\i::nat\1. f i) = f 0 + f 1" by simp have h3g2 : "(\i::nat\1. isolate_variable_sparse p2 var i * Var var ^ i) = isolate_variable_sparse p2 var 0 * Var var ^ 0 + isolate_variable_sparse p2 var 1 * Var var ^ 1" using h3g2a by blast have h3g : "MPoly_Type.degree p2 var = 1 \ p2 = A + B * Var var" apply(simp add: sum_over_zero A_def B_def) using h3g1 h3g2 by (metis (no_types, lifting) One_nat_def mult_cancel_left2 power_0 power_one_right) have h3h : "p2 = A + B * Var var" using h3c h3f h3g by auto have h4a : "\x::real. \y::real. insertion (nth_default 0 (list_update L var y)) A = x" using not_contains_insertion not_in_isovarspar A_def by blast have h4b : "\x::real. \y::real. insertion (nth_default 0 (list_update L var y)) B = x" using not_contains_insertion not_in_isovarspar B_def by blast have h4c : "aEval (Eq (A + B * Var var)) (list_update L var (sqrt Cv)) = eval (And (Atom(Leq (A*B))) (Atom (Eq (A^2-B^2*c)))) (list_update L var (sqrt Cv))" using quad_equality_helper hc detGreater0 h4a h4b lLength by blast have h4d : "aEval (Neq (A + B * Var var)) (list_update L var (sqrt Cv)) = (\ (eval (And (Atom(Leq (A*B))) (Atom (Eq (A^2-B^2*c)))) (list_update L var (sqrt Cv))))" using aEval.simps(1) aEval.simps(4) h4c by blast have h4e : "(\ (eval (And (Atom(Leq (A*B))) (Atom (Eq (A^2-B^2*c)))) (list_update L var (sqrt Cv)))) = eval (Or (Atom(Less(-A*B))) (Atom (Neq(A^2-B^2*c)))) (list_update L var (sqrt Cv))" by (metis aNeg.simps(2) aNeg.simps(3) aNeg_aEval eval.simps(1) eval.simps(4) eval.simps(5) mult_minus_left) have "aEval (Neq p) (list_update L var ((Av+Bv*sqrt(Cv))/Dv)) = aEval (Neq p1) (list_update L var (sqrt Cv))" using quad_part_1_neq nonzero ha hb hd p1_def lLength by blast also have "... = aEval (Neq p2) (list_update L var (sqrt Cv))" using p2_def quad_part_2_neq hc detGreater0 lLength by metis also have "... = aEval (Neq (A + B * Var var)) (list_update L var (sqrt Cv))" using h3h by auto also have "... = eval (Or (Atom(Less(-A*B))) (Atom (Neq(A^2-B^2*c)))) (list_update L var (sqrt Cv))" using h4c h4d h4e by auto also have "... = eval (quadratic_sub var a b c d (Neq p)) (list_update L var (sqrt Cv))" using p2_def A_def B_def p1_def quadratic_sub.simps(4) quadratic_part_1.simps(1) quadratic_part_1.simps(4) by (metis (no_types, lifting)) finally show ?thesis by blast qed (*-----------------------------------------------------------------------------------------------*) theorem free_in_quad : assumes freeA : "var\ vars a" assumes freeB : "var\ vars b" assumes freeC : "var\ vars c" assumes freeD : "var\ vars d" shows "freeIn var (quadratic_sub var a b c d A)" proof(cases A) case (Less p) define p1 where "(p1::real mpoly) = quadratic_part_1 var a b d (Less p)" define p2 where "(p2::real mpoly) = quadratic_part_2 var c p1" define A where "A = isolate_variable_sparse p2 var 0" define B where "B = isolate_variable_sparse p2 var 1" have h1 : "freeIn var (quadratic_sub var a b c d (Less p)) = freeIn var (Or (And (fm.Atom (Less A)) (fm.Atom (Less (B\<^sup>2 * c - A\<^sup>2)))) (And (fm.Atom (Leq B)) (Or (fm.Atom (Less A)) (fm.Atom (Less (A\<^sup>2 - B\<^sup>2 * c))))))" using p2_def A_def B_def p1_def quadratic_sub.simps(2) by metis have h2d : "var\vars(4::real mpoly)" by (metis freeB not_in_add not_in_pow numeral_Bit0 one_add_one power_0) have h2 : "freeIn var ((Or (And (fm.Atom (Less A)) (fm.Atom (Less (B\<^sup>2 * c - A\<^sup>2)))) (And (fm.Atom (Leq B)) (Or (fm.Atom (Less A)) (fm.Atom (Less (A\<^sup>2 - B\<^sup>2 * c)))))))" using vars_mult not_in_isovarspar A_def B_def not_in_sub not_in_mult not_in_neg not_in_pow not_in_isovarspar h2d freeC by (simp) show ?thesis using h1 h2 Less by blast next case (Eq p) define p1 where "(p1::real mpoly) = quadratic_part_1 var a b d (Eq p)" define p2 where "(p2::real mpoly) = quadratic_part_2 var c p1" define A where "A = isolate_variable_sparse p2 var 0" define B where "B = isolate_variable_sparse p2 var 1" have h1 : "freeIn var (quadratic_sub var a b c d (Eq p)) = freeIn var (And (Atom(Leq (A*B))) (Atom (Eq (A\<^sup>2 - B\<^sup>2 * c))))" using p2_def A_def B_def p1_def quadratic_sub.simps(1) by metis have h2d : "var\vars(4::real mpoly)" by (metis freeB not_in_add not_in_pow numeral_Bit0 one_add_one power_0) have h2 : "freeIn var (And (Atom(Leq (A*B))) (Atom (Eq (A\<^sup>2 - B\<^sup>2 * c))))" using vars_mult not_in_isovarspar A_def B_def not_in_sub not_in_mult not_in_neg not_in_pow not_in_isovarspar h2d freeC by (simp) show ?thesis using h1 h2 Eq by blast next case (Leq p) define p1 where "(p1::real mpoly) = quadratic_part_1 var a b d (Leq p)" define p2 where "(p2::real mpoly) = quadratic_part_2 var c p1" define A where "A = isolate_variable_sparse p2 var 0" define B where "B = isolate_variable_sparse p2 var 1" have h1 : "freeIn var (quadratic_sub var a b c d (Leq p)) = freeIn var (Or(And(Atom(Leq(A)))(Atom (Leq(B^2*c-A^2))))(And(Atom(Leq B))(Atom(Leq (A^2-B^2*c)))))" using p2_def A_def B_def p1_def quadratic_sub.simps(3) by metis have h2d : "var\vars(4::real mpoly)" by (metis freeB not_in_add not_in_pow numeral_Bit0 one_add_one power_0) have h2 : "freeIn var (Or(And(Atom(Leq(A)))(Atom (Leq(B^2*c-A^2))))(And(Atom(Leq B))(Atom(Leq (A^2-B^2*c)))))" using vars_mult not_in_isovarspar A_def B_def not_in_sub not_in_mult not_in_neg not_in_pow not_in_isovarspar h2d freeC by (simp) show ?thesis using h1 h2 Leq by blast next case (Neq p) define p1 where "(p1::real mpoly) = quadratic_part_1 var a b d (Neq p)" define p2 where "(p2::real mpoly) = quadratic_part_2 var c p1" define A where "A = isolate_variable_sparse p2 var 0" define B where "B = isolate_variable_sparse p2 var 1" have h1 : "freeIn var (quadratic_sub var a b c d (Neq p)) = freeIn var (Or (Atom(Less(-A*B))) (Atom (Neq(A^2-B^2*c))))" using p2_def A_def B_def p1_def quadratic_sub.simps(4) by metis have h2d : "var\vars(4::real mpoly)" by (metis freeB not_in_add not_in_pow numeral_Bit0 one_add_one power_0) have h2 : "freeIn var (Or (Atom(Less(-A*B))) (Atom (Neq(A^2-B^2*c))))" using vars_mult not_in_isovarspar A_def B_def not_in_sub not_in_mult not_in_neg not_in_pow not_in_isovarspar h2d freeC by (simp) show ?thesis using h1 h2 Neq by blast qed theorem quadratic_sub : assumes lLength : "length L > var" assumes nonzero : "Dv \ 0" assumes detGreater0 : "Cv \ 0" assumes freeC : "var \ vars c" assumes ha : "\x. insertion (nth_default 0 (list_update L var x)) (a::real mpoly) = (Av :: real)" assumes hb : "\x. insertion (nth_default 0 (list_update L var x)) (b::real mpoly) = (Bv :: real)" assumes hc : "\x. insertion (nth_default 0 (list_update L var x)) (c::real mpoly) = (Cv :: real)" assumes hd : "\x. insertion (nth_default 0 (list_update L var x)) (d::real mpoly) = (Dv :: real)" shows "aEval A (list_update L var ((Av+Bv*sqrt(Cv))/Dv)) = eval (quadratic_sub var a b c d A) (list_update L var (sqrt Cv))" proof(cases A) case (Less x1) then show ?thesis using quadratic_sub_less assms by blast next case (Eq x2) then show ?thesis using quadratic_sub_eq assms by blast next case (Leq x3) then show ?thesis using quadratic_sub_leq assms by blast next case (Neq x4) then show ?thesis using quadratic_sub_neq assms by blast qed lemma free_in_quad_fm_helper : assumes freeA : "var\ vars a" assumes freeB : "var\ vars b" assumes freeC : "var\ vars c" assumes freeD : "var\ vars d" shows "freeIn (var+z) (quadratic_sub_fm_helper var a b c d F z)" proof(induction F arbitrary: z) case TrueF then show ?case by auto next case FalseF then show ?case by auto next case (Atom x) then show ?case using free_in_quad[OF not_in_lift[OF assms(1)] not_in_lift[OF assms(2)] not_in_lift[OF assms(3)] not_in_lift[OF assms(4)], of z] by auto next case (And F1 F2) then show ?case by auto next case (Or F1 F2) then show ?case by auto next case (Neg F) then show ?case by auto next case (ExQ F) show ?case using ExQ[of "z+1"] by simp next case (AllQ F) show ?case using AllQ[of "z+1"] by simp next case (ExN x1 F) then show ?case by (metis (no_types, lifting) add.assoc freeIn.simps(13) liftmap.simps(10) quadratic_sub_fm_helper.simps) next case (AllN x1 F) then show ?case by (metis (no_types, lifting) freeIn.simps(12) group_cancel.add1 liftmap.simps(9) quadratic_sub_fm_helper.simps) qed theorem free_in_quad_fm : assumes freeA : "var\ vars a" assumes freeB : "var\ vars b" assumes freeC : "var\ vars c" assumes freeD : "var\ vars d" shows "freeIn var (quadratic_sub_fm var a b c d A)" using free_in_quad_fm_helper[OF assms, of 0] by auto lemma quadratic_sub_fm_helper : assumes nonzero : "Dv \ 0" assumes detGreater0 : "Cv \ 0" assumes freeC : "var \ vars c" assumes lLength : "length L > var+z" assumes ha : "\x. insertion (nth_default 0 (list_update (drop z L) var x)) (a::real mpoly) = (Av :: real)" assumes hb : "\x. insertion (nth_default 0 (list_update (drop z L) var x)) (b::real mpoly) = (Bv :: real)" assumes hc : "\x. insertion (nth_default 0 (list_update (drop z L) var x)) (c::real mpoly) = (Cv :: real)" assumes hd : "\x. insertion (nth_default 0 (list_update (drop z L) var x)) (d::real mpoly) = (Dv :: real)" shows "eval F (list_update L (var+z) ((Av+Bv*sqrt(Cv))/Dv)) = eval (quadratic_sub_fm_helper var a b c d F z) (list_update L (var+z) (sqrt Cv))" using assms proof(induction F arbitrary: z L) case TrueF then show ?case by auto next case FalseF then show ?case by auto next case (Atom x) define L1 where "L1 = drop z L" define L2 where "L2 = take z L" have L_def : "L = L2 @ L1" using L1_def L2_def by auto have lengthl2 : "length L2 = z" using L2_def using Atom.prems(4) by auto have "eval (Atom(Eq (a-Const Av))) ([] @ L1) = eval (liftFm 0 z (Atom(Eq (a- Const Av)))) ([] @ L2 @ L1)" by (metis eval_liftFm_helper lengthl2 list.size(3)) then have "(insertion (nth_default 0 (L2 @ L1)) (liftPoly 0 z (a - Const Av)) = 0)" apply(simp add: insertion_sub insertion_const) using Atom(5) unfolding L1_def by (metis list_update_id) then have "insertion (nth_default 0 (L2 @ L1)) (liftPoly 0 z a) = Av" using lift_minus by blast then have a1 : "\x. insertion (nth_default 0 (L[var + z := x])) (liftPoly 0 z a) = Av" unfolding L_def by (metis (no_types, lifting) Atom.prems(5) L1_def add.right_neutral add_diff_cancel_right' append_eq_append_conv append_eq_append_conv2 length_append lengthl2 lift_insertion list.size(3) list_update_append not_add_less2) have "eval (Atom(Eq (b-Const Bv))) ([] @ L1) = eval (liftFm 0 z (Atom(Eq (b- Const Bv)))) ([] @ L2 @ L1)" by (metis eval_liftFm_helper lengthl2 list.size(3)) then have "(insertion (nth_default 0 (L2 @ L1)) (liftPoly 0 z (b - Const Bv)) = 0)" apply(simp add: insertion_sub insertion_const) using Atom(6) unfolding L1_def by (metis list_update_id) then have "insertion (nth_default 0 (L2 @ L1)) (liftPoly 0 z b) = Bv" using lift_minus by blast then have a2 : "\x. insertion (nth_default 0 (L[var + z := x])) (liftPoly 0 z b) = Bv" unfolding L_def using Atom(6) L1_def by (metis L_def add_diff_cancel_right' append.simps(1) lengthl2 lift_insertion list.size(3) list_update_append not_add_less2) have "eval (Atom(Eq (c-Const Cv))) ([] @ L1) = eval (liftFm 0 z (Atom(Eq (c- Const Cv)))) ([] @ L2 @ L1)" by (metis eval_liftFm_helper lengthl2 list.size(3)) then have "(insertion (nth_default 0 (L2 @ L1)) (liftPoly 0 z (c - Const Cv)) = 0)" apply(simp add: insertion_sub insertion_const) using Atom(7) unfolding L1_def by (metis list_update_id) then have "insertion (nth_default 0 (L2 @ L1)) (liftPoly 0 z c) = Cv" using lift_minus by blast then have a3 : "\x. insertion (nth_default 0 (L[var + z := x])) (liftPoly 0 z c) = Cv" - unfolding L_def - proof - - obtain nn :: "(nat \ real) \ (nat \ real) \ real mpoly \ nat" where - "\x0 x1 x2. (\v3. v3 \ vars x2 \ x1 v3 \ x0 v3) = (nn x0 x1 x2 \ vars x2 \ x1 (nn x0 x1 x2) \ x0 (nn x0 x1 x2))" - by moura - then have f1: "\m f fa. nn fa f m \ vars m \ f (nn fa f m) \ fa (nn fa f m) \ insertion f m = insertion fa m" - by (meson insertion_irrelevant_vars) - obtain rr :: real where - "(\v0. insertion (nth_default 0 ((L2 @ L1)[var + z := v0])) (liftPoly 0 z c) \ Cv) = (insertion (nth_default 0 ((L2 @ L1)[var + z := rr])) (liftPoly 0 z c) \ Cv)" - by blast - moreover - { assume "var + z \ nn (nth_default 0 ((L2 @ L1)[var + z := rr])) (nth_default 0 (L2 @ L1)) (liftPoly 0 z c)" - moreover - { assume "(nth_default 0 (L2 @ L1) (nn (nth_default 0 ((L2 @ L1)[var + z := rr])) (nth_default 0 (L2 @ L1)) (liftPoly 0 z c)) = nth_default 0 ((L2 @ L1)[var + z := rr]) (nn (nth_default 0 ((L2 @ L1)[var + z := rr])) (nth_default 0 (L2 @ L1)) (liftPoly 0 z c))) \ ((L2 @ L1) ! nn (nth_default 0 ((L2 @ L1)[var + z := rr])) (nth_default 0 (L2 @ L1)) (liftPoly 0 z c) = (L2 @ L1)[var + z := rr] ! nn (nth_default 0 ((L2 @ L1)[var + z := rr])) (nth_default 0 (L2 @ L1)) (liftPoly 0 z c))" - then have "nth_default 0 ((L2 @ L1)[var + z := rr]) (nn (nth_default 0 ((L2 @ L1)[var + z := rr])) (nth_default 0 (L2 @ L1)) (liftPoly 0 z c)) \ (L2 @ L1)[var + z := rr] ! nn (nth_default 0 ((L2 @ L1)[var + z := rr])) (nth_default 0 (L2 @ L1)) (liftPoly 0 z c) \ nth_default 0 (L2 @ L1) (nn (nth_default 0 ((L2 @ L1)[var + z := rr])) (nth_default 0 (L2 @ L1)) (liftPoly 0 z c)) \ (L2 @ L1) ! nn (nth_default 0 ((L2 @ L1)[var + z := rr])) (nth_default 0 (L2 @ L1)) (liftPoly 0 z c)" - by linarith - then have "nn (nth_default 0 ((L2 @ L1)[var + z := rr])) (nth_default 0 (L2 @ L1)) (liftPoly 0 z c) \ vars (liftPoly 0 z c) \ nth_default 0 (L2 @ L1) (nn (nth_default 0 ((L2 @ L1)[var + z := rr])) (nth_default 0 (L2 @ L1)) (liftPoly 0 z c)) = nth_default 0 ((L2 @ L1)[var + z := rr]) (nn (nth_default 0 ((L2 @ L1)[var + z := rr])) (nth_default 0 (L2 @ L1)) (liftPoly 0 z c))" - by (metis (no_types) append_Nil2 length_list_update nth_default_append) } - ultimately have "nn (nth_default 0 ((L2 @ L1)[var + z := rr])) (nth_default 0 (L2 @ L1)) (liftPoly 0 z c) \ vars (liftPoly 0 z c) \ nth_default 0 (L2 @ L1) (nn (nth_default 0 ((L2 @ L1)[var + z := rr])) (nth_default 0 (L2 @ L1)) (liftPoly 0 z c)) = nth_default 0 ((L2 @ L1)[var + z := rr]) (nn (nth_default 0 ((L2 @ L1)[var + z := rr])) (nth_default 0 (L2 @ L1)) (liftPoly 0 z c))" - by force } - ultimately show "\r. insertion (nth_default 0 ((L2 @ L1)[var + z := r])) (liftPoly 0 z c) = Cv" - using f1 by (metis (full_types) Atom.prems(3) \insertion (nth_default 0 (L2 @ L1)) (liftPoly 0 z c) = Cv\ not_in_lift) - qed + unfolding L_def by (metis freeC list_update_id not_contains_insertion not_in_lift) have "eval (Atom(Eq (d-Const Dv))) ([] @ L1) = eval (liftFm 0 z (Atom(Eq (d- Const Dv)))) ([] @ L2 @ L1)" by (metis eval_liftFm_helper lengthl2 list.size(3)) then have "(insertion (nth_default 0 (L2 @ L1)) (liftPoly 0 z (d - Const Dv)) = 0)" apply(simp add: insertion_sub insertion_const) using Atom(8) unfolding L1_def by (metis list_update_id) then have "insertion (nth_default 0 (L2 @ L1)) (liftPoly 0 z d) = Dv" using lift_minus by blast then have a4 : "\x. insertion (nth_default 0 (L[var + z := x])) (liftPoly 0 z d) = Dv" unfolding L_def by (metis Atom(8) L1_def L_def add_diff_cancel_right' append.simps(1) lengthl2 lift_insertion list.size(3) list_update_append not_add_less2) then show ?case apply(simp) using quadratic_sub[OF Atom(4) Atom(1) Atom(2) not_in_lift[OF Atom(3)], of "(liftPoly 0 z a)" Av "(liftPoly 0 z b)" Bv "(liftPoly 0 z d)" x , OF a1 a2 a3 a4] . next case (And F1 F2) then show ?case by auto next case (Or F1 F2) then show ?case by auto next case (Neg F) then show ?case by auto next case (ExQ F) have lengthG : "var + (z + 1) < length (x#L)" for x using ExQ(5) by auto have forall : "\x. insertion (nth_default 0 ((drop z L)[var := x])) a = Av \ \x. insertion (nth_default 0 ((drop (z + 1) (x1 # L))[var := x])) a = Av" for x1 a Av by auto have l : "x # L[var + z := v] = ((x#L)[var+(z+1):=v])" for x v by auto have "eval (ExQ F) (L[var + z := (Av + Bv * sqrt Cv) / Dv]) = (\x. eval F (x # L[var + z := (Av + Bv * sqrt Cv) / Dv]))" by(simp) also have "... = (\x. eval (liftmap (\x. quadratic_sub (var + x) (liftPoly 0 x a) (liftPoly 0 x b) (liftPoly 0 x c) (liftPoly 0 x d)) F (z + 1)) (x # L[var + z := sqrt Cv]))" apply(rule ex_cong1) unfolding l using ExQ(1)[OF ExQ(2) ExQ(3) ExQ(4) lengthG forall[OF ExQ(6)] forall[OF ExQ(7)] forall[OF ExQ(8)] forall[OF ExQ(9)]] unfolding quadratic_sub_fm_helper.simps liftmap.simps by simp also have "... = eval (quadratic_sub_fm_helper var a b c d (ExQ F) z) (L[var + z := sqrt Cv])" unfolding quadratic_sub_fm_helper.simps liftmap.simps eval.simps by auto finally show ?case by simp next case (AllQ F) have lengthG : "var + (z + 1) < length (x#L)" for x using AllQ(5) by auto have forall : "\x. insertion (nth_default 0 ((drop z L)[var := x])) a = Av \ \x. insertion (nth_default 0 ((drop (z + 1) (x1 # L))[var := x])) a = Av" for x1 a Av by auto have l : "x # L[var + z := v] = ((x#L)[var+(z+1):=v])" for x v by auto have "eval (AllQ F) (L[var + z := (Av + Bv * sqrt Cv) / Dv]) = (\x. eval F (x # L[var + z := (Av + Bv * sqrt Cv) / Dv]))" by(simp) also have "... = (\x. eval (liftmap (\x. quadratic_sub (var + x) (liftPoly 0 x a) (liftPoly 0 x b) (liftPoly 0 x c) (liftPoly 0 x d)) F (z + 1)) (x # L[var + z := sqrt Cv]))" apply(rule all_cong1) unfolding l using AllQ(1)[OF AllQ(2) AllQ(3) AllQ(4) lengthG forall[OF AllQ(6)] forall[OF AllQ(7)] forall[OF AllQ(8)] forall[OF AllQ(9)]] unfolding quadratic_sub_fm_helper.simps liftmap.simps by simp also have "... = eval (quadratic_sub_fm_helper var a b c d (AllQ F) z) (L[var + z := sqrt Cv])" unfolding quadratic_sub_fm_helper.simps liftmap.simps eval.simps by auto finally show ?case by simp next case (ExN x1 F) have list : "\l. length l=x1 \ ((drop (z + x1) l @ drop (z + x1 - length l) L)) = ((drop z L))" by auto have map : "\ z L. eval (liftmap (\x A. (quadratic_sub (var + x) (liftPoly 0 x a) (liftPoly 0 x b) (liftPoly 0 x c) (liftPoly 0 x d) A)) F (z + x1)) L = eval (liftmap (\x A. (quadratic_sub (var + x1 + x) (liftPoly 0 (x+x1) a) (liftPoly 0 (x+x1) b) (liftPoly 0 (x+x1) c) (liftPoly 0 (x+x1) d) A)) F z) L" apply(induction F) apply(simp_all add:add.commute add.left_commute) apply force apply force by (metis (mono_tags, lifting) ab_semigroup_add_class.add_ac(1))+ show ?case apply simp apply(rule ex_cong1) subgoal for l using map[of z] list[of l] ExN(1)[OF ExN(2-4), of "z+x1" "l@L"] ExN(5-9) list_update_append apply auto by (simp add: list_update_append) + done next case (AllN x1 F) have list : "\l. length l=x1 \ ((drop (z + x1) l @ drop (z + x1 - length l) L)) = ((drop z L))" by auto have map : "\ z L. eval (liftmap (\x A. (quadratic_sub (var + x) (liftPoly 0 x a) (liftPoly 0 x b) (liftPoly 0 x c) (liftPoly 0 x d) A)) F (z + x1)) L = eval (liftmap (\x A. (quadratic_sub (var + x1 + x) (liftPoly 0 (x+x1) a) (liftPoly 0 (x+x1) b) (liftPoly 0 (x+x1) c) (liftPoly 0 (x+x1) d) A)) F z) L" apply(induction F) apply(simp_all add:add.commute add.left_commute) apply force apply force by (metis (mono_tags, lifting) ab_semigroup_add_class.add_ac(1))+ show ?case apply simp apply(rule all_cong1) subgoal for l using map[of z] list[of l] AllN(1)[OF AllN(2-4), of "z+x1" "l@L"] AllN(5-9) apply auto by (simp add: list_update_append) + done qed theorem quadratic_sub_fm : assumes lLength : "length L > var" assumes nonzero : "Dv \ 0" assumes detGreater0 : "Cv \ 0" assumes freeC : "var \ vars c" assumes ha : "\x. insertion (nth_default 0 (list_update L var x)) (a::real mpoly) = (Av :: real)" assumes hb : "\x. insertion (nth_default 0 (list_update L var x)) (b::real mpoly) = (Bv :: real)" assumes hc : "\x. insertion (nth_default 0 (list_update L var x)) (c::real mpoly) = (Cv :: real)" assumes hd : "\x. insertion (nth_default 0 (list_update L var x)) (d::real mpoly) = (Dv :: real)" shows "eval F (list_update L var ((Av+Bv*sqrt(Cv))/Dv)) = eval (quadratic_sub_fm var a b c d F) (list_update L var (sqrt Cv))" unfolding quadratic_sub_fm.simps using quadratic_sub_fm_helper[OF assms(2) assms(3) assms(4), of 0 L a Av b Bv d F] assms(1) assms(5) assms(6) assms(7) assms(8) by (simp add: lLength) end \ No newline at end of file