diff --git a/thys/Functional_Ordered_Resolution_Prover/IsaFoR_Term.thy b/thys/Functional_Ordered_Resolution_Prover/IsaFoR_Term.thy --- a/thys/Functional_Ordered_Resolution_Prover/IsaFoR_Term.thy +++ b/thys/Functional_Ordered_Resolution_Prover/IsaFoR_Term.thy @@ -1,634 +1,634 @@ (* Title: Integration of IsaFoR Terms Author: Dmitriy Traytel , 2014 Author: Anders Schlichtkrull , 2017 Maintainer: Anders Schlichtkrull *) section \Integration of \textsf{IsaFoR} Terms\ text \ This theory implements the abstract interface for atoms and substitutions using the \textsf{IsaFoR} library (part of the AFP entry \First_Order_Terms\). \ theory IsaFoR_Term imports Deriving.Derive Ordered_Resolution_Prover.Abstract_Substitution First_Order_Terms.Unification First_Order_Terms.Subsumption "HOL-Cardinals.Wellorder_Extension" Open_Induction.Restricted_Predicates begin hide_const (open) mgu abbreviation subst_apply_literal :: "('f, 'v) term literal \ ('f, 'v, 'w) gsubst \ ('f, 'w) term literal" (infixl "\lit" 60) where "L \lit \ \ map_literal (\A. A \ \) L" definition subst_apply_clause :: "('f, 'v) term clause \ ('f, 'v, 'w) gsubst \ ('f, 'w) term clause" (infixl "\cls" 60) where "C \cls \ = image_mset (\L. L \lit \) C" abbreviation vars_lit :: "('f, 'v) term literal \ 'v set" where "vars_lit L \ vars_term (atm_of L)" definition vars_clause :: "('f, 'v) term clause \ 'v set" where "vars_clause C = Union (set_mset (image_mset vars_lit C))" definition vars_clause_list :: "('f, 'v) term clause list \ 'v set" where "vars_clause_list Cs = Union (vars_clause ` set Cs) " definition vars_partitioned :: "('f,'v) term clause list \ bool" where "vars_partitioned Cs \ (\i < length Cs. \j < length Cs. i \ j \ (vars_clause (Cs ! i) \ vars_clause (Cs ! j)) = {})" lemma vars_clause_mono: "S \# C \ vars_clause S \ vars_clause C" unfolding vars_clause_def by auto interpretation substitution_ops "(\)" Var "(\\<^sub>s)" . lemma is_ground_atm_is_ground_on_var: assumes "is_ground_atm (A \ \)" and "v \ vars_term A" shows "is_ground_atm (\ v)" using assms proof (induction A) case (Var x) then show ?case by auto next case (Fun f ts) then show ?case unfolding is_ground_atm_def by auto qed lemma is_ground_lit_is_ground_on_var: assumes ground_lit: "is_ground_lit (subst_lit L \)" and v_in_L: "v \ vars_lit L" shows "is_ground_atm (\ v)" proof - let ?A = "atm_of L" from v_in_L have A_p: "v \ vars_term ?A" by auto then have "is_ground_atm (?A \ \)" using ground_lit unfolding is_ground_lit_def by auto then show ?thesis using A_p is_ground_atm_is_ground_on_var by metis qed lemma is_ground_cls_is_ground_on_var: assumes ground_clause: "is_ground_cls (subst_cls C \)" and v_in_C: "v \ vars_clause C" shows "is_ground_atm (\ v)" proof - from v_in_C obtain L where L_p: "L \# C" "v \ vars_lit L" unfolding vars_clause_def by auto then have "is_ground_lit (subst_lit L \)" using ground_clause unfolding is_ground_cls_def subst_cls_def by auto then show ?thesis using L_p is_ground_lit_is_ground_on_var by metis qed lemma is_ground_cls_list_is_ground_on_var: assumes ground_list: "is_ground_cls_list (subst_cls_list Cs \)" and v_in_Cs: "v \ vars_clause_list Cs" shows "is_ground_atm (\ v)" proof - from v_in_Cs obtain C where C_p: "C \ set Cs" "v \ vars_clause C" unfolding vars_clause_list_def by auto then have "is_ground_cls (subst_cls C \)" using ground_list unfolding is_ground_cls_list_def subst_cls_list_def by auto then show ?thesis using C_p is_ground_cls_is_ground_on_var by metis qed lemma same_on_vars_lit: assumes "\v \ vars_lit L. \ v = \ v" shows "subst_lit L \ = subst_lit L \" using assms proof (induction L) case (Pos x) then have "\v \ vars_term x. \ v = \ v \ subst_atm_abbrev x \ = subst_atm_abbrev x \" using term_subst_eq by metis+ then show ?case unfolding subst_lit_def using Pos by auto next case (Neg x) then have "\v \ vars_term x. \ v = \ v \ subst_atm_abbrev x \ = subst_atm_abbrev x \" using term_subst_eq by metis+ then show ?case unfolding subst_lit_def using Neg by auto qed lemma in_list_of_mset_in_S: assumes "i < length (list_of_mset S)" shows "list_of_mset S ! i \# S" proof - from assms have "list_of_mset S ! i \ set (list_of_mset S)" by auto then have "list_of_mset S ! i \# mset (list_of_mset S)" by (meson in_multiset_in_set) then show ?thesis by auto qed lemma same_on_vars_clause: assumes "\v \ vars_clause S. \ v = \ v" shows "subst_cls S \ = subst_cls S \" by (smt assms image_eqI image_mset_cong2 mem_simps(9) same_on_vars_lit set_image_mset subst_cls_def vars_clause_def) lemma vars_partitioned_var_disjoint: assumes "vars_partitioned Cs" shows "var_disjoint Cs" unfolding var_disjoint_def proof (intro allI impI) fix \s :: \('b \ ('a, 'b) term) list\ assume "length \s = length Cs" with assms[unfolded vars_partitioned_def] Fun_More.fun_merge[of "map vars_clause Cs" "nth \s"] obtain \ where \_p: "\i < length (map vars_clause Cs). \x \ map vars_clause Cs ! i. \ x = (\s ! i) x" by auto have "\i < length Cs. \S. S \# Cs ! i \ subst_cls S (\s ! i) = subst_cls S \" proof (rule, rule, rule, rule) fix i :: nat and S :: "('a, 'b) term literal multiset" assume "i < length Cs" and "S \# Cs ! i" then have "\v \ vars_clause S. (\s ! i) v = \ v" using vars_clause_mono[of S "Cs ! i"] \_p by auto then show "subst_cls S (\s ! i) = subst_cls S \" using same_on_vars_clause by auto qed then show "\\. \iS. S \# Cs ! i \ subst_cls S (\s ! i) = subst_cls S \" by auto qed lemma vars_in_instance_in_range_term: "vars_term (subst_atm_abbrev A \) \ Union (image vars_term (range \))" by (induction A) auto lemma vars_in_instance_in_range_lit: "vars_lit (subst_lit L \) \ Union (image vars_term (range \))" proof (induction L) case (Pos A) have "vars_term (A \ \) \ Union (image vars_term (range \))" using vars_in_instance_in_range_term[of A \] by blast then show ?case by auto next case (Neg A) have "vars_term (A \ \) \ Union (image vars_term (range \))" using vars_in_instance_in_range_term[of A \] by blast then show ?case by auto qed lemma vars_in_instance_in_range_cls: "vars_clause (subst_cls C \) \ Union (image vars_term (range \))" unfolding vars_clause_def subst_cls_def using vars_in_instance_in_range_lit[of _ \] by auto primrec renamings_apart :: "('f, nat) term clause list \ (('f, nat) subst) list" where "renamings_apart [] = []" | "renamings_apart (C # Cs) = (let \s = renamings_apart Cs in (\v. Var (v + Max (vars_clause_list (subst_cls_lists Cs \s) \ {0}) + 1)) # \s)" definition var_map_of_subst :: "('f, nat) subst \ nat \ nat" where "var_map_of_subst \ v = the_Var (\ v)" lemma len_renamings_apart: "length (renamings_apart Cs) = length Cs" by (induction Cs) (auto simp: Let_def) lemma renamings_apart_is_Var: "\\ \ set (renamings_apart Cs). \x. is_Var (\ x)" by (induction Cs) (auto simp: Let_def) lemma renamings_apart_inj: "\\ \ set (renamings_apart Cs). inj \" proof (induction Cs) case (Cons a Cs) then have "inj (\v. Var (Suc (v + Max (vars_clause_list (subst_cls_lists Cs (renamings_apart Cs)) \ {0}))))" by (meson add_right_imp_eq injI nat.inject term.inject(1)) then show ?case using Cons by (auto simp: Let_def) qed auto lemma finite_vars_clause[simp]: "finite (vars_clause x)" unfolding vars_clause_def by auto lemma finite_vars_clause_list[simp]: "finite (vars_clause_list Cs)" unfolding vars_clause_list_def by (induction Cs) auto lemma Suc_Max_notin_set: "finite X \ Suc (v + Max (insert 0 X)) \ X" by (metis Max.boundedE Suc_n_not_le_n empty_iff finite.insertI le_add2 vimageE vimageI vimage_Suc_insert_0) lemma vars_partitioned_Nil[simp]: "vars_partitioned []" unfolding vars_partitioned_def by auto lemma subst_cls_lists_Nil[simp]: "subst_cls_lists Cs [] = []" unfolding subst_cls_lists_def by auto lemma vars_clause_hd_partitioned_from_tl: assumes "Cs \[]" shows "vars_clause (hd (subst_cls_lists Cs (renamings_apart Cs))) \ vars_clause_list (tl (subst_cls_lists Cs (renamings_apart Cs))) = {}" using assms proof (induction Cs) case (Cons C Cs) define \' :: "nat \ nat" where "\' = (\v. (Suc (v + Max ((vars_clause_list (subst_cls_lists Cs (renamings_apart Cs))) \ {0}))))" define \ :: "nat \ ('a, nat) term" where "\ = (\v. Var (\' v))" - have "vars_clause (subst_cls C \) \ UNION (range \) vars_term" + have "vars_clause (subst_cls C \) \ \ (vars_term ` range \)" using vars_in_instance_in_range_cls[of C "hd (renamings_apart (C # Cs))"] \_def \'_def by (auto simp: Let_def) - moreover have "UNION (range \) vars_term + moreover have "\ (vars_term ` range \) \ vars_clause_list (subst_cls_lists Cs (renamings_apart Cs)) = {}" proof - have "range \' \ vars_clause_list (subst_cls_lists Cs (renamings_apart Cs)) = {}" unfolding \'_def using Suc_Max_notin_set by auto then show ?thesis unfolding \_def \'_def by auto qed ultimately have "vars_clause (subst_cls C \) \ vars_clause_list (subst_cls_lists Cs (renamings_apart Cs)) = {}" by auto then show ?case unfolding \_def \'_def unfolding subst_cls_lists_def by (simp add: Let_def subst_cls_lists_def) qed auto lemma vars_partitioned_renamings_apart: "vars_partitioned (subst_cls_lists Cs (renamings_apart Cs))" proof (induction Cs) case (Cons C Cs) { fix i :: nat and j :: nat assume ij: "i < Suc (length Cs)" "j < i" have "vars_clause (subst_cls_lists (C # Cs) (renamings_apart (C # Cs)) ! i) \ vars_clause (subst_cls_lists (C # Cs) (renamings_apart (C # Cs)) ! j) = {}" proof (cases i; cases j) fix j' :: nat assume i'j': "i = 0" "j = Suc j'" then show "vars_clause (subst_cls_lists (C # Cs) (renamings_apart (C # Cs)) ! i) \ vars_clause (subst_cls_lists (C # Cs) (renamings_apart (C # Cs)) ! j) = {}" using ij by auto next fix i' :: nat assume i'j': "i = Suc i'" "j = 0" have disjoin_C_Cs: "vars_clause (subst_cls_lists (C # Cs) (renamings_apart (C # Cs)) ! 0) \ vars_clause_list ((subst_cls_lists Cs (renamings_apart Cs))) = {}" using vars_clause_hd_partitioned_from_tl[of "C # Cs"] by (simp add: Let_def subst_cls_lists_def) { fix x assume asm: "x \ vars_clause (subst_cls_lists Cs (renamings_apart Cs) ! i')" then have "(subst_cls_lists Cs (renamings_apart Cs) ! i') \ set (subst_cls_lists Cs (renamings_apart Cs))" using i'j' ij unfolding subst_cls_lists_def by (metis Suc_less_SucD length_map len_renamings_apart length_zip min_less_iff_conj nth_mem) moreover from asm have "x \ vars_clause (subst_cls_lists Cs (renamings_apart Cs) ! i')" using i'j' ij unfolding subst_cls_lists_def by simp ultimately have "\D \ set (subst_cls_lists Cs (renamings_apart Cs)). x \ vars_clause D" by auto } then have "vars_clause (subst_cls_lists Cs (renamings_apart Cs) ! i') \ Union (set (map vars_clause ((subst_cls_lists Cs (renamings_apart Cs)))))" by auto then have "vars_clause (subst_cls_lists (C # Cs) (renamings_apart (C # Cs)) ! 0) \ vars_clause (subst_cls_lists Cs (renamings_apart Cs) ! i') = {}" using disjoin_C_Cs unfolding vars_clause_list_def by auto moreover have "subst_cls_lists Cs (renamings_apart Cs) ! i' = subst_cls_lists (C # Cs) (renamings_apart (C # Cs)) ! i" using i'j' ij unfolding subst_cls_lists_def by (simp add: Let_def) ultimately show "vars_clause (subst_cls_lists (C # Cs) (renamings_apart (C # Cs)) ! i) \ vars_clause (subst_cls_lists (C # Cs) (renamings_apart (C # Cs)) ! j) = {}" using i'j' by (simp add: Int_commute) next fix i' :: nat and j' :: nat assume i'j': "i = Suc i'" "j = Suc j'" have "i' j'" using \i = Suc i'\ \j = Suc j'\ ij by blast ultimately have "vars_clause (subst_cls_lists Cs (renamings_apart Cs) ! i') \ vars_clause (subst_cls_lists Cs (renamings_apart Cs) ! j') = {}" using Cons unfolding vars_partitioned_def by auto then show "vars_clause (subst_cls_lists (C # Cs) (renamings_apart (C # Cs)) ! i) \ vars_clause (subst_cls_lists (C # Cs) (renamings_apart (C # Cs)) ! j) = {}" unfolding i'j' by (simp add: subst_cls_lists_def Let_def) next assume \i = 0\ and \j = 0\ then show \vars_clause (subst_cls_lists (C # Cs) (renamings_apart (C # Cs)) ! i) \ vars_clause (subst_cls_lists (C # Cs) (renamings_apart (C # Cs)) ! j) = {}\ using ij by auto qed } then show ?case unfolding vars_partitioned_def by (metis (no_types, lifting) Int_commute Suc_lessI len_renamings_apart length_map length_nth_simps(2) length_zip min.idem nat.inject not_less_eq subst_cls_lists_def) qed auto interpretation substitution "(\)" "Var :: _ \ ('f, nat) term" "(\\<^sub>s)" renamings_apart "Fun undefined" proof (standard) show "\A. A \ Var = A" by auto next show "\A \ \. A \ \ \\<^sub>s \ = A \ \ \ \" by auto next show "\\ \. (\A. A \ \ = A \ \) \ \ = \" by (simp add: subst_term_eqI) next fix C :: "('f, nat) term clause" fix \ assume "is_ground_cls (subst_cls C \)" then have ground_atms_\: "\v. v \ vars_clause C \ is_ground_atm (\ v)" by (meson is_ground_cls_is_ground_on_var) define some_ground_trm :: "('f, nat) term" where "some_ground_trm = (Fun undefined [])" have ground_trm: "is_ground_atm some_ground_trm" unfolding is_ground_atm_def some_ground_trm_def by auto define \ where "\ = (\v. if v \ vars_clause C then \ v else some_ground_trm)" then have \_\: "\v \ vars_clause C. \ v = \ v" unfolding \_def by auto have all_ground_\: "is_ground_atm (\ v)" for v proof (cases "v \ vars_clause C") case True then show ?thesis using ground_atms_\ \_\ by auto next case False then show ?thesis unfolding \_def using ground_trm by auto qed have "is_ground_subst \" unfolding is_ground_subst_def proof fix A show "is_ground_atm (subst_atm_abbrev A \)" proof (induction A) case (Var v) then show ?case using all_ground_\ by auto next case (Fun f As) then show ?case using all_ground_\ by (simp add: is_ground_atm_def) qed qed moreover have "\v \ vars_clause C. \ v = \ v" using \_\ unfolding vars_clause_list_def by blast then have "subst_cls C \ = subst_cls C \" using same_on_vars_clause by auto ultimately show "\\. is_ground_subst \ \ subst_cls C \ = subst_cls C \" by auto next fix Cs :: "('f, nat) term clause list" show "length (renamings_apart Cs) = length Cs" using len_renamings_apart by auto next fix Cs :: "('f, nat) term clause list" fix \ :: "nat \ ('f, nat) Term.term" assume \_renaming: "\ \ set (renamings_apart Cs)" { have inj_is_renaming: "\\ :: ('f, nat) subst. (\x. is_Var (\ x)) \ inj \ \ is_renaming \" proof - fix \ :: "('f, nat) subst" fix x assume is_var_\: "\x. is_Var (\ x)" assume inj_\: "inj \" define \' where "\' = var_map_of_subst \" have \: "\ = Var \ \'" unfolding \'_def var_map_of_subst_def using is_var_\ by auto from is_var_\ inj_\ have "inj \'" unfolding is_renaming_def unfolding subst_domain_def inj_on_def \'_def var_map_of_subst_def by (metis term.collapse(1)) then have "inv \' \ \' = id" using inv_o_cancel[of \'] by simp then have "Var \ (inv \' \ \') = Var" by simp then have "\x. (Var \ (inv \' \ \')) x = Var x" by metis then have "\x. ((Var \ \') \\<^sub>s (Var \ (inv \'))) x = Var x" unfolding subst_compose_def by auto then have "\ \\<^sub>s (Var \ (inv \')) = Var" using \ by auto then show "is_renaming \" unfolding is_renaming_def by blast qed then have "\\ \ (set (renamings_apart Cs)). is_renaming \" using renamings_apart_is_Var renamings_apart_inj by blast } then show "is_renaming \" using \_renaming by auto next fix Cs :: "('f, nat) term clause list" have "vars_partitioned (subst_cls_lists Cs (renamings_apart Cs))" using vars_partitioned_renamings_apart by auto then show "var_disjoint (subst_cls_lists Cs (renamings_apart Cs))" using vars_partitioned_var_disjoint by auto next show "\\ As Bs. Fun undefined As \ \ = Fun undefined Bs \ map (\A. A \ \) As = Bs" by simp next show "wfP (strictly_generalizes_atm :: ('f, 'v) term \ _ \ _)" unfolding wfP_def by (rule wf_subset[OF wf_subsumes]) (auto simp: strictly_generalizes_atm_def generalizes_atm_def term_subsumable.subsumes_def subsumeseq_term.simps) qed fun pairs :: "'a list \ ('a \ 'a) list" where "pairs (x # y # xs) = (x, y) # pairs (y # xs)" | "pairs _ = []" derive compare "term" derive compare "literal" lemma class_linorder_compare: "class.linorder (le_of_comp compare) (lt_of_comp compare)" apply standard apply (simp_all add: lt_of_comp_def le_of_comp_def split: order.splits) apply (metis comparator.sym comparator_compare invert_order.simps(1) order.distinct(5)) apply (metis comparator_compare comparator_def order.distinct(5)) apply (metis comparator.sym comparator_compare invert_order.simps(1) order.distinct(5)) by (metis comparator.sym comparator_compare invert_order.simps(2) order.distinct(5)) context begin interpretation compare_linorder: linorder "le_of_comp compare" "lt_of_comp compare" by (rule class_linorder_compare) definition Pairs where "Pairs AAA = concat (compare_linorder.sorted_list_of_set ((pairs \ compare_linorder.sorted_list_of_set) ` AAA))" lemma unifies_all_pairs_iff: "(\p \ set (pairs xs). fst p \ \ = snd p \ \) \ (\a \ set xs. \b \ set xs. a \ \ = b \ \)" proof (induct xs rule: pairs.induct) case (1 x y xs) then show ?case unfolding pairs.simps list.set ball_Un ball_simps simp_thms fst_conv snd_conv by metis qed simp_all lemma in_pair_in_set: assumes "(A,B) \ set ((pairs As))" shows "A \ set As \ B \ set As" using assms proof (induction As) case (Cons A As) note Cons_outer = this show ?case proof (cases As) case Nil then show ?thesis using Cons_outer by auto next case (Cons B As') then show ?thesis using Cons_outer by auto qed qed auto lemma in_pairs_sorted_list_of_set_in_set: assumes "finite AAA" "\AA \ AAA. finite AA" "AB_pairs \ (pairs \ compare_linorder.sorted_list_of_set) ` AAA" and "(A :: _ :: compare, B) \ set AB_pairs" shows "\AA. AA \ AAA \ A \ AA \ B \ AA" proof - from assms have "AB_pairs \ (pairs \ compare_linorder.sorted_list_of_set) ` AAA" by auto then obtain AA where AA_p: "AA \ AAA \ (pairs \ compare_linorder.sorted_list_of_set) AA = AB_pairs" by auto have "(A, B) \ set (pairs (compare_linorder.sorted_list_of_set AA))" using AA_p[] assms(4) by auto then have "A \ set (compare_linorder.sorted_list_of_set AA)" and "B \ set (compare_linorder.sorted_list_of_set AA)" using in_pair_in_set[of A] by auto then show ?thesis using assms(2) AA_p by auto qed lemma unifiers_Pairs: assumes "finite AAA" and "\AA \ AAA. finite AA" shows "unifiers (set (Pairs AAA)) = {\. is_unifiers \ AAA}" proof (rule; rule) fix \ :: "('a, 'b) subst" assume asm: "\ \ unifiers (set (Pairs AAA))" have "\AA. AA \ AAA \ card (AA \\<^sub>s\<^sub>e\<^sub>t \) \ Suc 0" proof - fix AA :: "('a, 'b) term set" assume asm': "AA \ AAA" then have "\p \ set (pairs (compare_linorder.sorted_list_of_set AA)). subst_atm_abbrev (fst p) \ = subst_atm_abbrev (snd p) \" using assms asm unfolding Pairs_def by auto then have "\A \ AA. \B \ AA. subst_atm_abbrev A \ = subst_atm_abbrev B \" using assms asm' unfolding unifies_all_pairs_iff using compare_linorder.sorted_list_of_set by blast then show "card (AA \\<^sub>s\<^sub>e\<^sub>t \) \ Suc 0" by (smt imageE card.empty card_Suc_eq card_mono finite.intros(1) finite_insert le_SucI singletonI subsetI) qed then show "\ \ {\. is_unifiers \ AAA}" using assms by (auto simp: is_unifiers_def is_unifier_def subst_atms_def) next fix \ :: "('a, 'b) subst" assume asm: "\ \ {\. is_unifiers \ AAA}" { fix AB_pairs A B assume "AB_pairs \ set (compare_linorder.sorted_list_of_set ((pairs \ compare_linorder.sorted_list_of_set) ` AAA))" and "(A, B) \ set AB_pairs" then have "\AA. AA \ AAA \ A \ AA \ B \ AA" using assms by (simp add: in_pairs_sorted_list_of_set_in_set) then obtain AA where a: "AA \ AAA" "A \ AA" "B \ AA" by blast from a assms asm have card_AA_\: "card (AA \\<^sub>s\<^sub>e\<^sub>t \) \ Suc 0" unfolding is_unifiers_def is_unifier_def subst_atms_def by auto have "subst_atm_abbrev A \ = subst_atm_abbrev B \" proof (cases "card (AA \\<^sub>s\<^sub>e\<^sub>t \) = Suc 0") case True moreover have "subst_atm_abbrev A \ \ AA \\<^sub>s\<^sub>e\<^sub>t \" using a assms asm card_AA_\ by auto moreover have "subst_atm_abbrev B \ \ AA \\<^sub>s\<^sub>e\<^sub>t \" using a assms asm card_AA_\ by auto ultimately show ?thesis using a assms asm card_AA_\ by (metis (no_types, lifting) card_Suc_eq singletonD) next case False then have "card (AA \\<^sub>s\<^sub>e\<^sub>t \) = 0" using a assms asm card_AA_\ by arith then show ?thesis using a assms asm card_AA_\ by auto qed } then show "\ \ unifiers (set (Pairs AAA))" unfolding Pairs_def unifiers_def by auto qed end definition "mgu_sets AAA = map_option subst_of (unify (Pairs AAA) [])" interpretation mgu "(\)" "Var :: _ \ ('f :: compare, nat) term" "(\\<^sub>s)" "Fun undefined" renamings_apart mgu_sets proof fix AAA :: "('a :: compare, nat) term set set" and \ :: "('a, nat) subst" assume fin: "finite AAA" "\AA \ AAA. finite AA" and "mgu_sets AAA = Some \" then have "is_imgu \ (set (Pairs AAA))" using unify_sound unfolding mgu_sets_def by blast then show "is_mgu \ AAA" unfolding is_imgu_def is_mgu_def unifiers_Pairs[OF fin] by auto next fix AAA :: "('a :: compare, nat) term set set" and \ :: "('a, nat) subst" assume fin: "finite AAA" "\AA \ AAA. finite AA" and "is_unifiers \ AAA" then have "\ \ unifiers (set (Pairs AAA))" unfolding is_mgu_def unifiers_Pairs[OF fin] by auto then show "\\. mgu_sets AAA = Some \" using unify_complete unfolding mgu_sets_def by blast qed derive linorder prod derive linorder list end diff --git a/thys/Graph_Saturation/CombinedCorrectness.thy b/thys/Graph_Saturation/CombinedCorrectness.thy --- a/thys/Graph_Saturation/CombinedCorrectness.thy +++ b/thys/Graph_Saturation/CombinedCorrectness.thy @@ -1,234 +1,234 @@ section \Combined correctness\ text \This section does not correspond to any theorems in the paper. However, the main correctness proof is not a theorem in the paper either. As the paper sets out to prove that we can decide entailment and consistency, this file shows how to combine the results so far and indeed establish those properties. \ theory CombinedCorrectness imports GraphRewriting StandardRules begin (* a somewhat concrete function to get the model if one exists *) definition the_model where "the_model C Rs - \ let L = fst ` UNION Rs (edges o snd) \ {S_Bot,S_Top,S_Idt} \ S_Const ` C; + \ let L = fst ` \ ((edges o snd) ` Rs) \ {S_Bot,S_Top,S_Idt} \ S_Const ` C; Rules = Rs \ (standard_rules C L); sel = non_constructive_selector Rules in the_lcg sel Rules (0,{})" definition entailment_model where "entailment_model C Rs init - \ let L = fst ` UNION Rs (edges o snd) \ {S_Bot,S_Top,S_Idt} \ S_Const ` C \ fst ` edges init; + \ let L = fst ` \ ((edges o snd) ` Rs) \ {S_Bot,S_Top,S_Idt} \ S_Const ` C \ fst ` edges init; Rules = Rs \ (standard_rules C L); sel = non_constructive_selector Rules in the_lcg sel Rules (card (vertices init),edges init)" abbreviation check_consistency where "check_consistency C Rs \ getRel S_Bot (the_model C Rs) = {}" abbreviation check_entailment where "check_entailment C Rs R \ let mdl = entailment_model C Rs (translation (fst R)) in (0,1) \ :mdl:\snd R\ \ getRel S_Bot mdl \ {}" definition transl_rules where - "transl_rules T = UNION T (\ (x,y). {(translation x, translation (A_Int x y)),(translation y, translation (A_Int y x))})" + "transl_rules T = (\(x, y)\T. {(translation x, translation (A_Int x y)), (translation y, translation (A_Int y x))})" lemma gr_transl_rules: "x \ transl_rules T \ graph_rule x" using graph_rule_translation unfolding transl_rules_def by blast term entails lemma check_consistency: assumes "finite T" "finite C" shows "check_consistency C (transl_rules T) \ consistent (t::nat itself) C T" (is "?lhs = ?rhs") proof - from assms(1) have fin_t:"finite (transl_rules T)" unfolding transl_rules_def by fast define L where - "L = fst ` UNION (transl_rules T) (edges \ snd) \ {S_Bot,S_Top,S_Idt} \ S_Const ` C" - have "finite (UNION (transl_rules T) (edges \ snd))" using fin_t gr_transl_rules by auto + "L = fst ` \ ((edges \ snd) ` transl_rules T) \ {S_Bot,S_Top,S_Idt} \ S_Const ` C" + have "finite (\ ((edges \ snd) ` transl_rules T))" using fin_t gr_transl_rules by auto hence fin_l:"finite L" unfolding L_def using assms(2) by auto define Rules where "Rules = transl_rules T \ standard_rules C L" hence fin_r:"finite Rules" using assms(2) fin_t fin_l unfolding standard_rules_def by auto - have incl_L:"fst ` UNION Rules (edges o snd) \ L" + have incl_L:"fst ` (\ ((edges o snd) ` Rules)) \ L" unfolding L_def Rules_def by (auto elim:standard_rules_edges) have "\R\transl_rules T. graph_rule R" using gr_transl_rules by blast moreover have "\R\ constant_rules C. graph_rule R" using constant_rules_graph_rule by auto moreover have "\R\ identity_rules L. graph_rule R" using identity_rules_graph_rule by auto moreover have "\R\ {top_rule S_Top,nonempty_rule}. graph_rule R" using are_rules(1,2) by fastforce ultimately have gr:"set_of_graph_rules Rules" unfolding set_of_graph_rules_def Rules_def ball_Un standard_rules_def by blast define sel where "sel = non_constructive_selector Rules" hence sel:"valid_selector Rules sel" using gr non_constructive_selector by auto define cfg where "cfg = the_lcg sel Rules (0, {})" have cfg:"cfg = the_model C (transl_rules T)" unfolding cfg_def sel_def Rules_def L_def the_model_def Let_def.. have cfg_c:"least_consequence_graph TYPE('a + nat) Rules (graph_of (0,{})) cfg" unfolding cfg_def by (rule lcg_through_make_step[OF fin_r gr _ sel],auto) hence cfg_sdt:"maintainedA (standard_rules C L) cfg" and cfg_g: "graph cfg" and cfg_l:"least TYPE('a + nat) Rules (graph_of (0, {})) cfg" and cfg_m:"r \ transl_rules T \ maintained r cfg" for r unfolding Rules_def least_consequence_graph_def by auto have cfg_lbl:"fst ` edges cfg \ L" unfolding cfg_def by (auto intro!: the_lcg_edges[OF sel incl_L]) have d1: "?lhs \ ?rhs" proof - assume ?lhs from maintained_standard_noconstants[OF cfg_sdt cfg_g cfg_lbl this[folded cfg]] obtain G :: "('a Standard_Constant, 'a + nat) labeled_graph" where G_std:"standard' C G" and m:"maintained r cfg \ maintained r G" for r :: "('a Standard_Constant, nat) Graph_PreRule" by blast hence g:"graph G" unfolding standard_def by auto have "(a,b)\T \ G \ (a,b)" for a b proof(subst eq_as_subsets,standard) assume a:"(a,b)\T" from a cfg_m[unfolded transl_rules_def,THEN m] show "G \ a \ b" by (subst maintained_holds_iff[OF g,symmetric]) blast from a cfg_m[unfolded transl_rules_def,THEN m] show "G \ b \ a" by (subst maintained_holds_iff[OF g,symmetric]) blast qed hence h:"(\S\T. G \ S)" by auto with G_std show ?rhs unfolding model_def by blast qed have d2: "\ ?lhs \ ?rhs \ False" proof - assume "\ ?lhs" then obtain a b where ab:"(S_Bot,a,b) \ edges cfg" "a \ vertices cfg" "b \ vertices cfg" using cfg_g unfolding cfg getRel_def by auto assume ?rhs then obtain G :: "('a Standard_Constant, 'a + nat) labeled_graph" where G:"model C G T" by auto with model_def have std:"standard' C G" and holds:"\S\T. G \ S" by fast+ hence g:"graph G" unfolding standard_def by auto from maintained_holds_iff[OF g] holds have "maintainedA (transl_rules T) G" unfolding transl_rules_def by auto hence mnt:"maintainedA Rules G" unfolding Rules_def using standard_maintains_rules[OF std] by auto from consequence_graphI[OF _ _ g] gr[unfolded set_of_graph_rules_def] mnt have cg:"consequence_graph Rules G" by fast with cfg_l[unfolded least_def] have mtd:"maintained (graph_of (0, {}), cfg) G" by blast have "graph_homomorphism (fst (graph_of (0::nat, {}), cfg)) G {}" unfolding graph_homomorphism_def using g by auto with mtd maintained_def have "extensible (graph_of (0, {}), cfg) G {}" by auto then obtain g where "edges (map_graph g cfg) \ edges G" "vertices cfg = Domain g" unfolding extensible_def graph_homomorphism_def2 graph_union_iff by auto hence "\ a b. (S_Bot,a,b) \ edges G" using ab unfolding edge_preserving by auto thus False using std unfolding standard_def getRel_def by auto qed from d1 d2 show ?thesis by metis qed lemma check_entailment: assumes "finite T" "finite C" shows "check_entailment C (transl_rules T) S \ entails (t::nat itself) C T (fst S, (A_Int (fst S) (snd S)))" (is "?lhs = ?rhs") proof - from assms(1) have fin_t:"finite (transl_rules T)" unfolding transl_rules_def by fast define R where "R = transl_rule S" define init where "init = (card (vertices (fst R)), edges (fst R))" have gi[intro]:"graph (graph_of init)" and init:"graph_of init = translation (fst S)" using verts_in_translation[of "fst S"] unfolding inv_translation_def init_def R_def by auto define Rs where "Rs = transl_rules T" define L where - "L = fst ` UNION Rs (edges \ snd) \ {S_Bot,S_Top,S_Idt} \ S_Const ` C \ fst ` edges (fst R)" - have "finite (UNION (transl_rules T) (edges \ snd))" using fin_t gr_transl_rules by auto + "L = fst ` (\ ((edges o snd) ` Rs)) \ {S_Bot,S_Top,S_Idt} \ S_Const ` C \ fst ` edges (fst R)" + have "finite (\ ((edges \ snd) ` transl_rules T))" using fin_t gr_transl_rules by auto hence fin_l:"finite L" unfolding L_def Rs_def R_def using assms(2) by auto have fin_t:"finite Rs" using fin_t Rs_def by auto define Rules where "Rules = Rs \ standard_rules C L" hence fin_r:"finite Rules" using assms(2) fin_t fin_l unfolding standard_rules_def by auto - have incl_L:"fst ` UNION Rules (edges o snd) \ L" "fst ` snd init \ L" + have incl_L:"fst ` (\ ((edges o snd) ` Rules)) \ L" "fst ` snd init \ L" unfolding L_def Rules_def init_def by (auto elim:standard_rules_edges) have "\R\transl_rules T. graph_rule R" using gr_transl_rules by blast moreover have "\R\ constant_rules C. graph_rule R" using constant_rules_graph_rule by auto moreover have "\R\ identity_rules L. graph_rule R" using identity_rules_graph_rule by auto moreover have "\R\ {top_rule S_Top,nonempty_rule}. graph_rule R" using are_rules(1,2) by fastforce ultimately have gr:"set_of_graph_rules Rules" unfolding set_of_graph_rules_def Rules_def ball_Un standard_rules_def Rs_def by blast define sel where "sel = non_constructive_selector Rules" hence sel:"valid_selector Rules sel" using gr non_constructive_selector by auto define cfg where "cfg = the_lcg sel Rules init" have cfg:"cfg = entailment_model C Rs (fst R)" unfolding cfg_def sel_def Rules_def L_def entailment_model_def Let_def init_def.. have cfg_c:"least_consequence_graph TYPE('a + nat) Rules (graph_of init) cfg" unfolding cfg_def by (rule lcg_through_make_step[OF fin_r gr gi sel]) hence cfg_sdt:"maintainedA (standard_rules C L) cfg" and cfg_g: "graph cfg" and cfg_l:"least TYPE('a + nat) Rules (graph_of init) cfg" and cfg_m:"r \ Rs \ maintained r cfg" for r unfolding Rules_def least_consequence_graph_def by auto have cfg_lbl:"fst ` edges cfg \ L" unfolding cfg_def by (auto intro!: the_lcg_edges[OF sel incl_L]) have "(0,1) \ :translation (fst S):\fst S\" by (fact translation_self) hence "(0,1) \ :graph_of init:\fst S\" unfolding init by auto from subgraph_semantics[OF _ this] cfg_l[unfolded least_def] have cfg_fst:"(0,1) \ :cfg:\fst S\" unfolding cfg_def by auto from semantics_in_vertices[OF cfg_g this] have cfg_01:"0 \ vertices cfg" "1 \ vertices cfg" "(0,1)\vertices cfg\vertices cfg" by auto have d1: "\ ?lhs \ ?rhs \ False" proof - assume "\ ?lhs" hence gr:"(0,1) \ :cfg:\snd S\" "getRel S_Bot cfg = {}" unfolding entailment_model_def cfg R_def Rs_def Let_def by auto from maintained_standard_noconstants[OF cfg_sdt cfg_g cfg_lbl gr(2)] obtain G :: "('a Standard_Constant, 'a + nat) labeled_graph" and f g where fg:"G = map_graph_fn G (f \ g)" and f:"G = map_graph_fn cfg f" "subgraph (map_graph_fn G g) cfg" and G_std:"standard' C G" and m:"\ r:: ('a Standard_Constant, nat) Graph_PreRule. maintained r cfg \ maintained r G" and e:"\ x y e. x \ vertices cfg \ y \ vertices cfg \ (g (f x), g (f y)) \ :map_graph_fn G g:\e\ \ (x,y) \ :cfg:\e\" by clarify blast (* just blast also works, but clarify blast is much faster *) hence g:"graph G" unfolding standard_def by auto have "(a,b)\T \ G \ (a,b)" for a b apply(subst eq_as_subsets) using cfg_m[unfolded transl_rules_def Rs_def,THEN m] unfolding maintained_holds_iff[OF g,symmetric] by blast hence h:"(\S\T. G \ S)" by auto assume "?rhs" from this[unfolded entails_def model_def] G_std h have "G \ fst S \ snd S" by blast with cfg_fst cfg_g f(1) have "(f 0, f 1) \ :G:\snd S\" by auto then have "(g (f 0), g (f 1)) \ :map_graph_fn G g:\snd S\" using map_graph_in[OF g] by auto with e cfg_01(1,2) gr(1) show "False" by auto qed have "?lhs \ model C G T \ (a,b) \ :G:\fst S\ \ (a,b) \ :G:\snd S\" for G :: "('a Standard_Constant, 'a + nat) labeled_graph" and a b proof - assume mod:"model C G T" from mod model_def have std:"standard' C G" and holds:"\S\T. G \ S" by fast+ hence g:"graph G" unfolding standard_def by auto with maintained_holds_iff[OF g] holds have "maintainedA Rs G" unfolding transl_rules_def Rs_def by auto hence mnt:"maintainedA Rules G" unfolding Rules_def using standard_maintains_rules[OF std] by auto from consequence_graphI[OF _ _ g] gr[unfolded set_of_graph_rules_def] mnt have cg:"consequence_graph Rules G" by fast with cfg_l[unfolded least_def] have mtd:"maintained (graph_of init, cfg) G" by auto assume ab:"(a,b) \ :G:\fst S\" hence abv:"a \ vertices G" "b \ vertices G" using semantics_in_vertices[OF g] by auto from ab translation[OF g] init obtain f where f:"graph_homomorphism (graph_of init) G f" "(0, a) \ f \ (1, b) \ f" by auto from maintainedD2[OF mtd f(1)] obtain g where g:"graph_homomorphism cfg G g" and "f \ g" by blast with f have g01:"(0, a) \ g" "(1, b) \ g" by auto assume ?lhs then consider (maintained) "(0,1) \ :cfg:\snd S\" | (no_models) ":cfg:\\\ \ {}" using cfg_g unfolding cfg entailment_model_def Let_def Rs_def R_def by auto thus "(a,b) \ :G:\snd S\" proof(cases) case maintained from graph_homomorphism_semantics[OF g maintained g01] show ?thesis. next case no_models from graph_homomorphism_nonempty[OF g no_models] have "getRel S_Bot G \ {}" by auto hence False using std unfolding standard_def by auto thus ?thesis by auto qed qed hence d2:"?lhs \ ?rhs" unfolding entails_def by auto from d1 d2 show ?thesis by metis qed end \ No newline at end of file diff --git a/thys/Graph_Saturation/GraphRewriting.thy b/thys/Graph_Saturation/GraphRewriting.thy --- a/thys/Graph_Saturation/GraphRewriting.thy +++ b/thys/Graph_Saturation/GraphRewriting.thy @@ -1,552 +1,552 @@ section \Graph rewriting and saturation\ text \Here we describe graph rewriting, again without connecting it to semantics.\ theory GraphRewriting imports RulesAndChains "HOL-Library.Infinite_Set" begin (* Algorithm 1 on page 16 *) text \To describe Algorithm 1, we give a single step instead of the recursive call. This allows us to reason about its effect without dealing with non-termination. We define a worklist, saying what work can be done. A valid selection needs to be made in order to ensure fairness. To do a step, we define the function extend, and use it in @{term make_step}. A function that always makes a valid selection is used in this step. \ abbreviation graph_of where "graph_of \ \ X. LG (snd X) {0.. nat" where "nextMax x \ if x = {} then 0 else Suc (Max x)" lemma nextMax_max[intro]: assumes "finite x" "v \ x" shows "v < nextMax x" "v \ nextMax x" using Max.coboundedI[OF assms] assms(2) unfolding nextMax_def by auto definition worklist :: "nat \ ('a \ nat \ nat) set \ (('a, 'b) labeled_graph \ ('a, 'b) labeled_graph) set \ (nat \ ('a, 'b) Graph_PreRule \ ('b \ nat) set) set" where "worklist G Rs \ let G = graph_of G in {(N,R,f). R\ Rs \ graph_homomorphism (fst R) G f \ N = nextMax (Range f) \ \ extensible R G f }" definition valid_selection where "valid_selection Rs G R f \ let wl = worklist G Rs in (nextMax (Range f), R,f) \ wl \ (\ (N,_) \ wl. N \ nextMax (Range f)) \ graph_rule R" lemma valid_selection_exists: assumes "worklist G Rs \ {}" "set_of_graph_rules Rs" shows "\L R f. valid_selection Rs G R f" proof - define wl where "wl = worklist G Rs" hence wl_ne:"wl \ {}" using assms(1) by auto let ?N = "LEAST N. N \ Domain wl" from wl_ne have "\ N. N \ Domain wl" by auto with LeastI2 have "?N \ Domain wl" by metis then obtain L R f where NLRf:"(?N,(L,R),f)\wl" by auto hence N_def:"?N = nextMax (Range f)" and in_Rs: "(L,R) \ Rs" unfolding wl_def worklist_def Let_def by auto from Least_le wl_ne Domain.intros case_prodI2 have min:"(\ (N',_) \ wl. N' \ ?N)" by (metis (no_types, lifting)) from in_Rs have "finite_graph R" "subgraph L R" using assms(2)[unfolded set_of_graph_rules_def] by auto with min NLRf N_def show ?thesis unfolding wl_def[symmetric] valid_selection_def by auto qed definition valid_selector where "valid_selector Rs selector \ \ G. (worklist G Rs \ {} \ (\ (R,f)\UNIV. selector G = Some (R,f) \ valid_selection Rs G R f)) \ (worklist G Rs = {} \ selector G = None)" lemma valid_selectorD[dest]: assumes "valid_selector Rs selector" shows "worklist G Rs = {} \ selector G = None" "selector G = Some (R,f) \ valid_selection Rs G R f" using assms[unfolded valid_selector_def,rule_format,of G] by (cases "worklist G Rs = {}",auto) text \The following gives a valid selector. This selector is not useful as concrete implementation, because it used the choice operation.\ definition non_constructive_selector where "non_constructive_selector Rs G \ let wl = worklist G Rs in if wl = {} then None else Some (SOME (R,f). valid_selection Rs G R f) " lemma non_constructive_selector: assumes "set_of_graph_rules Rs" shows "valid_selector Rs (non_constructive_selector Rs)" unfolding valid_selector_def proof((clarify,standard;clarify),goal_cases) case (1 n E) let ?x = "(SOME (R, f). valid_selection Rs (n, E) R f)" from valid_selection_exists[OF 1 assms] have "\ R f. valid_selection Rs (n, E) R f" by auto hence "\ x. valid_selection Rs (n, E) (fst x) (snd x)" by auto from this prod.case_eq_if tfl_some have "\ valid_selection Rs (n, E) (fst ?x) (snd ?x) \ False" by (metis (mono_tags, lifting)) thus ?case unfolding non_constructive_selector_def Let_def using 1 by (auto simp:prod_eq_iff) qed (auto simp:non_constructive_selector_def) text \The following is used to make a weak pushout step. In the paper, we aren't too specific on how this should be done. Here we are. We work on natural numbers in order to be able to pick fresh elements easily. \ definition extend :: "nat \ ('b, 'a::linorder) Graph_PreRule \ ('a \ nat) set \ ('a \ nat) set" where "extend n R f \ f \ (let V_new = sorted_list_of_set (vertices (snd R) - vertices (fst R)) in set (zip V_new [n..<(n+length V_new)]))" lemma nextMax_set[simp]: assumes "sorted xs" shows "nextMax (set xs) = (if xs = Nil then 0 else Suc (last xs))" using assms proof(induct xs) case Nil show ?case unfolding nextMax_def by auto next case (Cons a list) hence "list \ [] \ fold max list a = last list" using list_sorted_max by (metis last.simps) thus ?case unfolding nextMax_def Max.set_eq_fold by auto qed lemma nextMax_Un_eq[simp]: "finite x \ finite y \ nextMax (x \ y) = max (nextMax x) (nextMax y)" unfolding nextMax_def using Max_Un by auto lemma extend: (* extensible into the new graph *) assumes "graph_homomorphism (fst R) (LG E {0.. extend n R f" defines "G' \ LG ((on_triple g `` (edges (snd R))) \ E) {0.. g" "subgraph (LG E {0.. vertices (snd R)" and gr_R:"graph (snd R)" unfolding subgraph_def graph_union_iff by auto hence fin_R_L[simp]:"finite ?R_L" and fin_L:"finite (vertices (fst R))" using finite_subset by auto from assms have f_dom:"Domain f = vertices (fst R)" and f_uni:"univalent f" unfolding graph_homomorphism_def by auto from assms[unfolded graph_homomorphism_def] have "f `` vertices (fst R) \ vertices (LG E {0.. {0.. n" using f_ran Max_in[OF fin_f] by (simp add:nextMax_def Suc_leI subset_eq) have "x \ Domain ?g \ x \ Domain f" for x unfolding f_dom Let_def by auto hence g_not_f:"(x,y) \ ?g \ (x,z) \ f" for x y z by auto have uni_g':"univalent ?g" "univalent (converse ?g)" unfolding Let_def by auto with f_uni have uni_g:"univalent g" by (auto simp:g_def extend_def g_not_f) from fin_g have "(a,b) \ g \ b < Suc (Max (Range g))" for a b unfolding less_Suc_eq_le by (rule Max.coboundedI) force hence "(a,b) \ g \ b < nextMax (Range g)" for a b unfolding nextMax_def by (cases "Range g = {}",auto) hence in_g:"(a,b) \ g \ b < max n (nextMax (Range g))" for a b by fastforce let ?G = "LG E {0.. E \ b < max n c" "(a, aa, b) \ E \ aa < max n c" for a aa b c by fastforce+ hence gr_G':"graph G'" unfolding G'_def restrict_def using in_g by auto show "subgraph (LG E {0.. g" by (auto simp:g_def extend_def) thus "agree_on (fst R) f g" using f_dom uni_g agree_on_subset equalityE by metis show "weak_universal t R ?G G' f g" proof fix a:: "('b \ 'x) set" fix b G assume a:"graph_homomorphism (snd R) G a" "graph_homomorphism ?G G b" "f O b \ a" hence univ_b:"univalent b" and univ_a:"univalent a" and rng_b:"Range b \ vertices G" and rng_a:"Range a \ vertices G" and ep_b:"edge_preserving b (edges (LG E {0.. b \ n \ y \ False" for y z using dom_b by (metis Domain.DomainI atLeastLessThan_iff not_less) have disj_doms:"Domain b \ Domain (?g\ O a) = {}" using help_dom_b unfolding Let_def by (auto dest!:set_zip_leftD) have "max n (nextMax (Range ?g)) = n + length (sorted_list_of_set ?R_L)" (is "_ = ?len") unfolding Let_def Range_snd set_map[symmetric] map_snd_zip[OF ln] nextMax_set[OF sorted_upt] by fastforce hence n_eq:"?len = max n (nextMax (Range g))" unfolding Range_snd[symmetric] g_def extend_def Range_Un_eq nextMax_Un_eq[OF fin_f fin_g'(2)] max.assoc[symmetric] max_absorb1[OF nextMax_f] by auto let ?h = "b \ ?g\ O a" have dg:"Domain (?g\) = {n.. vertices (fst R))" using dom_a subsLR by auto also have "\ = ?g `` ?R_L \ ?g `` vertices (fst R)" by auto also have "?g `` vertices (fst R) = {}" apply(rule Image_outside_Domain) unfolding Let_def Domain_set_zip[OF ln] by auto also have "?g `` ?R_L = Range ?g" apply(rule Image_Domain) unfolding Let_def Domain_set_zip[OF ln] by auto finally have dg2:"?g `` Domain a = {n.. O a) = {n.. vertices G" "(?g\ O a) `` vertices G' \ vertices G" using rng_a rng_b by auto hence v2: "?h `` vertices G' \ vertices G" by blast have v3: "univalent ?h" using disj_doms univalent_union[OF univ_b univalent_composes[OF uni_g'(2) univ_a]] by blast (* showing edge preservation *) { fix l x y x' y' assume a2:"(l,x,y) \ edges G'" "(x,x') \ ?h" "(y,y') \ ?h" have "(l,x',y') \ edges G" proof(cases "(l,x,y) \ edges ?G") case True with gr_G[THEN restrictD] have "x \ Domain b" "y \ Domain b" unfolding dom_b by auto hence "x \ Domain (converse ?g O a)" "y \ Domain (converse ?g O a)" using disj_doms by blast+ hence "(x,x') \ b" "(y,y') \ b" using a2 by auto with ep_b True show ?thesis unfolding edge_preserving by auto next have "g O ?h = f O b \ ?g O b \ ((f O ?g\) O a \ (?g O ?g\) O a)" unfolding g_def extend_def by blast also have "(?g O ?g\) = Id_on ?R_L" unfolding univalent_O_converse[OF uni_g'(2)] unfolding Let_def by auto also have "(f O ?g\) = {}" using f_ran unfolding Let_def by (auto dest!:set_zip_leftD) also have "?g O b = {}" using help_dom_b unfolding Let_def by (auto dest!:set_zip_rightD) finally have gOh:"g O ?h \ a" using a(3) by blast case False hence "(l,x,y) \ on_triple g `` edges (snd R)" using a2(1) unfolding G'_def by auto then obtain r_x r_y where r:"(l,r_x,r_y) \ edges (snd R)" "(r_x,x) \ g" "(r_y,y) \ g" by auto hence "(r_x,x') \ a" "(r_y,y') \ a" using gOh a2(2,3) by auto hence "(l,x',y') \ on_triple a `` edges (snd R)" using r(1) unfolding on_triple_def by auto thus ?thesis using ep_a unfolding edge_preserving by auto qed } hence v4: "edge_preserving ?h (edges G') (edges G)" by auto have "graph_homomorphism G' G ?h" by(fact graph_homomorphismI[OF v1 v2 v3 v4 gr_G' v6]) thus "\h. graph_homomorphism G' G h \ b \ h" by auto qed qed text \Showing that the extend function indeed creates a valid pushout.\ lemma selector_pushout: assumes "valid_selector Rs selector" "selector G'' = Some (R,f)" defines "G \ graph_of G''" assumes "graph G" defines "g \ extend (fst G'') R f" defines "G' \ LG (on_triple g `` edges (snd R) \ (snd G'')) {0.. g" "weak_universal t R G G' f g" using extend[OF igh[unfolded G_def],folded g_def,folded G'_def,folded G_def] igh(1) by auto thus ?thesis unfolding pushout_step_def by auto qed text \Making a single step in Algorithm 1. A prerequisite is that its first argument is a @{term valid_selector}.\ definition make_step where "make_step selector S \ case selector S of None \ S | Some (R,f) \ (let g = extend (fst S) R f in (max (fst S) (nextMax (Range g)), (on_triple g `` (edges (snd R))) \ (snd S)))" lemma WPC_through_make_step: assumes "set_of_graph_rules Rs" "graph (graph_of (X 0))" and makestep: "\ i. X (Suc i) = make_step selector (X i)" and selector: "valid_selector Rs selector" shows "Simple_WPC t Rs (\ i. graph_of (X i))" "chain (\ i. graph_of (X i))" proof note ms = makestep[unfolded make_step_def,rule_format] have gr:"graph (graph_of (X i))" for i proof(induct i) case (Suc i) then show ?case proof(cases "selector (X i)") case None then show ?thesis using ms Suc by auto next case (Some a) then obtain R f where Some:"selector (X i) = Some (R,f)" by fastforce then show ?thesis using ms[of i,unfolded Some Let_def] selector_pushout[OF selector Some Suc,of t ,unfolded pushout_step_def subgraph_def] by auto qed qed (fact assms) show "chain (\ i. graph_of (X i))" unfolding chain_def proof(clarify) fix i show "subgraph (graph_of (X i)) (graph_of (X (i + 1)))" proof(cases "selector (X i)") case None then show ?thesis using ms gr by (auto intro!:graph_homomorphismI) next case Some then obtain R f where Some:"selector (X i) = Some (R,f)" by fastforce then show ?thesis using ms selector_pushout[OF selector Some gr,of t] unfolding pushout_step_def Let_def by simp qed qed show "graph_of (X i) = graph_of (X (Suc i)) \ (\R\Rs. pushout_step t R (graph_of (X i)) (graph_of (X (Suc i))))" for i proof(cases "selector (X i)") case None thus ?thesis using ms by auto next case Some then obtain R f where Some:"selector (X i) = Some (R,f)" by fastforce hence "R \ Rs" using valid_selectorD(2)[OF selector,unfolded valid_selection_def worklist_def Let_def] by(cases R,blast) then show ?thesis using ms[of i,unfolded Some Let_def] selector_pushout[OF selector Some gr] unfolding make_step_def by auto qed qed (fact assms)+ lemma N_occurs_finitely_often: assumes "finite Rs" "set_of_graph_rules Rs" "graph (graph_of (X 0))" and makestep: "\ i. X (Suc i) = make_step selector (X i)" and selector: "valid_selector Rs selector" shows "finite {(R,f). \ i. R\ Rs \ graph_homomorphism (fst R) (graph_of (X i)) f \ nextMax (Range f) \ N}" (is "finite {(R,f).?P R f}") proof - have prod_eq : "(\ x \ {(x, y). A x y}. B x) \ (\ x. A (fst x) (snd x) \ B x)" "(x \ {(x, y). A x y}) \ (A (fst x) (snd x))" for A B x by auto let ?S = "{(R,f).?P R f}" let "?Q R f" = "Domain f = vertices (fst (R::('a, 'b) Graph_PreRule)) \ univalent f \ nextMax (Range f) \ N" have seteq:"(\R\Rs. {(R', f). R' = R \ ?Q R f}) = {(R,f). R \ Rs \ ?Q R f}" by auto have "\ R \ Rs. finite {(R',f). R' = R \ ?Q R f}" proof fix R assume "R \ Rs" hence fin:"finite (vertices (fst R))" using assms by auto hence fin2:"finite (Pow (vertices (fst R) \ {0..N}))" by auto have fin:"Domain x = vertices (fst R) \ univalent x \ finite (snd ` x)" for x:: "('b \ nat) set" using fin univalent_finite[of x] by simp hence "Domain f = vertices (fst R) \ univalent f \ (a,b) \ f \ nextMax (Range f) \ N \ b \ N" for f a b unfolding Range_snd using image_eqI nextMax_max(2) snd_conv order.trans by metis hence sub:"{f. ?Q R f} \ Pow (vertices (fst R) \ {0..N})" using nextMax_max[OF fin] by (auto simp:Range_snd image_def) from finite_subset[OF sub fin2] show "finite {(R',f). R' = R \ ?Q R f}" by auto qed from this[folded finite_UN[OF assms(1)],unfolded seteq] have fin:"finite {(R,f). R \ Rs \ ?Q R f}". have "?P R f \ R \ Rs \ ?Q R f" for R f unfolding graph_homomorphism_def by auto hence "?S \ {(R,f). R \ Rs \ ?Q R f}" unfolding subset_eq prod_eq by blast from finite_subset[OF this fin] show ?thesis by auto qed lemma inj_on_infinite: assumes "infinite A" "inj_on f A" "range f \ B" shows "infinite B" proof - from assms[unfolded infinite_iff_countable_subset] obtain g::"nat \ 'a" where g:"inj g \ range g \ A" by blast hence i:"inj (f o g)" using assms(2) using comp_inj_on inj_on_subset by blast have "range (f o g) \ B" using assms(3) by auto with i show ?thesis unfolding infinite_iff_countable_subset by blast qed lemma makestep_makes_selector_inj: assumes "selector (X y) = Some (R,f)" "selector (X x) = Some (R,f)" "valid_selector Rs selector" and step: "\ i. X (Suc i) = make_step selector (X i)" and chain:"chain (\ i. graph_of (X i))" shows "x = y" proof(rule ccontr) assume a:"x \ y" define x' y' where "x' \ min x y" "y' \ max x y" hence xy:"selector (X x') = Some (R, f)" "selector (X y') = Some (R, f)" "x' < y'" using assms(1,2) a unfolding min_def max_def by auto with valid_selectorD assms have "valid_selection Rs (X x') R f" "valid_selection Rs (X y') R f" by auto hence not_ex:"\ extensible R (graph_of (X y')) f" and hom:"graph_homomorphism (fst R) (graph_of (X x')) f" "graph_rule R" unfolding valid_selection_def Let_def worklist_def by auto have X:"X (Suc x') = (max (fst (X x')) (nextMax (Range (extend (fst (X x')) R f))), on_triple (extend (fst (X x')) R f) `` edges (snd R) \ snd (X x'))" unfolding step[unfolded make_step_def Let_def,rule_format] xy by auto let ?ex = "extend (fst (X x')) R f" have hom:"graph_homomorphism (snd R) (graph_of (X (Suc x'))) ?ex" and agr:"agree_on (fst R) f ?ex" using extend(1,2)[OF hom] unfolding X by auto from xy have "Suc x' \ y'" by auto with chain[unfolded chain_def2] have "subgraph (graph_of (X (Suc x'))) (graph_of (X y'))" by auto from subgraph_preserves_hom[OF this hom] have hom:"graph_homomorphism (snd R) (graph_of (X y')) ?ex". with agr have "extensible R (graph_of (X y')) f" unfolding extensible_def by auto thus False using not_ex by auto qed lemma fair_through_make_step: assumes "finite Rs" "set_of_graph_rules Rs" "graph (graph_of (X 0))" (* It should suffice to take infinitely many make_steps, rather than having every step be a make_step, but we focus on the algorithm as in the paper here *) and makestep: "\ i. X (Suc i) = make_step selector (X i)" and selector: "valid_selector Rs selector" shows "fair_chain Rs (\ i. graph_of (X i))" proof show chn:"chain (\i. graph_of (X i))" using WPC_through_make_step assms by blast fix R f i assume Rs:"R \ Rs" and h:"graph_homomorphism (fst R) (graph_of (X i)) f" hence R:"finite (vertices (snd R))" "subgraph (fst R) (snd R)" "finite (vertices (fst R))" using assms by auto hence f:"finite f" "finite (Range f)" "finite (Domain f)" "univalent f" using h unfolding graph_homomorphism_def Range_snd by auto define N where "N \ nextMax (Range f)" fix S let "?Q X' j" = " fst X' \ Rs \ graph_homomorphism (fst (fst X')) (graph_of (X (j+i))) (snd X') \ nextMax (Range (snd X')) \ N" let ?S = "{(R,f). \j. ?Q (R,f) j}" from assms(4) have "\ia. X (Suc ia + i) = make_step selector (X (ia + i))" by auto note r = assms(1,2) chain_then_restrict[OF chn] this assms(5) from N_occurs_finitely_often[of Rs "\ j. X (j + i)",OF r] have fin_S:"finite ?S" by auto { assume a:"\j. \ extensible R (graph_of (X j)) f" let "?P X' j" = "?Q X' j \ Some X' = selector (X (j+i))" { fix j let ?j = "j+i" have "?j \ i" by auto from subgraph_preserves_hom[OF chain[OF chn this] h] have h:"graph_homomorphism (fst R) (graph_of (X ?j)) f". have "\ extensible R (graph_of (X ?j)) f" using a by blast with h Rs have wl:"(nextMax (Range f),R,f) \ worklist (X ?j) Rs" unfolding worklist_def Let_def set_eq_iff by auto hence "worklist (X ?j) Rs \ {}" by auto with valid_selectorD[OF selector] obtain R' f' where sel:"Some (R',f') = selector (X ?j)" "valid_selection Rs (X ?j) R' f'" by auto hence max:"(nextMax (Range f'), R', f') \ worklist (X ?j) Rs" "(\(N, _)\worklist (X ?j) Rs. nextMax (Range f') \ N)" unfolding valid_selection_def Let_def by auto with wl have "nextMax (Range f') \ N" unfolding N_def by auto with max(1)[unfolded worklist_def Let_def mem_Collect_eq prod.case] sel(1) have "\ X'. ?P X' j" by (metis fst_conv snd_conv) } then obtain ch where ch:"\ j. ?P (ch j) j" by metis (* uses 'choice' internally *) have inj:"inj ch" proof fix x y assume "ch x = ch y" with ch[of x] ch[of y] have "selector (X (x + i)) = Some (ch x)" "selector (X (y + i)) = Some (ch x)" by auto with makestep_makes_selector_inj[OF _ _ selector makestep chn] have "x + i = y + i" by (cases "ch x",metis (full_types)) thus "x = y" by auto qed have "ch x \ ?S" for x using ch[of x] unfolding mem_Collect_eq by(intro case_prodI2) metis hence "range ch \ ?S" unfolding UNIV_def by(rule image_Collect_subsetI) with infinite_iff_countable_subset inj have "infinite ?S" by blast with fin_S have "False" by auto } thus "\j. extensible R (graph_of (X j)) f" by auto qed fun mk_chain where "mk_chain sel Rs init 0 = init" | "mk_chain sel Rs init (Suc n) = mk_chain sel Rs (make_step sel init) n" lemma mk_chain: "\ i. mk_chain sel Rs init (Suc i) = make_step sel (mk_chain sel Rs init i)" proof fix i show "mk_chain sel Rs init (Suc i) = make_step sel (mk_chain sel Rs init i)" by (induct i arbitrary:init,auto) qed text \Algorithm 1, abstractly.\ abbreviation the_lcg where "the_lcg sel Rs init \ chain_sup (\i. graph_of (mk_chain sel Rs init i))" lemma mk_chain_edges: assumes "valid_selector Rules sel" - "UNION Rules (edges \ snd) \ L \ UNIV" + "\ ((edges o snd) ` Rules) \ L \ UNIV" "edges (graph_of G) \ L \ UNIV" shows "edges (graph_of (mk_chain sel Rules G i)) \ L \ UNIV" using assms(3) proof(induct i arbitrary:G) case 0 then show ?case using assms(2) by auto next case (Suc i G) hence "edges (graph_of (make_step sel G)) \ L \ UNIV" proof(cases "sel G") case None show ?thesis unfolding None make_step_def using Suc by auto next case (Some a) then obtain R f where Some:"sel G = Some (R, f)" by fastforce hence "(a, x, y) \ edges (snd R) \ a \ L" for a x y using assms(2) valid_selectorD(2)[OF assms(1) Some] unfolding valid_selection_def Let_def worklist_def by auto then show ?thesis unfolding Some make_step_def Let_def using Suc by auto qed thus ?case unfolding mk_chain.simps by(rule Suc) qed lemma the_lcg_edges: assumes "valid_selector Rules sel" - "fst ` UNION Rules (edges \ snd) \ L" (is "fst `?fR \ _") + "fst ` (\ ((edges o snd) ` Rules)) \ L" (is "fst `?fR \ _") "fst ` snd G \ L" shows "fst ` edges (the_lcg sel Rules G) \ L" proof - from assms have "fst `?fR \ UNIV \ L \ UNIV" "fst `(edges (graph_of G)) \ UNIV \ L \ UNIV" by auto - hence "UNION Rules (edges \ snd) \ L \ UNIV" "edges (graph_of G) \ L \ UNIV" + hence "(\ ((edges o snd) ` Rules)) \ L \ UNIV" "edges (graph_of G) \ L \ UNIV" using fst_UNIV[of ?fR] fst_UNIV[of "(edges (graph_of G))"] by blast+ note assms = assms(1) this have "edges (graph_of (mk_chain sel Rules G i)) \ L \ UNIV" for i using mk_chain_edges[OF assms,unfolded Times_subset_cancel2[OF UNIV_I]]. hence "edges (the_lcg sel Rules G) \ L \ UNIV" unfolding chain_sup_def by auto thus ?thesis by auto qed text \Lemma 9.\ lemma lcg_through_make_step: assumes "finite Rs" "set_of_graph_rules Rs" "graph (graph_of init)" "valid_selector Rs sel" shows "least_consequence_graph t Rs (graph_of init) (the_lcg sel Rs init)" proof - from assms have gr:"graph (graph_of (mk_chain sel Rs init 0))" by auto note assms = assms(1,2) this mk_chain assms(4) from set_of_graph_rulesD[OF assms(2)] have "(\R. R \ Rs \ subgraph (fst R) (snd R) \ finite_graph (fst R))" by auto from fair_chain_impl_consequence_graph[OF fair_through_make_step[OF assms] this] wpc_simpl[OF WPC_through_make_step(1)[OF assms(2-)],THEN wpc_least] show ?thesis unfolding least_consequence_graph_def by auto qed end diff --git a/thys/HOL-CSP/Process.thy b/thys/HOL-CSP/Process.thy --- a/thys/HOL-CSP/Process.thy +++ b/thys/HOL-CSP/Process.thy @@ -1,1282 +1,1282 @@ (*<*) \\ ******************************************************************** * Project : HOL-CSP - A Shallow Embedding of CSP in Isabelle/HOL * Version : 2.0 * * Author : Burkhart Wolff, Safouan Taha, Lina Ye. * (Based on HOL-CSP 1.0 by Haykal Tej and Burkhart Wolff) * * This file : A Combined CSP Theory * * Copyright (c) 2009 Université Paris-Sud, France * * All rights reserved. * * Redistribution and use in source and binary forms, with or without * modification, are permitted provided that the following conditions are * met: * * * Redistributions of source code must retain the above copyright * notice, this list of conditions and the following disclaimer. * * * Redistributions in binary form must reproduce the above * copyright notice, this list of conditions and the following * disclaimer in the documentation and/or other materials provided * with the distribution. * * * Neither the name of the copyright holders nor the names of its * contributors may be used to endorse or promote products derived * from this software without specific prior written permission. * * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS * "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT * LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR * A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT * OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, * SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT * LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, * DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY * THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT * (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************\ (*>*) chapter\The Notion of Processes\ text\As mentioned earlier, we base the theory of CSP on HOLCF, a Isabelle/HOL library providing a theory of continuous functions, fixpoint induction and recursion.\ theory Process imports HOLCF begin text\Since HOLCF sets the default type class to @{class "cpo"}, while our Process theory establishes links between standard types and @{class "pcpo"} types. Consequently, we reset the default type class to the default in HOL.\ default_sort type subsection\Pre-Requisite: Basic Traces and tick-Freeness\ text\The denotational semantics of CSP assumes a distinguishable special event, called \verb+tick+ and written $\checkmark$, that is required to occur only in the end of traces in order to signalize successful termination of a process. (In the original text of Hoare, this treatment was more liberal and lead to foundational problems: the process invariant could not be established for the sequential composition operator of CSP; see @{cite "tej.ea:corrected:1997"} for details.)\ datatype '\ event = ev '\ | tick type_synonym '\ trace = "('\ event) list" text\We chose as standard ordering on traces the prefix ordering.\ instantiation list :: (type) order begin definition le_list_def : "(s::'a list) \ t \ (\ r. s @ r = t)" definition less_list_def: "(s::'a list) < t \ s \ t \ s \ t" lemma A : "((x::'a list) < y) = (x \ y \ \ y \ x)" by(auto simp: le_list_def less_list_def) instance proof fix x y z ::"'a list" show "(x < y) = (x \ y \ \ y \ x)" by(auto simp: le_list_def less_list_def) show "x \ x" by(simp add: le_list_def) assume A:"x \ y" and B:"y \ z" thus "x \ z" apply(insert A B, simp add: le_list_def, safe) apply(rename_tac r ra, rule_tac x="r@ra" in exI, simp) done assume A:"x \ y" and B:"y \ x" thus "x = y" by(insert A B, auto simp: le_list_def) qed end text\Some facts on the prefix ordering.\ lemma nil_le[simp]: "[] \ s" by(induct "s", simp_all, auto simp: le_list_def) lemma nil_le2[simp]: "s \ [] = (s = [])" by(induct "s", auto simp:le_list_def) lemma nil_less[simp]: "\ t < []" by(simp add: less_list_def) lemma nil_less2[simp]: "[] < t @ [a]" by(simp add: less_list_def) lemma less_self[simp]: "t < t@[a]" by(simp add:less_list_def le_list_def) lemma le_length_mono: "s \ t \ length s \ length t" by(auto simp: le_list_def) lemma less_length_mono: "s < t \ length s < length t" by(auto simp: less_list_def le_list_def) lemma less_cons: "s < t \ a # s < a # t" by (simp add: le_list_def less_list_def) lemma less_append: "s < t \ a @ s < a @ t" by (simp add: le_list_def less_list_def) lemma less_tail: "s \ [] \ s < t \ tl s < tl t" by (auto simp add: le_list_def less_list_def) \\should be in the List library\ lemma list_nonMt_append: "s \ [] \ \ a t. s = t @ [a]" by(erule rev_mp,induct "s",simp_all,case_tac "s = []",auto) lemma append_eq_first_pref_spec[rule_format]: "s @ t = r @ [x] \ t \ [] \ s \ r" apply(rule_tac x=s in spec) apply(induct r,auto) apply(erule rev_mp) apply(rename_tac xa, rule_tac list=xa in list.induct, simp_all) apply(simp add: le_list_def) apply(drule list_nonMt_append, auto) done lemma prefixes_fin: "let prefixes = {t. \t2. x = t @ t2} in finite prefixes \ card prefixes = length x + 1" proof(induct x, simp) case (Cons a x) hence A:"finite {t. (\t2. x = t @ t2)}" using not_add_less2 by fastforce have B:"inj_on Cons UNIV" by (metis injI list.inject) from Cons A B inj_on_iff_eq_card have C:"card ((\x. a#x)`{t. (\t2. x = t @ t2)}) = length x + 1" by fastforce have D:"{t. \t2. a # x = t @ t2} = {[]} \ (\x. a#x)`{t. (\t2. x = t @ t2)}" by(intro set_eqI iffI, auto simp add:Cons_eq_append_conv) from C D card_insert_if[of "(\x. a#x)`{t. (\t2. x = t @ t2)}"] show ?case by (metis (no_types, lifting) One_nat_def Suc_eq_plus1 Un_insert_left finite.insertI finite_imageI image_iff list.distinct(1) list.size(4) local.A sup_bot.left_neutral) qed lemma sublists_fin: "finite {t. \t1 t2. x = t1 @ t @ t2}" proof(induct x, simp) case (Cons a x) have "{t. \t1 t2. a # x = t1 @ t @ t2} \ {t. \t1 t2. x = t1 @ t @ t2} \ {t. \t2. a#x = t @ t2}" apply(auto) by (metis Cons_eq_append_conv) with Cons prefixes_fin[of "a#x"] show ?case by (meson finite_UnI finite_subset) qed lemma suffixes_fin: "finite {t. \t1. x = t1 @ t}" apply(subgoal_tac "{t. \t1. x = t1 @ t} \ {t. \t1 t2. x = t1 @ t @ t2}") using infinite_super sublists_fin apply blast by blast text\For the process invariant, it is a key element to reduce the notion of traces to traces that may only contain one tick event at the very end. This is captured by the definition of the predicate \verb+front_tickFree+ and its stronger version \verb+tickFree+. Here is the theory of this concept.\ definition tickFree :: "'\ trace \ bool" where "tickFree s = (tick \ set s)" definition front_tickFree :: "'\ trace \ bool" where "front_tickFree s = (s =[] \ tickFree(tl(rev s)))" lemma tickFree_Nil [simp]: "tickFree []" by(simp add: tickFree_def) lemma tickFree_Cons [simp]: "tickFree (a # t) = (a \ tick \ tickFree t)" by(auto simp add: tickFree_def) lemma tickFree_tl : "[|s ~= [] ; tickFree s|] ==> tickFree(tl s)" by(case_tac s, simp_all) lemma tickFree_append[simp]: "tickFree(s@t) = (tickFree s \ tickFree t)" by(simp add: tickFree_def member_def) lemma non_tickFree_tick [simp]: "\ tickFree [tick]" by(simp add: tickFree_def) lemma non_tickFree_implies_nonMt: "\ tickFree s \ s \ []" by(simp add:tickFree_def,erule rev_mp, induct s, simp_all) lemma tickFree_rev : "tickFree(rev t) = (tickFree t)" by(simp add: tickFree_def member_def) lemma front_tickFree_Nil[simp]: "front_tickFree []" by(simp add: front_tickFree_def) lemma front_tickFree_single[simp]: "front_tickFree [a]" by(simp add: front_tickFree_def) lemma tickFree_implies_front_tickFree: "tickFree s \ front_tickFree s" apply(simp add: tickFree_def front_tickFree_def member_def,safe) apply(erule contrapos_np, simp,(erule rev_mp)+) apply(rule_tac xs=s in List.rev_induct,simp_all) done lemma front_tickFree_charn: "front_tickFree s = (s = [] \ (\a t. s = t @ [a] \ tickFree t))" apply(simp add: front_tickFree_def) apply(cases "s=[]", simp_all) apply(drule list_nonMt_append, auto simp: tickFree_rev) done lemma front_tickFree_implies_tickFree: "front_tickFree (t @ [a]) \ tickFree t" by(simp add: tickFree_def front_tickFree_def member_def) lemma tickFree_implies_front_tickFree_single: "tickFree t \ front_tickFree (t @ [a])" by(simp add:front_tickFree_charn) lemma nonTickFree_n_frontTickFree: "\\ tickFree s; front_tickFree s \ \ \t. s = t @ [tick]" apply(frule non_tickFree_implies_nonMt) apply(drule front_tickFree_charn[THEN iffD1], auto) done lemma front_tickFree_dw_closed : "front_tickFree (s @ t) \ front_tickFree s" apply(erule rev_mp, rule_tac x= s in spec) apply(rule_tac xs=t in List.rev_induct, simp, safe) apply(rename_tac x xs xa) apply(simp only: append_assoc[symmetric]) apply(rename_tac x xs xa, erule_tac x="xa @ xs" in all_dupE) apply(drule front_tickFree_implies_tickFree) apply(erule_tac x="xa" in allE, auto) apply(auto dest!:tickFree_implies_front_tickFree) done lemma front_tickFree_append: "\ tickFree s; front_tickFree t\ \ front_tickFree (s @ t)" apply(drule front_tickFree_charn[THEN iffD1], auto) apply(erule tickFree_implies_front_tickFree) apply(subst append_assoc[symmetric]) apply(rule tickFree_implies_front_tickFree_single) apply(auto intro: tickFree_append) done lemma front_tickFree_mono: "front_tickFree (t @ r) \ r \ [] \ tickFree t \ front_tickFree r" by(metis append_assoc append_butlast_last_id front_tickFree_charn front_tickFree_implies_tickFree tickFree_append) subsection\ Basic Types, Traces, Failures and Divergences \ type_synonym '\ refusal = "('\ event) set" type_synonym '\ failure = "'\ trace \ '\ refusal" type_synonym '\ divergence = "'\ trace set" type_synonym '\ process\<^sub>0 = "'\ failure set \ '\ divergence" definition FAILURES :: "'\ process\<^sub>0 \ ('\ failure set)" where "FAILURES P = fst P" definition TRACES :: "'\ process\<^sub>0 \ ('\ trace set)" where "TRACES P = {tr. \ a. a \ FAILURES P \ tr = fst a}" definition DIVERGENCES :: "'\ process\<^sub>0 \ '\ divergence" where "DIVERGENCES P = snd P" definition REFUSALS :: "'\ process\<^sub>0 \ ('\ refusal set)" where "REFUSALS P = {ref. \ F. F \ FAILURES P \ F = ([],ref)}" subsection\ The Process Type Invariant \ definition is_process :: "'\ process\<^sub>0 \ bool" where "is_process P = (([],{}) \ FAILURES P \ (\ s X. (s,X) \ FAILURES P \ front_tickFree s) \ (\ s t . (s@t,{}) \ FAILURES P \ (s,{}) \ FAILURES P) \ (\ s X Y. (s,Y) \ FAILURES P & X <= Y \ (s,X) \ FAILURES P) \ (\ s X Y. (s,X) \ FAILURES P \ (\ c. c \ Y \ ((s@[c],{})\FAILURES P)) \ (s,X \ Y)\FAILURES P) \ (\ s X. (s@[tick],{}) : FAILURES P \ (s,X-{tick}) \ FAILURES P) \ (\ s t. s \ DIVERGENCES P \ tickFree s \ front_tickFree t \ s@t \ DIVERGENCES P) \ (\ s X. s \ DIVERGENCES P \ (s,X) \ FAILURES P) \ (\ s. s @ [tick] : DIVERGENCES P \ s \ DIVERGENCES P))" lemma is_process_spec: "is_process P = (([],{}) \ FAILURES P \ (\ s X. (s,X) \ FAILURES P \ front_tickFree s) \ (\ s t . (s @ t,{}) \ FAILURES P \ (s,{}) \ FAILURES P) \ (\ s X Y. (s,Y) \ FAILURES P \ \(X\Y) | (s,X) \ FAILURES P) \ (\ s X Y.(s,X) \ FAILURES P \ (\ c. c \ Y \ ((s@[c],{}) \ FAILURES P)) \(s,X \ Y) \ FAILURES P) \ (\ s X. (s@[tick],{}) \ FAILURES P \ (s,X - {tick}) \ FAILURES P) \ (\ s t. s \ DIVERGENCES P \ \tickFree s \ \front_tickFree t \ s @ t \ DIVERGENCES P) \ (\ s X. s \ DIVERGENCES P \ (s,X) \ FAILURES P) \ (\ s. s @ [tick] \ DIVERGENCES P \ s \ DIVERGENCES P))" by(simp only: is_process_def HOL.nnf_simps(1) HOL.nnf_simps(3) [symmetric] HOL.imp_conjL[symmetric]) lemma Process_eqI : assumes A: "FAILURES P = FAILURES Q " assumes B: "DIVERGENCES P = DIVERGENCES Q" shows "(P::'\ process\<^sub>0) = Q" apply(insert A B, unfold FAILURES_def DIVERGENCES_def) apply(rule_tac t=P in surjective_pairing[symmetric,THEN subst]) apply(rule_tac t=Q in surjective_pairing[symmetric,THEN subst]) apply(simp) done lemma process_eq_spec: "((P::'a process\<^sub>0) = Q) = (FAILURES P = FAILURES Q \ DIVERGENCES P = DIVERGENCES Q)" apply(auto simp: FAILURES_def DIVERGENCES_def) apply(rule_tac t=P in surjective_pairing[symmetric,THEN subst]) apply(rule_tac t=Q in surjective_pairing[symmetric,THEN subst]) apply(simp) done lemma process_surj_pair: "(FAILURES P,DIVERGENCES P) = P" by(auto simp:FAILURES_def DIVERGENCES_def) lemma Fa_eq_imp_Tr_eq: "FAILURES P = FAILURES Q \ TRACES P = TRACES Q" by(auto simp:FAILURES_def DIVERGENCES_def TRACES_def) lemma is_process1: "is_process P \ ([],{})\ FAILURES P " by(auto simp: is_process_def) lemma is_process2: "is_process P \ \ s X. (s,X) \ FAILURES P \ front_tickFree s " by(simp only: is_process_spec, metis) lemma is_process3: "is_process P \ \ s t. (s @ t,{}) \ FAILURES P \ (s, {}) \ FAILURES P" by(simp only: is_process_spec, metis) lemma is_process3_S_pref: "\is_process P; (t, {}) \ FAILURES P; s \ t\ \ (s, {}) \ FAILURES P" by(auto simp: le_list_def intro: is_process3 [rule_format]) lemma is_process4: "is_process P \ \s X Y. (s, Y) \ FAILURES P \ \ X \ Y \ (s, X) \ FAILURES P" by(simp only: is_process_spec, simp) lemma is_process4_S: "\is_process P; (s, Y) \ FAILURES P; X \ Y\ \ (s, X) \ FAILURES P" by(drule is_process4, auto) lemma is_process4_S1: "\is_process P; x \ FAILURES P; X \ snd x\ \ (fst x, X) \ FAILURES P" by(drule is_process4_S, auto) lemma is_process5: "is_process P \ \sa X Y. (sa, X) \ FAILURES P \ (\c. c \ Y \ (sa @ [c], {}) \ FAILURES P) \ (sa, X \ Y) \ FAILURES P" by(drule is_process_spec[THEN iffD1],metis) lemma is_process5_S: "\is_process P; (sa, X) \ FAILURES P; \c. c \ Y \ (sa @ [c], {}) \ FAILURES P\ \ (sa, X \ Y) \ FAILURES P" by(drule is_process5, metis) lemma is_process5_S1: "\is_process P; (sa, X) \ FAILURES P; (sa, X \ Y) \ FAILURES P\ \ \c. c \ Y \ (sa @ [c], {}) \ FAILURES P" by(erule contrapos_np, drule is_process5_S, simp_all) lemma is_process6: "is_process P \ \ s X. (s@[tick],{}) \ FAILURES P \ (s,X-{tick}) \ FAILURES P" by(drule is_process_spec[THEN iffD1], metis) lemma is_process6_S: "\is_process P ;(s@[tick],{}) \ FAILURES P\ \ (s,X-{tick}) \ FAILURES P" by(drule is_process6, metis) lemma is_process7: "is_process P \ \ s t. s \ DIVERGENCES P \ \ tickFree s \ \ front_tickFree t \ s @ t \ DIVERGENCES P" by(drule is_process_spec[THEN iffD1], metis) lemma is_process7_S: "\ is_process P;s : DIVERGENCES P;tickFree s;front_tickFree t\ \ s @ t \ DIVERGENCES P" by(drule is_process7, metis) lemma is_process8: "is_process P \ \ s X. s \ DIVERGENCES P \ (s,X) \ FAILURES P" by(drule is_process_spec[THEN iffD1], metis) lemma is_process8_S: "\ is_process P; s \ DIVERGENCES P \ \ (s,X) \ FAILURES P" by(drule is_process8, metis) lemma is_process9: "is_process P \ \ s. s@[tick] \ DIVERGENCES P \ s \ DIVERGENCES P" by(drule is_process_spec[THEN iffD1], metis) lemma is_process9_S: "\ is_process P;s@[tick] \ DIVERGENCES P \ \ s \ DIVERGENCES P" by(drule is_process9, metis) lemma Failures_implies_Traces: " \is_process P; (s, X) \ FAILURES P\ \ s \ TRACES P" by(simp add: TRACES_def, metis) lemma is_process5_sing: "\ is_process P ; (s,{x}) \ FAILURES P;(s,{}) \ FAILURES P\ \ (s @ [x],{}) \ FAILURES P" by(drule_tac X="{}" in is_process5_S1, auto) lemma is_process5_singT: "\ is_process P ; (s,{x}) \ FAILURES P;(s,{}) \ FAILURES P\ \ s @ [x] \ TRACES P" apply(drule is_process5_sing, auto) by(simp add: TRACES_def, auto) lemma front_trace_is_tickfree: assumes A: "is_process P" and B: "(t @ [tick],X) \ FAILURES P" shows "tickFree t" proof - have C: "front_tickFree(t @ [tick])" by(insert A B, drule is_process2, metis) show ?thesis by(rule front_tickFree_implies_tickFree[OF C]) qed lemma trace_with_Tick_implies_tickFree_front : "\is_process P; t @ [tick] \ TRACES P\ \ tickFree t" by(auto simp: TRACES_def intro: front_trace_is_tickfree) subsection \ The Abstraction to the process-Type \ typedef '\ process = "{p :: '\ process\<^sub>0 . is_process p}" proof - have "({(s, X). s = []},{}) \ {p::'\ process\<^sub>0. is_process p}" by(simp add: is_process_def front_tickFree_def FAILURES_def TRACES_def DIVERGENCES_def ) thus ?thesis by auto qed definition F :: "'\ process \ ('\ failure set)" where "F P = FAILURES (Rep_process P)" definition T :: "'\ process \ ('\ trace set)" where "T P = TRACES (Rep_process P)" definition D :: "'\ process \ '\ divergence" where "D P = DIVERGENCES (Rep_process P)" definition R :: "'\ process \ ('\ refusal set)" where "R P = REFUSALS (Rep_process P)" lemma is_process_Rep : "is_process (Rep_process P)" by(rule_tac P=is_process in CollectD, rule Rep_process) lemma Process_spec: "Abs_process((F P , D P)) = P" by(simp add: F_def FAILURES_def D_def DIVERGENCES_def Rep_process_inverse) lemma Process_eq_spec: "(P = Q)=(F P = F Q \ D P = D Q)" apply(rule iffI,simp) apply(rule_tac t=P in Process_spec[THEN subst]) apply(rule_tac t=Q in Process_spec[THEN subst]) apply simp done lemma Process_eq_spec_optimized: "(P = Q) = (D P = D Q \ (D P = D Q \ F P = F Q))" using Process_eq_spec by auto lemma is_processT: "([],{}) \ F P \ (\ s X. (s,X) \ F P \ front_tickFree s) \ (\ s t .(s@t,{}) \ F P \ (s,{}) \ F P) \ (\ s X Y.(s,Y) \ F P \ (X\Y) \ (s,X) \ F P) \ (\ s X Y.(s,X) \ F P \ (\c. c \ Y \((s@[c],{})\F P)) \ (s,X \ Y) \ F P) \ (\ s X. (s@[tick],{}) \ F P \ (s, X-{tick}) \ F P) \ (\ s t. s \ D P \ tickFree s \ front_tickFree t \ s @ t \ D P) \ (\ s X. s \ D P \ (s,X) \ F P) \ (\ s. s@[tick] \ D P \ s \ D P)" apply(simp only: F_def D_def T_def) apply(rule is_process_def[THEN iffD1]) apply(rule is_process_Rep) done lemma process_charn: "([], {}) \ F P \ (\s X. (s, X) \ F P \ front_tickFree s) \ (\s t. (s @ t, {}) \ F P \ (s, {}) \ F P) \ (\s X Y. (s, Y) \ F P \ \ X \ Y \ (s, X) \ F P) \ (\s X Y. (s, X) \ F P \ (\c. c \ Y \ (s @ [c], {}) \ F P) \ (s, X \ Y) \ F P) \ (\s X. (s @ [tick], {}) \ F P \ (s, X - {tick}) \ F P) \ (\s t. s \ D P \ \ tickFree s \ \ front_tickFree t \ s @ t \ D P) \ (\s X. s \ D P \ (s, X) \ F P) \ (\s. s @ [tick] \ D P \ s \ D P)" proof - { have A: "(\s t. (s @ t, {}) \ F P \ (s, {}) \ F P) = (\s t. (s @ t, {}) \ F P \ (s, {}) \ F P)" by metis have B : "(\s X Y. (s, Y) \ F P \ \ X \ Y \ (s, X) \ F P) = (\s X Y. (s, Y) \ F P \ X \ Y \ (s, X) \ F P) " by metis have C : "(\s t. s \ D P \ \ tickFree s \ \ front_tickFree t \ s @ t \ D P) = (\s t. s \ D P \ tickFree s \ front_tickFree t \ s @ t \ D P) " by metis have D:"(\s X. s \ D P \ (s, X) \ F P) = (\s X. s \ D P \ (s, X) \ F P)" by metis have E:"(\s. s @ [tick] \ D P \ s \ D P) = (\s. s @ [tick] \ D P \ s \ D P)" by metis note A B C D E } note a = this then show ?thesis apply(simp only: a) apply(rule is_processT) done qed text\ split of \verb+is_processT+: \ lemma is_processT1: "([],{}) \ F P" by(simp add:process_charn) lemma is_processT2: "\s X. (s, X) \ F P \ front_tickFree s" by(simp add:process_charn) lemma is_processT2_TR : "\s. s \ T P \ front_tickFree s" apply(simp add: F_def [symmetric] T_def TRACES_def, safe) apply (drule is_processT2[rule_format], assumption) done lemma is_proT2: assumes A : " (s, X) \ F P" and B : " s \ []" shows "tick \ set (tl (rev s))" proof - have C: "front_tickFree s" by(insert A B, simp add: is_processT2) show ?thesis by(insert C,simp add: B tickFree_def front_tickFree_def) qed lemma is_processT3 : "\s t. (s @ t, {}) \ F P \ (s, {}) \ F P" by(simp only: process_charn HOL.nnf_simps(3), simp) lemma is_processT3_S_pref : "\(t, {}) \ F P; s \ t\ \ (s, {}) \ F P" apply(simp only: le_list_def, safe) apply(erule is_processT3[rule_format]) done lemma is_processT4 : "\s X Y. (s, Y) \ F P \ X \ Y \ (s, X) \ F P" by(insert process_charn [of P], metis) lemma is_processT4_S1 : "\x \ F P; X \ snd x\ \ (fst x, X) \ F P" apply(rule_tac Y = "snd x" in is_processT4[rule_format]) apply simp done lemma is_processT5: "\s X Y.(s,X) \ F P \ (\c. c\Y \ (s@[c],{}) \ F P) \ (s,X\Y)\F P " by(simp add: process_charn) lemma is_processT5_S1: "\(s, X) \ F P; (s, X \ Y) \ F P\ \ \c. c \ Y \ (s @ [c], {}) \ F P" by(erule contrapos_np, simp add: is_processT5[rule_format]) lemma is_processT5_S2: "\(s, X) \ F P; (s @ [c], {}) \ F P\ \ (s, X \ {c}) \ F P" by(rule is_processT5[rule_format,OF conjI], metis, safe) lemma is_processT5_S2a: "\(s, X) \ F P; (s, X \ {c}) \ F P\ \ (s @ [c], {}) \ F P" apply(erule contrapos_np) apply(rule is_processT5_S2) apply(simp_all) done lemma is_processT5_S3: assumes A: "(s, {}) \ F P" and B: "(s @ [c], {}) \ F P" shows "(s, {c}) \ F P" proof - have C : " {c} = ({} Un {c})" by simp show ?thesis by(subst C, rule is_processT5_S2, simp_all add: A B) qed lemma is_processT5_S4: "\(s, {}) \ F P; (s, {c}) \ F P\ \ (s @ [c], {}) \ F P" by(erule contrapos_np, simp add: is_processT5_S3) lemma is_processT5_S5: "\(s, X) \ F P; \c. c \ Y \ (s, X \ {c}) \ F P\ \ \c. c \ Y \ (s @ [c], {}) \ F P" by(erule_tac Q = "\c. c \ Y \ (s, X \ {c}) \ F P" in contrapos_pp, metis is_processT5_S2) lemma is_processT5_S6: "([], {c}) \ F P \ ([c], {}) \ F P" apply(rule_tac t="[c]" and s="[]@[c]" in subst, simp) apply(rule is_processT5_S4, simp_all add: is_processT1) done lemma is_processT6: "\s X. (s @ [tick], {}) \ F P \ (s, X - {tick}) \ F P" by(simp add: process_charn) lemma is_processT7: "\s t. s \ D P \ tickFree s \ front_tickFree t \ s @ t \ D P" by(insert process_charn[of P], metis) lemmas is_processT7_S = is_processT7[rule_format,OF conjI[THEN conjI, THEN conj_commute[THEN iffD1]]] lemma is_processT8: "\s X. s \ D P \ (s, X) \ F P " by(insert process_charn[of P], metis) lemmas is_processT8_S = is_processT8[rule_format] lemma is_processT8_Pair: "fst s \ D P \ s \ F P" apply(subst surjective_pairing) apply(rule is_processT8_S, simp) done lemma is_processT9: "\s. s @ [tick] \ D P \ s \ D P" by(insert process_charn[of P], metis) lemma is_processT9_S_swap: "s \ D P \ s @ [tick] \ D P" by(erule contrapos_nn,simp add: is_processT9[rule_format]) subsection\ Some Consequences of the Process Characterization\ lemma no_Trace_implies_no_Failure: "s \ T P \ (s, {}) \ F P" by(simp add: T_def TRACES_def F_def) lemmas NT_NF = no_Trace_implies_no_Failure lemma T_def_spec: "T P = {tr. \a. a \ F P \ tr = fst a}" by(simp add: T_def TRACES_def F_def) lemma F_T: "(s, X) \ F P \ s \ T P" by(simp add: T_def_spec split_def, metis) lemma F_T1: "a \ F P \ fst a \ T P" by(rule_tac X="snd a" in F_T,simp) lemma T_F: "s \ T P \ (s, {}) \ F P" apply(auto simp: T_def_spec) apply(drule is_processT4_S1, simp_all) done lemmas is_processT4_empty [elim!]= F_T [THEN T_F] lemma NF_NT: "(s, {}) \ F P \ s \ T P" by(erule contrapos_nn, simp only: T_F) lemma is_processT6_S1: "\ tick \ X;(s @ [tick], {}) \ F P \ \ (s::'a event list, X) \ F P" by(subst Diff_triv[of X "{tick}", symmetric], simp, erule is_processT6[rule_format]) lemmas is_processT3_ST = T_F [THEN is_processT3[rule_format,THEN F_T]] lemmas is_processT3_ST_pref = T_F [THEN is_processT3_S_pref [THEN F_T]] lemmas is_processT3_SR = F_T [THEN T_F [THEN is_processT3[rule_format]]] lemmas D_T = is_processT8_S [THEN F_T] lemma D_T_subset : "D P \ T P" by(auto intro!:D_T) lemma NF_ND : "(s, X) \ F P \ s \ D P" by(erule contrapos_nn, simp add: is_processT8_S) lemmas NT_ND = D_T_subset[THEN Set.contra_subsetD] lemma T_F_spec : "((t, {}) \ F P) = (t \ T P)" by(auto simp:T_F F_T) lemma is_processT5_S7: " \t \ T P; (t, A) \ F P\ \ \x. x \ A \ t @ [x] \ T P" apply(erule contrapos_np, simp) apply(rule is_processT5[rule_format, OF conjI,of _ "{}", simplified]) apply(auto simp: T_F_spec) done lemma is_processT5_S7': " \(t, X) \ F P; (t, X \ A) \ F P\ \ \x. x \ A \ x \ X \ t @ [x] \ T P" apply(erule contrapos_np, simp, subst Un_Diff_cancel[of X A, symmetric]) apply(rule is_processT5[rule_format]) apply(auto simp: T_F_spec) done lemma Nil_subset_T: " {[]} \ T P" by(auto simp: T_F_spec[symmetric] is_processT1) lemma Nil_elem_T: "[] \ T P" by(simp add: Nil_subset_T[THEN subsetD]) lemmas D_imp_front_tickFree = is_processT8_S[THEN is_processT2[rule_format]] lemma D_front_tickFree_subset : "D P \ Collect front_tickFree" by(auto simp: D_imp_front_tickFree) lemma F_D_part : "F P = {(s, x). s \ D P} \ {(s, x). s \ D P \ (s, x) \ F P}" by(auto intro:is_processT8_Pair) lemma D_F : "{(s, x). s \ D P} \ F P" by(auto intro:is_processT8_Pair) lemma append_T_imp_tickFree: "\t @ s \ T P; s \ []\ \ tickFree t" by(frule is_processT2_TR[rule_format], simp add: front_tickFree_def tickFree_rev) corollary append_single_T_imp_tickFree : "t @ [a] \ T P \ tickFree t" by (simp add: append_T_imp_tickFree) lemma F_subset_imp_T_subset: "F P \ F Q \ T P \ T Q" by(auto simp: subsetD T_F_spec[symmetric]) lemma is_processT6_S2: "\tick \ X; [tick] \ T P\ \ ([], X) \ F P" by(erule is_processT6_S1, simp add: T_F_spec) lemma is_processT9_tick: "\[tick] \ D P; front_tickFree s\ \ s \ D P" apply(rule append.simps(1) [THEN subst, of _ s]) apply(rule is_processT7_S, simp_all) apply(rule is_processT9 [rule_format], simp) done lemma T_nonTickFree_imp_decomp: "\t \ T P; \ tickFree t\ \ \s. t = s @ [tick]" by(auto elim: is_processT2_TR[rule_format] nonTickFree_n_frontTickFree) subsection\ Process Approximation is a Partial Ordering, a Cpo, and a Pcpo \ text\The Failure/Divergence Model of CSP Semantics provides two orderings: The \emph{approximation ordering} (also called \emph{process ordering}) will be used for giving semantics to recursion (fixpoints) over processes, the \emph{refinement ordering} captures our intuition that a more concrete process is more deterministic and more defined than an abstract one. We start with the key-concepts of the approximation ordering, namely the predicates $min\_elems$ and $Ra$ (abbreviating \emph{refusals after}). The former provides just a set of minimal elements from a given set of elements of type-class $ord$ \ldots \ definition min_elems :: "('s::ord) set \ 's set" where "min_elems X = {s \ X. \t. t \ X \ \ (t < s)}" lemma Nil_min_elems : "[] \ A \ [] \ min_elems A" by(simp add: min_elems_def) lemma min_elems_le_self[simp] : "(min_elems A) \ A" by(auto simp: min_elems_def) lemmas elem_min_elems = Set.set_mp[OF min_elems_le_self] lemma min_elems_Collect_ftF_is_Nil : "min_elems (Collect front_tickFree) = {[]}" apply(auto simp: min_elems_def le_list_def) apply(drule front_tickFree_charn[THEN iffD1]) apply(auto dest!: tickFree_implies_front_tickFree) done lemma min_elems5 : assumes A: "(x::'a list) \ A" shows "\s\x. s \ min_elems A" proof - have * : "!! (x::'a list) (A::'a list set) (n::nat). x \ A \ length x \ n \ (\s\x. s \ min_elems A)" apply(rule_tac x=x in spec) apply(rule_tac n=n in nat_induct) (* quirk in induct *) apply(auto simp: Nil_min_elems) apply(case_tac "\ y. y \ A \ y < x",auto) apply(rename_tac A na x y, erule_tac x=y in allE, simp) apply(erule impE,drule less_length_mono, arith) apply(safe, rename_tac s, rule_tac x=s in exI,simp) apply(rule_tac x=x in exI, simp add:min_elems_def) done show ?thesis by(rule_tac n="length x" in *[rule_format],simp add:A) qed lemma min_elems4: "A \ {} \ \s. (s :: 'a trace) \ min_elems A" by(auto dest: min_elems5) lemma min_elems_charn: "t \ A \ \ t' r. t = (t' @ r) \ t' \ min_elems A" by(drule min_elems5[simplified le_list_def], auto) lemmas min_elems_ex = min_elems_charn (* Legacy *) lemma min_elems_no: "(x::'a list) \ min_elems A \ t \ A \ t \ x \ x = t" by (metis (no_types, lifting) less_list_def mem_Collect_eq min_elems_def) text\ \ldots while the second returns the set of possible refusal sets after a given trace $s$ and a given process $P$: \ definition Ra :: "['\ process, '\ trace] \ ('\ refusal set)" where "Ra P s = {X. (s, X) \ F P}" text\ In the following, we link the process theory to the underlying fixpoint/domain theory of HOLCF by identifying the approximation ordering with HOLCF's pcpo's. \ instantiation process :: ("type") below begin text\ declares approximation ordering $\_ \sqsubseteq \_$ also written \verb+_ << _+. \ definition le_approx_def : "P \ Q \ D Q \ D P \ (\s. s \ D P \ Ra P s = Ra Q s) \ min_elems (D P) \ T Q" text\ The approximation ordering captures the fact that more concrete processes should be more defined by ordering the divergence sets appropriately. For defined positions in a process, the failure sets must coincide pointwise; moreover, the minimal elements (wrt.~prefix ordering on traces, i.e.~lists) must be contained in the trace set of the more concrete process.\ instance .. end lemma le_approx1: "P\Q \ D Q \ D P" by(simp add: le_approx_def) lemma le_approx2: "\ P\Q; s \ D P\ \ (s,X) \ F Q = ((s,X) \ F P)" by(auto simp: Ra_def le_approx_def) lemma le_approx3: "P \ Q \ min_elems(D P) \ T Q" by(simp add: le_approx_def) lemma le_approx2T: "\ P\Q; s \ D P\ \ s \ T Q = (s \ T P)" by(auto simp: le_approx2 T_F_spec[symmetric]) lemma le_approx_lemma_F : "P\Q \ F Q \ F P" apply(subst F_D_part[of Q], subst F_D_part[of P]) apply(auto simp:le_approx_def Ra_def min_elems_def) done lemmas order_lemma = le_approx_lemma_F lemma le_approx_lemma_T: "P\Q \ T Q \ T P" by(auto dest!:le_approx_lemma_F simp: T_F_spec[symmetric]) lemma proc_ord2a : "\P \ Q; s \ D P\ \ ((s, X) \ F P) = ((s, X) \ F Q)" by(auto simp: le_approx_def Ra_def) instance process :: ("type") po proof fix P::"'\ process" show "P \ P" by(auto simp: le_approx_def min_elems_def elim: Process.D_T) next fix P Q ::"'\ process" assume A:"P \ Q" and B:"Q \ P" thus "P = Q" apply(insert A[THEN le_approx1] B[THEN le_approx1]) apply(insert A[THEN le_approx_lemma_F] B[THEN le_approx_lemma_F]) by(auto simp: Process_eq_spec) next fix P Q R ::"'\ process" assume A: "P \ Q" and B: "Q \ R" thus "P \ R" proof - have C : "D R \ D P" by(insert A[THEN le_approx1] B[THEN le_approx1], auto) have D : "\ s. s \ D P \ {X. (s, X) \ F P} = {X. (s, X) \ F R}" apply(rule allI, rule impI, rule set_eqI, simp) apply(frule A[THEN le_approx1, THEN Set.contra_subsetD]) apply(frule B[THEN le_approx1, THEN Set.contra_subsetD]) apply(drule A[THEN le_approx2], drule B[THEN le_approx2]) apply auto done have E : "min_elems (D P) \ T R" apply(insert B[THEN le_approx3] A[THEN le_approx3] ) apply(insert B[THEN le_approx_lemma_T] A[THEN le_approx1] ) apply(rule subsetI, simp add: min_elems_def, auto) apply(rename_tac x, case_tac "x \ D Q") apply(drule_tac B = "T R" and t=x in subset_iff[THEN iffD1,rule_format], auto) apply(subst B [THEN le_approx2T],simp) apply(rename_tac x, drule_tac B = "T Q" and t=x in subset_iff[THEN iffD1,rule_format],auto) done show ?thesis by(insert C D E, simp add: le_approx_def Ra_def) qed qed text\ At this point, we inherit quite a number of facts from the underlying HOLCF theory, which comprises a library of facts such as \verb+chain+, \verb+directed+(sets), upper bounds and least upper bounds, etc. \ text\ Some facts from the theory of complete partial orders: \begin{itemize} \item \verb+Porder.chainE+ : @{thm "Porder.chainE"} \item \verb+Porder.chain_mono+ : @{thm "Porder.chain_mono"} \item \verb+Porder.is_ubD+ : @{thm "Porder.is_ubD"} \item \verb+Porder.ub_rangeI+ : \\ @{thm "Porder.ub_rangeI"} \item \verb+Porder.ub_imageD+ : @{thm "Porder.ub_imageD"} \item \verb+Porder.is_ub_upward+ : @{thm "Porder.is_ub_upward"} \item \verb+Porder.is_lubD1+ : @{thm "Porder.is_lubD1"} \item \verb+Porder.is_lubI+ : @{thm "Porder.is_lubI"} \item \verb+Porder.is_lub_maximal+ : @{thm "Porder.is_lub_maximal"} \item \verb+Porder.is_lub_lub+ : @{thm "Porder.is_lub_lub"} \item \verb+Porder.is_lub_range_shift+: \\ @{thm "Porder.is_lub_range_shift"} \item \verb+Porder.is_lub_rangeD1+: @{thm "Porder.is_lub_rangeD1"} \item \verb+Porder.lub_eqI+: @{thm "Porder.lub_eqI"} \item \verb+Porder.is_lub_unique+:@{thm "Porder.is_lub_unique"} \end{itemize} \ definition lim_proc :: "('\ process) set \ '\ process" - where "lim_proc (X) = Abs_process (INTER X F, INTER X D)" + where "lim_proc (X) = Abs_process (\ (F ` X), \ (D ` X))" lemma min_elems3: "\s @ [c] \ D P; s @ [c] \ min_elems (D P)\ \ s \ D P" apply(auto simp: min_elems_def le_list_def less_list_def) apply(rename_tac t r) apply(subgoal_tac "t \ s") apply(subgoal_tac "r \ []") apply(simp add: le_list_def) apply(auto intro!: is_processT7_S append_eq_first_pref_spec) apply(auto dest!: D_T) apply(simp_all only: append_assoc[symmetric], drule append_T_imp_tickFree, simp_all add: tickFree_implies_front_tickFree)+ done lemma min_elems1: "\s \ D P; s @ [c] \ D P\ \ s @ [c] \ min_elems (D P)" by(erule contrapos_np, auto elim!: min_elems3) lemma min_elems2: "\s \ D P; s @ [c] \ D P; P \ S; Q \ S\ \ (s @ [c], {}) \ F Q" apply(frule_tac P=Q and Q=S in le_approx_lemma_T) apply(simp add: T_F_spec) apply(rule_tac A="T S" in subsetD, assumption) apply(rule_tac A="min_elems(D P)" in subsetD) apply(simp_all add: le_approx_def min_elems1) done lemma min_elems6: "\s \ D P; s @ [c] \ D P; P \ S\ \ (s @ [c], {}) \ F S" by(auto intro!: min_elems2) lemma ND_F_dir2: "\s \ D P; (s, {}) \ F P; P \ S; Q \ S\ \ (s, {}) \ F Q" apply(frule_tac P=Q and Q=S in le_approx_lemma_T) apply(simp add: le_approx_def Ra_def T_F_spec, safe) apply((erule_tac x=s in allE)+,simp) apply(drule_tac x="{}" in eqset_imp_iff, auto simp: T_F_spec) done \\orig version\ lemma ND_F_dir2': "\s \ D P; s \ T P; P \ S; Q \ S\ \ s \ T Q" apply(frule_tac P=Q and Q=S in le_approx_lemma_T) apply(simp add: le_approx_def Ra_def T_F_spec, safe) apply((erule_tac x=s in allE)+,simp) apply(drule_tac x="{}" in eqset_imp_iff, auto simp: T_F_spec) done lemma chain_lemma: "\chain S\ \ S i \ S k \ S k \ S i" by(case_tac "i \ k", auto intro:chain_mono chain_mono_less) lemma is_process_REP_LUB: assumes chain: "chain S" - shows "is_process(INTER (range S) F,INTER (range S) D)" + shows "is_process (\ (F ` range S), \ (D ` range S))" proof (auto simp: is_process_def) show "([], {}) \ FAILURES (\ a::nat. F (S a), \ a::nat. D (S a))" by(auto simp: FAILURES_def is_processT) next fix s::"'a trace" fix X::"'a event set" assume A : "(s, X) \ (FAILURES (\ a :: nat. F (S a), \ a :: nat. D (S a)))" thus "front_tickFree s" by(auto simp: DIVERGENCES_def FAILURES_def intro!: is_processT2[rule_format]) next fix s t::"'a trace" assume " (s @ t, {}) \ FAILURES (\ a::nat. F (S a), \ a::nat. D (S a)) " thus "(s, {}) \ FAILURES (\ a::nat. F (S a), \ a::nat. D (S a))" by(auto simp: DIVERGENCES_def FAILURES_def intro: is_processT3[rule_format]) next fix s::"'a trace" fix X Y ::"'a event set" assume "(s, Y) \ FAILURES (\ a::nat. F (S a), \ a::nat. D (S a))" and "X \ Y" thus "(s, X) \ FAILURES (\ a::nat. F (S a), \ a::nat. D (S a))" by(auto simp: DIVERGENCES_def FAILURES_def intro: is_processT4[rule_format]) next fix s::"'a trace" fix X Y ::" 'a event set" assume A:"(s, X) \ FAILURES (\ a::nat. F (S a), \ a::nat. D (S a))" assume B:"\ c. c\Y \ (s@[c],{})\FAILURES(\ a::nat. F(S a),\ a::nat. D(S a))" thus "(s, X \ Y) \ FAILURES (\ a::nat. F (S a), \ a::nat. D (S a))" txt\ What does this mean: All trace prolongations $c$ in all $Y$, which are blocking in the limit, will also occur in the refusal set of the limit. \ using A B chain proof (auto simp: DIVERGENCES_def FAILURES_def, case_tac "\ x. x \ (range S) \ (s, X \ Y) \ F x", simp_all add:DIVERGENCES_def FAILURES_def,rename_tac a, case_tac "s \ D (S a)",simp_all add: is_processT8) fix a::nat assume X: "\a. (s, X) \ F (S a)" have X_ref_at_a: "(s, X) \ F (S a)" using X by auto assume Y: "\c. c \ Y \ (\a. (s @ [c], {}) \ F (S a))" assume defined: "s \ D (S a)" show "(s::'a trace, X \ Y) \ F (S a)" proof(auto simp:X_ref_at_a intro!: is_processT5[rule_format], frule Y[THEN spec, THEN mp], erule exE, erule_tac Q="(s @ [c], {}) \ F (S a)" in contrapos_pp) fix c::"'a event" fix a' :: nat assume s_c_trace_not_trace_somewhere: "(s @ [c], {}) \ F (S a')" show "(s @ [c], {}) \ F (S a)" proof(insert chain_lemma[OF chain, of "a" "a'"],erule disjE) assume before: "S a \ S a'" show "(s @ [c], {}) \ F (S a)" using s_c_trace_not_trace_somewhere before apply(case_tac "s @ [c] \ D (S a)", simp_all add: T_F_spec before[THEN le_approx2T,symmetric]) apply(erule contrapos_nn) apply(simp only: T_F_spec[symmetric]) apply(auto dest!:min_elems6[OF defined]) done next assume after:"S a' \ S a" show "(s @ [c], {}) \ F (S a)" using s_c_trace_not_trace_somewhere by(simp add:T_F_spec after[THEN le_approx2T] s_c_trace_not_trace_somewhere[THEN NF_ND]) qed qed qed next fix s::"'a trace" fix X::"'a event set" assume "(s @ [tick], {}) \ FAILURES (\ a::nat. F (S a), \ a::nat. D (S a))" thus "(s, X - {tick}) \ FAILURES (\ a::nat. F (S a), \ a::nat. D (S a))" by(auto simp: DIVERGENCES_def FAILURES_def intro! : is_processT6[rule_format]) next fix s t ::"'a trace" assume "s : DIVERGENCES (\ a::nat. F (S a), \ a::nat. D (S a))" and "tickFree s" and " front_tickFree t" thus "s @ t \ DIVERGENCES (\ a::nat. F (S a), \ a::nat. D (S a))" by(auto simp: DIVERGENCES_def FAILURES_def intro: is_processT7[rule_format]) next fix s::"'a trace" fix X::"'a event set" assume "s \ DIVERGENCES (\ a::nat. F (S a), \ a::nat. D (S a)) " thus "(s, X) \ FAILURES (\ a::nat. F (S a), \ a::nat. D (S a))" by(auto simp: DIVERGENCES_def FAILURES_def intro: is_processT8[rule_format]) next fix s::"'a trace" assume "s @ [tick] \ DIVERGENCES (\ a::nat. F (S a), \ a::nat. D (S a)) " thus "s \ DIVERGENCES (\ a::nat. F (S a), \ a::nat. D (S a))" by(auto simp: DIVERGENCES_def FAILURES_def intro: is_processT9[rule_format]) qed lemmas Rep_Abs_LUB = Abs_process_inverse[simplified Rep_process, simplified, OF is_process_REP_LUB, simplified] -lemma F_LUB: "chain S \ F(lim_proc(range S)) = INTER (range S) F" +lemma F_LUB: "chain S \ F(lim_proc(range S)) = \ (F ` range S)" by(simp add: lim_proc_def , subst F_def, auto simp: FAILURES_def Rep_Abs_LUB) -lemma D_LUB: "chain S \ D(lim_proc(range S)) = INTER (range S) D" +lemma D_LUB: "chain S \ D(lim_proc(range S)) = \ (D ` range S)" by(simp add: lim_proc_def , subst D_def, auto simp: DIVERGENCES_def Rep_Abs_LUB) -lemma T_LUB: "chain S \ T(lim_proc(range S)) = INTER (range S) T" +lemma T_LUB: "chain S \ T(lim_proc(range S)) = \ (T ` range S)" apply(simp add: lim_proc_def , subst T_def) apply(simp add: TRACES_def FAILURES_def Rep_Abs_LUB) apply(auto intro: F_T, rule_tac x="{}" in exI, auto intro: T_F) done schematic_goal D_LUB_2: "chain S \ t \ D(lim_proc(range S)) = ?X" apply(subst D_LUB, simp) apply(rule trans, simp) apply(simp) done schematic_goal T_LUB_2: "chain S \ (t \ T (lim_proc (range S))) = ?X" apply(subst T_LUB, simp) apply(rule trans, simp) apply(simp) done subsection\ Process Refinement is a Partial Ordering\ text\ The following type instantiation declares the refinement order $\_ \le \_ $ written \verb+_ <= _+. It captures the intuition that more concrete processes should be more deterministic and more defined.\ instantiation process :: (type) ord begin definition le_ref_def : "P \ Q \ D Q \ D P \ F Q \ F P" definition less_ref_def : "(P::'a process) < Q \ P \ Q \ P \ Q" instance .. end lemma le_approx_implies_le_ref: "(P::'\ process) \ Q \ P \ Q" by(simp add: le_ref_def le_approx1 le_approx_lemma_F) lemma le_ref1: "P \ Q \ D Q \ D P" by(simp add: le_ref_def) lemma le_ref2: "P\Q \ F Q \ F P" by(simp add: le_ref_def) lemma le_ref2T : "P\Q \ T Q \ T P" by (rule subsetI) (simp add: T_F_spec[symmetric] le_ref2[THEN subsetD]) instance process :: (type) order proof fix P Q R :: "'\ process" { show "(P < Q) = (P \ Q \ \ Q \ P)" by(auto simp: le_ref_def less_ref_def Process_eq_spec) next show "P \ P" by(simp add: le_ref_def) next assume "P \ Q" and "Q \ R" then show "P \ R" by (simp add: le_ref_def, auto) next assume "P \ Q" and "Q \ P" then show "P = Q" by(auto simp: le_ref_def Process_eq_spec) } qed lemma lim_proc_is_ub:"chain S \ range S <| lim_proc (range S)" apply(auto simp: is_ub_def le_approx_def F_LUB D_LUB T_LUB Ra_def) using chain_lemma is_processT8 le_approx2 apply blast using D_T chain_lemma le_approx2T le_approx_def by blast lemma lim_proc_is_lub1: "chain S \ \ u . (range S <| u \ D u \ D (lim_proc (range S)))" by(auto simp: D_LUB, frule_tac i=a in Porder.ub_rangeD, auto dest: le_approx1) lemma lim_proc_is_lub2: "chain S \ \ u . range S <| u \ (\ s. s \ D (lim_proc (range S)) \ Ra (lim_proc (range S)) s = Ra u s)" apply(auto simp: is_ub_def D_LUB F_LUB Ra_def) using proc_ord2a apply blast using is_processT8_S proc_ord2a by blast lemma lim_proc_is_lub3a: "front_tickFree s \ s \ D P \ (\t. t \ D P \ \ t < s @ [c])" apply(rule impI, erule contrapos_np, auto) apply(auto simp: le_list_def less_list_def) by (metis D_def butlast_append butlast_snoc front_tickFree_mono is_process7 is_process_Rep self_append_conv) lemma lim_proc_is_lub3b: assumes 1 : "\x. x \ X \ (\xa. xa \ D x \ (\t. t \ D x \ \ t < xa) \ xa \ T u)" and 2 : "xa \ X" and 3 : "\xa. xa \ X \ x \ D xa" and 4 : "\t. (\x. x \ X \ t \ D x) \ \ t < x" shows "x \ T u" proof (cases "\t. t \ D xa \ \ t < x") case True from \\t. t \ D xa \ \ t < x\ show ?thesis using 1 2 3 by simp next case False from \\(\t. t \ D xa \ \ t < x)\ and 3 4 have A: "\y r c. y \ X \ r \ D y \ r \ D xa \ r @ [c] = x" by (metis D_imp_front_tickFree front_tickFree_charn less_self lim_proc_is_lub3a nil_less tickFree_implies_front_tickFree) show ?thesis apply(insert A) apply(erule exE)+ using "1" "3" D_imp_front_tickFree lim_proc_is_lub3a by blast qed lemma lim_proc_is_lub3c: assumes *:"chain S" and **:"X = range S" \\protection for range - otherwise auto unfolds and gets lost\ shows "\ u. X <| u \ min_elems(D(lim_proc X)) \ T u" proof - - have B : "D (lim_proc X) = INTER X D" by(simp add: * ** D_LUB) + have B : "D (lim_proc X) = \ (D ` X)" by(simp add: * ** D_LUB) show ?thesis apply(auto simp: is_ub_def * **) apply(auto simp: B min_elems_def le_approx_def HOL.imp_conjR HOL.all_conj_distrib Ball_def) apply(simp add: subset_iff imp_conjL[symmetric]) apply(rule lim_proc_is_lub3b[of "range S", simplified]) using "**" B by auto qed lemma limproc_is_lub: "chain S \ range S <<| lim_proc (range S)" apply (auto simp: is_lub_def lim_proc_is_ub) apply (simp add: le_approx_def is_lub_def lim_proc_is_ub) by (simp add: lim_proc_is_lub1 lim_proc_is_lub2 lim_proc_is_lub3c) lemma limproc_is_thelub: "chain S \ Lub S = lim_proc (range S)" by (frule limproc_is_lub,frule Porder.po_class.lub_eqI, simp) instance process :: (type) cpo proof fix S ::"nat \ '\ process" assume C:"chain S" thus "\ x. range S <<| x" using limproc_is_lub by blast qed instance process :: (type) pcpo proof show "\ x::'a process. \ y::'a process. x \ y" proof - have is_process_witness : "is_process({(s,X). front_tickFree s},{d. front_tickFree d})" apply(auto simp:is_process_def FAILURES_def DIVERGENCES_def) apply(auto elim!: tickFree_implies_front_tickFree front_tickFree_dw_closed front_tickFree_append) done have bot_inverse : "Rep_process(Abs_process({(s, X). front_tickFree s},Collect front_tickFree))= ({(s, X). front_tickFree s}, Collect front_tickFree)" by(subst Abs_process_inverse, simp_all add: Rep_process is_process_witness) have divergences_frontTickFree: "\y x. x \ snd (Rep_process y) \ front_tickFree x" by(rule D_imp_front_tickFree, simp add: D_def DIVERGENCES_def) have failures_frontTickFree: "\y s x. (s, x) \ fst (Rep_process y) \ front_tickFree s" by(rule is_processT2[rule_format], simp add: F_def FAILURES_def) have minelems_contains_mt: "\y x. x \ min_elems (Collect front_tickFree) \ x = []" by(simp add: min_elems_def front_tickFree_charn,safe, auto simp: Nil_elem_T) show ?thesis apply(rule_tac x="Abs_process ({(s,X). front_tickFree s},{d. front_tickFree d})" in exI) apply(auto simp: le_approx_def bot_inverse Ra_def F_def D_def FAILURES_def DIVERGENCES_def divergences_frontTickFree failures_frontTickFree) apply (metis minelems_contains_mt Nil_elem_T ) done qed qed subsection\ Process Refinement is admissible \ lemma le_adm[simp]: "cont (u::('a::cpo) \ 'b process) \ monofun v \ adm(\x. u x \ v x)" proof(auto simp add:le_ref_def cont2contlubE adm_def, goal_cases) case (1 Y x) hence "v (Y i) \ v (\i. Y i)" for i by (simp add: is_ub_thelub monofunE) hence "D (v (\i. Y i)) \ D (u (Y i))" for i using "1"(4) le_approx1 by blast then show ?case using D_LUB[OF ch2ch_cont[OF 1(1) 1(3)]] limproc_is_thelub[OF ch2ch_cont[OF 1(1) 1(3)]] "1"(5) by force next case (2 Y a b) hence "v (Y i) \ v (\i. Y i)" for i by (simp add: is_ub_thelub monofunE) hence "F (v (\i. Y i)) \ F (u (Y i))" for i using "2"(4) le_approx_lemma_F by blast then show ?case using F_LUB[OF ch2ch_cont[OF 2(1) 2(3)]] limproc_is_thelub[OF ch2ch_cont[OF 2(1) 2(3)]] "2"(5) by force qed lemmas le_adm_cont[simp] = le_adm[OF _ cont2mono] subsection\ Conditional statement is cont \ text\The conditional operator of CSP is obtained by a direct shallow embedding. Here we prove it continuous\ lemma if_then_else_cont[simp]: assumes *:"(\x. P x \ cont ((f::'c \ ('a::cpo) \ 'b process) x))" and **:"(\x. \ P x \ cont (g x))" shows "\x. cont(\y. if P x then f x y else g x y)" using * ** by (auto simp:cont_def) end diff --git a/thys/KD_Tree/Balanced.thy b/thys/KD_Tree/Balanced.thy --- a/thys/KD_Tree/Balanced.thy +++ b/thys/KD_Tree/Balanced.thy @@ -1,594 +1,596 @@ (* File: Balanced.thy Author: Martin Rau, TU München *) section "Building a balanced \k\-d Tree from a List of Points" theory Balanced imports KDTree Median_Of_Medians_Selection.Median_Of_Medians_Selection begin text \ Build a balanced \k\-d Tree by recursively partition the points into two lists. The partitioning criteria will be the median at a particular axis \ws\. The left list will contain all points \p\ with @{term "p!ws \ median"}. The right list will contain all points with median at axis @{term "ws \ p!a"}. The left and right list differ in length by one or none. The axis \ws\ will the widest spread axis. The axis which has the widest spread of point values. \ subsection "Widest Spread of a List of Points" definition spread_set :: "axis \ point set \ real" where "spread_set a ps = ( let as = (\p. p!a) ` ps in Max as - Min as )" definition spread :: "axis \ point list \ real" where "spread a ps = ( let as = map (\p. p!a) (tl ps) in fold max as (hd ps !a) - fold min as (hd ps!a) )" definition is_widest_spread :: "axis \ dimension \ point set \ bool" where "is_widest_spread a k ps = (\a' < k. spread_set a' ps \ spread_set a ps)" fun widest_spread_rec :: "axis \ point list \ axis * real" where "widest_spread_rec 0 ps = (0, spread 0 ps)" | "widest_spread_rec a ps = ( let (a', s') = widest_spread_rec (a - 1) ps in let s = spread a ps in if s \ s' then (a', s') else (a, s) )" definition widest_spread :: "point list \ axis" where "widest_spread ps = ( let (a, _) = widest_spread_rec (dim (hd ps) - 1) ps in a )" fun widest_spread_invar :: "dimension \ kdt \ bool" where "widest_spread_invar _ (Leaf _) \ True" | "widest_spread_invar k (Node a s l r) \ is_widest_spread a k (set_kdt l \ set_kdt r) \ widest_spread_invar k l \ widest_spread_invar k r" lemma spread_is_spread_set: "ps \ [] \ spread_set a (set ps) = spread a ps" using Max.set_eq_fold[of "hd ps !a" _] Min.set_eq_fold[of "hd ps !a"] apply (auto simp add: Let_def spread_def spread_set_def) by (metis (no_types, lifting) hd_Cons_tl list.simps(15) list.simps(9) set_map) lemma widest_spread_rec_is_spread: "(ws, s) = widest_spread_rec a ps \ s = spread ws ps" by (induction a) (auto simp add: Let_def split: prod.splits if_splits) lemma is_widest_spread_k_le_ws: "is_widest_spread ws k ps \ spread_set k ps \ spread_set ws ps \ is_widest_spread ws (k+1) ps" using is_widest_spread_def less_Suc_eq by auto lemma is_widest_spread_k_gt_ws: "is_widest_spread ws k ps \ \ (spread_set k ps \ spread_set ws ps) \ is_widest_spread k (k+1) ps" using is_widest_spread_def less_Suc_eq by auto lemma widest_spread_rec_le_a: "ps \ [] \ (ws, s) = widest_spread_rec a ps \ ws \ a" by (induction a arbitrary: ws s) (auto simp add: Let_def le_Suc_eq split: prod.splits if_splits) lemma widest_spread_rec_is_widest_spread: "ps \ [] \ (ws, s) = widest_spread_rec a ps \ is_widest_spread ws (a+1) (set ps)" proof (induction a arbitrary: ws s) case 0 thus ?case using is_widest_spread_def by simp next case (Suc a) then obtain ws' s' where *: "(ws', s') = widest_spread_rec a ps" by (metis surj_pair) hence "is_widest_spread ws' (Suc a) (set ps)" using Suc.IH Suc.prems(1) by simp then show ?case using Suc.prems * spread_is_spread_set widest_spread_rec_is_spread is_widest_spread_k_le_ws[of ws' "Suc a" "set ps"] is_widest_spread_k_gt_ws[of ws' "Suc a" "set ps"] by (auto simp add: Let_def split: prod.splits if_splits) qed lemma widest_spread_lt_k: "\p \ set ps. dim p = k \ 0 < k \ ps \ [] \ widest_spread ps < k" using widest_spread_def widest_spread_rec_le_a apply (auto split: prod.splits) by (metis Suc_le_lessD Suc_le_mono Suc_pred) lemma widest_spread_is_widest_spread: assumes "ps \ []" "\p \ set ps. dim p = k" shows "is_widest_spread (widest_spread ps) k (set ps)" proof (cases k) case 0 then show ?thesis using is_widest_spread_def by simp next case (Suc n) obtain ws s where *: "(ws, s) = widest_spread_rec (dim (hd ps) - 1) ps" using prod.collapse by blast moreover have "dim (hd ps) = k" using assms(1,2) by simp ultimately have "is_widest_spread ws (k - 1 + 1) (set ps)" using widest_spread_rec_is_widest_spread assms(1) by simp hence "is_widest_spread ws k (set ps)" using Suc by simp thus ?thesis using * widest_spread_def by (auto split: prod.split) qed subsection "Partitioning a List of Points by Median" definition axis_median :: "axis \ point list \ real" where "axis_median a ps = ( let n = (length ps - 1) div 2 in let ps' = map (\p. p!a) ps in fast_select n ps' )" lemma size_mset_map_P: "size {# y \# mset (map f xs). P y #} = size {# x \# mset xs. P (f x) #}" by (induction xs) auto lemma size_axis_median_length: assumes "0 < length ps" shows "size {# p \# mset ps. p!a < axis_median a ps #} \ (length ps - 1) div 2" (is "?thesis1") and "size {# p \# mset ps. axis_median a ps < p!a #} \ length ps div 2" (is "?thesis2") proof - let ?n = "(length ps - 1) div 2" let ?ps' = "map (\p. p!a) ps" let ?m = "fast_select ?n ?ps'" have "length ps = length ?ps'" by simp moreover have "?n < length ?ps'" using assms calculation by linarith ultimately have *: "median ?ps' = ?m" using median_def fast_select_correct by metis have "size {# a \# mset ?ps'. a < ?m #} \ (length ps - 1) div 2" using * size_less_than_median[of ?ps'] by simp hence "size {# p \# mset ps. p!a < ?m #} \ (length ps - 1) div 2" using size_mset_map_P[of "\a. a < ?m"] by metis thus ?thesis1 using axis_median_def by metis have "size {# a \# mset ?ps'. ?m < a #} \ length ps div 2" using * size_greater_than_median[of ?ps'] by simp hence "size {# p \# mset ps. ?m < p!a #} \ length ps div 2" using size_mset_map_P[of "\a. ?m < a"] by metis thus ?thesis2 using axis_median_def by metis qed subsubsection "Threeway Partitioning a List of Points" fun partition :: "axis \ real \ point list \ point list * point list * point list" where "partition _ _ [] = ([], [], [])" | "partition a m (p # ps) = ( let (lt, eq, gt) = partition a m ps in if p!a < m then (p # lt, eq, gt) else if p!a = m then (lt, p # eq, gt) else (lt, eq, p # gt) )" lemma partition_set: assumes "(lt, eq, gt) = partition a m ps" shows "set ps = set lt \ set eq \ set gt" using assms by (induction ps arbitrary: lt eq gt) (auto split: prod.splits if_splits) lemma partition_mset: assumes "(lt, eq, gt) = partition a m ps" shows "mset lt = {# p \# mset ps. p!a < m #}" and "mset eq = {# p \# mset ps. p!a = m #}" and "mset gt = {# p \# mset ps. m < p!a #}" using assms by (induction ps arbitrary: lt eq gt) (auto split: prod.splits if_splits) lemma partition_filter: assumes "(lt, eq, gt) = partition a m ps" shows "\p \ set lt. p!a < m" and "\p \ set eq. p!a = m" and "\p \ set gt. m < p!a" using assms by (induction ps arbitrary: lt eq gt) (auto split: prod.splits if_splits) lemma partition_length: assumes "(lt, eq, gt) = partition a m ps" shows "length ps = length lt + length eq + length gt" and "length lt = size {# p \# mset ps. p!a < m #}" and "length eq = size {# p \# mset ps. p!a = m #}" and "length gt = size {# p \# mset ps. m < p!a #}" using assms by (induction ps arbitrary: lt eq gt) (auto split: prod.splits if_splits) subsubsection "Partitioning a List of Points balanced by Median" definition partition_by_median :: "axis \ point list \ point list * real * point list" where "partition_by_median a ps = ( let m = axis_median a ps in let (lt, eq, gt) = partition a m ps in let rem = length ps div 2 - length lt in (lt @ take rem eq, m, drop rem eq @ gt) )" lemma set_take_drop: "set xs = set (take n xs) \ set (drop n xs)" by (metis append_take_drop_id set_append) lemma partition_by_median_set: assumes "(l, m, r) = partition_by_median a ps" shows "set ps = set l \ set r" using assms unfolding partition_by_median_def apply (simp add: Let_def split: prod.splits) using set_take_drop by (metis partition_set inf_sup_aci(6)) lemma partition_by_median_filter: assumes "(l, m, r) = partition_by_median a ps" shows "\p \ set l. p!a \ m" and "\p \ set r. m \ p!a" using assms unfolding partition_by_median_def apply (auto simp add: Let_def split: prod.splits) by (metis le_less dual_order.refl in_set_takeD in_set_dropD partition_filter)+ lemma partition_by_median_length_1: assumes "(l, m, r) = partition_by_median a ps" shows "length ps = length l + length r" using assms unfolding partition_by_median_def apply (simp add: Let_def min_def split: prod.splits) by (metis (no_types, lifting) add.assoc partition_length(1)) lemma partition_by_median_length_2: assumes "(l, m, r) = partition_by_median a ps" "0 < length ps" shows "length r - length l \ 1" and "length l \ length r" proof - let ?m = "axis_median a ps" let ?leg = "partition a ?m ps" let ?lt = "fst ?leg" let ?eq = "fst (snd ?leg)" let ?gt = "snd (snd ?leg)" let ?rem = "length ps div 2 - length ?lt" let ?l = "?lt @ take ?rem ?eq" let ?r = "drop ?rem ?eq @ ?gt" have *: "(?lt, ?eq, ?gt) = partition a ?m ps" by simp have "length ?lt \ (length ps - 1) div 2" using assms * partition_length(2) size_axis_median_length(1) by presburger moreover have "length ?gt \ length ps div 2" using assms * partition_length(4) size_axis_median_length(2) by presburger moreover have "length ps = length ?lt + length ?eq + length ?gt" using * partition_length(1) by simp ultimately have L: "length ?l = length ps div 2" by simp have **: "(?l, ?m, ?r) = partition_by_median a ps" by (auto simp add: Let_def partition_by_median_def split: prod.splits) hence "length ps = length ?l + length ?r" using partition_by_median_length_1 by blast hence "length ?l \ length ?r" "length ?r - length ?l \ 1" using L by linarith+ thus "length l \ length r" "length r - length l \ 1" using ** by (metis Pair_inject assms(1))+ qed lemma partition_by_median_length_3: assumes "(l, m, r) = partition_by_median a ps" "0 < length ps" shows "length l < length ps" proof (rule ccontr) assume *: "\ (length l < length ps)" have "length ps = length l + length r" using partition_by_median_length_1 assms(1) by simp hence "length l = length ps" "length r = 0" using * by simp_all moreover have "length l \ length r" using partition_by_median_length_2 assms by fastforce+ ultimately show "False" using assms(2) by simp qed lemma partition_by_median_length_4: assumes "(l, m, r) = partition_by_median a ps" "1 < length ps" shows "length r < length ps" proof (rule ccontr) assume *: "\ (length r < length ps)" have "length ps = length l + length r" using partition_by_median_length_1 assms(1) by simp hence "length r = length ps" "length l = 0" using * by simp_all moreover have "length r - length l \ 1" using partition_by_median_length_2 assms by fastforce+ ultimately show "False" using assms(2) by simp qed lemma partition_by_median_length_5: assumes "(l, m, r) = partition_by_median a ps" "1 < length ps" shows "0 < length l" and "0 < length r" using assms partition_by_median_length_1 partition_by_median_length_4 apply simp using assms partition_by_median_length_1 partition_by_median_length_3 by fastforce lemmas partition_by_median_length = partition_by_median_length_1 partition_by_median_length_2 partition_by_median_length_3 partition_by_median_length_4 partition_by_median_length_5 subsection \Building the Tree\ function (sequential) build :: "dimension \ point list \ kdt" where "build _ [] = undefined" (* We never hit this case recursively. Only if the original input is really [].*) | "build k [p] = Leaf p" | "build k ps = ( let a = widest_spread ps in let (l, m, r) = partition_by_median a ps in Node a m (build k l) (build k r) )" by pat_completeness auto termination build using partition_by_median_length(4,5) apply (relation "measure (\( _, ps). length ps)") apply (auto) apply fastforce+ done text \ Setting up different build simps for length induct. \ lemma build_simp_1: "ps = [p] \ build k ps = Leaf p" by simp lemma build_simp_2: "ps = p\<^sub>0 # p\<^sub>1 # ps' \ a = widest_spread ps \ (l, m, r) = partition_by_median a ps \ build k ps = Node a m (build k l) (build k r)" using build.simps(3) by (auto simp add: Let_def split: prod.splits) lemma length_ps_gt_1: "1 < length ps \ \p\<^sub>0 p\<^sub>1 ps'. ps = p\<^sub>0 # p\<^sub>1 # ps'" by (induction ps) (auto simp add: neq_Nil_conv) lemma build_simp_3: "1 < length ps \ a = widest_spread ps \ (l, m, r) = partition_by_median a ps \ build k ps = Node a m (build k l) (build k r)" using build_simp_2 length_ps_gt_1 by fast lemmas build_simps[simp] = build_simp_1 build_simp_3 declare build.simps[simp del] text \ The main theorems. \ theorem build_set: assumes "0 < length ps" shows "set ps = set_kdt (build k ps)" using assms proof (induction ps rule: length_induct) case (1 ps) then show ?case proof (cases "length ps \ 1") case True then obtain p where "ps = [p]" using "1.prems" by (cases ps) auto thus ?thesis by simp next case False let ?a = "widest_spread ps" let ?lmr = "partition_by_median ?a ps" let ?l = "fst ?lmr" let ?m = "fst (snd ?lmr)" let ?r = "snd (snd ?lmr)" have "set ?l = set_kdt (build k ?l)" "set ?r = set_kdt (build k ?r)" using False partition_by_median_length(4,5,6,7)[of ?l ?m ?r ?a ps] "1.IH" by force+ moreover have "set ps = set ?l \ set ?r" using partition_by_median_set by (metis prod.collapse) moreover have "build k ps = Node ?a ?m (build k ?l) (build k ?r)" using False by simp ultimately show ?thesis by auto qed qed theorem build_invar: assumes "0 < length ps" "\p \ set ps. dim p = k" "distinct ps" "0 < k" shows "invar k (build k ps)" using assms proof (induction ps rule: length_induct) case (1 ps) then show ?case proof (cases "length ps \ 1") case True then obtain p where P: "ps = [p]" using "1.prems"(1) by (cases ps) auto hence "dim p = k" using "1.prems"(2) by simp thus ?thesis using P by simp next case False let ?a = "widest_spread ps" let ?lmr = "partition_by_median ?a ps" let ?l = "fst ?lmr" let ?m = "fst (snd ?lmr)" let ?r = "snd (snd ?lmr)" have 1: "length ps = length ?l + length ?r" using partition_by_median_length(1) by (metis prod.collapse)+ hence 2: "length ?l < length ps" "length ?r < length ps" using False partition_by_median_length(4,5) not_le_imp_less "1.prems" by (metis prod.collapse)+ hence 3: "0 < length ?l" "0 < length ?r" using 1 False partition_by_median_length(6,7) by simp_all moreover have 4: "set ps = set ?l \ set ?r" using partition_by_median_set by (metis prod.collapse) moreover have "distinct ?l" "distinct ?r" and 5: "set ?l \ set ?r = {}" using "1.prems"(3) 1 4 by (metis card_distinct distinct_append distinct_card length_append set_append)+ moreover have "\p \ set ?l .dim p = k" "\p \ set ?r .dim p = k" using "1.prems"(2) 4 5 by simp_all ultimately have "invar k (build k ?l)" "invar k (build k ?r)" using "1.IH" 2 by (simp_all add: "1.prems"(4)) moreover have "\p \ set ?l. p ! ?a \ ?m" "\p \ set ?r. ?m \ p ! ?a" using partition_by_median_filter by (metis prod.collapse)+ moreover have "build k ps = Node ?a ?m (build k ?l) (build k ?r)" using False by simp moreover have "?a < k" using "1.prems"(1,2,4) widest_spread_lt_k by auto ultimately show ?thesis using 3 5 build_set by auto qed qed theorem build_widest_spread: assumes "0 < length ps" "\p \ set ps. dim p = k" shows "widest_spread_invar k (build k ps)" using assms proof (induction ps rule: length_induct) case (1 ps) then show ?case proof (cases "length ps \ 1") case True then obtain p where "ps = [p]" using "1.prems" by (cases ps) auto thus ?thesis by simp next case False let ?a = "widest_spread ps" let ?lmr = "partition_by_median ?a ps" let ?l = "fst ?lmr" let ?m = "fst (snd ?lmr)" let ?r = "snd (snd ?lmr)" have 1: "length ps = length ?l + length ?r" using partition_by_median_length(1) by (metis prod.collapse)+ hence 2: "length ?l < length ps" "length ?r < length ps" - using False partition_by_median_length(4,5) not_le_imp_less "1.prems" by (metis prod.collapse)+ + using False partition_by_median_length(4,5) not_le_imp_less "1.prems" + apply auto + by (metis prod.collapse)+ hence 3: "0 < length ?l" "0 < length ?r" using 1 False partition_by_median_length(6,7) by simp_all moreover have 4: "set ps = set ?l \ set ?r" using partition_by_median_set by (metis prod.collapse) moreover have "\p \ set ?l .dim p = k" "\p \ set ?r .dim p = k" using "1.prems"(2) 4 by simp_all ultimately have "widest_spread_invar k (build k ?l)" "widest_spread_invar k (build k ?r)" using "1.IH" 2 by simp_all moreover have "is_widest_spread ?a k (set_kdt (build k ?l) \ set_kdt (build k ?r))" using widest_spread_is_widest_spread "1.prems" 3 4 build_set by fastforce moreover have "build k ps = Node ?a ?m (build k ?l) (build k ?r)" using False by simp ultimately show ?thesis by auto qed qed theorem build_size: assumes "0 < length ps" shows "size_kdt (build k ps) = length ps" using assms proof (induction ps rule: length_induct) case (1 ps) then show ?case proof (cases "length ps \ 1") case True then obtain p where "ps = [p]" using "1.prems" by (cases ps) auto thus ?thesis by simp next case False let ?a = "widest_spread ps" let ?lmr = "partition_by_median ?a ps" let ?l = "fst ?lmr" let ?m = "fst (snd ?lmr)" let ?r = "snd (snd ?lmr)" have "size_kdt (build k ?l) = length ?l" "size_kdt (build k ?r) = length ?r" using False partition_by_median_length(4,5,6,7)[of ?l ?m ?r ?a ps] "1.IH" by force+ moreover have "build k ps = Node ?a ?m (build k ?l) (build k ?r)" using False by simp ultimately show ?thesis using partition_by_median_length(1) by (metis prod.collapse size_kdt.simps(2)) qed qed theorem build_balanced: assumes "0 < length ps" shows "balanced (build k ps)" using assms proof (induction ps rule: length_induct) case (1 ps) show ?case proof (cases "length ps \ 1") case True then obtain p where "ps = [p]" using "1.prems" by (cases ps) auto thus ?thesis unfolding balanced_def by simp next case False let ?a = "widest_spread ps" let ?lmr = "partition_by_median ?a ps" let ?l = "fst ?lmr" let ?m = "fst (snd ?lmr)" let ?r = "snd (snd ?lmr)" have 0: "length ps = length ?l + length ?r" "length ?r - length ?l \ 1" "length ?l \ length ?r" using partition_by_median_length(1,2,3) "1.prems" by (metis prod.collapse)+ hence 1: "length ?l + 1 = length ?r \ length ?l = length ?r" by linarith moreover have 2: "length ?l < length ps" "length ?r < length ps" using False 0 by auto moreover have 3: "0 < length ?l" "0 < length ?r" using "1.prems" 0 1 2 by linarith+ ultimately have 4: "balanced (build k ?l)" "balanced (build k ?r)" using "1.IH" by simp_all have "build k ps = Node ?a ?m (build k ?l) (build k ?r)" using False by simp moreover have "size_kdt (build k ?l) + 1 = size_kdt (build k ?r) \ size_kdt (build k ?l) = size_kdt (build k ?r)" using 1 3 build_size by simp ultimately show ?thesis using 4 balanced_Node_if_wbal2 by auto qed qed lemma complete_if_balanced_size_2powh: assumes "balanced kdt" "size_kdt kdt = 2 ^ h" shows "complete kdt" proof (rule ccontr) assume "\ complete kdt" hence "2 ^ (min_height kdt) < size_kdt kdt" "size_kdt kdt < 2 ^ height kdt" by (simp_all add: min_height_size_if_incomplete size_height_if_incomplete) hence "height kdt - min_height kdt > 1" using assms(2) by simp hence "\ balanced kdt" using balanced_def by simp thus "False" using assms(1) by simp qed theorem build_complete: assumes "length ps = 2 ^ h" shows "complete (build k ps)" using assms complete_if_balanced_size_2powh by (simp add: build_balanced build_size) corollary build_height: "length ps = 2 ^ h \ length ps = 2 ^ (height (build k ps))" using build_complete build_size complete_iff_size by auto end \ No newline at end of file diff --git a/thys/Kruskal/UGraph_Impl.thy b/thys/Kruskal/UGraph_Impl.thy --- a/thys/Kruskal/UGraph_Impl.thy +++ b/thys/Kruskal/UGraph_Impl.thy @@ -1,349 +1,349 @@ section "Kruskal on UGraphs" theory UGraph_Impl imports Kruskal_Impl UGraph begin definition "\ = (\(u,w,v). Upair u v)" subsection "Interpreting \Kruskl_Impl\ with a UGraph" abbreviation (in uGraph) "getEdges_SPEC csuper_E \ (SPEC (\L. distinct (map \ L) \ \ ` set L = E \ (\(a, wv, b)\set L. w (\ (a, wv, b)) = wv) \ set L \ csuper_E))" locale uGraph_impl = uGraph E w for E :: "nat uprod set" and w :: "nat uprod \ int" + fixes getEdges_impl :: "(nat \ int \ nat) list Heap" and csuper_E :: "(nat \ int \ nat) set" assumes getEdges_impl: "(uncurry0 getEdges_impl, uncurry0 (getEdges_SPEC csuper_E)) \ unit_assn\<^sup>k \\<^sub>a list_assn (nat_assn \\<^sub>a int_assn \\<^sub>a nat_assn)" begin - abbreviation "V \ (UNION E set_uprod)" + abbreviation "V \ \ (set_uprod ` E)" lemma max_node_is_Max_V: " E = \ ` set la \ max_node la = Max (insert 0 V)" proof - assume E: "E = \ ` set la" have *: "fst ` set la \ (snd \ snd) ` set la = (\x\set la. case x of (x1, x1a, x2a) \ {x1, x2a})" by auto force show ?thesis unfolding E using * by (auto simp add: \_def max_node_def prod.case_distrib) qed sublocale s: Kruskal_Impl E "\(set_uprod ` E)" set_uprod "\u v e. Upair u v = e" "subforest" "uconnectedV" w \ "PR_CONST (\(u,w,v). RETURN (u,v))" "PR_CONST (getEdges_SPEC csuper_E)" getEdges_impl "csuper_E" " (\(u,w,v). return (u,v))" unfolding subforest_def proof (unfold_locales, goal_cases) show "finite E" by simp next fix E' assume "forest E' \ E' \ E" then show "E' \ E" by auto next show "forest {} \ {} \ E" apply (auto simp: decycle_def forest_def) using epath.elims(2) by fastforce next fix X Y assume "forest X \ X \ E" "Y \ X" then show "forest Y \ Y \ E" using forest_mono by auto next case (5 u v) then show ?case unfolding uconnected_def apply auto using epath.elims(2) by force next - case (6 E1 E2 u v) - then have "(u, v) \ (uconnected E1)" and uv: "u \ (UNION E set_uprod)" "v\(UNION E set_uprod)" + case (6 E1 E2 u v) + then have "(u, v) \ (uconnected E1)" and uv: "u \ V" "v \ V" by auto then obtain p where 1: "epath E1 u p v" unfolding uconnected_def by auto from 6 uv have 2: "\(\p. epath E2 u p v)" unfolding uconnected_def by auto from 1 2 have "\a b. (a, b) \ uconnected E2 \ Upair a b \ E2 \ Upair a b \ E1" by(rule findaugmenting_edge) then show ?case by auto next case (7 F e u v) note f = \forest F \ F \ E\ note notin = \e \ E - F\ \Upair u v = e\ from notin ecard2 have unv: "u\v" by fastforce show "(forest (insert e F) \ insert e F \ E) = ((u, v) \ uconnectedV F)" proof assume a: "forest (insert e F) \ insert e F \ E " have "(u, v) \ uconnected F" apply(rule insert_stays_forest_means_not_connected) using notin a unv by auto - then show "((u, v) \ Restr (uconnected F) (UNION E set_uprod))" by auto + then show "((u, v) \ Restr (uconnected F) V)" by auto next - assume a: "(u, v) \ Restr (uconnected F) (UNION E set_uprod)" + assume a: "(u, v) \ Restr (uconnected F) V" have "forest (insert (Upair u v) F)" apply(rule augment_forest_overedges[where E=E]) using notin f a unv by auto moreover have "insert e F \ E" using notin f by auto ultimately show "forest (insert e F) \ insert e F \ E" using notin by auto qed next fix F assume "F\E" show "equiv V (uconnectedV F)" by(rule equiv_vert_uconnected) next case (9 F) then show ?case by auto next case (10 x y F) then show ?case using insert_uconnectedV_per' by metis next case (11 x) then show ?case apply(cases x) by auto next case (12 u v e) then show ?case by auto next case (13 u v e) then show ?case by auto next case (14 a F e) then show ?case using ecard2 by force next case (15 v) then show ?case using ecard2 by auto next case 16 show "V \ V" by auto next case 17 show "finite V" by simp next case (18 a b e T) then show ?case apply auto subgoal unfolding uconnected_def apply auto apply(rule exI[where x="[e]"]) apply simp using ecard2 by force subgoal by force subgoal by force done next case (19 xi x) then show ?case by (auto split: prod.splits simp: \_def) next case 20 show ?case by auto next case 21 show ?case using getEdges_impl by simp next case (22 l) from max_node_is_Max_V[OF 22] show "max_node l = Max (insert 0 V)" . next case (23) then show ?case apply sepref_to_hoare by sep_auto qed lemma spanningForest_eq_basis: "spanningForest = s.basis" unfolding spanningForest_def s.basis_def by auto lemma minSpanningForest_eq_minbasis: "minSpanningForest = s.minBasis" unfolding minSpanningForest_def s.MSF_def spanningForest_eq_basis by auto lemma kruskal_correct': " kruskal getEdges_impl (\(u,w,v). return (u,v)) () <\r. \ (distinct r \ set r \ csuper_E \ s.MSF (set (map \ r)))>\<^sub>t" using s.kruskal_correct_forest by auto lemma kruskal_correct: " kruskal getEdges_impl (\(u,w,v). return (u,v)) () <\r. \ (distinct r \ set r \ csuper_E \ minSpanningForest (set (map \ r)))>\<^sub>t" using s.kruskal_correct_forest minSpanningForest_eq_minbasis by auto end subsection "Kruskal on UGraph from list of concrete edges" definition "uGraph_from_list_\_weight L e = (THE w. \a' b'. Upair a' b' = e \ (a', w, b') \ set L)" abbreviation "uGraph_from_list_\_edges L \ \ ` set L" locale fromlist = fixes L :: "(nat \ int \ nat) list" assumes dist: "distinct (map \ L)" and no_selfloop: "\u w v. (u,w,v)\set L \ u\v" begin lemma not_distinct_map: "a\set l \ b\set l \ a\b \ \ a = \ b \ \ distinct (map \ l)" by (meson distinct_map_eq) lemma ii: "(a, aa, b) \ set L \ uGraph_from_list_\_weight L (Upair a b) = aa" unfolding uGraph_from_list_\_weight_def apply rule subgoal by auto apply clarify subgoal for w a' b' apply(auto) subgoal using distinct_map_eq[OF dist, of "(a, aa, b)" "(a, w, b)"] unfolding \_def by auto subgoal using distinct_map_eq[OF dist, of "(a, aa, b)" "(a', w, b')"] unfolding \_def by fastforce done done sublocale uGraph_impl "\ ` set L" "uGraph_from_list_\_weight L" "return L" "set L" proof (unfold_locales) fix e assume *: "e \ \ ` set L" from * obtain u w v where "(u,w,v) \ set L" "e = \ (u, w, v)" by auto then show "proper_uprod e" using no_selfloop unfolding \_def by auto next show "finite (\ ` set L)" by auto next show "(uncurry0 (return L),uncurry0((SPEC (\La. distinct (map \ La) \ \ ` set La = \ ` set L \ (\(aa, wv, ba)\set La. uGraph_from_list_\_weight L (\ (aa, wv, ba)) = wv) \ set La \ set L)))) \ unit_assn\<^sup>k \\<^sub>a list_assn (nat_assn \\<^sub>a int_assn \\<^sub>a nat_assn)" apply sepref_to_hoare using dist apply sep_auto subgoal using ii unfolding \_def by auto subgoal by simp subgoal by (auto simp: pure_fold list_assn_emp) done qed lemmas kruskal_correct = kruskal_correct definition (in -) "kruskal_algo L = kruskal (return L) (\(u,w,v). return (u,v)) ()" end subsection \Outside the locale\ definition uGraph_from_list_invar :: "(nat\int\nat) list \ bool" where "uGraph_from_list_invar L = (distinct (map \ L) \ (\p\set L. case p of (u,w,v) \u\v))" lemma uGraph_from_list_invar_conv: "uGraph_from_list_invar L = fromlist L" by(auto simp add: uGraph_from_list_invar_def fromlist_def) lemma uGraph_from_list_invar_subset: "uGraph_from_list_invar L \ set L'\ set L \ distinct L' \ uGraph_from_list_invar L'" unfolding uGraph_from_list_invar_def by (auto simp: distinct_map inj_on_subset) lemma uGraph_from_list_\_inj_on: "uGraph_from_list_invar E \ inj_on \ (set E)" by(auto simp: distinct_map uGraph_from_list_invar_def ) lemma sum_easier: "uGraph_from_list_invar L \ set E \ set L \ sum (uGraph_from_list_\_weight L) (uGraph_from_list_\_edges E) = sum (\(u,w,v). w) (set E)" proof - assume a: "uGraph_from_list_invar L" assume b: "set E \ set L" have *: "\e. e\set E \ ((\e. THE w. \a' b'. Upair a' b' = e \ (a', w, b') \ set L) \ \) e = (case e of (u, w, v) \ w)" apply simp apply(rule the_equality) subgoal using b by(auto simp: \_def split: prod.splits) subgoal using a b apply(auto simp: uGraph_from_list_invar_def distinct_map split: prod.splits) using \_def by (smt \_def inj_onD old.prod.case prod.inject set_mp) done have inj_on_E: "inj_on \ (set E)" apply(rule inj_on_subset) apply(rule uGraph_from_list_\_inj_on) by fact+ show ?thesis unfolding uGraph_from_list_\_weight_def apply(subst sum.reindex[OF inj_on_E] ) using * by auto qed lemma corr: "uGraph_from_list_invar L \ kruskal_algo L <\F. \ (uGraph_from_list_invar F \ set F \ set L \ uGraph.minSpanningForest (uGraph_from_list_\_edges L) (uGraph_from_list_\_weight L) (uGraph_from_list_\_edges F))>\<^sub>t" apply(sep_auto heap: fromlist.kruskal_correct simp: uGraph_from_list_invar_conv kruskal_algo_def ) using uGraph_from_list_invar_subset uGraph_from_list_invar_conv by simp lemma "uGraph_from_list_invar L \ kruskal_algo L <\F. \ (uGraph_from_list_invar F \ set F \ set L \ uGraph.spanningForest (uGraph_from_list_\_edges L) (uGraph_from_list_\_edges F) \ (\F'. uGraph.spanningForest (uGraph_from_list_\_edges L) (uGraph_from_list_\_edges F') \ set F' \ set L \ sum (\(u,w,v). w) (set F) \ sum (\(u,w,v). w) (set F')))>\<^sub>t" proof - assume a: "uGraph_from_list_invar L" then interpret fromlist L apply unfold_locales by (auto simp: uGraph_from_list_invar_def) from a show ?thesis by(sep_auto heap: corr simp: minSpanningForest_def sum_easier) qed subsection \Kruskal with input check\ definition "kruskal' L = kruskal (return L) (\(u,w,v). return (u,v)) ()" definition "kruskal_checked L = (if uGraph_from_list_invar L then do { F \ kruskal' L; return (Some F) } else return None)" lemma " kruskal_checked L <\ Some F \ \ (uGraph_from_list_invar L \ set F \ set L \ uGraph.minSpanningForest (uGraph_from_list_\_edges L) (uGraph_from_list_\_weight L) (uGraph_from_list_\_edges F)) | None \ \ (\ uGraph_from_list_invar L)>\<^sub>t" unfolding kruskal_checked_def apply(cases "uGraph_from_list_invar L") apply simp_all subgoal proof - assume [simp]: "uGraph_from_list_invar L" then interpret fromlist L apply unfold_locales by(auto simp: uGraph_from_list_invar_def) show ?thesis unfolding kruskal'_def by (sep_auto heap: kruskal_correct) qed subgoal by sep_auto done subsection \Code export\ export_code uGraph_from_list_invar checking SML_imp export_code kruskal_checked checking SML_imp ML_val \ val export_nat = @{code integer_of_nat} val import_nat = @{code nat_of_integer} val export_int = @{code integer_of_int} val import_int = @{code int_of_integer} val import_list = map (fn (a,b,c) => (import_nat a, (import_int b, import_nat c))) val export_list = map (fn (a,(b,c)) => (export_nat a, export_int b, export_nat c)) val export_Some_list = (fn SOME l => SOME (export_list l) | NONE => NONE) fun kruskal l = @{code kruskal} (fn () => import_list l) (fn (a,(_,c)) => fn () => (a,c)) () () |> export_list fun kruskal_checked l = @{code kruskal_checked} (import_list l) () |> export_Some_list val result = kruskal [(1,~9,2),(2,~3,3),(3,~4,1)] val result4 = kruskal [(1,~100,4), (3,64,5), (1,13,2), (3,20,2), (2,5,5), (4,80,3), (4,40,5)] val result' = kruskal_checked [(1,~9,2),(2,~3,3),(3,~4,1)] val result1' = kruskal_checked [(1,~9,2),(2,~3,3),(3,~4,1),(1,5,3)] val result2' = kruskal_checked [(1,~9,2),(2,~3,3),(3,~4,1),(3,~4,1)] val result3' = kruskal_checked [(1,~9,2),(2,~3,3),(3,~4,1),(1,~4,1)] val result4' = kruskal_checked [(1,~100,4), (3,64,5), (1,13,2), (3,20,2), (2,5,5), (4,80,3), (4,40,5)] \ end \ No newline at end of file diff --git a/thys/LTL_Master_Theorem/LTL_to_DRA/DRA_Construction.thy b/thys/LTL_Master_Theorem/LTL_to_DRA/DRA_Construction.thy --- a/thys/LTL_Master_Theorem/LTL_to_DRA/DRA_Construction.thy +++ b/thys/LTL_Master_Theorem/LTL_to_DRA/DRA_Construction.thy @@ -1,440 +1,443 @@ (* Author: Benedikt Seidl License: BSD *) section \Constructing DRAs for LTL Formulas\ theory DRA_Construction imports Transition_Functions "../Quotient_Type" "../Omega_Words_Fun_Stream" "../Logical_Characterization/Master_Theorem" Transition_Systems_and_Automata.DBA_Combine Transition_Systems_and_Automata.DCA_Combine Transition_Systems_and_Automata.DRA_Combine begin \ \We use prefix and suffix on infinite words.\ hide_const Sublist.prefix Sublist.suffix locale dra_construction = transition_functions eq + quotient eq Rep Abs for eq :: "'a ltln \ 'a ltln \ bool" (infix "\" 75) and Rep :: "'ltlq \ 'a ltln" and Abs :: "'a ltln \ 'ltlq" begin subsection \Lifting Setup\ abbreviation true\<^sub>n_lifted :: "'ltlq" ("\true\<^sub>n") where "\true\<^sub>n \ Abs true\<^sub>n" abbreviation false\<^sub>n_lifted :: "'ltlq" ("\false\<^sub>n") where "\false\<^sub>n \ Abs false\<^sub>n" abbreviation af_letter_lifted :: "'a set \ 'ltlq \ 'ltlq" ("\afletter") where "\afletter \ \ \ Abs (af_letter (Rep \) \)" abbreviation af_lifted :: "'ltlq \ 'a set list \ 'ltlq" ("\af") where "\af \ w \ fold \afletter w \" abbreviation GF_advice_lifted :: "'ltlq \ 'a ltln set \ 'ltlq" ("_\[_]\<^sub>\" [90,60] 89) where "\\[X]\<^sub>\ \ Abs ((Rep \)[X]\<^sub>\)" lemma af_letter_lifted_semantics: "\afletter \ (Abs \) = Abs (af_letter \ \)" by (metis Rep_Abs_eq af_letter_congruent Abs_eq) lemma af_lifted_semantics: "\af (Abs \) w = Abs (af \ w)" by (induction w rule: rev_induct) (auto simp: Abs_eq, insert Rep_Abs_eq af_letter_congruent eq_sym, blast) lemma af_lifted_range: "range (\af (Abs \)) \ {Abs \ | \. nested_prop_atoms \ \ nested_prop_atoms \}" using af_lifted_semantics af_nested_prop_atoms by blast definition af_letter\<^sub>F_lifted :: "'a ltln \ 'a set \ 'ltlq \ 'ltlq" ("\afletter\<^sub>F") where "\afletter\<^sub>F \ \ \ \ Abs (af_letter\<^sub>F \ (Rep \) \)" definition af_letter\<^sub>G_lifted :: "'a ltln \ 'a set \ 'ltlq \ 'ltlq" ("\afletter\<^sub>G") where "\afletter\<^sub>G \ \ \ \ Abs (af_letter\<^sub>G \ (Rep \) \)" lemma af_letter\<^sub>F_lifted_semantics: "\afletter\<^sub>F \ \ (Abs \) = Abs (af_letter\<^sub>F \ \ \)" by (metis af_letter\<^sub>F_lifted_def Rep_inverse af_letter\<^sub>F_def af_letter_congruent Abs_eq) lemma af_letter\<^sub>G_lifted_semantics: "\afletter\<^sub>G \ \ (Abs \) = Abs (af_letter\<^sub>G \ \ \)" by (metis af_letter\<^sub>G_lifted_def Rep_inverse af_letter\<^sub>G_def af_letter_congruent Abs_eq) abbreviation af\<^sub>F_lifted :: "'a ltln \ 'ltlq \ 'a set list \ 'ltlq" ("\af\<^sub>F") where "\af\<^sub>F \ \ w \ fold (\afletter\<^sub>F \) w \" abbreviation af\<^sub>G_lifted :: "'a ltln \ 'ltlq \ 'a set list \ 'ltlq" ("\af\<^sub>G") where "\af\<^sub>G \ \ w \ fold (\afletter\<^sub>G \) w \" lemma af\<^sub>F_lifted_semantics: "\af\<^sub>F \ (Abs \) w = Abs (af\<^sub>F \ \ w)" by (induction w rule: rev_induct) (auto simp: af_letter\<^sub>F_lifted_semantics) lemma af\<^sub>G_lifted_semantics: "\af\<^sub>G \ (Abs \) w = Abs (af\<^sub>G \ \ w)" by (induction w rule: rev_induct) (auto simp: af_letter\<^sub>G_lifted_semantics) lemma af\<^sub>F_lifted_range: "nested_prop_atoms \ \ nested_prop_atoms (F\<^sub>n \) \ range (\af\<^sub>F \ (Abs \)) \ {Abs \ | \. nested_prop_atoms \ \ nested_prop_atoms (F\<^sub>n \)}" using af\<^sub>F_lifted_semantics af\<^sub>F_nested_prop_atoms by blast lemma af\<^sub>G_lifted_range: "nested_prop_atoms \ \ nested_prop_atoms (G\<^sub>n \) \ range (\af\<^sub>G \ (Abs \)) \ {Abs \ | \. nested_prop_atoms \ \ nested_prop_atoms (G\<^sub>n \)}" using af\<^sub>G_lifted_semantics af\<^sub>G_nested_prop_atoms by blast definition af_letter\<^sub>\_lifted :: "'a ltln set \ 'a set \ 'ltlq \ 'ltlq \ 'ltlq \ 'ltlq" ("\afletter\<^sub>\") where "\afletter\<^sub>\ X \ p \ (Abs (fst (af_letter\<^sub>\ X (Rep (fst p), Rep (snd p)) \)), Abs (snd (af_letter\<^sub>\ X (Rep (fst p), Rep (snd p)) \)))" abbreviation af\<^sub>\_lifted :: "'a ltln set \ 'ltlq \ 'ltlq \ 'a set list \ 'ltlq \ 'ltlq" ("\af\<^sub>\") where "\af\<^sub>\ X p w \ fold (\afletter\<^sub>\ X) w p" lemma af_letter\<^sub>\_lifted_semantics: "\afletter\<^sub>\ X \ (Abs x, Abs y) = (Abs (fst (af_letter\<^sub>\ X (x, y) \)), Abs (snd (af_letter\<^sub>\ X (x, y) \)))" by (simp add: af_letter\<^sub>\_def af_letter\<^sub>\_lifted_def) (insert GF_advice_congruent Rep_Abs_eq Rep_inverse af_letter_lifted_semantics eq_trans Abs_eq, blast) lemma af\<^sub>\_lifted_semantics: "\af\<^sub>\ X (Abs \, Abs \) w = (Abs (fst (af\<^sub>\ X (\, \) w)), Abs (snd (af\<^sub>\ X (\, \) w)))" apply (induction w rule: rev_induct) apply (auto simp: af_letter\<^sub>\_lifted_def af_letter\<^sub>\_lifted_semantics af_letter_lifted_semantics) by (metis (no_types, hide_lams) af_letter\<^sub>\_lifted_def af\<^sub>\_fst af_letter\<^sub>\_lifted_semantics eq_fst_iff prod.sel(2)) subsection \Büchi automata for basic languages\ definition \\<^sub>\ :: "'a ltln \ ('a set, 'ltlq) dba" where "\\<^sub>\ \ = dba UNIV (Abs \) \afletter (\\. \ = \true\<^sub>n)" definition \\<^sub>\_GF :: "'a ltln \ ('a set, 'ltlq) dba" where "\\<^sub>\_GF \ = dba UNIV (Abs (F\<^sub>n \)) (\afletter\<^sub>F \) (\\. \ = \true\<^sub>n)" definition \\<^sub>\ :: "'a ltln \ ('a set, 'ltlq) dca" where "\\<^sub>\ \ = dca UNIV (Abs \) \afletter (\\. \ = \false\<^sub>n)" definition \\<^sub>\_FG :: "'a ltln \ ('a set, 'ltlq) dca" where "\\<^sub>\_FG \ = dca UNIV (Abs (G\<^sub>n \)) (\afletter\<^sub>G \) (\\. \ = \false\<^sub>n)" lemma dba_run: "DBA.run (dba UNIV p \ \) (to_stream w) p" unfolding DBA.run_def by (rule transition_system.snth_run) fastforce lemma dca_run: "DCA.run (dca UNIV p \ \) (to_stream w) p" unfolding DCA.run_def by (rule transition_system.snth_run) fastforce lemma \\<^sub>\_language: "\ \ \LTL \ to_stream w \ DBA.language (\\<^sub>\ \) \ w \\<^sub>n \" proof - assume "\ \ \LTL" then have "w \\<^sub>n \ \ (\n. \k\n. af \ (w[0 \ k]) \ true\<^sub>n)" by (meson af_\LTL af_prefix_true le_cases) also have "\ \ (\n. \k\n. af \ (w[0 \ Suc k]) \ true\<^sub>n)" by (meson af_prefix_true le_SucI order_refl) also have "\ \ infs (\\. \ = \true\<^sub>n) (DBA.trace (\\<^sub>\ \) (to_stream w) (Abs \))" by (simp add: infs_snth \\<^sub>\_def DBA.transition_def af_lifted_semantics Abs_eq[symmetric] af_letter_lifted_semantics) also have "\ \ to_stream w \ DBA.language (\\<^sub>\ \)" unfolding \\<^sub>\_def dba.initial_def dba.accepting_def by (metis dba_run DBA.language DBA.language_elim dba.sel(2) dba.sel(4)) finally show ?thesis by simp qed lemma \\<^sub>\_GF_language: "\ \ \LTL \ to_stream w \ DBA.language (\\<^sub>\_GF \) \ w \\<^sub>n G\<^sub>n (F\<^sub>n \)" proof - assume "\ \ \LTL" then have "w \\<^sub>n G\<^sub>n (F\<^sub>n \) \ (\n. \k. af (F\<^sub>n \) (w[n \ k]) \\<^sub>L true\<^sub>n)" using ltl_lang_equivalence.af_\LTL_GF by blast also have "\ \ (\n. \k>n. af\<^sub>F \ (F\<^sub>n \) (w[0 \ k]) \ true\<^sub>n)" using af\<^sub>F_semantics_ltr af\<^sub>F_semantics_rtl using \\ \ \LTL\ af_\LTL_GF calculation by blast also have "\ \ (\n. \k\n. af\<^sub>F \ (F\<^sub>n \) (w[0 \ Suc k]) \ true\<^sub>n)" by (metis less_Suc_eq_le less_imp_Suc_add) also have "\ \ infs (\\. \ = \true\<^sub>n) (DBA.trace (\\<^sub>\_GF \) (to_stream w) (Abs (F\<^sub>n \)))" by (simp add: infs_snth \\<^sub>\_GF_def DBA.transition_def af\<^sub>F_lifted_semantics Abs_eq[symmetric] af_letter\<^sub>F_lifted_semantics) also have "\ \ to_stream w \ DBA.language (\\<^sub>\_GF \)" unfolding \\<^sub>\_GF_def dba.initial_def dba.accepting_def by (metis dba_run DBA.language DBA.language_elim dba.sel(2) dba.sel(4)) finally show ?thesis by simp qed lemma \\<^sub>\_language: "\ \ \LTL \ to_stream w \ DCA.language (\\<^sub>\ \) \ w \\<^sub>n \" proof - assume "\ \ \LTL" then have "w \\<^sub>n \ \ (\n. \k\n. \ af \ (w[0 \ k]) \ false\<^sub>n)" by (meson af_\LTL af_prefix_false le_cases order_refl) also have "\ \ (\n. \k\n. \ af \ (w[0 \ Suc k]) \ false\<^sub>n)" by (meson af_prefix_false le_SucI order_refl) also have "\ \ fins (\\. \ = \false\<^sub>n) (DCA.trace (\\<^sub>\ \) (to_stream w) (Abs \))" by (simp add: infs_snth \\<^sub>\_def DBA.transition_def af_lifted_semantics Abs_eq[symmetric] af_letter_lifted_semantics) also have "\ \ to_stream w \ DCA.language (\\<^sub>\ \)" unfolding \\<^sub>\_def dca.initial_def dca.rejecting_def by (metis dca_run DCA.language DCA.language_elim dca.sel(2) dca.sel(4)) finally show ?thesis by simp qed lemma \\<^sub>\_FG_language: "\ \ \LTL \ to_stream w \ DCA.language (\\<^sub>\_FG \) \ w \\<^sub>n F\<^sub>n (G\<^sub>n \)" proof - assume "\ \ \LTL" then have "w \\<^sub>n F\<^sub>n (G\<^sub>n \) \ (\k. \j. \ af (G\<^sub>n \) (w[k \ j]) \\<^sub>L false\<^sub>n)" using ltl_lang_equivalence.af_\LTL_FG by blast also have "\ \ (\n. \k>n. \ af\<^sub>G \ (G\<^sub>n \) (w[0 \ k]) \ false\<^sub>n)" using af\<^sub>G_semantics_ltr af\<^sub>G_semantics_rtl using \\ \ \LTL\ af_\LTL_FG calculation by blast also have "\ \ (\n. \k\n. \ af\<^sub>G \ (G\<^sub>n \) (w[0 \ Suc k]) \ false\<^sub>n)" by (metis less_Suc_eq_le less_imp_Suc_add) also have "\ \ fins (\\. \ = \false\<^sub>n) (DCA.trace (\\<^sub>\_FG \) (to_stream w) (Abs (G\<^sub>n \)))" by (simp add: infs_snth \\<^sub>\_FG_def DBA.transition_def af\<^sub>G_lifted_semantics Abs_eq[symmetric] af_letter\<^sub>G_lifted_semantics) also have "\ \ to_stream w \ DCA.language (\\<^sub>\_FG \)" unfolding \\<^sub>\_FG_def dca.initial_def dca.rejecting_def by (metis dca_run DCA.language DCA.language_elim dca.sel(2) dca.sel(4)) finally show ?thesis by simp qed lemma \\<^sub>\_nodes: "DBA.nodes (\\<^sub>\ \) \ {Abs \ | \. nested_prop_atoms \ \ nested_prop_atoms \}" unfolding \\<^sub>\_def transition_system_initial.nodes_alt_def using af_lifted_semantics af_nested_prop_atoms by fastforce lemma \\<^sub>\_GF_nodes: "DBA.nodes (\\<^sub>\_GF \) \ {Abs \ | \. nested_prop_atoms \ \ nested_prop_atoms (F\<^sub>n \)}" unfolding \\<^sub>\_GF_def DBA.nodes_def transition_system_initial.nodes_alt_def transition_system.reachable_alt_def using af\<^sub>F_nested_prop_atoms[of "F\<^sub>n \"] by (auto simp: af\<^sub>F_lifted_semantics) lemma \\<^sub>\_nodes: "DCA.nodes (\\<^sub>\ \) \ {Abs \ | \. nested_prop_atoms \ \ nested_prop_atoms \}" unfolding \\<^sub>\_def transition_system_initial.nodes_alt_def using af_lifted_semantics af_nested_prop_atoms by fastforce lemma \\<^sub>\_FG_nodes: "DCA.nodes (\\<^sub>\_FG \) \ {Abs \ | \. nested_prop_atoms \ \ nested_prop_atoms (G\<^sub>n \)}" unfolding \\<^sub>\_FG_def DCA.nodes_def transition_system_initial.nodes_alt_def transition_system.reachable_alt_def using af\<^sub>G_nested_prop_atoms[of "G\<^sub>n \"] by (auto simp: af\<^sub>G_lifted_semantics) subsection \A DCA checking the GF-advice Function\ definition \ :: "'a ltln \ 'a ltln set \ ('a set, 'ltlq \ 'ltlq) dca" where "\ \ X = dca UNIV (Abs \, Abs (\[X]\<^sub>\)) (\afletter\<^sub>\ X) (\p. snd p = \false\<^sub>n)" lemma \_language: "to_stream w \ DCA.language (\ \ X) \ (\i. suffix i w \\<^sub>n af \ (prefix i w)[X]\<^sub>\)" proof - have "(\i. suffix i w \\<^sub>n af \ (prefix i w)[X]\<^sub>\) \ (\m. \k\m. \ snd (af\<^sub>\ X (\, \[X]\<^sub>\) (prefix (Suc k) w)) \ false\<^sub>n)" using af\<^sub>\_semantics_ltr af\<^sub>\_semantics_rtl by blast also have "\ \ fins (\p. snd p = \false\<^sub>n) (DCA.trace (\ \ X) (to_stream w) (Abs \, Abs (\[X]\<^sub>\)))" by(simp add: infs_snth \_def DCA.transition_def af\<^sub>\_lifted_semantics af_letter\<^sub>\_lifted_semantics Abs_eq) also have "\ \ to_stream w \ DCA.language (\ \ X)" by (simp add: \_def dca.initial_def dca.rejecting_def DCA.language_def dca_run) finally show ?thesis by blast qed lemma \_nodes: "DCA.nodes (\ \ X) \ {Abs \ | \. nested_prop_atoms \ \ nested_prop_atoms \} \ {Abs \ | \. nested_prop_atoms \ \ nested_prop_atoms\<^sub>\ \ X}" unfolding \_def DCA.nodes_def transition_system_initial.nodes_alt_def transition_system.reachable_alt_def apply (auto simp add: af\<^sub>\_lifted_semantics af_letter\<^sub>\_lifted_semantics) using af\<^sub>\_fst_nested_prop_atoms apply force using GF_advice_nested_prop_atoms\<^sub>\ af\<^sub>\_snd_nested_prop_atoms dra_construction_axioms by fastforce subsection \A DRA for each combination of sets X and Y\ lemma dba_language: "(\w. to_stream w \ DBA.language \ \ w \\<^sub>n \) \ DBA.language \ = {w. to_omega w \\<^sub>n \}" by (metis (mono_tags, lifting) Collect_cong DBA.language_def mem_Collect_eq to_stream_to_omega) lemma dca_language: "(\w. to_stream w \ DCA.language \ \ w \\<^sub>n \) \ DCA.language \ = {w. to_omega w \\<^sub>n \}" by (metis (mono_tags, lifting) Collect_cong DCA.language_def mem_Collect_eq to_stream_to_omega) definition \\<^sub>1 :: "'a ltln \ 'a ltln list \ ('a set, 'ltlq \ 'ltlq) dca" where "\\<^sub>1 \ xs = \ \ (set xs)" lemma \\<^sub>1_language: "to_omega ` DCA.language (\\<^sub>1 \ xs) = L\<^sub>1 \ (set xs)" by (simp add: \\<^sub>1_def L\<^sub>1_def set_eq_iff \_language) lemma \\<^sub>1_alphabet: "DCA.alphabet (\\<^sub>1 \ xs) = UNIV" unfolding \\<^sub>1_def \_def by simp definition \\<^sub>2 :: "'a ltln list \ 'a ltln list \ ('a set, 'ltlq list degen) dba" where "\\<^sub>2 xs ys = dbail (map (\\. \\<^sub>\_GF (\[set ys]\<^sub>\)) xs)" lemma \\<^sub>2_language: "to_omega ` DBA.language (\\<^sub>2 xs ys) = L\<^sub>2 (set xs) (set ys)" by (simp add: \\<^sub>2_def L\<^sub>2_def set_eq_iff dba_language[OF \\<^sub>\_GF_language[OF FG_advice_\LTL(1)]]) lemma \\<^sub>2_alphabet: "DBA.alphabet (\\<^sub>2 xs ys) = UNIV" by (simp add: \\<^sub>2_def \\<^sub>\_GF_def dbail_def dbgail_def) definition \\<^sub>3 :: "'a ltln list \ 'a ltln list \ ('a set, 'ltlq list) dca" where "\\<^sub>3 xs ys = dcail (map (\\. \\<^sub>\_FG (\[set xs]\<^sub>\)) ys)" lemma \\<^sub>3_language: "to_omega ` DCA.language (\\<^sub>3 xs ys) = L\<^sub>3 (set xs) (set ys)" by (simp add: \\<^sub>3_def L\<^sub>3_def set_eq_iff dca_language[OF \\<^sub>\_FG_language[OF GF_advice_\LTL(1)]]) lemma \\<^sub>3_alphabet: "DCA.alphabet (\\<^sub>3 xs ys) = UNIV" by (simp add: \\<^sub>3_def \\<^sub>\_FG_def dcail_def) definition "\' \ xs ys = dbcrai (\\<^sub>2 xs ys) (dcai (\\<^sub>1 \ xs) (\\<^sub>3 xs ys))" lemma \'_language: "to_omega ` DRA.language (\' \ xs ys) = (L\<^sub>1 \ (set xs) \ L\<^sub>2 (set xs) (set ys) \ L\<^sub>3 (set xs) (set ys))" by (simp add: \'_def \\<^sub>1_language \\<^sub>2_language \\<^sub>3_language) fastforce lemma \'_alphabet: "DRA.alphabet (\' \ xs ys) = UNIV" by (simp add: \'_def dbcrai_def dcai_def \\<^sub>1_alphabet \\<^sub>2_alphabet \\<^sub>3_alphabet) subsection \A DRA for @{term "L(\)"}\ definition "ltl_to_dra \ = draul (map (\(xs, ys). \' \ xs ys) (advice_sets \))" lemma ltl_to_dra_language: "to_omega ` DRA.language (ltl_to_dra \) = language_ltln \" proof - - have 1: "INTER (set (map (\(x, y). \' \ x y) (advice_sets \))) dra.alphabet = UNION (set (map (\(x, y). \' \ x y) (advice_sets \))) dra.alphabet" - by (induction "advice_sets \") (metis advice_sets_not_empty, simp add: \'_alphabet split_def advice_sets_not_empty) - + have "(\(a, b)\set (advice_sets \). dra.alphabet (\' \ a b)) = + (\(a, b)\set (advice_sets \). dra.alphabet (\' \ a b))" + using advice_sets_not_empty by (simp add: \'_alphabet) + then have *: "DRA.language (draul (map (\(x, y). \' \ x y) (advice_sets \))) = + \ (DRA.language ` set (map (\(x, y). \' \ x y) (advice_sets \)))" + by (simp add: split_def) have "language_ltln \ = \ {(L\<^sub>1 \ X \ L\<^sub>2 X Y \ L\<^sub>3 X Y) | X Y. X \ subformulas\<^sub>\ \ \ Y \ subformulas\<^sub>\ \}" unfolding master_theorem_language by auto also have "\ = \ {L\<^sub>1 \ (set xs) \ L\<^sub>2 (set xs) (set ys) \ L\<^sub>3 (set xs) (set ys) | xs ys. (xs, ys) \ set (advice_sets \)}" unfolding advice_sets_subformulas by (metis (no_types, lifting)) also have "\ = \ {to_omega ` DRA.language (\' \ xs ys) | xs ys. (xs, ys) \ set (advice_sets \)}" by (simp add: \'_language) finally show ?thesis - unfolding ltl_to_dra_def draul_language[OF 1] by auto + using * by (auto simp add: ltl_to_dra_def) qed lemma ltl_to_dra_alphabet: "alphabet (ltl_to_dra \) = UNIV" by (auto simp: ltl_to_dra_def draul_def \'_alphabet split: prod.split) (insert advice_sets_subformulas, blast) subsection \Setting the Alphabet of a DRA\ definition dra_set_alphabet :: "('a set, 'b) dra \ 'a set set \ ('a set, 'b) dra" where "dra_set_alphabet \ \ = dra \ (initial \) (transition \) (accepting \)" lemma dra_set_alphabet_language: "\ \ alphabet \ \ language (dra_set_alphabet \ \) = language \ \ {s. sset s \ \}" by (auto simp add: dra_set_alphabet_def language_def set_eq_iff DRA.run_alt_def) lemma dra_set_alphabet_alphabet[simp]: "alphabet (dra_set_alphabet \ \) = \" unfolding dra_set_alphabet_def by simp lemma dra_set_alphabet_nodes: "\ \ alphabet \ \ DRA.nodes (dra_set_alphabet \ \) \ DRA.nodes \" unfolding dra_set_alphabet_def DRA.nodes_def transition_system_initial.nodes_alt_def transition_system.reachable_alt_def by auto (metis DRA.path_alt_def DRA.path_def dra.sel(1) dra.sel(3) subset_trans) subsection \A DRA for @{term "L(\)"} with a finite alphabet\ definition "ltl_to_dra_alphabet \ Ap = dra_set_alphabet (ltl_to_dra \) (Pow Ap)" lemma ltl_to_dra_alphabet_language: assumes "atoms_ltln \ \ Ap" shows "to_omega ` language (ltl_to_dra_alphabet \ Ap) = language_ltln \ \ {w. range w \ Pow Ap}" proof - have 1: "Pow Ap \ alphabet (ltl_to_dra \)" unfolding ltl_to_dra_alphabet by simp show ?thesis unfolding ltl_to_dra_alphabet_def dra_set_alphabet_language[OF 1] by (simp add: ltl_to_dra_language sset_range) force qed lemma ltl_to_dra_alphabet_alphabet[simp]: "alphabet (ltl_to_dra_alphabet \ Ap) = Pow Ap" unfolding ltl_to_dra_alphabet_def by simp lemma ltl_to_dra_alphabet_nodes: "DRA.nodes (ltl_to_dra_alphabet \ Ap) \ DRA.nodes (ltl_to_dra \)" unfolding ltl_to_dra_alphabet_def by (rule dra_set_alphabet_nodes) (simp add: ltl_to_dra_alphabet) end end diff --git a/thys/Polynomials/MPoly_PM.thy b/thys/Polynomials/MPoly_PM.thy --- a/thys/Polynomials/MPoly_PM.thy +++ b/thys/Polynomials/MPoly_PM.thy @@ -1,4394 +1,4394 @@ (* Author: Alexander Maletzky *) section \Multivariate Polynomials with Power-Products Represented by Polynomial Mappings\ theory MPoly_PM imports Quasi_PM_Power_Products begin text \Many notions introduced in this theory for type @{typ "('x \\<^sub>0 'a) \\<^sub>0 'b"} closely resemble those introduced in @{theory Polynomials.MPoly_Type} for type @{typ "'a mpoly"}.\ lemma monomial_single_power: "(monomial c (Poly_Mapping.single x k)) ^ n = monomial (c ^ n) (Poly_Mapping.single x (k * n))" proof - have eq: "(\i = 0.. t)" proof - have "(\i = 0..i = 0.. t" by (simp only: map_scale_sum_distrib_right map_scale_one_left) thus ?thesis by (simp add: punit.monomial_power) qed lemma times_canc_left: assumes "h * p = h * q" and "h \ (0::('x::linorder \\<^sub>0 nat) \\<^sub>0 'a::ring_no_zero_divisors)" shows "p = q" proof (rule ccontr) assume "p \ q" hence "p - q \ 0" by simp with assms(2) have "h * (p - q) \ 0" by simp hence "h * p \ h * q" by (simp add: algebra_simps) thus False using assms(1) .. qed lemma times_canc_right: assumes "p * h = q * h" and "h \ (0::('x::linorder \\<^sub>0 nat) \\<^sub>0 'a::ring_no_zero_divisors)" shows "p = q" proof (rule ccontr) assume "p \ q" hence "p - q \ 0" by simp hence "(p - q) * h \ 0" using assms(2) by simp hence "p * h \ q * h" by (simp add: algebra_simps) thus False using assms(1) .. qed subsection \Degree\ lemma plus_minus_assoc_pm_nat_1: "s + t - u = (s - (u - t)) + (t - (u::_ \\<^sub>0 nat))" by (rule poly_mapping_eqI, simp add: lookup_add lookup_minus) lemma plus_minus_assoc_pm_nat_2: "s + (t - u) = (s + (except (u - t) (- keys s))) + t - (u::_ \\<^sub>0 nat)" proof (rule poly_mapping_eqI) fix x show "lookup (s + (t - u)) x = lookup (s + except (u - t) (- keys s) + t - u) x" proof (cases "x \ keys s") case True thus ?thesis by (simp add: plus_minus_assoc_pm_nat_1 lookup_add lookup_minus lookup_except) next case False hence "lookup s x = 0" by (simp add: in_keys_iff) with False show ?thesis by (simp add: lookup_add lookup_minus lookup_except) qed qed lemma deg_pm_sum: "deg_pm (sum t A) = (\a\A. deg_pm (t a))" by (induct A rule: infinite_finite_induct) (auto simp: deg_pm_plus) lemma deg_pm_mono: "s adds t \ deg_pm s \ deg_pm (t::_ \\<^sub>0 _::add_linorder_min)" by (metis addsE deg_pm_plus le_iff_add) lemma adds_deg_pm_antisym: "s adds t \ deg_pm t \ deg_pm (s::_ \\<^sub>0 _::add_linorder_min) \ s = t" by (metis (no_types, lifting) add.right_neutral add.right_neutral add_left_cancel addsE deg_pm_eq_0_iff deg_pm_mono deg_pm_plus dual_order.antisym) lemma deg_pm_minus: assumes "s adds (t::_ \\<^sub>0 _::comm_monoid_add)" shows "deg_pm (t - s) = deg_pm t - deg_pm s" proof - from assms have "(t - s) + s = t" by (rule adds_minus) hence "deg_pm t = deg_pm ((t - s) + s)" by simp also have "\ = deg_pm (t - s) + deg_pm s" by (simp only: deg_pm_plus) finally show ?thesis by simp qed lemma adds_group [simp]: "s adds (t::'a \\<^sub>0 'b::ab_group_add)" proof (rule addsI) show "t = s + (t - s)" by simp qed lemmas deg_pm_minus_group = deg_pm_minus[OF adds_group] lemma deg_pm_minus_le: "deg_pm (t - s) \ deg_pm (t::_ \\<^sub>0 nat)" proof - have "keys (t - s) \ keys t" by (rule, simp add: lookup_minus in_keys_iff) hence "deg_pm (t - s) = (\x\keys t. lookup (t - s) x)" using finite_keys by (rule deg_pm_superset) also have "\ \ (\x\keys t. lookup t x)" by (rule sum_mono) (simp add: lookup_minus) also have "\ = deg_pm t" by (rule sym, rule deg_pm_superset, fact subset_refl, fact finite_keys) finally show ?thesis . qed lemma minus_id_iff: "t - s = t \ keys t \ keys (s::_ \\<^sub>0 nat) = {}" proof assume "t - s = t" { fix x assume "x \ keys t" and "x \ keys s" hence "0 < lookup t x" and "0 < lookup s x" by (simp_all add: in_keys_iff) hence "lookup (t - s) x \ lookup t x" by (simp add: lookup_minus) with \t - s = t\ have False by simp } thus "keys t \ keys s = {}" by blast next assume *: "keys t \ keys s = {}" show "t - s = t" proof (rule poly_mapping_eqI) fix x have "lookup t x - lookup s x = lookup t x" proof (cases "x \ keys t") case True with * have "x \ keys s" by blast thus ?thesis by (simp add: in_keys_iff) next case False thus ?thesis by (simp add: in_keys_iff) qed thus "lookup (t - s) x = lookup t x" by (simp only: lookup_minus) qed qed lemma deg_pm_minus_id_iff: "deg_pm (t - s) = deg_pm t \ keys t \ keys (s::_ \\<^sub>0 nat) = {}" proof assume eq: "deg_pm (t - s) = deg_pm t" { fix x assume "x \ keys t" and "x \ keys s" hence "0 < lookup t x" and "0 < lookup s x" by (simp_all add: in_keys_iff) hence *: "lookup (t - s) x < lookup t x" by (simp add: lookup_minus) have "keys (t - s) \ keys t" by (rule, simp add: lookup_minus in_keys_iff) hence "deg_pm (t - s) = (\x\keys t. lookup (t - s) x)" using finite_keys by (rule deg_pm_superset) also from finite_keys have "\ < (\x\keys t. lookup t x)" proof (rule sum_strict_mono_ex1) show "\x\keys t. lookup (t - s) x \ lookup t x" by (simp add: lookup_minus) next from \x \ keys t\ * show "\x\keys t. lookup (t - s) x < lookup t x" .. qed also have "\ = deg_pm t" by (rule sym, rule deg_pm_superset, fact subset_refl, fact finite_keys) finally have False by (simp add: eq) } thus "keys t \ keys s = {}" by blast next assume "keys t \ keys s = {}" hence "t - s = t" by (simp only: minus_id_iff) thus "deg_pm (t - s) = deg_pm t" by (simp only:) qed definition poly_deg :: "(('x \\<^sub>0 'a::add_linorder) \\<^sub>0 'b::zero) \ 'a" where "poly_deg p = (if keys p = {} then 0 else Max (deg_pm ` keys p))" definition maxdeg :: "(('x \\<^sub>0 'a::add_linorder) \\<^sub>0 'b::zero) set \ 'a" where "maxdeg A = Max (poly_deg ` A)" definition mindeg :: "(('x \\<^sub>0 'a::add_linorder) \\<^sub>0 'b::zero) set \ 'a" where "mindeg A = Min (poly_deg ` A)" lemma poly_deg_monomial: "poly_deg (monomial c t) = (if c = 0 then 0 else deg_pm t)" by (simp add: poly_deg_def) lemma poly_deg_monomial_zero [simp]: "poly_deg (monomial c 0) = 0" by (simp add: poly_deg_monomial) lemma poly_deg_zero [simp]: "poly_deg 0 = 0" by (simp only: single_zero[of 0, symmetric] poly_deg_monomial_zero) lemma poly_deg_one [simp]: "poly_deg 1 = 0" by (simp only: single_one[symmetric] poly_deg_monomial_zero) lemma poly_degE: assumes "p \ 0" obtains t where "t \ keys p" and "poly_deg p = deg_pm t" proof - from assms have "poly_deg p = Max (deg_pm ` keys p)" by (simp add: poly_deg_def) also have "\ \ deg_pm ` keys p" proof (rule Max_in) from assms show "deg_pm ` keys p \ {}" by simp qed simp finally obtain t where "t \ keys p" and "poly_deg p = deg_pm t" .. thus ?thesis .. qed lemma poly_deg_max_keys: "t \ keys p \ deg_pm t \ poly_deg p" using finite_keys by (auto simp: poly_deg_def) lemma poly_deg_leI: "(\t. t \ keys p \ deg_pm t \ (d::'a::add_linorder_min)) \ poly_deg p \ d" using finite_keys by (auto simp: poly_deg_def) lemma poly_deg_lessI: "p \ 0 \ (\t. t \ keys p \ deg_pm t < (d::'a::add_linorder_min)) \ poly_deg p < d" using finite_keys by (auto simp: poly_deg_def) lemma poly_deg_zero_imp_monomial: assumes "poly_deg p = (0::'a::add_linorder_min)" shows "monomial (lookup p 0) 0 = p" proof (rule keys_subset_singleton_imp_monomial, rule) fix t assume "t \ keys p" have "t = 0" proof (rule ccontr) assume "t \ 0" hence "deg_pm t \ 0" by simp hence "0 < deg_pm t" using not_gr_zero by blast also from \t \ keys p\ have "... \ poly_deg p" by (rule poly_deg_max_keys) finally have "poly_deg p \ 0" by simp from this assms show False .. qed thus "t \ {0}" by simp qed lemma poly_deg_plus_le: "poly_deg (p + q) \ max (poly_deg p) (poly_deg (q::(_ \\<^sub>0 'a::add_linorder_min) \\<^sub>0 _))" proof (rule poly_deg_leI) fix t assume "t \ keys (p + q)" also have "... \ keys p \ keys q" by (fact Poly_Mapping.keys_add) finally show "deg_pm t \ max (poly_deg p) (poly_deg q)" proof assume "t \ keys p" hence "deg_pm t \ poly_deg p" by (rule poly_deg_max_keys) thus ?thesis by (simp add: le_max_iff_disj) next assume "t \ keys q" hence "deg_pm t \ poly_deg q" by (rule poly_deg_max_keys) thus ?thesis by (simp add: le_max_iff_disj) qed qed lemma poly_deg_uminus [simp]: "poly_deg (-p) = poly_deg p" by (simp add: poly_deg_def keys_uminus) lemma poly_deg_minus_le: "poly_deg (p - q) \ max (poly_deg p) (poly_deg (q::(_ \\<^sub>0 'a::add_linorder_min) \\<^sub>0 _))" proof (rule poly_deg_leI) fix t assume "t \ keys (p - q)" also have "... \ keys p \ keys q" by (fact keys_minus) finally show "deg_pm t \ max (poly_deg p) (poly_deg q)" proof assume "t \ keys p" hence "deg_pm t \ poly_deg p" by (rule poly_deg_max_keys) thus ?thesis by (simp add: le_max_iff_disj) next assume "t \ keys q" hence "deg_pm t \ poly_deg q" by (rule poly_deg_max_keys) thus ?thesis by (simp add: le_max_iff_disj) qed qed lemma poly_deg_times_le: "poly_deg (p * q) \ poly_deg p + poly_deg (q::(_ \\<^sub>0 'a::add_linorder_min) \\<^sub>0 _)" proof (rule poly_deg_leI) fix t assume "t \ keys (p * q)" then obtain u v where "u \ keys p" and "v \ keys q" and "t = u + v" by (rule in_keys_timesE) from \u \ keys p\ have "deg_pm u \ poly_deg p" by (rule poly_deg_max_keys) moreover from \v \ keys q\ have "deg_pm v \ poly_deg q" by (rule poly_deg_max_keys) ultimately show "deg_pm t \ poly_deg p + poly_deg q" by (simp add: \t = u + v\ deg_pm_plus add_mono) qed lemma poly_deg_times: assumes "p \ 0" and "q \ (0::('x::linorder \\<^sub>0 'a::add_linorder_min) \\<^sub>0 'b::semiring_no_zero_divisors)" shows "poly_deg (p * q) = poly_deg p + poly_deg q" using poly_deg_times_le proof (rule antisym) let ?A = "\f. {u. deg_pm u < poly_deg f}" define p1 where "p1 = except p (?A p)" define p2 where "p2 = except p (- ?A p)" define q1 where "q1 = except q (?A q)" define q2 where "q2 = except q (- ?A q)" have deg_p1: "deg_pm t = poly_deg p" if "t \ keys p1" for t proof - from that have "t \ keys p" and "poly_deg p \ deg_pm t" by (simp_all add: p1_def keys_except not_less) from this(1) have "deg_pm t \ poly_deg p" by (rule poly_deg_max_keys) thus ?thesis using \poly_deg p \ deg_pm t\ by (rule antisym) qed have deg_p2: "t \ keys p2 \ deg_pm t < poly_deg p" for t by (simp add: p2_def keys_except) have deg_q1: "deg_pm t = poly_deg q" if "t \ keys q1" for t proof - from that have "t \ keys q" and "poly_deg q \ deg_pm t" by (simp_all add: q1_def keys_except not_less) from this(1) have "deg_pm t \ poly_deg q" by (rule poly_deg_max_keys) thus ?thesis using \poly_deg q \ deg_pm t\ by (rule antisym) qed have deg_q2: "t \ keys q2 \ deg_pm t < poly_deg q" for t by (simp add: q2_def keys_except) have p: "p = p1 + p2" unfolding p1_def p2_def by (fact except_decomp) have "p1 \ 0" proof - from assms(1) obtain t where "t \ keys p" and "poly_deg p = deg_pm t" by (rule poly_degE) hence "t \ keys p1" by (simp add: p1_def keys_except) thus ?thesis by auto qed have q: "q = q1 + q2" unfolding q1_def q2_def by (fact except_decomp) have "q1 \ 0" proof - from assms(2) obtain t where "t \ keys q" and "poly_deg q = deg_pm t" by (rule poly_degE) hence "t \ keys q1" by (simp add: q1_def keys_except) thus ?thesis by auto qed with \p1 \ 0\ have "p1 * q1 \ 0" by simp hence "keys (p1 * q1) \ {}" by simp then obtain u where "u \ keys (p1 * q1)" by blast then obtain s t where "s \ keys p1" and "t \ keys q1" and u: "u = s + t" by (rule in_keys_timesE) from \s \ keys p1\ have "deg_pm s = poly_deg p" by (rule deg_p1) moreover from \t \ keys q1\ have "deg_pm t = poly_deg q" by (rule deg_q1) ultimately have eq: "poly_deg p + poly_deg q = deg_pm u" by (simp only: u deg_pm_plus) also have "\ \ poly_deg (p * q)" proof (rule poly_deg_max_keys) have "u \ keys (p1 * q2 + p2 * q)" proof assume "u \ keys (p1 * q2 + p2 * q)" also have "\ \ keys (p1 * q2) \ keys (p2 * q)" by (rule Poly_Mapping.keys_add) finally have "deg_pm u < poly_deg p + poly_deg q" proof assume "u \ keys (p1 * q2)" then obtain s' t' where "s' \ keys p1" and "t' \ keys q2" and u: "u = s' + t'" by (rule in_keys_timesE) from \s' \ keys p1\ have "deg_pm s' = poly_deg p" by (rule deg_p1) moreover from \t' \ keys q2\ have "deg_pm t' < poly_deg q" by (rule deg_q2) ultimately show ?thesis by (simp add: u deg_pm_plus) next assume "u \ keys (p2 * q)" then obtain s' t' where "s' \ keys p2" and "t' \ keys q" and u: "u = s' + t'" by (rule in_keys_timesE) from \s' \ keys p2\ have "deg_pm s' < poly_deg p" by (rule deg_p2) moreover from \t' \ keys q\ have "deg_pm t' \ poly_deg q" by (rule poly_deg_max_keys) ultimately show ?thesis by (simp add: u deg_pm_plus add_less_le_mono) qed thus False by (simp only: eq) qed with \u \ keys (p1 * q1)\ have "u \ keys (p1 * q1 + (p1 * q2 + p2 * q))" by (rule in_keys_plusI1) thus "u \ keys (p * q)" by (simp only: p q algebra_simps) qed finally show "poly_deg p + poly_deg q \ poly_deg (p * q)" . qed corollary poly_deg_monom_mult_le: "poly_deg (punit.monom_mult c (t::_ \\<^sub>0 'a::add_linorder_min) p) \ deg_pm t + poly_deg p" proof - have "poly_deg (punit.monom_mult c t p) \ poly_deg (monomial c t) + poly_deg p" by (simp only: times_monomial_left[symmetric] poly_deg_times_le) also have "... \ deg_pm t + poly_deg p" by (simp add: poly_deg_monomial) finally show ?thesis . qed lemma poly_deg_monom_mult: assumes "c \ 0" and "p \ (0::(_ \\<^sub>0 'a::add_linorder_min) \\<^sub>0 'b::semiring_no_zero_divisors)" shows "poly_deg (punit.monom_mult c t p) = deg_pm t + poly_deg p" proof (rule order.antisym, fact poly_deg_monom_mult_le) from assms(2) obtain s where "s \ keys p" and "poly_deg p = deg_pm s" by (rule poly_degE) have "deg_pm t + poly_deg p = deg_pm (t + s)" by (simp add: \poly_deg p = deg_pm s\ deg_pm_plus) also have "... \ poly_deg (punit.monom_mult c t p)" proof (rule poly_deg_max_keys) from \s \ keys p\ show "t + s \ keys (punit.monom_mult c t p)" unfolding punit.keys_monom_mult[OF assms(1)] by fastforce qed finally show "deg_pm t + poly_deg p \ poly_deg (punit.monom_mult c t p)" . qed lemma poly_deg_map_scale: "poly_deg (c \ p) = (if c = (0::_::semiring_no_zero_divisors) then 0 else poly_deg p)" by (simp add: poly_deg_def keys_map_scale) lemma poly_deg_sum_le: "((poly_deg (sum f A))::'a::add_linorder_min) \ Max (poly_deg ` f ` A)" proof (cases "finite A") case True thus ?thesis proof (induct A) case empty show ?case by simp next case (insert a A) show ?case proof (cases "A = {}") case True thus ?thesis by simp next case False have "poly_deg (sum f (insert a A)) \ max (poly_deg (f a)) (poly_deg (sum f A))" by (simp only: comm_monoid_add_class.sum.insert[OF insert(1) insert(2)] poly_deg_plus_le) also have "... \ max (poly_deg (f a)) (Max (poly_deg ` f ` A))" using insert(3) max.mono by blast also have "... = (Max (poly_deg ` f ` (insert a A)))" using False by (simp add: insert(1)) finally show ?thesis . qed qed next case False thus ?thesis by simp qed lemma poly_deg_prod_le: "((poly_deg (prod f A))::'a::add_linorder_min) \ (\a\A. poly_deg (f a))" proof (cases "finite A") case True thus ?thesis proof (induct A) case empty show ?case by simp next case (insert a A) have "poly_deg (prod f (insert a A)) \ (poly_deg (f a)) + (poly_deg (prod f A))" by (simp only: comm_monoid_mult_class.prod.insert[OF insert(1) insert(2)] poly_deg_times_le) also have "... \ (poly_deg (f a)) + (\a\A. poly_deg (f a))" using insert(3) add_le_cancel_left by blast also have "... = (\a\insert a A. poly_deg (f a))" by (simp add: insert(1) insert(2)) finally show ?case . qed next case False thus ?thesis by simp qed lemma maxdeg_max: assumes "finite A" and "p \ A" shows "poly_deg p \ maxdeg A" unfolding maxdeg_def using assms by auto lemma mindeg_min: assumes "finite A" and "p \ A" shows "mindeg A \ poly_deg p" unfolding mindeg_def using assms by auto subsection \Indeterminates\ definition indets :: "(('x \\<^sub>0 nat) \\<^sub>0 'b::zero) \ 'x set" - where "indets p = UNION (keys p) keys" + where "indets p = \ (keys ` keys p)" definition PPs :: "'x set \ ('x \\<^sub>0 nat) set" (".[(_)]") where "PPs X = {t. keys t \ X}" definition Polys :: "'x set \ (('x \\<^sub>0 nat) \\<^sub>0 'b::zero) set" ("P[(_)]") where "Polys X = {p. keys p \ .[X]}" subsubsection \@{const indets}\ lemma in_indetsI: assumes "x \ keys t" and "t \ keys p" shows "x \ indets p" using assms by (auto simp add: indets_def) lemma in_indetsE: assumes "x \ indets p" obtains t where "t \ keys p" and "x \ keys t" using assms by (auto simp add: indets_def) lemma keys_subset_indets: "t \ keys p \ keys t \ indets p" by (auto dest: in_indetsI) lemma indets_empty_imp_monomial: assumes "indets p = {}" shows "monomial (lookup p 0) 0 = p" proof (rule keys_subset_singleton_imp_monomial, rule) fix t assume "t \ keys p" have "t = 0" proof (rule ccontr) assume "t \ 0" hence "keys t \ {}" by simp then obtain x where "x \ keys t" by blast from this \t \ keys p\ have "x \ indets p" by (rule in_indetsI) with assms show False by simp qed thus "t \ {0}" by simp qed lemma finite_indets: "finite (indets p)" by (simp only: indets_def, rule finite_UN_I, (rule finite_keys)+) lemma indets_zero [simp]: "indets 0 = {}" by (simp add: indets_def) lemma indets_one [simp]: "indets 1 = {}" by (simp add: indets_def) lemma indets_monomial_single_subset: "indets (monomial c (Poly_Mapping.single v k)) \ {v}" proof fix x assume "x \ indets (monomial c (Poly_Mapping.single v k))" then have "x = v" unfolding indets_def by (metis UN_E lookup_eq_zero_in_keys_contradict lookup_single_not_eq) thus "x \ {v}" by simp qed lemma indets_monomial_single: assumes "c \ 0" and "k \ 0" shows "indets (monomial c (Poly_Mapping.single v k)) = {v}" proof (rule, fact indets_monomial_single_subset, simp) from assms show "v \ indets (monomial c (monomial k v))" by (simp add: indets_def) qed lemma indets_monomial: assumes "c \ 0" shows "indets (monomial c t) = keys t" proof (rule antisym; rule subsetI) fix x assume "x \ indets (monomial c t)" then have "lookup t x \ 0" unfolding indets_def by (metis UN_E lookup_eq_zero_in_keys_contradict lookup_single_not_eq) thus "x \ keys t" by (meson lookup_not_eq_zero_eq_in_keys) next fix x assume "x \ keys t" then have "lookup t x \ 0" by (meson lookup_not_eq_zero_eq_in_keys) thus "x \ indets (monomial c t)" unfolding indets_def using assms by (metis UN_iff lookup_not_eq_zero_eq_in_keys lookup_single_eq) qed lemma indets_monomial_subset: "indets (monomial c t) \ keys t" by (cases "c = 0", simp_all add: indets_def) lemma indets_monomial_zero [simp]: "indets (monomial c 0) = {}" by (simp add: indets_def) lemma indets_plus_subset: "indets (p + q) \ indets p \ indets q" proof fix x assume "x \ indets (p + q)" then obtain t where "x \ keys t" and "t \ keys (p + q)" by (metis UN_E indets_def) hence "t \ keys p \ keys q" by (metis Poly_Mapping.keys_add subsetCE) thus "x \ indets p \ indets q" using indets_def \x \ keys t\ by fastforce qed lemma indets_uminus [simp]: "indets (-p) = indets p" by (simp add: indets_def keys_uminus) lemma indets_minus_subset: "indets (p - q) \ indets p \ indets q" proof fix x assume "x \ indets (p - q)" then obtain t where "x \ keys t" and "t \ keys (p - q)" by (metis UN_E indets_def) hence "t \ keys p \ keys q" by (metis keys_minus subsetCE) thus "x \ indets p \ indets q" using indets_def \x \ keys t\ by fastforce qed lemma indets_times_subset: "indets (p * q) \ indets p \ indets (q::(_ \\<^sub>0 _::cancel_comm_monoid_add) \\<^sub>0 _)" proof fix x assume "x \ indets (p * q)" then obtain t where "t \ keys (p * q)" and "x \ keys t" unfolding indets_def by blast from this(1) obtain u v where "u \ keys p" "v \ keys q" and "t = u + v" by (rule in_keys_timesE) hence "x \ keys u \ keys v" by (metis \x \ keys t\ Poly_Mapping.keys_add subsetCE) thus "x \ indets p \ indets q" unfolding indets_def using \u \ keys p\ \v \ keys q\ by blast qed corollary indets_monom_mult_subset: "indets (punit.monom_mult c t p) \ keys t \ indets p" proof - have "indets (punit.monom_mult c t p) \ indets (monomial c t) \ indets p" by (simp only: times_monomial_left[symmetric] indets_times_subset) also have "... \ keys t \ indets p" using indets_monomial_subset[of t c] by blast finally show ?thesis . qed lemma indets_monom_mult: assumes "c \ 0" and "p \ (0::('x \\<^sub>0 nat) \\<^sub>0 'b::semiring_no_zero_divisors)" shows "indets (punit.monom_mult c t p) = keys t \ indets p" proof (rule, fact indets_monom_mult_subset, rule) fix x assume "x \ keys t \ indets p" thus "x \ indets (punit.monom_mult c t p)" proof assume "x \ keys t" from assms(2) have "keys p \ {}" by simp then obtain s where "s \ keys p" by blast hence "t + s \ (+) t ` keys p" by fastforce also from assms(1) have "... = keys (punit.monom_mult c t p)" by (simp add: punit.keys_monom_mult) finally have "t + s \ keys (punit.monom_mult c t p)" . show ?thesis proof (rule in_indetsI) from \x \ keys t\ show "x \ keys (t + s)" by (simp add: keys_plus_ninv_comm_monoid_add) qed fact next assume "x \ indets p" then obtain s where "s \ keys p" and "x \ keys s" by (rule in_indetsE) from this(1) have "t + s \ (+) t ` keys p" by fastforce also from assms(1) have "... = keys (punit.monom_mult c t p)" by (simp add: punit.keys_monom_mult) finally have "t + s \ keys (punit.monom_mult c t p)" . show ?thesis proof (rule in_indetsI) from \x \ keys s\ show "x \ keys (t + s)" by (simp add: keys_plus_ninv_comm_monoid_add) qed fact qed qed lemma indets_sum_subset: "indets (sum f A) \ (\a\A. indets (f a))" proof (cases "finite A") case True thus ?thesis proof (induct A) case empty show ?case by simp next case (insert a A) have "indets (sum f (insert a A)) \ indets (f a) \ indets (sum f A)" by (simp only: comm_monoid_add_class.sum.insert[OF insert(1) insert(2)] indets_plus_subset) also have "... \ indets (f a) \ (\a\A. indets (f a))" using insert(3) by blast also have "... = (\a\insert a A. indets (f a))" by simp finally show ?case . qed next case False thus ?thesis by simp qed lemma indets_prod_subset: "indets (prod (f::_ \ ((_ \\<^sub>0 _::cancel_comm_monoid_add) \\<^sub>0 _)) A) \ (\a\A. indets (f a))" proof (cases "finite A") case True thus ?thesis proof (induct A) case empty show ?case by simp next case (insert a A) have "indets (prod f (insert a A)) \ indets (f a) \ indets (prod f A)" by (simp only: comm_monoid_mult_class.prod.insert[OF insert(1) insert(2)] indets_times_subset) also have "... \ indets (f a) \ (\a\A. indets (f a))" using insert(3) by blast also have "... = (\a\insert a A. indets (f a))" by simp finally show ?case . qed next case False thus ?thesis by simp qed lemma indets_power_subset: "indets (p ^ n) \ indets (p::('x \\<^sub>0 nat) \\<^sub>0 'b::comm_semiring_1)" proof - have "p ^ n = (\i=0.. (\i\{0.. indets p" by simp finally show ?thesis . qed lemma indets_empty_iff_poly_deg_zero: "indets p = {} \ poly_deg p = 0" proof assume "indets p = {}" hence "monomial (lookup p 0) 0 = p" by (rule indets_empty_imp_monomial) moreover have "poly_deg (monomial (lookup p 0) 0) = 0" by simp ultimately show "poly_deg p = 0" by metis next assume "poly_deg p = 0" hence "monomial (lookup p 0) 0 = p" by (rule poly_deg_zero_imp_monomial) moreover have "indets (monomial (lookup p 0) 0) = {}" by simp ultimately show "indets p = {}" by metis qed subsubsection \@{const PPs}\ lemma PPsI: "keys t \ X \ t \ .[X]" by (simp add: PPs_def) lemma PPsD: "t \ .[X] \ keys t \ X" by (simp add: PPs_def) lemma PPs_empty [simp]: ".[{}] = {0}" by (simp add: PPs_def) lemma PPs_UNIV [simp]: ".[UNIV] = UNIV" by (simp add: PPs_def) lemma PPs_singleton: ".[{x}] = range (Poly_Mapping.single x)" proof (rule set_eqI) fix t show "t \ .[{x}] \ t \ range (Poly_Mapping.single x)" proof assume "t \ .[{x}]" hence "keys t \ {x}" by (rule PPsD) hence "Poly_Mapping.single x (lookup t x) = t" by (rule keys_subset_singleton_imp_monomial) from this[symmetric] UNIV_I show "t \ range (Poly_Mapping.single x)" .. next assume "t \ range (Poly_Mapping.single x)" then obtain e where "t = Poly_Mapping.single x e" .. thus "t \ .[{x}]" by (simp add: PPs_def) qed qed lemma zero_in_PPs: "0 \ .[X]" by (simp add: PPs_def) lemma PPs_mono: "X \ Y \ .[X] \ .[Y]" by (auto simp: PPs_def) lemma PPs_closed_single: assumes "x \ X" shows "Poly_Mapping.single x e \ .[X]" proof (rule PPsI) have "keys (Poly_Mapping.single x e) \ {x}" by simp also from assms have "... \ X" by simp finally show "keys (Poly_Mapping.single x e) \ X" . qed lemma PPs_closed_plus: assumes "s \ .[X]" and "t \ .[X]" shows "s + t \ .[X]" proof - have "keys (s + t) \ keys s \ keys t" by (fact Poly_Mapping.keys_add) also from assms have "... \ X" by (simp add: PPs_def) finally show ?thesis by (rule PPsI) qed lemma PPs_closed_minus: assumes "s \ .[X]" shows "s - t \ .[X]" proof - have "keys (s - t) \ keys s" by (metis lookup_minus lookup_not_eq_zero_eq_in_keys subsetI zero_diff) also from assms have "... \ X" by (rule PPsD) finally show ?thesis by (rule PPsI) qed lemma PPs_closed_adds: assumes "s \ .[X]" and "t adds s" shows "t \ .[X]" proof - from assms(2) have "s - (s - t) = t" by (metis add_minus_2 adds_minus) moreover from assms(1) have "s - (s - t) \ .[X]" by (rule PPs_closed_minus) ultimately show ?thesis by simp qed lemma PPs_closed_gcs: assumes "s \ .[X]" shows "gcs s t \ .[X]" using assms gcs_adds by (rule PPs_closed_adds) lemma PPs_closed_lcs: assumes "s \ .[X]" and "t \ .[X]" shows "lcs s t \ .[X]" proof - from assms have "s + t \ .[X]" by (rule PPs_closed_plus) hence "(s + t) - gcs s t \ .[X]" by (rule PPs_closed_minus) thus ?thesis by (simp add: gcs_plus_lcs[of s t, symmetric]) qed lemma PPs_closed_except': "t \ .[X] \ except t Y \ .[X - Y]" by (auto simp: keys_except PPs_def) lemma PPs_closed_except: "t \ .[X] \ except t Y \ .[X]" by (auto simp: keys_except PPs_def) lemma PPs_UnI: assumes "tx \ .[X]" and "ty \ .[Y]" and "t = tx + ty" shows "t \ .[X \ Y]" proof - from assms(1) have "tx \ .[X \ Y]" by rule (simp add: PPs_mono) moreover from assms(2) have "ty \ .[X \ Y]" by rule (simp add: PPs_mono) ultimately show ?thesis unfolding assms(3) by (rule PPs_closed_plus) qed lemma PPs_UnE: assumes "t \ .[X \ Y]" obtains tx ty where "tx \ .[X]" and "ty \ .[Y]" and "t = tx + ty" proof - from assms have "keys t \ X \ Y" by (rule PPsD) define tx where "tx = except t (- X)" have "keys tx \ X" by (simp add: tx_def keys_except) hence "tx \ .[X]" by (simp add: PPs_def) have "tx adds t" by (simp add: tx_def adds_poly_mappingI le_fun_def lookup_except) from adds_minus[OF this] have "t = tx + (t - tx)" by (simp only: ac_simps) have "t - tx \ .[Y]" proof (rule PPsI, rule) fix x assume "x \ keys (t - tx)" also have "... \ keys t \ keys tx" by (rule keys_minus) also from \keys t \ X \ Y\ \keys tx \ X\ have "... \ X \ Y" by blast finally show "x \ Y" proof assume "x \ X" hence "x \ keys (t - tx)" by (simp add: tx_def lookup_except lookup_minus in_keys_iff) thus ?thesis using \x \ keys (t - tx)\ .. qed qed with \tx \ .[X]\ show ?thesis using \t = tx + (t - tx)\ .. qed lemma PPs_Un: ".[X \ Y] = (\t\.[X]. (+) t ` .[Y])" (is "?A = ?B") proof (rule set_eqI) fix t show "t \ ?A \ t \ ?B" proof assume "t \ ?A" then obtain tx ty where "tx \ .[X]" and "ty \ .[Y]" and "t = tx + ty" by (rule PPs_UnE) from this(2) have "t \ (+) tx ` .[Y]" unfolding \t = tx + ty\ by (rule imageI) with \tx \ .[X]\ show "t \ ?B" .. next assume "t \ ?B" then obtain tx where "tx \ .[X]" and "t \ (+) tx ` .[Y]" .. from this(2) obtain ty where "ty \ .[Y]" and "t = tx + ty" .. with \tx \ .[X]\ show "t \ ?A" by (rule PPs_UnI) qed qed corollary PPs_insert: ".[insert x X] = (\e. (+) (Poly_Mapping.single x e) ` .[X])" proof - have ".[insert x X] = .[{x} \ X]" by simp also have "... = (\t\.[{x}]. (+) t ` .[X])" by (fact PPs_Un) also have "... = (\e. (+) (Poly_Mapping.single x e) ` .[X])" by (simp add: PPs_singleton) finally show ?thesis . qed corollary PPs_insertI: assumes "tx \ .[X]" and "t = Poly_Mapping.single x e + tx" shows "t \ .[insert x X]" proof - from assms(1) have "t \ (+) (Poly_Mapping.single x e) ` .[X]" unfolding assms(2) by (rule imageI) with UNIV_I show ?thesis unfolding PPs_insert by (rule UN_I) qed corollary PPs_insertE: assumes "t \ .[insert x X]" obtains e tx where "tx \ .[X]" and "t = Poly_Mapping.single x e + tx" proof - from assms obtain e where "t \ (+) (Poly_Mapping.single x e) ` .[X]" unfolding PPs_insert .. then obtain tx where "tx \ .[X]" and "t = Poly_Mapping.single x e + tx" .. thus ?thesis .. qed lemma PPs_Int: ".[X \ Y] = .[X] \ .[Y]" by (auto simp: PPs_def) lemma PPs_INT: ".[\ X] = \ (PPs ` X)" by (auto simp: PPs_def) subsubsection \@{const Polys}\ lemma Polys_alt: "P[X] = {p. indets p \ X}" by (auto simp: Polys_def PPs_def indets_def) lemma PolysI: "keys p \ .[X] \ p \ P[X]" by (simp add: Polys_def) lemma PolysI_alt: "indets p \ X \ p \ P[X]" by (simp add: Polys_alt) lemma PolysD: assumes "p \ P[X]" shows "keys p \ .[X]" and "indets p \ X" using assms by (simp add: Polys_def, simp add: Polys_alt) lemma Polys_empty: "P[{}] = ((range (Poly_Mapping.single 0))::(('x \\<^sub>0 nat) \\<^sub>0 'b::zero) set)" proof (rule set_eqI) fix p :: "('x \\<^sub>0 nat) \\<^sub>0 'b::zero" show "p \ P[{}] \ p \ range (Poly_Mapping.single 0)" proof assume "p \ P[{}]" hence "keys p \ .[{}]" by (rule PolysD) also have "... = {0}" by simp finally have "keys p \ {0}" . hence "Poly_Mapping.single 0 (lookup p 0) = p" by (rule keys_subset_singleton_imp_monomial) from this[symmetric] UNIV_I show "p \ range (Poly_Mapping.single 0)" .. next assume "p \ range (Poly_Mapping.single 0)" then obtain c where "p = monomial c 0" .. thus "p \ P[{}]" by (simp add: Polys_def) qed qed lemma Polys_UNIV [simp]: "P[UNIV] = UNIV" by (simp add: Polys_def) lemma zero_in_Polys: "0 \ P[X]" by (simp add: Polys_def) lemma one_in_Polys: "1 \ P[X]" by (simp add: Polys_def zero_in_PPs) lemma Polys_mono: "X \ Y \ P[X] \ P[Y]" by (auto simp: Polys_alt) lemma Polys_closed_monomial: "t \ .[X] \ monomial c t \ P[X]" using indets_monomial_subset[where c=c and t=t] by (auto simp: Polys_alt PPs_def) lemma Polys_closed_plus: "p \ P[X] \ q \ P[X] \ p + q \ P[X]" using indets_plus_subset[of p q] by (auto simp: Polys_alt PPs_def) lemma Polys_closed_uminus: "p \ P[X] \ -p \ P[X]" by (simp add: Polys_def keys_uminus) lemma Polys_closed_minus: "p \ P[X] \ q \ P[X] \ p - q \ P[X]" using indets_minus_subset[of p q] by (auto simp: Polys_alt PPs_def) lemma Polys_closed_monom_mult: "t \ .[X] \ p \ P[X] \ punit.monom_mult c t p \ P[X]" using indets_monom_mult_subset[of c t p] by (auto simp: Polys_alt PPs_def) corollary Polys_closed_map_scale: "p \ P[X] \ (c::_::semiring_0) \ p \ P[X]" unfolding punit.map_scale_eq_monom_mult using zero_in_PPs by (rule Polys_closed_monom_mult) lemma Polys_closed_times: "p \ P[X] \ q \ P[X] \ p * q \ P[X]" using indets_times_subset[of p q] by (auto simp: Polys_alt PPs_def) lemma Polys_closed_power: "p \ P[X] \ p ^ m \ P[X]" by (induct m) (auto intro: one_in_Polys Polys_closed_times) lemma Polys_closed_sum: "(\a. a \ A \ f a \ P[X]) \ sum f A \ P[X]" by (induct A rule: infinite_finite_induct) (auto intro: zero_in_Polys Polys_closed_plus) lemma Polys_closed_prod: "(\a. a \ A \ f a \ P[X]) \ prod f A \ P[X]" by (induct A rule: infinite_finite_induct) (auto intro: one_in_Polys Polys_closed_times) lemma Polys_closed_sum_list: "(\x. x \ set xs \ x \ P[X]) \ sum_list xs \ P[X]" by (induct xs) (auto intro: zero_in_Polys Polys_closed_plus) lemma Polys_closed_except: "p \ P[X] \ except p T \ P[X]" by (auto intro!: PolysI simp: keys_except dest!: PolysD(1)) lemma times_in_PolysD: assumes "p * q \ P[X]" and "p \ P[X]" and "p \ (0::('x::linorder \\<^sub>0 nat) \\<^sub>0 'a::semiring_no_zero_divisors)" shows "q \ P[X]" proof - define qX where "qX = except q (- .[X])" define qY where "qY = except q .[X]" have q: "q = qX + qY" by (simp only: qX_def qY_def add.commute flip: except_decomp) have "qX \ P[X]" by (rule PolysI) (simp add: qX_def keys_except) with assms(2) have "p * qX \ P[X]" by (rule Polys_closed_times) show ?thesis proof (cases "qY = 0") case True with \qX \ P[X]\ show ?thesis by (simp add: q) next case False with assms(3) have "p * qY \ 0" by simp hence "keys (p * qY) \ {}" by simp then obtain t where "t \ keys (p * qY)" by blast then obtain t1 t2 where "t2 \ keys qY" and t: "t = t1 + t2" by (rule in_keys_timesE) have "t \ .[X]" unfolding t proof assume "t1 + t2 \ .[X]" hence "t1 + t2 - t1 \ .[X]" by (rule PPs_closed_minus) hence "t2 \ .[X]" by simp with \t2 \ keys qY\ show False by (simp add: qY_def keys_except) qed have "t \ keys (p * qX)" proof assume "t \ keys (p * qX)" also from \p * qX \ P[X]\ have "\ \ .[X]" by (rule PolysD) finally have "t \ .[X]" . with \t \ .[X]\ show False .. qed with \t \ keys (p * qY)\ have "t \ keys (p * qX + p * qY)" by (rule in_keys_plusI2) also have "\ = keys (p * q)" by (simp only: q algebra_simps) finally have "p * q \ P[X]" using \t \ .[X]\ by (auto simp: Polys_def) thus ?thesis using assms(1) .. qed qed lemma poly_mapping_plus_induct_Polys [consumes 1, case_names 0 plus]: assumes "p \ P[X]" and "P 0" and "\p c t. t \ .[X] \ p \ P[X] \ c \ 0 \ t \ keys p \ P p \ P (monomial c t + p)" shows "P p" using assms(1) proof (induct p rule: poly_mapping_plus_induct) case 1 show ?case by (fact assms(2)) next case step: (2 p c t) from step.hyps(1) have 1: "keys (monomial c t) = {t}" by simp also from step.hyps(2) have "\ \ keys p = {}" by simp finally have "keys (monomial c t + p) = keys (monomial c t) \ keys p" by (rule keys_add[symmetric]) hence "keys (monomial c t + p) = insert t (keys p)" by (simp only: 1 flip: insert_is_Un) moreover from step.prems(1) have "keys (monomial c t + p) \ .[X]" by (rule PolysD) ultimately have "t \ .[X]" and "keys p \ .[X]" by blast+ from this(2) have "p \ P[X]" by (rule PolysI) hence "P p" by (rule step.hyps) with \t \ .[X]\ \p \ P[X]\ step.hyps(1, 2) show ?case by (rule assms(3)) qed lemma Polys_Int: "P[X \ Y] = P[X] \ P[Y]" by (auto simp: Polys_def PPs_Int) lemma Polys_INT: "P[\ X] = \ (Polys ` X)" by (auto simp: Polys_def PPs_INT) subsection \Substitution Homomorphism\ text \The substitution homomorphism defined here is more general than @{const insertion}, since it replaces indeterminates by @{emph \polynomials\} rather than coefficients, and therefore constructs new polynomials.\ definition subst_pp :: "('x \ (('y \\<^sub>0 nat) \\<^sub>0 'a)) \ ('x \\<^sub>0 nat) \ (('y \\<^sub>0 nat) \\<^sub>0 'a::comm_semiring_1)" where "subst_pp f t = (\x\keys t. (f x) ^ (lookup t x))" definition poly_subst :: "('x \ (('y \\<^sub>0 nat) \\<^sub>0 'a)) \ (('x \\<^sub>0 nat) \\<^sub>0 'a) \ (('y \\<^sub>0 nat) \\<^sub>0 'a::comm_semiring_1)" where "poly_subst f p = (\t\keys p. punit.monom_mult (lookup p t) 0 (subst_pp f t))" lemma subst_pp_alt: "subst_pp f t = (\x. (f x) ^ (lookup t x))" proof - from finite_keys have "subst_pp f t = (\x. if x \ keys t then (f x) ^ (lookup t x) else 1)" unfolding subst_pp_def by (rule Prod_any.conditionalize) also have "... = (\x. (f x) ^ (lookup t x))" by (rule Prod_any.cong) (simp add: in_keys_iff) finally show ?thesis . qed lemma subst_pp_zero [simp]: "subst_pp f 0 = 1" by (simp add: subst_pp_def) lemma subst_pp_trivial_not_zero: assumes "t \ 0" shows "subst_pp (\_. 0) t = (0::(_ \\<^sub>0 'b::comm_semiring_1))" unfolding subst_pp_def using finite_keys proof (rule prod_zero) from assms have "keys t \ {}" by simp then obtain x where "x \ keys t" by blast thus "\x\keys t. 0 ^ lookup t x = (0::(_ \\<^sub>0 'b))" proof from \x \ keys t\ have "0 < lookup t x" by (simp add: in_keys_iff) thus "0 ^ lookup t x = (0::(_ \\<^sub>0 'b))" by (rule Power.semiring_1_class.zero_power) qed qed lemma subst_pp_single: "subst_pp f (Poly_Mapping.single x e) = (f x) ^ e" by (simp add: subst_pp_def) corollary subst_pp_trivial: "subst_pp (\_. 0) t = (if t = 0 then 1 else 0)" by (simp split: if_split add: subst_pp_trivial_not_zero) lemma power_lookup_not_one_subset_keys: "{x. f x ^ (lookup t x) \ 1} \ keys t" proof (rule, simp) fix x assume "f x ^ (lookup t x) \ 1" thus "x \ keys t" unfolding in_keys_iff by (metis power_0) qed corollary finite_power_lookup_not_one: "finite {x. f x ^ (lookup t x) \ 1}" by (rule finite_subset, fact power_lookup_not_one_subset_keys, fact finite_keys) lemma subst_pp_plus: "subst_pp f (s + t) = subst_pp f s * subst_pp f t" by (simp add: subst_pp_alt lookup_add power_add, rule Prod_any.distrib, (fact finite_power_lookup_not_one)+) lemma subst_pp_id: assumes "\x. x \ keys t \ f x = monomial 1 (Poly_Mapping.single x 1)" shows "subst_pp f t = monomial 1 t" proof - have "subst_pp f t = (\x\keys t. monomial 1 (Poly_Mapping.single x (lookup t x)))" proof (simp only: subst_pp_def, rule prod.cong, fact refl) fix x assume "x \ keys t" thus "f x ^ lookup t x = monomial 1 (Poly_Mapping.single x (lookup t x))" by (simp add: assms monomial_single_power) qed also have "... = monomial 1 t" by (simp add: punit.monomial_prod_sum[symmetric] poly_mapping_sum_monomials) finally show ?thesis . qed lemma in_indets_subst_ppE: assumes "x \ indets (subst_pp f t)" obtains y where "y \ keys t" and "x \ indets (f y)" proof - note assms also have "indets (subst_pp f t) \ (\y\keys t. indets ((f y) ^ (lookup t y)))" unfolding subst_pp_def by (rule indets_prod_subset) finally obtain y where "y \ keys t" and "x \ indets ((f y) ^ (lookup t y))" .. note this(2) also have "indets ((f y) ^ (lookup t y)) \ indets (f y)" by (rule indets_power_subset) finally have "x \ indets (f y)" . with \y \ keys t\ show ?thesis .. qed lemma subst_pp_by_monomials: assumes "\y. y \ keys t \ f y = monomial (c y) (s y)" shows "subst_pp f t = monomial (\y\keys t. (c y) ^ lookup t y) (\y\keys t. lookup t y \ s y)" by (simp add: subst_pp_def assms monomial_power_map_scale punit.monomial_prod_sum) lemma poly_deg_subst_pp_eq_zeroI: assumes "\x. x \ keys t \ poly_deg (f x) = 0" shows "poly_deg (subst_pp f t) = 0" proof - have "poly_deg (subst_pp f t) \ (\x\keys t. poly_deg ((f x) ^ (lookup t x)))" unfolding subst_pp_def by (fact poly_deg_prod_le) also have "... = 0" proof (rule sum.neutral, rule) fix x assume "x \ keys t" hence "poly_deg (f x) = 0" by (rule assms) have "f x ^ lookup t x = (\i=0.. (\i=0..poly_deg (f x) = 0\) finally show "poly_deg (f x ^ lookup t x) = 0" by simp qed finally show ?thesis by simp qed lemma poly_deg_subst_pp_le: assumes "\x. x \ keys t \ poly_deg (f x) \ 1" shows "poly_deg (subst_pp f t) \ deg_pm t" proof - have "poly_deg (subst_pp f t) \ (\x\keys t. poly_deg ((f x) ^ (lookup t x)))" unfolding subst_pp_def by (fact poly_deg_prod_le) also have "... \ (\x\keys t. lookup t x)" proof (rule sum_mono) fix x assume "x \ keys t" hence "poly_deg (f x) \ 1" by (rule assms) have "f x ^ lookup t x = (\i=0.. (\i=0..poly_deg (f x) \ 1\ have "... \ (\i=0.. lookup t x" by simp qed also have "... = deg_pm t" by (rule deg_pm_superset[symmetric], fact subset_refl, fact finite_keys) finally show ?thesis by simp qed lemma poly_subst_alt: "poly_subst f p = (\t. punit.monom_mult (lookup p t) 0 (subst_pp f t))" proof - from finite_keys have "poly_subst f p = (\t. if t \ keys p then punit.monom_mult (lookup p t) 0 (subst_pp f t) else 0)" unfolding poly_subst_def by (rule Sum_any.conditionalize) also have "\ = (\t. punit.monom_mult (lookup p t) 0 (subst_pp f t))" by (rule Sum_any.cong) (simp add: in_keys_iff) finally show ?thesis . qed lemma poly_subst_trivial [simp]: "poly_subst (\_. 0) p = monomial (lookup p 0) 0" by (simp add: poly_subst_def subst_pp_trivial if_distrib in_keys_iff cong: if_cong) (metis mult.right_neutral times_monomial_left) lemma poly_subst_zero [simp]: "poly_subst f 0 = 0" by (simp add: poly_subst_def) lemma monom_mult_lookup_not_zero_subset_keys: "{t. punit.monom_mult (lookup p t) 0 (subst_pp f t) \ 0} \ keys p" proof (rule, simp) fix t assume "punit.monom_mult (lookup p t) 0 (subst_pp f t) \ 0" thus "t \ keys p" unfolding in_keys_iff by (metis punit.monom_mult_zero_left) qed corollary finite_monom_mult_lookup_not_zero: "finite {t. punit.monom_mult (lookup p t) 0 (subst_pp f t) \ 0}" by (rule finite_subset, fact monom_mult_lookup_not_zero_subset_keys, fact finite_keys) lemma poly_subst_plus: "poly_subst f (p + q) = poly_subst f p + poly_subst f q" by (simp add: poly_subst_alt lookup_add punit.monom_mult_dist_left, rule Sum_any.distrib, (fact finite_monom_mult_lookup_not_zero)+) lemma poly_subst_uminus: "poly_subst f (-p) = - poly_subst f (p::('x \\<^sub>0 nat) \\<^sub>0 'b::comm_ring_1)" by (simp add: poly_subst_def keys_uminus punit.monom_mult_uminus_left sum_negf) lemma poly_subst_minus: "poly_subst f (p - q) = poly_subst f p - poly_subst f (q::('x \\<^sub>0 nat) \\<^sub>0 'b::comm_ring_1)" proof - have "poly_subst f (p + (-q)) = poly_subst f p + poly_subst f (-q)" by (fact poly_subst_plus) thus ?thesis by (simp add: poly_subst_uminus) qed lemma poly_subst_monomial: "poly_subst f (monomial c t) = punit.monom_mult c 0 (subst_pp f t)" by (simp add: poly_subst_def lookup_single) corollary poly_subst_one [simp]: "poly_subst f 1 = 1" by (simp add: single_one[symmetric] poly_subst_monomial punit.monom_mult_monomial del: single_one) lemma poly_subst_times: "poly_subst f (p * q) = poly_subst f p * poly_subst f q" proof - have bij: "bij (\(l, n, m). (m, l, n))" by (auto intro!: bijI injI simp add: image_def) let ?P = "keys p" let ?Q = "keys q" let ?PQ = "{s + t | s t. lookup p s \ 0 \ lookup q t \ 0}" have fin_PQ: "finite ?PQ" by (rule finite_not_eq_zero_sumI, simp_all) have fin_1: "finite {l. lookup p l * (\qa. lookup q qa when t = l + qa) \ 0}" for t proof (rule finite_subset) show "{l. lookup p l * (\qa. lookup q qa when t = l + qa) \ 0} \ keys p" by (rule, auto simp: in_keys_iff) qed (fact finite_keys) have fin_2: "finite {v. (lookup q v when t = u + v) \ 0}" for t u proof (rule finite_subset) show "{v. (lookup q v when t = u + v) \ 0} \ keys q" by (rule, auto simp: in_keys_iff) qed (fact finite_keys) have fin_3: "finite {v. (lookup p u * lookup q v when t = u + v) \ 0}" for t u proof (rule finite_subset) show "{v. (lookup p u * lookup q v when t = u + v) \ 0} \ keys q" by (rule, auto simp add: in_keys_iff simp del: lookup_not_eq_zero_eq_in_keys) qed (fact finite_keys) have "(\t. punit.monom_mult (lookup (p * q) t) 0 (subst_pp f t)) = (\t. \u. punit.monom_mult (lookup p u * (\v. lookup q v when t = u + v)) 0 (subst_pp f t))" by (simp add: times_poly_mapping.rep_eq prod_fun_def punit.monom_mult_Sum_any_left[OF fin_1]) also have "\ = (\t. \u. \v. (punit.monom_mult (lookup p u * lookup q v) 0 (subst_pp f t)) when t = u + v)" by (simp add: Sum_any_right_distrib[OF fin_2] punit.monom_mult_Sum_any_left[OF fin_3] mult_when punit.when_monom_mult) also have "\ = (\t. (\(u, v). (punit.monom_mult (lookup p u * lookup q v) 0 (subst_pp f t)) when t = u + v))" by (subst (2) Sum_any.cartesian_product [of "?P \ ?Q"]) (auto simp: in_keys_iff) also have "\ = (\(t, u, v). punit.monom_mult (lookup p u * lookup q v) 0 (subst_pp f t) when t = u + v)" apply (subst Sum_any.cartesian_product [of "?PQ \ (?P \ ?Q)"]) apply (auto simp: fin_PQ in_keys_iff) apply (metis monomial_0I mult_not_zero times_monomial_left) done also have "\ = (\(u, v, t). punit.monom_mult (lookup p u * lookup q v) 0 (subst_pp f t) when t = u + v)" using bij by (rule Sum_any.reindex_cong [of "\(u, v, t). (t, u, v)"]) (simp add: fun_eq_iff) also have "\ = (\(u, v). \t. punit.monom_mult (lookup p u * lookup q v) 0 (subst_pp f t) when t = u + v)" apply (subst Sum_any.cartesian_product2 [of "(?P \ ?Q) \ ?PQ"]) apply (auto simp: fin_PQ in_keys_iff) apply (metis monomial_0I mult_not_zero times_monomial_left) done also have "\ = (\(u, v). punit.monom_mult (lookup p u * lookup q v) 0 (subst_pp f u * subst_pp f v))" by (simp add: subst_pp_plus) also have "\ = (\u. \v. punit.monom_mult (lookup p u * lookup q v) 0 (subst_pp f u * subst_pp f v))" by (subst Sum_any.cartesian_product [of "?P \ ?Q"]) (auto simp: in_keys_iff) also have "\ = (\u. \v. (punit.monom_mult (lookup p u) 0 (subst_pp f u)) * (punit.monom_mult (lookup q v) 0 (subst_pp f v)))" by (simp add: times_monomial_left[symmetric] ac_simps mult_single) also have "\ = (\t. punit.monom_mult (lookup p t) 0 (subst_pp f t)) * (\t. punit.monom_mult (lookup q t) 0 (subst_pp f t))" by (rule Sum_any_product [symmetric], (fact finite_monom_mult_lookup_not_zero)+) finally show ?thesis by (simp add: poly_subst_alt) qed corollary poly_subst_monom_mult: "poly_subst f (punit.monom_mult c t p) = punit.monom_mult c 0 (subst_pp f t * poly_subst f p)" by (simp only: times_monomial_left[symmetric] poly_subst_times poly_subst_monomial mult.assoc) corollary poly_subst_monom_mult': "poly_subst f (punit.monom_mult c t p) = (punit.monom_mult c 0 (subst_pp f t)) * poly_subst f p" by (simp only: times_monomial_left[symmetric] poly_subst_times poly_subst_monomial) lemma poly_subst_sum: "poly_subst f (sum p A) = (\a\A. poly_subst f (p a))" by (rule fun_sum_commute, simp_all add: poly_subst_plus) lemma poly_subst_prod: "poly_subst f (prod p A) = (\a\A. poly_subst f (p a))" by (rule fun_prod_commute, simp_all add: poly_subst_times) lemma poly_subst_power: "poly_subst f (p ^ n) = (poly_subst f p) ^ n" by (induct n, simp_all add: poly_subst_times) lemma poly_subst_subst_pp: "poly_subst f (subst_pp g t) = subst_pp (\x. poly_subst f (g x)) t" by (simp only: subst_pp_def poly_subst_prod poly_subst_power) lemma poly_subst_poly_subst: "poly_subst f (poly_subst g p) = poly_subst (\x. poly_subst f (g x)) p" proof - have "poly_subst f (poly_subst g p) = poly_subst f (\t\keys p. punit.monom_mult (lookup p t) 0 (subst_pp g t))" by (simp only: poly_subst_def) also have "\ = (\t\keys p. punit.monom_mult (lookup p t) 0 (subst_pp (\x. poly_subst f (g x)) t))" by (simp add: poly_subst_sum poly_subst_monom_mult poly_subst_subst_pp) also have "\ = poly_subst (\x. poly_subst f (g x)) p" by (simp only: poly_subst_def) finally show ?thesis . qed lemma poly_subst_id: assumes "\x. x \ indets p \ f x = monomial 1 (Poly_Mapping.single x 1)" shows "poly_subst f p = p" proof - have "poly_subst f p = (\t\keys p. monomial (lookup p t) t)" proof (simp only: poly_subst_def, rule sum.cong, fact refl) fix t assume "t \ keys p" have eq: "subst_pp f t = monomial 1 t" by (rule subst_pp_id, rule assms, erule in_indetsI, fact \t \ keys p\) show "punit.monom_mult (lookup p t) 0 (subst_pp f t) = monomial (lookup p t) t" by (simp add: eq punit.monom_mult_monomial) qed also have "... = p" by (simp only: poly_mapping_sum_monomials) finally show ?thesis . qed lemma in_keys_poly_substE: assumes "t \ keys (poly_subst f p)" obtains s where "s \ keys p" and "t \ keys (subst_pp f s)" proof - note assms also have "keys (poly_subst f p) \ (\t\keys p. keys (punit.monom_mult (lookup p t) 0 (subst_pp f t)))" unfolding poly_subst_def by (rule keys_sum_subset) finally obtain s where "s \ keys p" and "t \ keys (punit.monom_mult (lookup p s) 0 (subst_pp f s))" .. note this(2) also have "\ \ (+) 0 ` keys (subst_pp f s)" by (rule punit.keys_monom_mult_subset[simplified]) also have "\ = keys (subst_pp f s)" by simp finally have "t \ keys (subst_pp f s)" . with \s \ keys p\ show ?thesis .. qed lemma in_indets_poly_substE: assumes "x \ indets (poly_subst f p)" obtains y where "y \ indets p" and "x \ indets (f y)" proof - note assms also have "indets (poly_subst f p) \ (\t\keys p. indets (punit.monom_mult (lookup p t) 0 (subst_pp f t)))" unfolding poly_subst_def by (rule indets_sum_subset) finally obtain t where "t \ keys p" and "x \ indets (punit.monom_mult (lookup p t) 0 (subst_pp f t))" .. note this(2) also have "indets (punit.monom_mult (lookup p t) 0 (subst_pp f t)) \ keys (0::('a \\<^sub>0 nat)) \ indets (subst_pp f t)" by (rule indets_monom_mult_subset) also have "... = indets (subst_pp f t)" by simp finally obtain y where "y \ keys t" and "x \ indets (f y)" by (rule in_indets_subst_ppE) from this(1) \t \ keys p\ have "y \ indets p" by (rule in_indetsI) from this \x \ indets (f y)\ show ?thesis .. qed lemma poly_deg_poly_subst_eq_zeroI: assumes "\x. x \ indets p \ poly_deg (f x) = 0" shows "poly_deg (poly_subst (f::_ \ (('y \\<^sub>0 _) \\<^sub>0 _)) (p::('x \\<^sub>0 _) \\<^sub>0 'b::comm_semiring_1)) = 0" proof (cases "p = 0") case True thus ?thesis by simp next case False have "poly_deg (poly_subst f p) \ Max (poly_deg ` (\t. punit.monom_mult (lookup p t) 0 (subst_pp f t)) ` keys p)" unfolding poly_subst_def by (fact poly_deg_sum_le) also have "... \ 0" proof (rule Max.boundedI) show "finite (poly_deg ` (\t. punit.monom_mult (lookup p t) 0 (subst_pp f t)) ` keys p)" by (simp add: finite_image_iff) next from False show "poly_deg ` (\t. punit.monom_mult (lookup p t) 0 (subst_pp f t)) ` keys p \ {}" by simp next fix d assume "d \ poly_deg ` (\t. punit.monom_mult (lookup p t) 0 (subst_pp f t)) ` keys p" then obtain t where "t \ keys p" and d: "d = poly_deg (punit.monom_mult (lookup p t) 0 (subst_pp f t))" by fastforce have "d \ deg_pm (0::'y \\<^sub>0 nat) + poly_deg (subst_pp f t)" unfolding d by (fact poly_deg_monom_mult_le) also have "... = poly_deg (subst_pp f t)" by simp also have "... = 0" by (rule poly_deg_subst_pp_eq_zeroI, rule assms, erule in_indetsI, fact) finally show "d \ 0" . qed finally show ?thesis by simp qed lemma poly_deg_poly_subst_le: assumes "\x. x \ indets p \ poly_deg (f x) \ 1" shows "poly_deg (poly_subst (f::_ \ (('y \\<^sub>0 _) \\<^sub>0 _)) (p::('x \\<^sub>0 nat) \\<^sub>0 'b::comm_semiring_1)) \ poly_deg p" proof (cases "p = 0") case True thus ?thesis by simp next case False have "poly_deg (poly_subst f p) \ Max (poly_deg ` (\t. punit.monom_mult (lookup p t) 0 (subst_pp f t)) ` keys p)" unfolding poly_subst_def by (fact poly_deg_sum_le) also have "... \ poly_deg p" proof (rule Max.boundedI) show "finite (poly_deg ` (\t. punit.monom_mult (lookup p t) 0 (subst_pp f t)) ` keys p)" by (simp add: finite_image_iff) next from False show "poly_deg ` (\t. punit.monom_mult (lookup p t) 0 (subst_pp f t)) ` keys p \ {}" by simp next fix d assume "d \ poly_deg ` (\t. punit.monom_mult (lookup p t) 0 (subst_pp f t)) ` keys p" then obtain t where "t \ keys p" and d: "d = poly_deg (punit.monom_mult (lookup p t) 0 (subst_pp f t))" by fastforce have "d \ deg_pm (0::'y \\<^sub>0 nat) + poly_deg (subst_pp f t)" unfolding d by (fact poly_deg_monom_mult_le) also have "... = poly_deg (subst_pp f t)" by simp also have "... \ deg_pm t" by (rule poly_deg_subst_pp_le, rule assms, erule in_indetsI, fact) also from \t \ keys p\ have "... \ poly_deg p" by (rule poly_deg_max_keys) finally show "d \ poly_deg p" . qed finally show ?thesis by simp qed lemma subst_pp_cong: "s = t \ (\x. x \ keys t \ f x = g x) \ subst_pp f s = subst_pp g t" by (simp add: subst_pp_def) lemma poly_subst_cong: assumes "p = q" and "\x. x \ indets q \ f x = g x" shows "poly_subst f p = poly_subst g q" proof (simp add: poly_subst_def assms(1), rule sum.cong) fix t assume "t \ keys q" { fix x assume "x \ keys t" with \t \ keys q\ have "x \ indets q" by (auto simp: indets_def) hence "f x = g x" by (rule assms(2)) } thus "punit.monom_mult (lookup q t) 0 (subst_pp f t) = punit.monom_mult (lookup q t) 0 (subst_pp g t)" by (simp cong: subst_pp_cong) qed (fact refl) lemma Polys_homomorphismE: obtains h where "\p q. h (p + q) = h p + h q" and "\p q. h (p * q) = h p * h q" and "\p::('x \\<^sub>0 nat) \\<^sub>0 'a::comm_ring_1. h (h p) = h p" and "range h = P[X]" proof - let ?f = "\x. if x \ X then monomial (1::'a) (Poly_Mapping.single x 1) else 1" have 1: "poly_subst ?f p = p" if "p \ P[X]" for p proof (rule poly_subst_id) fix x assume "x \ indets p" also from that have "\ \ X" by (rule PolysD) finally show "?f x = monomial 1 (Poly_Mapping.single x 1)" by simp qed have 2: "poly_subst ?f p \ P[X]" for p proof (intro PolysI_alt subsetI) fix x assume "x \ indets (poly_subst ?f p)" then obtain y where "x \ indets (?f y)" by (rule in_indets_poly_substE) thus "x \ X" by (simp add: indets_monomial split: if_split_asm) qed from poly_subst_plus poly_subst_times show ?thesis proof fix p from 2 show "poly_subst ?f (poly_subst ?f p) = poly_subst ?f p" by (rule 1) next show "range (poly_subst ?f) = P[X]" proof (intro set_eqI iffI) fix p :: "_ \\<^sub>0 'a" assume "p \ P[X]" hence "p = poly_subst ?f p" by (simp only: 1) thus "p \ range (poly_subst ?f)" by (rule image_eqI) simp qed (auto intro: 2) qed qed lemma in_idealE_Polys_finite: assumes "finite B" and "B \ P[X]" and "p \ P[X]" and "(p::('x \\<^sub>0 nat) \\<^sub>0 'a::comm_ring_1) \ ideal B" obtains q where "\b. q b \ P[X]" and "p = (\b\B. q b * b)" proof - obtain h where "\p q. h (p + q) = h p + h q" and "\p q. h (p * q) = h p * h q" and "\p::('x \\<^sub>0 nat) \\<^sub>0 'a. h (h p) = h p" and rng[symmetric]: "range h = P[X]" by (rule Polys_homomorphismE) blast from this(1-3) assms obtain q where "\b. q b \ P[X]" and "p = (\b\B. q b * b)" unfolding rng by (rule in_idealE_homomorphism_finite) blast thus ?thesis .. qed corollary in_idealE_Polys: assumes "B \ P[X]" and "p \ P[X]" and "p \ ideal B" obtains A q where "finite A" and "A \ B" and "\b. q b \ P[X]" and "p = (\b\A. q b * b)" proof - from assms(3) obtain A where "finite A" and "A \ B" and "p \ ideal A" by (rule ideal.span_finite_subset) from this(2) assms(1) have "A \ P[X]" by (rule subset_trans) with \finite A\ obtain q where "\b. q b \ P[X]" and "p = (\b\A. q b * b)" using assms(2) \p \ ideal A\ by (rule in_idealE_Polys_finite) blast with \finite A\ \A \ B\ show ?thesis .. qed lemma ideal_induct_Polys [consumes 3, case_names 0 plus]: assumes "F \ P[X]" and "p \ P[X]" and "p \ ideal F" assumes "P 0" and "\c q h. c \ P[X] \ q \ F \ P h \ h \ P[X] \ P (c * q + h)" shows "P (p::('x \\<^sub>0 nat) \\<^sub>0 'a::comm_ring_1)" proof - obtain h where "\p q. h (p + q) = h p + h q" and "\p q. h (p * q) = h p * h q" and "\p::('x \\<^sub>0 nat) \\<^sub>0 'a. h (h p) = h p" and rng[symmetric]: "range h = P[X]" by (rule Polys_homomorphismE) blast from this(1-3) assms show ?thesis unfolding rng by (rule ideal_induct_homomorphism) blast qed lemma image_poly_subst_ideal_subset: "poly_subst g ` ideal F \ ideal (poly_subst g ` F)" proof (intro subsetI, elim imageE) fix h f assume h: "h = poly_subst g f" assume "f \ ideal F" thus "h \ ideal (poly_subst g ` F)" unfolding h proof (induct f rule: ideal.span_induct_alt) case base show ?case by (simp add: ideal.span_zero) next case (step c f h) from step.hyps(1) have "poly_subst g f \ ideal (poly_subst g ` F)" by (intro ideal.span_base imageI) hence "poly_subst g c * poly_subst g f \ ideal (poly_subst g ` F)" by (rule ideal.span_scale) hence "poly_subst g c * poly_subst g f + poly_subst g h \ ideal (poly_subst g ` F)" using step.hyps(2) by (rule ideal.span_add) thus ?case by (simp only: poly_subst_plus poly_subst_times) qed qed subsection \Evaluating Polynomials\ lemma lookup_times_zero: "lookup (p * q) 0 = lookup p 0 * lookup q (0::'a::{comm_powerprod,ninv_comm_monoid_add})" proof - have eq: "(\v\keys q. lookup q v when t + v = 0) = (lookup q 0 when t = 0)" for t proof - have "(\v\keys q. lookup q v when t + v = 0) = (\v\keys q \ {0}. lookup q v when t + v = 0)" proof (intro sum.mono_neutral_right ballI) fix v assume "v \ keys q - keys q \ {0}" hence "v \ 0" by blast hence "t + v \ 0" using plus_eq_zero_2 by blast thus "(lookup q v when t + v = 0) = 0" by simp qed simp_all also have "\ = (lookup q 0 when t = 0)" by (cases "0 \ keys q") (simp_all add: in_keys_iff) finally show ?thesis . qed have "(\t\keys p. lookup p t * lookup q 0 when t = 0) = (\t\keys p \ {0}. lookup p t * lookup q 0 when t = 0)" proof (intro sum.mono_neutral_right ballI) fix t assume "t \ keys p - keys p \ {0}" hence "t \ 0" by blast thus "(lookup p t * lookup q 0 when t = 0) = 0" by simp qed simp_all also have "\ = lookup p 0 * lookup q 0" by (cases "0 \ keys p") (simp_all add: in_keys_iff) finally show ?thesis by (simp add: lookup_times eq when_distrib) qed corollary lookup_prod_zero: "lookup (prod f I) 0 = (\i\I. lookup (f i) (0::_::{comm_powerprod,ninv_comm_monoid_add}))" by (induct I rule: infinite_finite_induct) (simp_all add: lookup_times_zero) corollary lookup_power_zero: "lookup (p ^ k) 0 = lookup p (0::_::{comm_powerprod,ninv_comm_monoid_add}) ^ k" by (induct k) (simp_all add: lookup_times_zero) definition poly_eval :: "('x \ 'a) \ (('x \\<^sub>0 nat) \\<^sub>0 'a) \ 'a::comm_semiring_1" where "poly_eval a p = lookup (poly_subst (\y. monomial (a y) (0::'x \\<^sub>0 nat)) p) 0" lemma poly_eval_alt: "poly_eval a p = (\t\keys p. lookup p t * (\x\keys t. a x ^ lookup t x))" by (simp add: poly_eval_def poly_subst_def lookup_sum lookup_times_zero subst_pp_def lookup_prod_zero lookup_power_zero flip: times_monomial_left) lemma poly_eval_monomial: "poly_eval a (monomial c t) = c * (\x\keys t. a x ^ lookup t x)" by (simp add: poly_eval_def poly_subst_monomial subst_pp_def punit.lookup_monom_mult lookup_prod_zero lookup_power_zero) lemma poly_eval_zero [simp]: "poly_eval a 0 = 0" by (simp only: poly_eval_def poly_subst_zero lookup_zero) lemma poly_eval_zero_left [simp]: "poly_eval 0 p = lookup p 0" by (simp add: poly_eval_def) lemma poly_eval_plus: "poly_eval a (p + q) = poly_eval a p + poly_eval a q" by (simp only: poly_eval_def poly_subst_plus lookup_add) lemma poly_eval_uminus [simp]: "poly_eval a (- p) = - poly_eval (a::_::comm_ring_1) p" by (simp only: poly_eval_def poly_subst_uminus lookup_uminus) lemma poly_eval_minus: "poly_eval a (p - q) = poly_eval a p - poly_eval (a::_::comm_ring_1) q" by (simp only: poly_eval_def poly_subst_minus lookup_minus) lemma poly_eval_one [simp]: "poly_eval a 1 = 1" by (simp add: poly_eval_def lookup_one) lemma poly_eval_times: "poly_eval a (p * q) = poly_eval a p * poly_eval a q" by (simp only: poly_eval_def poly_subst_times lookup_times_zero) lemma poly_eval_power: "poly_eval a (p ^ m) = poly_eval a p ^ m" by (induct m) (simp_all add: poly_eval_times) lemma poly_eval_sum: "poly_eval a (sum f I) = (\i\I. poly_eval a (f i))" by (induct I rule: infinite_finite_induct) (simp_all add: poly_eval_plus) lemma poly_eval_prod: "poly_eval a (prod f I) = (\i\I. poly_eval a (f i))" by (induct I rule: infinite_finite_induct) (simp_all add: poly_eval_times) lemma poly_eval_cong: "p = q \ (\x. x \ indets q \ a x = b x) \ poly_eval a p = poly_eval b q" by (simp add: poly_eval_def cong: poly_subst_cong) lemma indets_poly_eval_subset: "indets (poly_eval a p) \ \ (indets ` a ` indets p) \ \ (indets ` lookup p ` keys p)" proof (induct p rule: poly_mapping_plus_induct) case 1 show ?case by simp next case (2 p c t) have "keys (monomial c t + p) = keys (monomial c t) \ keys p" by (rule keys_plus_eqI) (simp add: 2(2)) with 2(1) have eq1: "keys (monomial c t + p) = insert t (keys p)" by simp hence eq2: "indets (monomial c t + p) = keys t \ indets p" by (simp add: indets_def) from 2(2) have eq3: "lookup (monomial c t + p) t = c" by (simp add: lookup_add in_keys_iff) have eq4: "lookup (monomial c t + p) s = lookup p s" if "s \ keys p" for s using that 2(2) by (auto simp: lookup_add lookup_single when_def) have "indets (poly_eval a (monomial c t + p)) = indets (c * (\x\keys t. a x ^ lookup t x) + poly_eval a p)" by (simp only: poly_eval_plus poly_eval_monomial) also have "\ \ indets (c * (\x\keys t. a x ^ lookup t x)) \ indets (poly_eval a p)" by (fact indets_plus_subset) also have "\ \ indets c \ (\ (indets ` a ` keys t)) \ (\ (indets ` a ` indets p) \ \ (indets ` lookup p ` keys p))" proof (intro Un_mono 2(3)) have "indets (c * (\x\keys t. a x ^ lookup t x)) \ indets c \ indets (\x\keys t. a x ^ lookup t x)" by (fact indets_times_subset) also have "indets (\x\keys t. a x ^ lookup t x) \ (\x\keys t. indets (a x ^ lookup t x))" by (fact indets_prod_subset) also have "\ \ (\x\keys t. indets (a x))" by (intro UN_mono subset_refl indets_power_subset) also have "\ = \ (indets ` a ` keys t)" by simp finally show "indets (c * (\x\keys t. a x ^ lookup t x)) \ indets c \ \ (indets ` a ` keys t)" by blast qed also have "\ = \ (indets ` a ` indets (monomial c t + p)) \ \ (indets ` lookup (monomial c t + p) ` keys (monomial c t + p))" by (simp add: eq1 eq2 eq3 eq4 Un_commute Un_assoc Un_left_commute) finally show ?case . qed lemma image_poly_eval_ideal: "poly_eval a ` ideal F = ideal (poly_eval a ` F)" proof (intro image_ideal_eq_surj poly_eval_plus poly_eval_times surjI) fix x show "poly_eval a (monomial x 0) = x" by (simp add: poly_eval_monomial) qed subsection \Replacing Indeterminates\ definition map_indets where "map_indets f = poly_subst (\x. monomial 1 (Poly_Mapping.single (f x) 1))" lemma shows map_indets_zero [simp]: "map_indets f 0 = 0" and map_indets_one [simp]: "map_indets f 1 = 1" and map_indets_uminus [simp]: "map_indets f (- r) = - map_indets f (r::_ \\<^sub>0 _::comm_ring_1)" and map_indets_plus: "map_indets f (p + q) = map_indets f p + map_indets f q" and map_indets_minus: "map_indets f (r - s) = map_indets f r - map_indets f s" and map_indets_times: "map_indets f (p * q) = map_indets f p * map_indets f q" and map_indets_power [simp]: "map_indets f (p ^ m) = map_indets f p ^ m" and map_indets_sum: "map_indets f (sum g A) = (\a\A. map_indets f (g a))" and map_indets_prod: "map_indets f (prod g A) = (\a\A. map_indets f (g a))" by (simp_all add: map_indets_def poly_subst_uminus poly_subst_plus poly_subst_minus poly_subst_times poly_subst_power poly_subst_sum poly_subst_prod) lemma map_indets_monomial: "map_indets f (monomial c t) = monomial c (\x\keys t. Poly_Mapping.single (f x) (lookup t x))" by (simp add: map_indets_def poly_subst_monomial subst_pp_def monomial_power_map_scale punit.monom_mult_monomial flip: punit.monomial_prod_sum) lemma map_indets_id: "(\x. x \ indets p \ f x = x) \ map_indets f p = p" by (simp add: map_indets_def poly_subst_id) lemma map_indets_map_indets: "map_indets f (map_indets g p) = map_indets (f \ g) p" by (simp add: map_indets_def poly_subst_poly_subst poly_subst_monomial subst_pp_single) lemma map_indets_cong: "p = q \ (\x. x \ indets q \ f x = g x) \ map_indets f p = map_indets g q" unfolding map_indets_def by (simp cong: poly_subst_cong) lemma poly_subst_map_indets: "poly_subst f (map_indets g p) = poly_subst (f \ g) p" by (simp add: map_indets_def poly_subst_poly_subst poly_subst_monomial subst_pp_single comp_def) lemma poly_eval_map_indets: "poly_eval a (map_indets g p) = poly_eval (a \ g) p" by (simp add: poly_eval_def poly_subst_map_indets comp_def) (simp add: poly_subst_def lookup_sum lookup_times_zero subst_pp_def lookup_prod_zero lookup_power_zero flip: times_monomial_left) lemma map_indets_inverseE_Polys: assumes "inj_on f X" and "p \ P[X]" shows "map_indets (the_inv_into X f) (map_indets f p) = p" unfolding map_indets_map_indets proof (rule map_indets_id) fix x assume "x \ indets p" also from assms(2) have "\ \ X" by (rule PolysD) finally show "(the_inv_into X f \ f) x = x" using assms(1) by (auto intro: the_inv_into_f_f) qed lemma map_indets_inverseE: assumes "inj f" obtains g where "g = the_inv f" and "g \ f = id" and "map_indets g \ map_indets f = id" proof - define g where "g = the_inv f" moreover from assms have eq: "g \ f = id" by (auto intro!: ext the_inv_f_f simp: g_def) moreover have "map_indets g \ map_indets f = id" by (rule ext) (simp add: map_indets_map_indets eq map_indets_id) ultimately show ?thesis .. qed lemma indets_map_indets_subset: "indets (map_indets f (p::_ \\<^sub>0 'a::comm_semiring_1)) \ f ` indets p" proof fix x assume "x \ indets (map_indets f p)" then obtain y where "y \ indets p" and "x \ indets (monomial (1::'a) (Poly_Mapping.single (f y) 1))" unfolding map_indets_def by (rule in_indets_poly_substE) from this(2) have x: "x = f y" by (simp add: indets_monomial) from \y \ indets p\ show "x \ f ` indets p" unfolding x by (rule imageI) qed corollary map_indets_in_Polys: "map_indets f p \ P[f ` indets p]" using indets_map_indets_subset by (rule PolysI_alt) lemma indets_map_indets: assumes "inj_on f (indets p)" shows "indets (map_indets f p) = f ` indets p" using indets_map_indets_subset proof (rule subset_antisym) let ?g = "the_inv_into (indets p) f" have "p = map_indets ?g (map_indets f p)" unfolding map_indets_map_indets by (rule sym, rule map_indets_id) (simp add: assms the_inv_into_f_f) also have "indets \ \ ?g ` indets (map_indets f p)" by (fact indets_map_indets_subset) finally have "f ` indets p \ f ` ?g ` indets (map_indets f p)" by (rule image_mono) also have "\ = (\x. x) ` indets (map_indets f p)" unfolding image_image using refl proof (rule image_cong) fix x assume "x \ indets (map_indets f p)" with indets_map_indets_subset have "x \ f ` indets p" .. with assms show "f (?g x) = x" by (rule f_the_inv_into_f) qed finally show "f ` indets p \ indets (map_indets f p)" by simp qed lemma image_map_indets_Polys: "map_indets f ` P[X] = (P[f ` X]::(_ \\<^sub>0 'a::comm_semiring_1) set)" proof (intro set_eqI iffI) fix p :: "_ \\<^sub>0 'a" assume "p \ map_indets f ` P[X]" then obtain q where "q \ P[X]" and "p = map_indets f q" .. note this(2) also have "map_indets f q \ P[f ` indets q]" by (fact map_indets_in_Polys) also from \q \ _\ have "\ \ P[f ` X]" by (auto intro!: Polys_mono imageI dest: PolysD) finally show "p \ P[f ` X]" . next fix p :: "_ \\<^sub>0 'a" assume "p \ P[f ` X]" define g where "g = (\y. SOME x. x \ X \ f x = y)" have "g y \ X \ f (g y) = y" if "y \ indets p" for y proof - note that also from \p \ _\ have "indets p \ f ` X" by (rule PolysD) finally obtain x where "x \ X" and "y = f x" .. hence "x \ X \ f x = y" by simp thus ?thesis unfolding g_def by (rule someI) qed hence 1: "g y \ X" and 2: "f (g y) = y" if "y \ indets p" for y using that by simp_all show "p \ map_indets f ` P[X]" proof show "p = map_indets f (map_indets g p)" by (rule sym) (simp add: map_indets_map_indets map_indets_id 2) next have "map_indets g p \ P[g ` indets p]" by (fact map_indets_in_Polys) also have "\ \ P[X]" by (auto intro!: Polys_mono 1) finally show "map_indets g p \ P[X]" . qed qed corollary range_map_indets: "range (map_indets f) = P[range f]" proof - have "range (map_indets f) = map_indets f ` P[UNIV]" by simp also have "\ = P[range f]" by (simp only: image_map_indets_Polys) finally show ?thesis . qed lemma in_keys_map_indetsE: assumes "t \ keys (map_indets f (p::_ \\<^sub>0 'a::comm_semiring_1))" obtains s where "s \ keys p" and "t = (\x\keys s. Poly_Mapping.single (f x) (lookup s x))" proof - let ?f = "(\x. monomial (1::'a) (Poly_Mapping.single (f x) 1))" from assms obtain s where "s \ keys p" and "t \ keys (subst_pp ?f s)" unfolding map_indets_def by (rule in_keys_poly_substE) note this(2) also have "\ \ {\x\keys s. Poly_Mapping.single (f x) (lookup s x)}" by (simp add: subst_pp_def monomial_power_map_scale flip: punit.monomial_prod_sum) finally have "t = (\x\keys s. Poly_Mapping.single (f x) (lookup s x))" by simp with \s \ keys p\ show ?thesis .. qed lemma keys_map_indets_subset: "keys (map_indets f p) \ (\t. \x\keys t. Poly_Mapping.single (f x) (lookup t x)) ` keys p" by (auto elim: in_keys_map_indetsE) lemma keys_map_indets: assumes "inj_on f (indets p)" shows "keys (map_indets f p) = (\t. \x\keys t. Poly_Mapping.single (f x) (lookup t x)) ` keys p" using keys_map_indets_subset proof (rule subset_antisym) let ?g = "the_inv_into (indets p) f" have "p = map_indets ?g (map_indets f p)" unfolding map_indets_map_indets by (rule sym, rule map_indets_id) (simp add: assms the_inv_into_f_f) also have "keys \ \ (\t. \x\keys t. monomial (lookup t x) (?g x)) ` keys (map_indets f p)" by (rule keys_map_indets_subset) finally have "(\t. \x\keys t. Poly_Mapping.single (f x) (lookup t x)) ` keys p \ (\t. \x\keys t. Poly_Mapping.single (f x) (lookup t x)) ` (\t. \x\keys t. Poly_Mapping.single (?g x) (lookup t x)) ` keys (map_indets f p)" by (rule image_mono) also from refl have "\ = (\t. \x. Poly_Mapping.single (f x) (lookup t x)) ` (\t. \x\keys t. Poly_Mapping.single (?g x) (lookup t x)) ` keys (map_indets f p)" by (rule image_cong) (smt Sum_any.conditionalize Sum_any.cong finite_keys not_in_keys_iff_lookup_eq_zero single_zero) also have "\ = (\t. t) ` keys (map_indets f p)" unfolding image_image using refl proof (rule image_cong) fix t assume "t \ keys (map_indets f p)" have "(\x. monomial (lookup (\y\keys t. Poly_Mapping.single (?g y) (lookup t y)) x) (f x)) = (\x. \y\keys t. monomial (lookup t y when ?g y = x) (f x))" by (simp add: lookup_sum lookup_single monomial_sum) also have "\ = (\x\indets p. \y\keys t. Poly_Mapping.single (f x) (lookup t y when ?g y = x))" proof (intro Sum_any.expand_superset finite_indets subsetI) fix x assume "x \ {a. (\y\keys t. Poly_Mapping.single (f a) (lookup t y when ?g y = a)) \ 0}" hence "(\y\keys t. Poly_Mapping.single (f x) (lookup t y when ?g y = x)) \ 0" by simp then obtain y where "y \ keys t" and *: "Poly_Mapping.single (f x) (lookup t y when ?g y = x) \ 0" by (rule sum.not_neutral_contains_not_neutral) from this(1) have "y \ indets (map_indets f p)" using \t \ _\ by (rule in_indetsI) with indets_map_indets_subset have "y \ f ` indets p" .. from * have "x = ?g y" by (simp add: when_def split: if_split_asm) also from assms \y \ f ` indets p\ subset_refl have "\ \ indets p" by (rule the_inv_into_into) finally show "x \ indets p" . qed also have "\ = (\y\keys t. \x\indets p. Poly_Mapping.single (f x) (lookup t y when ?g y = x))" by (fact sum.swap) also from refl have "\ = (\y\keys t. Poly_Mapping.single y (lookup t y))" proof (rule sum.cong) fix x assume "x \ keys t" hence "x \ indets (map_indets f p)" using \t \ _\ by (rule in_indetsI) with indets_map_indets_subset have "x \ f ` indets p" .. with assms have "?g x \ indets p" using subset_refl by (rule the_inv_into_into) hence "{?g x} \ indets p" by simp with finite_indets have "(\y\indets p. Poly_Mapping.single (f y) (lookup t x when ?g x = y)) = (\y\{?g x}. Poly_Mapping.single (f y) (lookup t x when ?g x = y))" by (rule sum.mono_neutral_right) (simp add: monomial_0_iff when_def) also from assms \x \ f ` indets p\ have "\ = Poly_Mapping.single x (lookup t x)" by (simp add: f_the_inv_into_f) finally show "(\y\indets p. Poly_Mapping.single (f y) (lookup t x when ?g x = y)) = Poly_Mapping.single x (lookup t x)" . qed also have "\ = t" by (fact poly_mapping_sum_monomials) finally show "(\x. monomial (lookup (\y\keys t. Poly_Mapping.single (?g y) (lookup t y)) x) (f x)) = t" . qed also have "\ = keys (map_indets f p)" by simp finally show "(\t. \x\keys t. Poly_Mapping.single (f x) (lookup t x)) ` keys p \ keys (map_indets f p)" . qed lemma poly_deg_map_indets_le: "poly_deg (map_indets f p) \ poly_deg p" proof (rule poly_deg_leI) fix t assume "t \ keys (map_indets f p)" then obtain s where "s \ keys p" and t: "t = (\x\keys s. Poly_Mapping.single (f x) (lookup s x))" by (rule in_keys_map_indetsE) from this(1) have "deg_pm s \ poly_deg p" by (rule poly_deg_max_keys) thus "deg_pm t \ poly_deg p" by (simp add: t deg_pm_sum deg_pm_single deg_pm_superset[OF subset_refl]) qed lemma poly_deg_map_indets: assumes "inj_on f (indets p)" shows "poly_deg (map_indets f p) = poly_deg p" proof - from assms have "deg_pm ` keys (map_indets f p) = deg_pm ` keys p" by (simp add: keys_map_indets image_image deg_pm_sum deg_pm_single flip: deg_pm_superset[OF subset_refl]) thus ?thesis by (auto simp: poly_deg_def) qed lemma map_indets_inj_on_PolysI: assumes "inj_on (f::'x \ 'y) X" shows "inj_on ((map_indets f)::_ \ _ \\<^sub>0 'a::comm_semiring_1) P[X]" proof (rule inj_onI) fix p q :: "_ \\<^sub>0 'a" assume "p \ P[X]" with assms have 1: "map_indets (the_inv_into X f) (map_indets f p) = p" (is "map_indets ?g _ = _") by (rule map_indets_inverseE_Polys) assume "q \ P[X]" with assms have "map_indets ?g (map_indets f q) = q" by (rule map_indets_inverseE_Polys) moreover assume "map_indets f p = map_indets f q" ultimately show "p = q" using 1 by (simp add: map_indets_map_indets) qed lemma map_indets_injI: assumes "inj f" shows "inj (map_indets f)" proof - from assms have "inj_on (map_indets f) P[UNIV]" by (rule map_indets_inj_on_PolysI) thus ?thesis by simp qed lemma image_map_indets_ideal: assumes "inj f" shows "map_indets f ` ideal F = ideal (map_indets f ` (F::(_ \\<^sub>0 'a::comm_ring_1) set)) \ P[range f]" proof from map_indets_plus map_indets_times have "map_indets f ` ideal F \ ideal (map_indets f ` F)" by (rule image_ideal_subset) moreover from subset_UNIV have "map_indets f ` ideal F \ range (map_indets f)" by (rule image_mono) ultimately show "map_indets f ` ideal F \ ideal (map_indets f ` F) \ P[range f]" unfolding range_map_indets by blast next show "ideal (map_indets f ` F) \ P[range f] \ map_indets f ` ideal F" proof fix p assume "p \ ideal (map_indets f ` F) \ P[range f]" hence "p \ ideal (map_indets f ` F)" and "p \ range (map_indets f)" by (simp_all add: range_map_indets) from this(1) obtain F0 q where "F0 \ map_indets f ` F" and p: "p = (\f'\F0. q f' * f')" by (rule ideal.spanE) from this(1) obtain F' where "F' \ F" and F0: "F0 = map_indets f ` F'" by (rule subset_imageE) from assms obtain g where "map_indets g \ map_indets f = (id::_ \ _ \\<^sub>0 'a)" by (rule map_indets_inverseE) hence eq: "map_indets g (map_indets f p') = p'" for p'::"_ \\<^sub>0 'a" by (simp add: pointfree_idE) from assms have "inj (map_indets f)" by (rule map_indets_injI) from this subset_UNIV have "inj_on (map_indets f) F'" by (rule inj_on_subset) from \p \ range _\ obtain p' where "p = map_indets f p'" .. hence "p = map_indets f (map_indets g p)" by (simp add: eq) also from \inj_on _ F'\ have "\ = map_indets f (\f'\F'. map_indets g (q (map_indets f f')) * f')" by (simp add: p F0 sum.reindex map_indets_sum map_indets_times eq) finally have "p = map_indets f (\f'\F'. map_indets g (q (map_indets f f')) * f')" . moreover have "(\f'\F'. map_indets g (q (map_indets f f')) * f') \ ideal F" proof show "(\f'\F'. map_indets g (q (map_indets f f')) * f') \ ideal F'" by (rule ideal.sum_in_spanI) next from \F' \ F\ show "ideal F' \ ideal F" by (rule ideal.span_mono) qed ultimately show "p \ map_indets f ` ideal F" by (rule image_eqI) qed qed subsection \Homogeneity\ definition homogeneous :: "(('x \\<^sub>0 nat) \\<^sub>0 'a::zero) \ bool" where "homogeneous p \ (\s\keys p. \t\keys p. deg_pm s = deg_pm t)" definition hom_component :: "(('x \\<^sub>0 nat) \\<^sub>0 'a) \ nat \ (('x \\<^sub>0 nat) \\<^sub>0 'a::zero)" where "hom_component p n = except p {t. deg_pm t \ n}" definition hom_components :: "(('x \\<^sub>0 nat) \\<^sub>0 'a) \ (('x \\<^sub>0 nat) \\<^sub>0 'a::zero) set" where "hom_components p = hom_component p ` deg_pm ` keys p" definition homogeneous_set :: "(('x \\<^sub>0 nat) \\<^sub>0 'a::zero) set \ bool" where "homogeneous_set A \ (\a\A. \n. hom_component a n \ A)" lemma homogeneousI: "(\s t. s \ keys p \ t \ keys p \ deg_pm s = deg_pm t) \ homogeneous p" unfolding homogeneous_def by blast lemma homogeneousD: "homogeneous p \ s \ keys p \ t \ keys p \ deg_pm s = deg_pm t" unfolding homogeneous_def by blast lemma homogeneousD_poly_deg: assumes "homogeneous p" and "t \ keys p" shows "deg_pm t = poly_deg p" proof (rule antisym) from assms(2) show "deg_pm t \ poly_deg p" by (rule poly_deg_max_keys) next show "poly_deg p \ deg_pm t" proof (rule poly_deg_leI) fix s assume "s \ keys p" with assms(1) have "deg_pm s = deg_pm t" using assms(2) by (rule homogeneousD) thus "deg_pm s \ deg_pm t" by simp qed qed lemma homogeneous_monomial [simp]: "homogeneous (monomial c t)" by (auto split: if_split_asm intro: homogeneousI) corollary homogeneous_zero [simp]: "homogeneous 0" and homogeneous_one [simp]: "homogeneous 1" by (simp_all only: homogeneous_monomial flip: single_zero[of 0] single_one) lemma homogeneous_uminus_iff [simp]: "homogeneous (- p) \ homogeneous p" by (auto intro!: homogeneousI dest: homogeneousD simp: keys_uminus) lemma homogeneous_monom_mult: "homogeneous p \ homogeneous (punit.monom_mult c t p)" by (auto intro!: homogeneousI elim!: punit.keys_monom_multE simp: deg_pm_plus dest: homogeneousD) lemma homogeneous_monom_mult_rev: assumes "c \ (0::'a::semiring_no_zero_divisors)" and "homogeneous (punit.monom_mult c t p)" shows "homogeneous p" proof (rule homogeneousI) fix s s' assume "s \ keys p" hence 1: "t + s \ keys (punit.monom_mult c t p)" using assms(1) by (rule punit.keys_monom_multI[simplified]) assume "s' \ keys p" hence "t + s' \ keys (punit.monom_mult c t p)" using assms(1) by (rule punit.keys_monom_multI[simplified]) with assms(2) 1 have "deg_pm (t + s) = deg_pm (t + s')" by (rule homogeneousD) thus "deg_pm s = deg_pm s'" by (simp add: deg_pm_plus) qed lemma homogeneous_times: assumes "homogeneous p" and "homogeneous q" shows "homogeneous (p * q)" proof (rule homogeneousI) fix s t assume "s \ keys (p * q)" then obtain sp sq where sp: "sp \ keys p" and sq: "sq \ keys q" and s: "s = sp + sq" by (rule in_keys_timesE) assume "t \ keys (p * q)" then obtain tp tq where tp: "tp \ keys p" and tq: "tq \ keys q" and t: "t = tp + tq" by (rule in_keys_timesE) from assms(1) sp tp have "deg_pm sp = deg_pm tp" by (rule homogeneousD) moreover from assms(2) sq tq have "deg_pm sq = deg_pm tq" by (rule homogeneousD) ultimately show "deg_pm s = deg_pm t" by (simp only: s t deg_pm_plus) qed lemma lookup_hom_component: "lookup (hom_component p n) = (\t. lookup p t when deg_pm t = n)" by (rule ext) (simp add: hom_component_def lookup_except) lemma keys_hom_component: "keys (hom_component p n) = {t. t \ keys p \ deg_pm t = n}" by (auto simp: hom_component_def keys_except) lemma keys_hom_componentD: assumes "t \ keys (hom_component p n)" shows "t \ keys p" and "deg_pm t = n" using assms by (simp_all add: keys_hom_component) lemma homogeneous_hom_component: "homogeneous (hom_component p n)" by (auto dest: keys_hom_componentD intro: homogeneousI) lemma hom_component_zero [simp]: "hom_component 0 = 0" by (rule ext) (simp add: hom_component_def) lemma hom_component_zero_iff: "hom_component p n = 0 \ (\t\keys p. deg_pm t \ n)" by (metis (mono_tags, lifting) empty_iff keys_eq_empty_iff keys_hom_component mem_Collect_eq subsetI subset_antisym) lemma hom_component_uminus [simp]: "hom_component (- p) = - hom_component p" by (intro ext poly_mapping_eqI) (simp add: hom_component_def lookup_except) lemma hom_component_plus: "hom_component (p + q) n = hom_component p n + hom_component q n" by (rule poly_mapping_eqI) (simp add: hom_component_def lookup_except lookup_add) lemma hom_component_minus: "hom_component (p - q) n = hom_component p n - hom_component q n" by (rule poly_mapping_eqI) (simp add: hom_component_def lookup_except lookup_minus) lemma hom_component_monom_mult: "punit.monom_mult c t (hom_component p n) = hom_component (punit.monom_mult c t p) (deg_pm t + n)" by (auto simp: hom_component_def lookup_except punit.lookup_monom_mult deg_pm_minus deg_pm_mono intro!: poly_mapping_eqI) lemma hom_component_inject: assumes "t \ keys p" and "hom_component p (deg_pm t) = hom_component p n" shows "deg_pm t = n" proof - from assms(1) have "t \ keys (hom_component p (deg_pm t))" by (simp add: keys_hom_component) hence "0 \ lookup (hom_component p (deg_pm t)) t" by (simp add: in_keys_iff) also have "lookup (hom_component p (deg_pm t)) t = lookup (hom_component p n) t" by (simp only: assms(2)) also have "\ = (lookup p t when deg_pm t = n)" by (simp only: lookup_hom_component) finally show ?thesis by simp qed lemma hom_component_of_homogeneous: assumes "homogeneous p" shows "hom_component p n = (p when n = poly_deg p)" proof (cases "n = poly_deg p") case True have "hom_component p n = p" proof (rule poly_mapping_eqI) fix t show "lookup (hom_component p n) t = lookup p t" proof (cases "t \ keys p") case True with assms have "deg_pm t = n" unfolding \n = poly_deg p\ by (rule homogeneousD_poly_deg) thus ?thesis by (simp add: lookup_hom_component) next case False moreover from this have "t \ keys (hom_component p n)" by (simp add: keys_hom_component) ultimately show ?thesis by (simp add: in_keys_iff) qed qed with True show ?thesis by simp next case False have "hom_component p n = 0" unfolding hom_component_zero_iff proof (intro ballI notI) fix t assume "t \ keys p" with assms have "deg_pm t = poly_deg p" by (rule homogeneousD_poly_deg) moreover assume "deg_pm t = n" ultimately show False using False by simp qed with False show ?thesis by simp qed lemma hom_components_zero [simp]: "hom_components 0 = {}" by (simp add: hom_components_def) lemma hom_components_zero_iff [simp]: "hom_components p = {} \ p = 0" by (simp add: hom_components_def) lemma hom_components_uminus: "hom_components (- p) = uminus ` hom_components p" by (simp add: hom_components_def keys_uminus image_image) lemma hom_components_monom_mult: "hom_components (punit.monom_mult c t p) = (if c = 0 then {} else punit.monom_mult c t ` hom_components p)" for c::"'a::semiring_no_zero_divisors" by (simp add: hom_components_def punit.keys_monom_mult image_image deg_pm_plus hom_component_monom_mult) lemma hom_componentsI: "q = hom_component p (deg_pm t) \ t \ keys p \ q \ hom_components p" unfolding hom_components_def by blast lemma hom_componentsE: assumes "q \ hom_components p" obtains t where "t \ keys p" and "q = hom_component p (deg_pm t)" using assms unfolding hom_components_def by blast lemma hom_components_of_homogeneous: assumes "homogeneous p" shows "hom_components p = (if p = 0 then {} else {p})" proof (split if_split, intro conjI impI) assume "p \ 0" have "deg_pm ` keys p = {poly_deg p}" proof (rule set_eqI) fix n have "n \ deg_pm ` keys p \ n = poly_deg p" proof assume "n \ deg_pm ` keys p" then obtain t where "t \ keys p" and "n = deg_pm t" .. from assms this(1) have "deg_pm t = poly_deg p" by (rule homogeneousD_poly_deg) thus "n = poly_deg p" by (simp only: \n = deg_pm t\) next assume "n = poly_deg p" from \p \ 0\ have "keys p \ {}" by simp then obtain t where "t \ keys p" by blast with assms have "deg_pm t = poly_deg p" by (rule homogeneousD_poly_deg) hence "n = deg_pm t" by (simp only: \n = poly_deg p\) with \t \ keys p\ show "n \ deg_pm ` keys p" by (rule rev_image_eqI) qed thus "n \ deg_pm ` keys p \ n \ {poly_deg p}" by simp qed with assms show "hom_components p = {p}" by (simp add: hom_components_def hom_component_of_homogeneous) qed simp lemma finite_hom_components: "finite (hom_components p)" unfolding hom_components_def using finite_keys by (intro finite_imageI) lemma hom_components_homogeneous: "q \ hom_components p \ homogeneous q" by (elim hom_componentsE) (simp only: homogeneous_hom_component) lemma hom_components_nonzero: "q \ hom_components p \ q \ 0" by (auto elim!: hom_componentsE simp: hom_component_zero_iff) lemma deg_pm_hom_components: assumes "q1 \ hom_components p" and "q2 \ hom_components p" and "t1 \ keys q1" and "t2 \ keys q2" shows "deg_pm t1 = deg_pm t2 \ q1 = q2" proof - from assms(1) obtain s1 where "s1 \ keys p" and q1: "q1 = hom_component p (deg_pm s1)" by (rule hom_componentsE) from assms(3) have t1: "deg_pm t1 = deg_pm s1" unfolding q1 by (rule keys_hom_componentD) from assms(2) obtain s2 where "s2 \ keys p" and q2: "q2 = hom_component p (deg_pm s2)" by (rule hom_componentsE) from assms(4) have t2: "deg_pm t2 = deg_pm s2" unfolding q2 by (rule keys_hom_componentD) from \s1 \ keys p\ show ?thesis by (auto simp: q1 q2 t1 t2 dest: hom_component_inject) qed lemma poly_deg_hom_components: assumes "q1 \ hom_components p" and "q2 \ hom_components p" shows "poly_deg q1 = poly_deg q2 \ q1 = q2" proof - from assms(1) have "homogeneous q1" and "q1 \ 0" by (rule hom_components_homogeneous, rule hom_components_nonzero) from this(2) have "keys q1 \ {}" by simp then obtain t1 where "t1 \ keys q1" by blast with \homogeneous q1\ have t1: "deg_pm t1 = poly_deg q1" by (rule homogeneousD_poly_deg) from assms(2) have "homogeneous q2" and "q2 \ 0" by (rule hom_components_homogeneous, rule hom_components_nonzero) from this(2) have "keys q2 \ {}" by simp then obtain t2 where "t2 \ keys q2" by blast with \homogeneous q2\ have t2: "deg_pm t2 = poly_deg q2" by (rule homogeneousD_poly_deg) from assms \t1 \ keys q1\ \t2 \ keys q2\ have "deg_pm t1 = deg_pm t2 \ q1 = q2" by (rule deg_pm_hom_components) thus ?thesis by (simp only: t1 t2) qed lemma hom_components_keys_disjoint: assumes "q1 \ hom_components p" and "q2 \ hom_components p" and "q1 \ q2" shows "keys q1 \ keys q2 = {}" proof (rule ccontr) assume "keys q1 \ keys q2 \ {}" then obtain t where "t \ keys q1" and "t \ keys q2" by blast with assms(1, 2) have "deg_pm t = deg_pm t \ q1 = q2" by (rule deg_pm_hom_components) with assms(3) show False by simp qed lemma Keys_hom_components: "Keys (hom_components p) = keys p" by (auto simp: Keys_def hom_components_def keys_hom_component) lemma lookup_hom_components: "q \ hom_components p \ t \ keys q \ lookup q t = lookup p t" by (auto elim!: hom_componentsE simp: keys_hom_component lookup_hom_component) lemma poly_deg_hom_components_le: assumes "q \ hom_components p" shows "poly_deg q \ poly_deg p" proof (rule poly_deg_leI) fix t assume "t \ keys q" also from assms have "\ \ Keys (hom_components p)" by (rule keys_subset_Keys) also have "\ = keys p" by (fact Keys_hom_components) finally show "deg_pm t \ poly_deg p" by (rule poly_deg_max_keys) qed lemma sum_hom_components: "\(hom_components p) = p" proof (rule poly_mapping_eqI) fix t show "lookup (\(hom_components p)) t = lookup p t" unfolding lookup_sum proof (cases "t \ keys p") case True also have "keys p = Keys (hom_components p)" by (simp only: Keys_hom_components) finally obtain q where q: "q \ hom_components p" and t: "t \ keys q" by (rule in_KeysE) from this(1) have "(\q0\hom_components p. lookup q0 t) = (\q0\insert q (hom_components p). lookup q0 t)" by (simp only: insert_absorb) also from finite_hom_components have "\ = lookup q t + (\q0\hom_components p - {q}. lookup q0 t)" by (rule sum.insert_remove) also from q t have "\ = lookup p t + (\q0\hom_components p - {q}. lookup q0 t)" by (simp only: lookup_hom_components) also have "(\q0\hom_components p - {q}. lookup q0 t) = 0" proof (intro sum.neutral ballI) fix q0 assume "q0 \ hom_components p - {q}" hence "q0 \ hom_components p" and "q \ q0" by blast+ with q have "keys q \ keys q0 = {}" by (rule hom_components_keys_disjoint) with t have "t \ keys q0" by blast thus "lookup q0 t = 0" by (simp add: in_keys_iff) qed finally show "(\q\hom_components p. lookup q t) = lookup p t" by simp next case False hence "t \ Keys (hom_components p)" by (simp add: Keys_hom_components) hence "\q\hom_components p. lookup q t = 0" by (simp add: Keys_def in_keys_iff) hence "(\q\hom_components p. lookup q t) = 0" by (rule sum.neutral) also from False have "\ = lookup p t" by (simp add: in_keys_iff) finally show "(\q\hom_components p. lookup q t) = lookup p t" . qed qed lemma homogeneous_setI: "(\a n. a \ A \ hom_component a n \ A) \ homogeneous_set A" by (simp add: homogeneous_set_def) lemma homogeneous_setD: "homogeneous_set A \ a \ A \ hom_component a n \ A" by (simp add: homogeneous_set_def) lemma homogeneous_set_Polys: "homogeneous_set (P[X]::(_ \\<^sub>0 'a::zero) set)" proof (intro homogeneous_setI PolysI subsetI) fix p::"_ \\<^sub>0 'a" and n t assume "p \ P[X]" assume "t \ keys (hom_component p n)" hence "t \ keys p" by (rule keys_hom_componentD) also from \p \ P[X]\ have "\ \ .[X]" by (rule PolysD) finally show "t \ .[X]" . qed lemma homogeneous_set_IntI: "homogeneous_set A \ homogeneous_set B \ homogeneous_set (A \ B)" by (simp add: homogeneous_set_def) lemma homogeneous_setD_hom_components: assumes "homogeneous_set A" and "a \ A" and "b \ hom_components a" shows "b \ A" proof - from assms(3) obtain t::"'a \\<^sub>0 nat" where "b = hom_component a (deg_pm t)" by (rule hom_componentsE) also from assms(1, 2) have "\ \ A" by (rule homogeneous_setD) finally show ?thesis . qed lemma zero_in_homogeneous_set: assumes "homogeneous_set A" and "A \ {}" shows "0 \ A" proof - from assms(2) obtain a where "a \ A" by blast have "lookup a t = 0" if "deg_pm t = Suc (poly_deg a)" for t proof (rule ccontr) assume "lookup a t \ 0" hence "t \ keys a" by (simp add: in_keys_iff) hence "deg_pm t \ poly_deg a" by (rule poly_deg_max_keys) thus False by (simp add: that) qed hence "0 = hom_component a (Suc (poly_deg a))" by (intro poly_mapping_eqI) (simp add: lookup_hom_component when_def) also from assms(1) \a \ A\ have "\ \ A" by (rule homogeneous_setD) finally show ?thesis . qed lemma homogeneous_ideal: assumes "\f. f \ F \ homogeneous f" and "p \ ideal F" shows "hom_component p n \ ideal F" proof - from assms(2) have "p \ punit.pmdl F" by simp thus ?thesis proof (induct p rule: punit.pmdl_induct) case module_0 show ?case by (simp add: ideal.span_zero) next case (module_plus a f c t) let ?f = "punit.monom_mult c t f" from module_plus.hyps(3) have "f \ punit.pmdl F" by (simp add: ideal.span_base) hence *: "?f \ punit.pmdl F" by (rule punit.pmdl_closed_monom_mult) from module_plus.hyps(3) have "homogeneous f" by (rule assms(1)) hence "homogeneous ?f" by (rule homogeneous_monom_mult) hence "hom_component ?f n = (?f when n = poly_deg ?f)" by (rule hom_component_of_homogeneous) also from * have "\ \ ideal F" by (simp add: when_def ideal.span_zero) finally have "hom_component ?f n \ ideal F" . with module_plus.hyps(2) show ?case unfolding hom_component_plus by (rule ideal.span_add) qed qed corollary homogeneous_set_homogeneous_ideal: "(\f. f \ F \ homogeneous f) \ homogeneous_set (ideal F)" by (auto intro: homogeneous_setI homogeneous_ideal) corollary homogeneous_ideal': assumes "\f. f \ F \ homogeneous f" and "p \ ideal F" and "q \ hom_components p" shows "q \ ideal F" using _ assms(2, 3) proof (rule homogeneous_setD_hom_components) from assms(1) show "homogeneous_set (ideal F)" by (rule homogeneous_set_homogeneous_ideal) qed lemma homogeneous_idealE_homogeneous: assumes "\f. f \ F \ homogeneous f" and "p \ ideal F" and "homogeneous p" obtains F' q where "finite F'" and "F' \ F" and "p = (\f\F'. q f * f)" and "\f. homogeneous (q f)" and "\f. f \ F' \ poly_deg (q f * f) = poly_deg p" and "\f. f \ F' \ q f = 0" proof - from assms(2) obtain F'' q' where "finite F''" and "F'' \ F" and p: "p = (\f\F''. q' f * f)" by (rule ideal.spanE) let ?A = "\f. {h \ hom_components (q' f). poly_deg h + poly_deg f = poly_deg p}" let ?B = "\f. {h \ hom_components (q' f). poly_deg h + poly_deg f \ poly_deg p}" define F' where "F' = {f \ F''. (\(?A f)) * f \ 0}" define q where "q = (\f. (\(?A f)) when f \ F')" have "F' \ F''" by (simp add: F'_def) hence "F' \ F" using \F'' \ F\ by (rule subset_trans) have 1: "deg_pm t + poly_deg f = poly_deg p" if "f \ F'" and "t \ keys (q f)" for f t proof - from that have "t \ keys (\(?A f))" by (simp add: q_def) also have "\ \ (\h\?A f. keys h)" by (fact keys_sum_subset) finally obtain h where "h \ ?A f" and "t \ keys h" .. from this(1) have "h \ hom_components (q' f)" and eq: "poly_deg h + poly_deg f = poly_deg p" by simp_all from this(1) have "homogeneous h" by (rule hom_components_homogeneous) hence "deg_pm t = poly_deg h" using \t \ keys h\ by (rule homogeneousD_poly_deg) thus ?thesis by (simp only: eq) qed have 2: "deg_pm t = poly_deg p" if "f \ F'" and "t \ keys (q f * f)" for f t proof - from that(1) \F' \ F\ have "f \ F" .. hence "homogeneous f" by (rule assms(1)) from that(2) obtain s1 s2 where "s1 \ keys (q f)" and "s2 \ keys f" and t: "t = s1 + s2" by (rule in_keys_timesE) from that(1) this(1) have "deg_pm s1 + poly_deg f = poly_deg p" by (rule 1) moreover from \homogeneous f\ \s2 \ keys f\ have "deg_pm s2 = poly_deg f" by (rule homogeneousD_poly_deg) ultimately show ?thesis by (simp add: t deg_pm_plus) qed from \F' \ F''\ \finite F''\ have "finite F'" by (rule finite_subset) thus ?thesis using \F' \ F\ proof note p also from refl have "(\f\F''. q' f * f) = (\f\F''. (\(?A f) * f) + (\(?B f) * f))" proof (rule sum.cong) fix f assume "f \ F''" from sum_hom_components have "q' f = (\(hom_components (q' f)))" by (rule sym) also have "\ = (\(?A f \ ?B f))" by (rule arg_cong[where f="sum (\x. x)"]) blast also have "\ = \(?A f) + \(?B f)" proof (rule sum.union_disjoint) have "?A f \ hom_components (q' f)" by blast thus "finite (?A f)" using finite_hom_components by (rule finite_subset) next have "?B f \ hom_components (q' f)" by blast thus "finite (?B f)" using finite_hom_components by (rule finite_subset) qed blast finally show "q' f * f = (\(?A f) * f) + (\(?B f) * f)" by (metis (no_types, lifting) distrib_right) qed also have "\ = (\f\F''. \(?A f) * f) + (\f\F''. \(?B f) * f)" by (rule sum.distrib) also from \finite F''\ \F' \ F''\ have "(\f\F''. \(?A f) * f) = (\f\F'. q f * f)" proof (intro sum.mono_neutral_cong_right ballI) fix f assume "f \ F'' - F'" thus "\(?A f) * f = 0" by (simp add: F'_def) next fix f assume "f \ F'" thus "\(?A f) * f = q f * f" by (simp add: q_def) qed finally have p[symmetric]: "p = (\f\F'. q f * f) + (\f\F''. \(?B f) * f)" . moreover have "keys (\f\F''. \(?B f) * f) = {}" proof (rule, rule) fix t assume t_in: "t \ keys (\f\F''. \(?B f) * f)" also have "\ \ (\f\F''. keys (\(?B f) * f))" by (fact keys_sum_subset) finally obtain f where "f \ F''" and "t \ keys (\(?B f) * f)" .. from this(2) obtain s1 s2 where "s1 \ keys (\(?B f))" and "s2 \ keys f" and t: "t = s1 + s2" by (rule in_keys_timesE) from \f \ F''\ \F'' \ F\ have "f \ F" .. hence "homogeneous f" by (rule assms(1)) note \s1 \ keys (\(?B f))\ also have "keys (\(?B f)) \ (\h\?B f. keys h)" by (fact keys_sum_subset) finally obtain h where "h \ ?B f" and "s1 \ keys h" .. from this(1) have "h \ hom_components (q' f)" and neq: "poly_deg h + poly_deg f \ poly_deg p" by simp_all from this(1) have "homogeneous h" by (rule hom_components_homogeneous) hence "deg_pm s1 = poly_deg h" using \s1 \ keys h\ by (rule homogeneousD_poly_deg) moreover from \homogeneous f\ \s2 \ keys f\ have "deg_pm s2 = poly_deg f" by (rule homogeneousD_poly_deg) ultimately have "deg_pm t \ poly_deg p" using neq by (simp add: t deg_pm_plus) have "t \ keys (\f\F'. q f * f)" proof assume "t \ keys (\f\F'. q f * f)" also have "\ \ (\f\F'. keys (q f * f))" by (fact keys_sum_subset) finally obtain f where "f \ F'" and "t \ keys (q f * f)" .. hence "deg_pm t = poly_deg p" by (rule 2) with \deg_pm t \ poly_deg p\ show False .. qed with t_in have "t \ keys ((\f\F'. q f * f) + (\f\F''. \(?B f) * f))" by (rule in_keys_plusI2) hence "t \ keys p" by (simp only: p) with assms(3) have "deg_pm t = poly_deg p" by (rule homogeneousD_poly_deg) with \deg_pm t \ poly_deg p\ show "t \ {}" .. qed (fact empty_subsetI) ultimately show "p = (\f\F'. q f * f)" by simp next fix f show "homogeneous (q f)" proof (cases "f \ F'") case True show ?thesis proof (rule homogeneousI) fix s t assume "s \ keys (q f)" with True have *: "deg_pm s + poly_deg f = poly_deg p" by (rule 1) assume "t \ keys (q f)" with True have "deg_pm t + poly_deg f = poly_deg p" by (rule 1) with * show "deg_pm s = deg_pm t" by simp qed next case False thus ?thesis by (simp add: q_def) qed assume "f \ F'" show "poly_deg (q f * f) = poly_deg p" proof (intro antisym) show "poly_deg (q f * f) \ poly_deg p" proof (rule poly_deg_leI) fix t assume "t \ keys (q f * f)" with \f \ F'\ have "deg_pm t = poly_deg p" by (rule 2) thus "deg_pm t \ poly_deg p" by simp qed next from \f \ F'\ have "q f * f \ 0" by (simp add: q_def F'_def) hence "keys (q f * f) \ {}" by simp then obtain t where "t \ keys (q f * f)" by blast with \f \ F'\ have "deg_pm t = poly_deg p" by (rule 2) moreover from \t \ keys (q f * f)\ have "deg_pm t \ poly_deg (q f * f)" by (rule poly_deg_max_keys) ultimately show "poly_deg p \ poly_deg (q f * f)" by simp qed qed (simp add: q_def) qed corollary homogeneous_idealE: assumes "\f. f \ F \ homogeneous f" and "p \ ideal F" obtains F' q where "finite F'" and "F' \ F" and "p = (\f\F'. q f * f)" and "\f. poly_deg (q f * f) \ poly_deg p" and "\f. f \ F' \ q f = 0" proof (cases "p = 0") case True show ?thesis proof show "p = (\f\{}. (\_. 0) f * f)" by (simp add: True) qed simp_all next case False define P where "P = (\h qf. finite (fst qf) \ fst qf \ F \ h = (\f\fst qf. snd qf f * f) \ (\f\fst qf. poly_deg (snd qf f * f) = poly_deg h) \ (\f. f \ fst qf \ snd qf f = 0))" define q0 where "q0 = (\h. SOME qf. P h qf)" have 1: "P h (q0 h)" if "h \ hom_components p" for h proof - note assms(1) moreover from assms that have "h \ ideal F" by (rule homogeneous_ideal') moreover from that have "homogeneous h" by (rule hom_components_homogeneous) ultimately obtain F' q where "finite F'" and "F' \ F" and "h = (\f\F'. q f * f)" and "\f. f \ F' \ poly_deg (q f * f) = poly_deg h" and "\f. f \ F' \ q f = 0" by (rule homogeneous_idealE_homogeneous) blast+ hence "P h (F', q)" by (simp add: P_def) thus ?thesis unfolding q0_def by (rule someI) qed define F' where "F' = (\h\hom_components p. fst (q0 h))" define q where "q = (\f. \h\hom_components p. snd (q0 h) f)" show ?thesis proof have "finite F' \ F' \ F" unfolding F'_def UN_subset_iff finite_UN[OF finite_hom_components] proof (intro conjI ballI) fix h assume "h \ hom_components p" hence "P h (q0 h)" by (rule 1) thus "finite (fst (q0 h))" and "fst (q0 h) \ F" by (simp_all only: P_def) qed thus "finite F'" and "F' \ F" by simp_all from sum_hom_components have "p = (\(hom_components p))" by (rule sym) also from refl have "\ = (\h\hom_components p. \f\F'. snd (q0 h) f * f)" proof (rule sum.cong) fix h assume "h \ hom_components p" hence "P h (q0 h)" by (rule 1) hence "h = (\f\fst (q0 h). snd (q0 h) f * f)" and 2: "\f. f \ fst (q0 h) \ snd (q0 h) f = 0" by (simp_all add: P_def) note this(1) also from \finite F'\ have "(\f\fst (q0 h). (snd (q0 h)) f * f) = (\f\F'. snd (q0 h) f * f)" proof (intro sum.mono_neutral_left ballI) show "fst (q0 h) \ F'" unfolding F'_def using \h \ hom_components p\ by blast next fix f assume "f \ F' - fst (q0 h)" hence "f \ fst (q0 h)" by simp hence "snd (q0 h) f = 0" by (rule 2) thus "snd (q0 h) f * f = 0" by simp qed finally show "h = (\f\F'. snd (q0 h) f * f)" . qed also have "\ = (\f\F'. \h\hom_components p. snd (q0 h) f * f)" by (rule sum.swap) also have "\ = (\f\F'. q f * f)" by (simp only: q_def sum_distrib_right) finally show "p = (\f\F'. q f * f)" . fix f have "poly_deg (q f * f) = poly_deg (\h\hom_components p. snd (q0 h) f * f)" by (simp only: q_def sum_distrib_right) also have "\ \ Max (poly_deg ` (\h. snd (q0 h) f * f) ` hom_components p)" by (rule poly_deg_sum_le) also have "\ = Max ((\h. poly_deg (snd (q0 h) f * f)) ` hom_components p)" (is "_ = Max (?f ` _)") by (simp only: image_image) also have "\ \ poly_deg p" proof (rule Max.boundedI) from finite_hom_components show "finite (?f ` hom_components p)" by (rule finite_imageI) next from False show "?f ` hom_components p \ {}" by simp next fix d assume "d \ ?f ` hom_components p" then obtain h where "h \ hom_components p" and d: "d = ?f h" .. from this(1) have "P h (q0 h)" by (rule 1) hence 2: "\f. f \ fst (q0 h) \ poly_deg (snd (q0 h) f * f) = poly_deg h" and 3: "\f. f \ fst (q0 h) \ snd (q0 h) f = 0" by (simp_all add: P_def) show "d \ poly_deg p" proof (cases "f \ fst (q0 h)") case True hence "poly_deg (snd (q0 h) f * f) = poly_deg h" by (rule 2) hence "d = poly_deg h" by (simp only: d) also from \h \ hom_components p\ have "\ \ poly_deg p" by (rule poly_deg_hom_components_le) finally show ?thesis . next case False hence "snd (q0 h) f = 0" by (rule 3) thus ?thesis by (simp add: d) qed qed finally show "poly_deg (q f * f) \ poly_deg p" . assume "f \ F'" show "q f = 0" unfolding q_def proof (intro sum.neutral ballI) fix h assume "h \ hom_components p" hence "P h (q0 h)" by (rule 1) hence 2: "\f. f \ fst (q0 h) \ snd (q0 h) f = 0" by (simp add: P_def) show "snd (q0 h) f = 0" proof (intro 2 notI) assume "f \ fst (q0 h)" hence "f \ F'" unfolding F'_def using \h \ hom_components p\ by blast with \f \ F'\ show False .. qed qed qed qed corollary homogeneous_idealE_finite: assumes "finite F" and "\f. f \ F \ homogeneous f" and "p \ ideal F" obtains q where "p = (\f\F. q f * f)" and "\f. poly_deg (q f * f) \ poly_deg p" and "\f. f \ F \ q f = 0" proof - from assms(2, 3) obtain F' q where "F' \ F" and p: "p = (\f\F'. q f * f)" and "\f. poly_deg (q f * f) \ poly_deg p" and 1: "\f. f \ F' \ q f = 0" by (rule homogeneous_idealE) blast+ show ?thesis proof from assms(1) \F' \ F\ have "(\f\F'. q f * f) = (\f\F. q f * f)" proof (intro sum.mono_neutral_left ballI) fix f assume "f \ F - F'" hence "f \ F'" by simp hence "q f = 0" by (rule 1) thus "q f * f = 0" by simp qed thus "p = (\f\F. q f * f)" by (simp only: p) next fix f show "poly_deg (q f * f) \ poly_deg p" by fact assume "f \ F" with \F' \ F\ have "f \ F'" by blast thus "q f = 0" by (rule 1) qed qed subsubsection \Homogenization and Dehomogenization\ definition homogenize :: "'x \ (('x \\<^sub>0 nat) \\<^sub>0 'a) \ (('x \\<^sub>0 nat) \\<^sub>0 'a::semiring_1)" where "homogenize x p = (\t\keys p. monomial (lookup p t) (Poly_Mapping.single x (poly_deg p - deg_pm t) + t))" definition dehomo_subst :: "'x \ 'x \ (('x \\<^sub>0 nat) \\<^sub>0 'a::zero_neq_one)" where "dehomo_subst x = (\y. if y = x then 1 else monomial 1 (Poly_Mapping.single y 1))" definition dehomogenize :: "'x \ (('x \\<^sub>0 nat) \\<^sub>0 'a) \ (('x \\<^sub>0 nat) \\<^sub>0 'a::comm_semiring_1)" where "dehomogenize x = poly_subst (dehomo_subst x)" lemma homogenize_zero [simp]: "homogenize x 0 = 0" by (simp add: homogenize_def) lemma homogenize_uminus [simp]: "homogenize x (- p) = - homogenize x (p::_ \\<^sub>0 'a::ring_1)" by (simp add: homogenize_def keys_uminus sum.reindex inj_on_def single_uminus sum_negf) lemma homogenize_monom_mult [simp]: "homogenize x (punit.monom_mult c t p) = punit.monom_mult c t (homogenize x p)" for c::"'a::{semiring_1,semiring_no_zero_divisors_cancel}" proof (cases "p = 0") case True thus ?thesis by simp next case False show ?thesis proof (cases "c = 0") case True thus ?thesis by simp next case False show ?thesis by (simp add: homogenize_def punit.keys_monom_mult \p \ 0\ False sum.reindex punit.lookup_monom_mult punit.monom_mult_sum_right poly_deg_monom_mult punit.monom_mult_monomial ac_simps deg_pm_plus) qed qed lemma homogenize_alt: "homogenize x p = (\q\hom_components p. punit.monom_mult 1 (Poly_Mapping.single x (poly_deg p - poly_deg q)) q)" proof - have "homogenize x p = (\t\Keys (hom_components p). monomial (lookup p t) (Poly_Mapping.single x (poly_deg p - deg_pm t) + t))" by (simp only: homogenize_def Keys_hom_components) also have "\ = (\t\(\ (keys ` hom_components p)). monomial (lookup p t) (Poly_Mapping.single x (poly_deg p - deg_pm t) + t))" by (simp only: Keys_def) also have "\ = (\q\hom_components p. (\t\keys q. monomial (lookup p t) (Poly_Mapping.single x (poly_deg p - deg_pm t) + t)))" by (auto intro!: sum.UNION_disjoint finite_hom_components finite_keys dest: hom_components_keys_disjoint) also have "\ = (\q\hom_components p. punit.monom_mult 1 (Poly_Mapping.single x (poly_deg p - poly_deg q)) q)" using refl proof (rule sum.cong) fix q assume q: "q \ hom_components p" hence "homogeneous q" by (rule hom_components_homogeneous) have "(\t\keys q. monomial (lookup p t) (Poly_Mapping.single x (poly_deg p - deg_pm t) + t)) = (\t\keys q. punit.monom_mult 1 (Poly_Mapping.single x (poly_deg p - poly_deg q)) (monomial (lookup q t) t))" using refl proof (rule sum.cong) fix t assume "t \ keys q" with \homogeneous q\ have "deg_pm t = poly_deg q" by (rule homogeneousD_poly_deg) moreover from q \t \ keys q\ have "lookup q t = lookup p t" by (rule lookup_hom_components) ultimately show "monomial (lookup p t) (Poly_Mapping.single x (poly_deg p - deg_pm t) + t) = punit.monom_mult 1 (Poly_Mapping.single x (poly_deg p - poly_deg q)) (monomial (lookup q t) t)" by (simp add: punit.monom_mult_monomial) qed also have "\ = punit.monom_mult 1 (Poly_Mapping.single x (poly_deg p - poly_deg q)) q" by (simp only: poly_mapping_sum_monomials flip: punit.monom_mult_sum_right) finally show "(\t\keys q. monomial (lookup p t) (Poly_Mapping.single x (poly_deg p - deg_pm t) + t)) = punit.monom_mult 1 (Poly_Mapping.single x (poly_deg p - poly_deg q)) q" . qed finally show ?thesis . qed lemma keys_homogenizeE: assumes "t \ keys (homogenize x p)" obtains t' where "t' \ keys p" and "t = Poly_Mapping.single x (poly_deg p - deg_pm t') + t'" proof - note assms also have "keys (homogenize x p) \ (\t\keys p. keys (monomial (lookup p t) (Poly_Mapping.single x (poly_deg p - deg_pm t) + t)))" unfolding homogenize_def by (rule keys_sum_subset) finally obtain t' where "t' \ keys p" and "t \ keys (monomial (lookup p t') (Poly_Mapping.single x (poly_deg p - deg_pm t') + t'))" .. from this(2) have "t = Poly_Mapping.single x (poly_deg p - deg_pm t') + t'" by (simp split: if_split_asm) with \t' \ keys p\ show ?thesis .. qed lemma keys_homogenizeE_alt: assumes "t \ keys (homogenize x p)" obtains q t' where "q \ hom_components p" and "t' \ keys q" and "t = Poly_Mapping.single x (poly_deg p - poly_deg q) + t'" proof - note assms also have "keys (homogenize x p) \ (\q\hom_components p. keys (punit.monom_mult 1 (Poly_Mapping.single x (poly_deg p - poly_deg q)) q))" unfolding homogenize_alt by (rule keys_sum_subset) finally obtain q where q: "q \ hom_components p" and "t \ keys (punit.monom_mult 1 (Poly_Mapping.single x (poly_deg p - poly_deg q)) q)" .. note this(2) also have "\ \ (+) (Poly_Mapping.single x (poly_deg p - poly_deg q)) ` keys q" by (rule punit.keys_monom_mult_subset[simplified]) finally obtain t' where "t' \ keys q" and "t = Poly_Mapping.single x (poly_deg p - poly_deg q) + t'" .. with q show ?thesis .. qed lemma deg_pm_homogenize: assumes "t \ keys (homogenize x p)" shows "deg_pm t = poly_deg p" proof - from assms obtain q t' where q: "q \ hom_components p" and "t' \ keys q" and t: "t = Poly_Mapping.single x (poly_deg p - poly_deg q) + t'" by (rule keys_homogenizeE_alt) from q have "homogeneous q" by (rule hom_components_homogeneous) hence "deg_pm t' = poly_deg q" using \t' \ keys q\ by (rule homogeneousD_poly_deg) moreover from q have "poly_deg q \ poly_deg p" by (rule poly_deg_hom_components_le) ultimately show ?thesis by (simp add: t deg_pm_plus deg_pm_single) qed corollary homogeneous_homogenize: "homogeneous (homogenize x p)" proof (rule homogeneousI) fix s t assume "s \ keys (homogenize x p)" hence *: "deg_pm s = poly_deg p" by (rule deg_pm_homogenize) assume "t \ keys (homogenize x p)" hence "deg_pm t = poly_deg p" by (rule deg_pm_homogenize) with * show "deg_pm s = deg_pm t" by simp qed corollary poly_deg_homogenize_le: "poly_deg (homogenize x p) \ poly_deg p" proof (rule poly_deg_leI) fix t assume "t \ keys (homogenize x p)" hence "deg_pm t = poly_deg p" by (rule deg_pm_homogenize) thus "deg_pm t \ poly_deg p" by simp qed lemma homogenize_id_iff [simp]: "homogenize x p = p \ homogeneous p" proof assume "homogenize x p = p" moreover have "homogeneous (homogenize x p)" by (fact homogeneous_homogenize) ultimately show "homogeneous p" by simp next assume "homogeneous p" hence "hom_components p = (if p = 0 then {} else {p})" by (rule hom_components_of_homogeneous) thus "homogenize x p = p" by (simp add: homogenize_alt split: if_split_asm) qed lemma homogenize_homogenize [simp]: "homogenize x (homogenize x p) = homogenize x p" by (simp add: homogeneous_homogenize) lemma homogenize_monomial: "homogenize x (monomial c t) = monomial c t" by (simp only: homogenize_id_iff homogeneous_monomial) lemma indets_homogenize_subset: "indets (homogenize x p) \ insert x (indets p)" proof fix y assume "y \ indets (homogenize x p)" then obtain t where "t \ keys (homogenize x p)" and "y \ keys t" by (rule in_indetsE) from this(1) obtain t' where "t' \ keys p" and t: "t = Poly_Mapping.single x (poly_deg p - deg_pm t') + t'" by (rule keys_homogenizeE) note \y \ keys t\ also have "keys t \ keys (Poly_Mapping.single x (poly_deg p - deg_pm t')) \ keys t'" unfolding t by (rule Poly_Mapping.keys_add) finally show "y \ insert x (indets p)" proof assume "y \ keys (Poly_Mapping.single x (poly_deg p - deg_pm t'))" thus ?thesis by (simp split: if_split_asm) next assume "y \ keys t'" hence "y \ indets p" using \t' \ keys p\ by (rule in_indetsI) thus ?thesis by simp qed qed lemma homogenize_in_Polys: "p \ P[X] \ homogenize x p \ P[insert x X]" using indets_homogenize_subset[of x p] by (auto simp: Polys_alt) lemma lookup_homogenize: assumes "x \ indets p" and "x \ keys t" shows "lookup (homogenize x p) (Poly_Mapping.single x (poly_deg p - deg_pm t) + t) = lookup p t" proof - let ?p = "homogenize x p" let ?t = "Poly_Mapping.single x (poly_deg p - deg_pm t) + t" have eq: "(\s\keys p - {t}. lookup (monomial (lookup p s) (Poly_Mapping.single x (poly_deg p - deg_pm s) + s)) ?t) = 0" proof (intro sum.neutral ballI) fix s assume "s \ keys p - {t}" hence "s \ keys p" and "s \ t" by simp_all from this(1) have "keys s \ indets p" by (simp add: in_indetsI subsetI) with assms(1) have "x \ keys s" by blast have "?t \ Poly_Mapping.single x (poly_deg p - deg_pm s) + s" proof assume a: "?t = Poly_Mapping.single x (poly_deg p - deg_pm s) + s" hence "lookup ?t x = lookup (Poly_Mapping.single x (poly_deg p - deg_pm s) + s) x" by simp moreover from assms(2) have "lookup t x = 0" by (simp add: in_keys_iff) moreover from \x \ keys s\ have "lookup s x = 0" by (simp add: in_keys_iff) ultimately have "poly_deg p - deg_pm t = poly_deg p - deg_pm s" by (simp add: lookup_add) with a have "s = t" by simp with \s \ t\ show False .. qed thus "lookup (monomial (lookup p s) (Poly_Mapping.single x (poly_deg p - deg_pm s) + s)) ?t = 0" by (simp add: lookup_single) qed show ?thesis proof (cases "t \ keys p") case True have "lookup ?p ?t = (\s\keys p. lookup (monomial (lookup p s) (Poly_Mapping.single x (poly_deg p - deg_pm s) + s)) ?t)" by (simp add: homogenize_def lookup_sum) also have "\ = lookup (monomial (lookup p t) ?t) ?t + (\s\keys p - {t}. lookup (monomial (lookup p s) (Poly_Mapping.single x (poly_deg p - deg_pm s) + s)) ?t)" using finite_keys True by (rule sum.remove) also have "\ = lookup p t" by (simp add: eq) finally show ?thesis . next case False hence 1: "keys p - {t} = keys p" by simp have "lookup ?p ?t = (\s\keys p - {t}. lookup (monomial (lookup p s) (Poly_Mapping.single x (poly_deg p - deg_pm s) + s)) ?t)" by (simp add: homogenize_def lookup_sum 1) also have "\ = 0" by (simp only: eq) also from False have "\ = lookup p t" by (simp add: in_keys_iff) finally show ?thesis . qed qed lemma keys_homogenizeI: assumes "x \ indets p" and "t \ keys p" shows "Poly_Mapping.single x (poly_deg p - deg_pm t) + t \ keys (homogenize x p)" (is "?t \ keys ?p") proof - from assms(2) have "keys t \ indets p" by (simp add: in_indetsI subsetI) with assms(1) have "x \ keys t" by blast with assms(1) have "lookup ?p ?t = lookup p t" by (rule lookup_homogenize) also from assms(2) have "\ \ 0" by (simp add: in_keys_iff) finally show ?thesis by (simp add: in_keys_iff) qed lemma keys_homogenize: "x \ indets p \ keys (homogenize x p) = (\t. Poly_Mapping.single x (poly_deg p - deg_pm t) + t) ` keys p" by (auto intro: keys_homogenizeI elim: keys_homogenizeE) lemma card_keys_homogenize: assumes "x \ indets p" shows "card (keys (homogenize x p)) = card (keys p)" unfolding keys_homogenize[OF assms] proof (intro card_image inj_onI) fix s t assume "s \ keys p" and "t \ keys p" with assms have "x \ keys s" and "x \ keys t" by (auto dest: in_indetsI simp only:) let ?s = "Poly_Mapping.single x (poly_deg p - deg_pm s)" let ?t = "Poly_Mapping.single x (poly_deg p - deg_pm t)" assume "?s + s = ?t + t" hence "lookup (?s + s) x = lookup (?t + t) x" by simp with \x \ keys s\ \x \ keys t\ have "?s = ?t" by (simp add: lookup_add in_keys_iff) with \?s + s = ?t + t\ show "s = t" by simp qed lemma poly_deg_homogenize: assumes "x \ indets p" shows "poly_deg (homogenize x p) = poly_deg p" proof (cases "p = 0") case True thus ?thesis by simp next case False then obtain t where "t \ keys p" and 1: "poly_deg p = deg_pm t" by (rule poly_degE) from assms this(1) have "Poly_Mapping.single x (poly_deg p - deg_pm t) + t \ keys (homogenize x p)" by (rule keys_homogenizeI) hence "t \ keys (homogenize x p)" by (simp add: 1) hence "poly_deg p \ poly_deg (homogenize x p)" unfolding 1 by (rule poly_deg_max_keys) with poly_deg_homogenize_le show ?thesis by (rule antisym) qed lemma maxdeg_homogenize: - assumes "x \ UNION F indets" + assumes "x \ \ (indets ` F)" shows "maxdeg (homogenize x ` F) = maxdeg F" unfolding maxdeg_def image_image proof (rule arg_cong[where f=Max], rule set_eqI) fix d show "d \ (\f. poly_deg (homogenize x f)) ` F \ d \ poly_deg ` F" proof assume "d \ (\f. poly_deg (homogenize x f)) ` F" then obtain f where "f \ F" and d: "d = poly_deg (homogenize x f)" .. from assms this(1) have "x \ indets f" by blast hence "d = poly_deg f" by (simp add: d poly_deg_homogenize) with \f \ F\ show "d \ poly_deg ` F" by (rule rev_image_eqI) next assume "d \ poly_deg ` F" then obtain f where "f \ F" and d: "d = poly_deg f" .. from assms this(1) have "x \ indets f" by blast hence "d = poly_deg (homogenize x f)" by (simp add: d poly_deg_homogenize) with \f \ F\ show "d \ (\f. poly_deg (homogenize x f)) ` F" by (rule rev_image_eqI) qed qed lemma homogeneous_ideal_homogenize: assumes "\f. f \ F \ homogeneous f" and "p \ ideal F" shows "homogenize x p \ ideal F" proof - have "homogenize x p = (\q\hom_components p. punit.monom_mult 1 (Poly_Mapping.single x (poly_deg p - poly_deg q)) q)" by (fact homogenize_alt) also have "\ \ ideal F" proof (rule ideal.span_sum) fix q assume "q \ hom_components p" with assms have "q \ ideal F" by (rule homogeneous_ideal') thus "punit.monom_mult 1 (Poly_Mapping.single x (poly_deg p - poly_deg q)) q \ ideal F" by (rule punit.pmdl_closed_monom_mult[simplified]) qed finally show ?thesis . qed lemma subst_pp_dehomo_subst [simp]: "subst_pp (dehomo_subst x) t = monomial (1::'b::comm_semiring_1) (except t {x})" proof - have "subst_pp (dehomo_subst x) t = ((\y\keys t. dehomo_subst x y ^ lookup t y)::_ \\<^sub>0 'b)" by (fact subst_pp_def) also have "\ = (\y\keys t - {y0. dehomo_subst x y0 ^ lookup t y0 = (1::_ \\<^sub>0 'b)}. dehomo_subst x y ^ lookup t y)" by (rule sym, rule prod.setdiff_irrelevant, fact finite_keys) also have "\ = (\y\keys t - {x}. monomial 1 (Poly_Mapping.single y 1) ^ lookup t y)" proof (rule prod.cong) have "dehomo_subst x x ^ lookup t x = 1" by (simp add: dehomo_subst_def) moreover { fix y assume "y \ x" hence "dehomo_subst x y ^ lookup t y = monomial 1 (Poly_Mapping.single y (lookup t y))" by (simp add: dehomo_subst_def monomial_single_power) moreover assume "dehomo_subst x y ^ lookup t y = 1" ultimately have "Poly_Mapping.single y (lookup t y) = 0" by (smt single_one monomial_inj zero_neq_one) hence "lookup t y = 0" by (rule monomial_0D) moreover assume "y \ keys t" ultimately have False by (simp add: in_keys_iff) } ultimately show "keys t - {y0. dehomo_subst x y0 ^ lookup t y0 = 1} = keys t - {x}" by auto qed (simp add: dehomo_subst_def) also have "\ = (\y\keys t - {x}. monomial 1 (Poly_Mapping.single y (lookup t y)))" by (simp add: monomial_single_power) also have "\ = monomial 1 (\y\keys t - {x}. Poly_Mapping.single y (lookup t y))" by (simp flip: punit.monomial_prod_sum) also have "(\y\keys t - {x}. Poly_Mapping.single y (lookup t y)) = except t {x}" proof (rule poly_mapping_eqI, simp add: lookup_sum lookup_except lookup_single, rule) fix y assume "y \ x" show "(\z\keys t - {x}. lookup t z when z = y) = lookup t y" proof (cases "y \ keys t") case True have "finite (keys t - {x})" by simp moreover from True \y \ x\ have "y \ keys t - {x}" by simp ultimately have "(\z\keys t - {x}. lookup t z when z = y) = (lookup t y when y = y) + (\z\keys t - {x} - {y}. lookup t z when z = y)" by (rule sum.remove) also have "(\z\keys t - {x} - {y}. lookup t z when z = y) = 0" by auto finally show ?thesis by simp next case False hence "(\z\keys t - {x}. lookup t z when z = y) = 0" by (auto simp: when_def) also from False have "\ = lookup t y" by (simp add: in_keys_iff) finally show ?thesis . qed qed finally show ?thesis . qed lemma shows dehomogenize_zero [simp]: "dehomogenize x 0 = 0" and dehomogenize_one [simp]: "dehomogenize x 1 = 1" and dehomogenize_monomial: "dehomogenize x (monomial c t) = monomial c (except t {x})" and dehomogenize_plus: "dehomogenize x (p + q) = dehomogenize x p + dehomogenize x q" and dehomogenize_uminus: "dehomogenize x (- r) = - dehomogenize x (r::_ \\<^sub>0 _::comm_ring_1)" and dehomogenize_minus: "dehomogenize x (r - r') = dehomogenize x r - dehomogenize x r'" and dehomogenize_times: "dehomogenize x (p * q) = dehomogenize x p * dehomogenize x q" and dehomogenize_power: "dehomogenize x (p ^ n) = dehomogenize x p ^ n" and dehomogenize_sum: "dehomogenize x (sum f A) = (\a\A. dehomogenize x (f a))" and dehomogenize_prod: "dehomogenize x (prod f A) = (\a\A. dehomogenize x (f a))" by (simp_all add: dehomogenize_def poly_subst_monomial poly_subst_plus poly_subst_uminus poly_subst_minus poly_subst_times poly_subst_power poly_subst_sum poly_subst_prod punit.monom_mult_monomial) corollary dehomogenize_monom_mult: "dehomogenize x (punit.monom_mult c t p) = punit.monom_mult c (except t {x}) (dehomogenize x p)" by (simp only: times_monomial_left[symmetric] dehomogenize_times dehomogenize_monomial) lemma poly_deg_dehomogenize_le: "poly_deg (dehomogenize x p) \ poly_deg p" unfolding dehomogenize_def dehomo_subst_def by (rule poly_deg_poly_subst_le) (simp add: poly_deg_monomial deg_pm_single) lemma indets_dehomogenize: "indets (dehomogenize x p) \ indets p - {x}" for p::"('x \\<^sub>0 nat) \\<^sub>0 'a::comm_semiring_1" proof fix y::'x assume "y \ indets (dehomogenize x p)" then obtain y' where "y' \ indets p" and "y \ indets ((dehomo_subst x y')::_ \\<^sub>0 'a)" unfolding dehomogenize_def by (rule in_indets_poly_substE) from this(2) have "y = y'" and "y' \ x" by (simp_all add: dehomo_subst_def indets_monomial split: if_split_asm) with \y' \ indets p\ show "y \ indets p - {x}" by simp qed lemma dehomogenize_id_iff [simp]: "dehomogenize x p = p \ x \ indets p" proof assume eq: "dehomogenize x p = p" from indets_dehomogenize[of x p] show "x \ indets p" by (auto simp: eq) next assume a: "x \ indets p" show "dehomogenize x p = p" unfolding dehomogenize_def proof (rule poly_subst_id) fix y assume "y \ indets p" with a have "y \ x" by blast thus "dehomo_subst x y = monomial 1 (Poly_Mapping.single y 1)" by (simp add: dehomo_subst_def) qed qed lemma dehomogenize_dehomogenize [simp]: "dehomogenize x (dehomogenize x p) = dehomogenize x p" proof - from indets_dehomogenize[of x p] have "x \ indets (dehomogenize x p)" by blast thus ?thesis by simp qed lemma dehomogenize_homogenize [simp]: "dehomogenize x (homogenize x p) = dehomogenize x p" proof - have "dehomogenize x (homogenize x p) = sum (dehomogenize x) (hom_components p)" by (simp add: homogenize_alt dehomogenize_sum dehomogenize_monom_mult except_single) also have "\ = dehomogenize x p" by (simp only: sum_hom_components flip: dehomogenize_sum) finally show ?thesis . qed corollary dehomogenize_homogenize_id: "x \ indets p \ dehomogenize x (homogenize x p) = p" by simp lemma range_dehomogenize: "range (dehomogenize x) = (P[- {x}] :: (_ \\<^sub>0 'a::comm_semiring_1) set)" proof (intro subset_antisym subsetI PolysI_alt range_eqI) fix p::"_ \\<^sub>0 'a" and y assume "p \ range (dehomogenize x)" then obtain q where p: "p = dehomogenize x q" .. assume "y \ indets p" hence "y \ indets (dehomogenize x q)" by (simp only: p) with indets_dehomogenize have "y \ indets q - {x}" .. thus "y \ - {x}" by simp next fix p::"_ \\<^sub>0 'a" assume "p \ P[- {x}]" hence "x \ indets p" by (auto dest: PolysD) thus "p = dehomogenize x (homogenize x p)" by (rule dehomogenize_homogenize_id[symmetric]) qed lemma dehomogenize_alt: "dehomogenize x p = (\t\keys p. monomial (lookup p t) (except t {x}))" proof - have "dehomogenize x p = dehomogenize x (\t\keys p. monomial (lookup p t) t)" by (simp only: poly_mapping_sum_monomials) also have "\ = (\t\keys p. monomial (lookup p t) (except t {x}))" by (simp only: dehomogenize_sum dehomogenize_monomial) finally show ?thesis . qed lemma keys_dehomogenizeE: assumes "t \ keys (dehomogenize x p)" obtains s where "s \ keys p" and "t = except s {x}" proof - note assms also have "keys (dehomogenize x p) \ (\s\keys p. keys (monomial (lookup p s) (except s {x})))" unfolding dehomogenize_alt by (rule keys_sum_subset) finally obtain s where "s \ keys p" and "t \ keys (monomial (lookup p s) (except s {x}))" .. from this(2) have "t = except s {x}" by (simp split: if_split_asm) with \s \ keys p\ show ?thesis .. qed lemma except_inj_on_keys_homogeneous: assumes "homogeneous p" shows "inj_on (\t. except t {x}) (keys p)" proof fix s t assume "s \ keys p" and "t \ keys p" from assms this(1) have "deg_pm s = poly_deg p" by (rule homogeneousD_poly_deg) moreover from assms \t \ keys p\ have "deg_pm t = poly_deg p" by (rule homogeneousD_poly_deg) ultimately have "deg_pm (Poly_Mapping.single x (lookup s x) + except s {x}) = deg_pm (Poly_Mapping.single x (lookup t x) + except t {x})" by (simp only: flip: plus_except) moreover assume 1: "except s {x} = except t {x}" ultimately have 2: "lookup s x = lookup t x" by (simp only: deg_pm_plus deg_pm_single) show "s = t" proof (rule poly_mapping_eqI) fix y show "lookup s y = lookup t y" proof (cases "y = x") case True with 2 show ?thesis by simp next case False hence "lookup s y = lookup (except s {x}) y" and "lookup t y = lookup (except t {x}) y" by (simp_all add: lookup_except) with 1 show ?thesis by simp qed qed qed lemma lookup_dehomogenize: assumes "homogeneous p" and "t \ keys p" shows "lookup (dehomogenize x p) (except t {x}) = lookup p t" proof - let ?t = "except t {x}" have eq: "(\s\keys p - {t}. lookup (monomial (lookup p s) (except s {x})) ?t) = 0" proof (intro sum.neutral ballI) fix s assume "s \ keys p - {t}" hence "s \ keys p" and "s \ t" by simp_all have "?t \ except s {x}" proof from assms(1) have "inj_on (\t. except t {x}) (keys p)" by (rule except_inj_on_keys_homogeneous) moreover assume "?t = except s {x}" ultimately have "t = s" using assms(2) \s \ keys p\ by (rule inj_onD) with \s \ t\ show False by simp qed thus "lookup (monomial (lookup p s) (except s {x})) ?t = 0" by (simp add: lookup_single) qed have "lookup (dehomogenize x p) ?t = (\s\keys p. lookup (monomial (lookup p s) (except s {x})) ?t)" by (simp only: dehomogenize_alt lookup_sum) also have "\ = lookup (monomial (lookup p t) ?t) ?t + (\s\keys p - {t}. lookup (monomial (lookup p s) (except s {x})) ?t)" using finite_keys assms(2) by (rule sum.remove) also have "\ = lookup p t" by (simp add: eq) finally show ?thesis . qed lemma keys_dehomogenizeI: assumes "homogeneous p" and "t \ keys p" shows "except t {x} \ keys (dehomogenize x p)" proof - from assms have "lookup (dehomogenize x p) (except t {x}) = lookup p t" by (rule lookup_dehomogenize) also from assms(2) have "\ \ 0" by (simp add: in_keys_iff) finally show ?thesis by (simp add: in_keys_iff) qed lemma homogeneous_homogenize_dehomogenize: assumes "homogeneous p" obtains d where "d = poly_deg p - poly_deg (homogenize x (dehomogenize x p))" and "punit.monom_mult 1 (Poly_Mapping.single x d) (homogenize x (dehomogenize x p)) = p" proof (cases "p = 0") case True hence "0 = poly_deg p - poly_deg (homogenize x (dehomogenize x p))" and "punit.monom_mult 1 (Poly_Mapping.single x 0) (homogenize x (dehomogenize x p)) = p" by simp_all thus ?thesis .. next case False let ?q = "dehomogenize x p" let ?p = "homogenize x ?q" define d where "d = poly_deg p - poly_deg ?p" show ?thesis proof have "punit.monom_mult 1 (Poly_Mapping.single x d) ?p = (\t\keys ?q. monomial (lookup ?q t) (Poly_Mapping.single x (d + (poly_deg ?q - deg_pm t)) + t))" by (simp add: homogenize_def punit.monom_mult_sum_right punit.monom_mult_monomial flip: add.assoc single_add) also have "\ = (\t\keys ?q. monomial (lookup ?q t) (Poly_Mapping.single x (poly_deg p - deg_pm t) + t))" using refl proof (rule sum.cong) fix t assume "t \ keys ?q" have "poly_deg ?p = poly_deg ?q" proof (rule poly_deg_homogenize) from indets_dehomogenize show "x \ indets ?q" by fastforce qed hence d: "d = poly_deg p - poly_deg ?q" by (simp only: d_def) thm poly_deg_dehomogenize_le from \t \ keys ?q\ have "d + (poly_deg ?q - deg_pm t) = (d + poly_deg ?q) - deg_pm t" by (intro add_diff_assoc poly_deg_max_keys) also have "d + poly_deg ?q = poly_deg p" by (simp add: d poly_deg_dehomogenize_le) finally show "monomial (lookup ?q t) (Poly_Mapping.single x (d + (poly_deg ?q - deg_pm t)) + t) = monomial (lookup ?q t) (Poly_Mapping.single x (poly_deg p - deg_pm t) + t)" by (simp only:) qed also have "\ = (\t\(\s. except s {x}) ` keys p. monomial (lookup ?q t) (Poly_Mapping.single x (poly_deg p - deg_pm t) + t))" proof (rule sum.mono_neutral_left) show "keys (dehomogenize x p) \ (\s. except s {x}) ` keys p" proof fix t assume "t \ keys (dehomogenize x p)" then obtain s where "s \ keys p" and "t = except s {x}" by (rule keys_dehomogenizeE) thus "t \ (\s. except s {x}) ` keys p" by (rule rev_image_eqI) qed qed (simp_all add: in_keys_iff) also from assms have "\ = (\t\keys p. monomial (lookup ?q (except t {x})) (Poly_Mapping.single x (poly_deg p - deg_pm (except t {x})) + except t {x}))" by (intro sum.reindex[unfolded comp_def] except_inj_on_keys_homogeneous) also from refl have "\ = (\t\keys p. monomial (lookup p t) t)" proof (rule sum.cong) fix t assume "t \ keys p" with assms have "lookup ?q (except t {x}) = lookup p t" by (rule lookup_dehomogenize) moreover have "Poly_Mapping.single x (poly_deg p - deg_pm (except t {x})) + except t {x} = t" (is "?l = _") proof (rule poly_mapping_eqI) fix y show "lookup ?l y = lookup t y" proof (cases "y = x") case True from assms \t \ keys p\ have "deg_pm t = poly_deg p" by (rule homogeneousD_poly_deg) also have "deg_pm t = deg_pm (Poly_Mapping.single x (lookup t x) + except t {x})" by (simp flip: plus_except) also have "\ = lookup t x + deg_pm (except t {x})" by (simp only: deg_pm_plus deg_pm_single) finally have "poly_deg p - deg_pm (except t {x}) = lookup t x" by simp thus ?thesis by (simp add: True lookup_add lookup_except lookup_single) next case False thus ?thesis by (simp add: lookup_add lookup_except lookup_single) qed qed ultimately show "monomial (lookup ?q (except t {x})) (Poly_Mapping.single x (poly_deg p - deg_pm (except t {x})) + except t {x}) = monomial (lookup p t) t" by (simp only:) qed also have "\ = p" by (fact poly_mapping_sum_monomials) finally show "punit.monom_mult 1 (Poly_Mapping.single x d) ?p = p" . qed (simp only: d_def) qed lemma dehomogenize_zeroD: assumes "dehomogenize x p = 0" and "homogeneous p" shows "p = 0" proof - from assms(2) obtain d where "punit.monom_mult 1 (Poly_Mapping.single x d) (homogenize x (dehomogenize x p)) = p" by (rule homogeneous_homogenize_dehomogenize) thus ?thesis by (simp add: assms(1)) qed lemma dehomogenize_ideal: "dehomogenize x ` ideal F = ideal (dehomogenize x ` F) \ P[- {x}]" unfolding range_dehomogenize[symmetric] using dehomogenize_plus dehomogenize_times dehomogenize_dehomogenize by (rule image_ideal_eq_Int) corollary dehomogenize_ideal_subset: "dehomogenize x ` ideal F \ ideal (dehomogenize x ` F)" by (simp add: dehomogenize_ideal) lemma ideal_dehomogenize: assumes "ideal G = ideal (homogenize x ` F)" and "F \ P[UNIV - {x}]" shows "ideal (dehomogenize x ` G) = ideal F" proof - have eq: "dehomogenize x (homogenize x f) = f" if "f \ F" for f proof (rule dehomogenize_homogenize_id) from that assms(2) have "f \ P[UNIV - {x}]" .. thus "x \ indets f" by (auto simp: Polys_alt) qed show ?thesis proof (intro Set.equalityI ideal.span_subset_spanI) show "dehomogenize x ` G \ ideal F" proof fix q assume "q \ dehomogenize x ` G" then obtain g where "g \ G" and q: "q = dehomogenize x g" .. from this(1) have "g \ ideal G" by (rule ideal.span_base) also have "\ = ideal (homogenize x ` F)" by fact finally have "q \ dehomogenize x ` ideal (homogenize x ` F)" using q by (rule rev_image_eqI) also have "\ \ ideal (dehomogenize x ` homogenize x ` F)" by (rule dehomogenize_ideal_subset) also have "dehomogenize x ` homogenize x ` F = F" by (auto simp: eq image_image simp del: dehomogenize_homogenize intro!: image_eqI) finally show "q \ ideal F" . qed next show "F \ ideal (dehomogenize x ` G)" proof fix f assume "f \ F" hence "homogenize x f \ homogenize x ` F" by (rule imageI) also have "\ \ ideal (homogenize x ` F)" by (rule ideal.span_superset) also from assms(1) have "\ = ideal G" by (rule sym) finally have "dehomogenize x (homogenize x f) \ dehomogenize x ` ideal G" by (rule imageI) with \f \ F\ have "f \ dehomogenize x ` ideal G" by (simp only: eq) also have "\ \ ideal (dehomogenize x ` G)" by (rule dehomogenize_ideal_subset) finally show "f \ ideal (dehomogenize x ` G)" . qed qed qed subsection \Embedding Polynomial Rings in Larger Polynomial Rings (With One Additional Indeterminate)\ text \We define a homomorphism for embedding a polynomial ring in a larger polynomial ring, and its inverse. This is mainly needed for homogenizing wrt. a fresh indeterminate.\ definition extend_indets_subst :: "'x \ ('x option \\<^sub>0 nat) \\<^sub>0 'a::comm_semiring_1" where "extend_indets_subst x = monomial 1 (Poly_Mapping.single (Some x) 1)" definition extend_indets :: "(('x \\<^sub>0 nat) \\<^sub>0 'a) \ ('x option \\<^sub>0 nat) \\<^sub>0 'a::comm_semiring_1" where "extend_indets = poly_subst extend_indets_subst" definition restrict_indets_subst :: "'x option \ 'x \\<^sub>0 nat" where "restrict_indets_subst x = (case x of Some y \ Poly_Mapping.single y 1 | _ \ 0)" definition restrict_indets :: "(('x option \\<^sub>0 nat) \\<^sub>0 'a) \ ('x \\<^sub>0 nat) \\<^sub>0 'a::comm_semiring_1" where "restrict_indets = poly_subst (\x. monomial 1 (restrict_indets_subst x))" definition restrict_indets_pp :: "('x option \\<^sub>0 nat) \ ('x \\<^sub>0 nat)" where "restrict_indets_pp t = (\x\keys t. lookup t x \ restrict_indets_subst x)" lemma lookup_extend_indets_subst_aux: "lookup (\y\keys t. Poly_Mapping.single (Some y) (lookup t y)) = (\x. case x of Some y \ lookup t y | _ \ 0)" proof - have "(\x\keys t. lookup t x when x = y) = lookup t y" for y proof (cases "y \ keys t") case True hence "(\x\keys t. lookup t x when x = y) = (\x\insert y (keys t). lookup t x when x = y)" by (simp only: insert_absorb) also have "\ = lookup t y + (\x\keys t - {y}. lookup t x when x = y)" by (simp add: sum.insert_remove) also have "(\x\keys t - {y}. lookup t x when x = y) = 0" by (auto simp: when_def intro: sum.neutral) finally show ?thesis by simp next case False hence "(\x\keys t. lookup t x when x = y) = 0" by (auto simp: when_def intro: sum.neutral) with False show ?thesis by (simp add: in_keys_iff) qed thus ?thesis by (auto simp: lookup_sum lookup_single split: option.split) qed lemma keys_extend_indets_subst_aux: "keys (\y\keys t. Poly_Mapping.single (Some y) (lookup t y)) = Some ` keys t" by (auto simp: lookup_extend_indets_subst_aux simp flip: lookup_not_eq_zero_eq_in_keys split: option.splits) lemma subst_pp_extend_indets_subst: "subst_pp extend_indets_subst t = monomial 1 (\y\keys t. Poly_Mapping.single (Some y) (lookup t y))" proof - have "subst_pp extend_indets_subst t = monomial (\y\keys t. 1 ^ lookup t y) (\y\keys t. lookup t y \ Poly_Mapping.single (Some y) 1)" by (rule subst_pp_by_monomials) (simp only: extend_indets_subst_def) also have "\ = monomial 1 (\y\keys t. Poly_Mapping.single (Some y) (lookup t y))" by simp finally show ?thesis . qed lemma keys_extend_indets: "keys (extend_indets p) = (\t. \y\keys t. Poly_Mapping.single (Some y) (lookup t y)) ` keys p" proof - have "keys (extend_indets p) = (\t\keys p. keys (punit.monom_mult (lookup p t) 0 (subst_pp extend_indets_subst t)))" unfolding extend_indets_def poly_subst_def using finite_keys proof (rule keys_sum) fix s t :: "'a \\<^sub>0 nat" assume "s \ t" then obtain x where "lookup s x \ lookup t x" by (meson poly_mapping_eqI) have "(\y\keys t. monomial (lookup t y) (Some y)) \ (\y\keys s. monomial (lookup s y) (Some y))" (is "?l \ ?r") proof assume "?l = ?r" hence "lookup ?l (Some x) = lookup ?r (Some x)" by (simp only:) hence "lookup s x = lookup t x" by (simp add: lookup_extend_indets_subst_aux) with \lookup s x \ lookup t x\ show False .. qed thus "keys (punit.monom_mult (lookup p s) 0 (subst_pp extend_indets_subst s)) \ keys (punit.monom_mult (lookup p t) 0 (subst_pp extend_indets_subst t)) = {}" by (simp add: subst_pp_extend_indets_subst punit.monom_mult_monomial) qed also have "\ = (\t. \y\keys t. monomial (lookup t y) (Some y)) ` keys p" by (auto simp: subst_pp_extend_indets_subst punit.monom_mult_monomial split: if_split_asm) finally show ?thesis . qed lemma indets_extend_indets: "indets (extend_indets p) = Some ` indets (p::_ \\<^sub>0 'a::comm_semiring_1)" proof (rule set_eqI) fix x show "x \ indets (extend_indets p) \ x \ Some ` indets p" proof assume "x \ indets (extend_indets p)" then obtain y where "y \ indets p" and "x \ indets (monomial (1::'a) (Poly_Mapping.single (Some y) 1))" unfolding extend_indets_def extend_indets_subst_def by (rule in_indets_poly_substE) from this(2) indets_monomial_single_subset have "x \ {Some y}" .. hence "x = Some y" by simp with \y \ indets p\ show "x \ Some ` indets p" by (rule rev_image_eqI) next assume "x \ Some ` indets p" then obtain y where "y \ indets p" and x: "x = Some y" .. from this(1) obtain t where "t \ keys p" and "y \ keys t" by (rule in_indetsE) from this(2) have "Some y \ keys (\y\keys t. Poly_Mapping.single (Some y) (lookup t y))" unfolding keys_extend_indets_subst_aux by (rule imageI) moreover have "(\y\keys t. Poly_Mapping.single (Some y) (lookup t y)) \ keys (extend_indets p)" unfolding keys_extend_indets using \t \ keys p\ by (rule imageI) ultimately show "x \ indets (extend_indets p)" unfolding x by (rule in_indetsI) qed qed lemma poly_deg_extend_indets [simp]: "poly_deg (extend_indets p) = poly_deg p" proof - have eq: "deg_pm ((\y\keys t. Poly_Mapping.single (Some y) (lookup t y))) = deg_pm t" for t::"'a \\<^sub>0 nat" proof - have "deg_pm ((\y\keys t. Poly_Mapping.single (Some y) (lookup t y))) = (\y\keys t. lookup t y)" by (simp add: deg_pm_sum deg_pm_single) also from subset_refl finite_keys have "\ = deg_pm t" by (rule deg_pm_superset[symmetric]) finally show ?thesis . qed show ?thesis proof (rule antisym) show "poly_deg (extend_indets p) \ poly_deg p" proof (rule poly_deg_leI) fix t assume "t \ keys (extend_indets p)" then obtain s where "s \ keys p" and "t = (\y\keys s. Poly_Mapping.single (Some y) (lookup s y))" unfolding keys_extend_indets .. from this(2) have "deg_pm t = deg_pm s" by (simp only: eq) also from \s \ keys p\ have "\ \ poly_deg p" by (rule poly_deg_max_keys) finally show "deg_pm t \ poly_deg p" . qed next show "poly_deg p \ poly_deg (extend_indets p)" proof (rule poly_deg_leI) fix t assume "t \ keys p" hence *: "(\y\keys t. Poly_Mapping.single (Some y) (lookup t y)) \ keys (extend_indets p)" unfolding keys_extend_indets by (rule imageI) have "deg_pm t = deg_pm (\y\keys t. Poly_Mapping.single (Some y) (lookup t y))" by (simp only: eq) also from * have "\ \ poly_deg (extend_indets p)" by (rule poly_deg_max_keys) finally show "deg_pm t \ poly_deg (extend_indets p)" . qed qed qed lemma shows extend_indets_zero [simp]: "extend_indets 0 = 0" and extend_indets_one [simp]: "extend_indets 1 = 1" and extend_indets_monomial: "extend_indets (monomial c t) = punit.monom_mult c 0 (subst_pp extend_indets_subst t)" and extend_indets_plus: "extend_indets (p + q) = extend_indets p + extend_indets q" and extend_indets_uminus: "extend_indets (- r) = - extend_indets (r::_ \\<^sub>0 _::comm_ring_1)" and extend_indets_minus: "extend_indets (r - r') = extend_indets r - extend_indets r'" and extend_indets_times: "extend_indets (p * q) = extend_indets p * extend_indets q" and extend_indets_power: "extend_indets (p ^ n) = extend_indets p ^ n" and extend_indets_sum: "extend_indets (sum f A) = (\a\A. extend_indets (f a))" and extend_indets_prod: "extend_indets (prod f A) = (\a\A. extend_indets (f a))" by (simp_all add: extend_indets_def poly_subst_monomial poly_subst_plus poly_subst_uminus poly_subst_minus poly_subst_times poly_subst_power poly_subst_sum poly_subst_prod) lemma extend_indets_zero_iff [simp]: "extend_indets p = 0 \ p = 0" by (metis (no_types, lifting) imageE imageI keys_extend_indets lookup_zero not_in_keys_iff_lookup_eq_zero poly_deg_extend_indets poly_deg_zero poly_deg_zero_imp_monomial) lemma extend_indets_inject: assumes "extend_indets p = extend_indets (q::_ \\<^sub>0 _::comm_ring_1)" shows "p = q" proof - from assms have "extend_indets (p - q) = 0" by (simp add: extend_indets_minus) thus ?thesis by simp qed corollary inj_extend_indets: "inj (extend_indets::_ \ _ \\<^sub>0 _::comm_ring_1)" using extend_indets_inject by (intro injI) lemma poly_subst_extend_indets: "poly_subst f (extend_indets p) = poly_subst (f \ Some) p" by (simp add: extend_indets_def poly_subst_poly_subst extend_indets_subst_def poly_subst_monomial subst_pp_single o_def) lemma poly_eval_extend_indets: "poly_eval a (extend_indets p) = poly_eval (a \ Some) p" proof - have eq: "poly_eval a (extend_indets (monomial c t)) = poly_eval (\x. a (Some x)) (monomial c t)" for c t by (simp add: extend_indets_monomial poly_eval_times poly_eval_monomial poly_eval_prod poly_eval_power subst_pp_def extend_indets_subst_def flip: times_monomial_left) show ?thesis by (induct p rule: poly_mapping_plus_induct) (simp_all add: extend_indets_plus poly_eval_plus eq) qed lemma lookup_restrict_indets_pp: "lookup (restrict_indets_pp t) = (\x. lookup t (Some x))" proof - let ?f = "\z x. lookup t x * lookup (case x of None \ 0 | Some y \ Poly_Mapping.single y 1) z" have "sum (?f z) (keys t) = lookup t (Some z)" for z proof (cases "Some z \ keys t") case True hence "sum (?f z) (keys t) = sum (?f z) (insert (Some z) (keys t))" by (simp only: insert_absorb) also have "\ = lookup t (Some z) + sum (?f z) (keys t - {Some z})" by (simp add: sum.insert_remove) also have "sum (?f z) (keys t - {Some z}) = 0" by (auto simp: when_def lookup_single intro: sum.neutral split: option.splits) finally show ?thesis by simp next case False hence "sum (?f z) (keys t) = 0" by (auto simp: when_def lookup_single intro: sum.neutral split: option.splits) with False show ?thesis by (simp add: in_keys_iff) qed thus ?thesis by (auto simp: restrict_indets_pp_def restrict_indets_subst_def lookup_sum) qed lemma keys_restrict_indets_pp: "keys (restrict_indets_pp t) = the ` (keys t - {None})" proof (rule set_eqI) fix x show "x \ keys (restrict_indets_pp t) \ x \ the ` (keys t - {None})" proof assume "x \ keys (restrict_indets_pp t)" hence "Some x \ keys t" by (simp add: lookup_restrict_indets_pp flip: lookup_not_eq_zero_eq_in_keys) hence "Some x \ keys t - {None}" by blast moreover have "x = the (Some x)" by simp ultimately show "x \ the ` (keys t - {None})" by (rule rev_image_eqI) next assume "x \ the ` (keys t - {None})" then obtain y where "y \ keys t - {None}" and "x = the y" .. hence "Some x \ keys t" by auto thus "x \ keys (restrict_indets_pp t)" by (simp add: lookup_restrict_indets_pp flip: lookup_not_eq_zero_eq_in_keys) qed qed lemma subst_pp_restrict_indets_subst: "subst_pp (\x. monomial 1 (restrict_indets_subst x)) t = monomial 1 (restrict_indets_pp t)" by (simp add: subst_pp_def monomial_power_map_scale restrict_indets_pp_def flip: punit.monomial_prod_sum) lemma restrict_indets_pp_zero [simp]: "restrict_indets_pp 0 = 0" by (simp add: restrict_indets_pp_def) lemma restrict_indets_pp_plus: "restrict_indets_pp (s + t) = restrict_indets_pp s + restrict_indets_pp t" by (rule poly_mapping_eqI) (simp add: lookup_add lookup_restrict_indets_pp) lemma restrict_indets_pp_except_None [simp]: "restrict_indets_pp (except t {None}) = restrict_indets_pp t" by (rule poly_mapping_eqI) (simp add: lookup_add lookup_restrict_indets_pp lookup_except) lemma deg_pm_restrict_indets_pp: "deg_pm (restrict_indets_pp t) + lookup t None = deg_pm t" proof - have "deg_pm t = sum (lookup t) (insert None (keys t))" by (rule deg_pm_superset) auto also from finite_keys have "\ = lookup t None + sum (lookup t) (keys t - {None})" by (rule sum.insert_remove) also have "sum (lookup t) (keys t - {None}) = (\x\keys t. lookup t x * deg_pm (restrict_indets_subst x))" by (intro sum.mono_neutral_cong_left) (auto simp: restrict_indets_subst_def deg_pm_single) also have "\ = deg_pm (restrict_indets_pp t)" by (simp only: restrict_indets_pp_def deg_pm_sum deg_pm_map_scale) finally show ?thesis by simp qed lemma keys_restrict_indets_subset: "keys (restrict_indets p) \ restrict_indets_pp ` keys p" proof fix t assume "t \ keys (restrict_indets p)" also have "\ = keys (\t\keys p. monomial (lookup p t) (restrict_indets_pp t))" by (simp add: restrict_indets_def poly_subst_def subst_pp_restrict_indets_subst punit.monom_mult_monomial) also have "\ \ (\t\keys p. keys (monomial (lookup p t) (restrict_indets_pp t)))" by (rule keys_sum_subset) also have "\ = restrict_indets_pp ` keys p" by (auto split: if_split_asm) finally show "t \ restrict_indets_pp ` keys p" . qed lemma keys_restrict_indets: assumes "None \ indets p" shows "keys (restrict_indets p) = restrict_indets_pp ` keys p" proof - have "keys (restrict_indets p) = keys (\t\keys p. monomial (lookup p t) (restrict_indets_pp t))" by (simp add: restrict_indets_def poly_subst_def subst_pp_restrict_indets_subst punit.monom_mult_monomial) also from finite_keys have "\ = (\t\keys p. keys (monomial (lookup p t) (restrict_indets_pp t)))" proof (rule keys_sum) fix s t assume "s \ keys p" hence "keys s \ indets p" by (rule keys_subset_indets) with assms have "None \ keys s" by blast assume "t \ keys p" hence "keys t \ indets p" by (rule keys_subset_indets) with assms have "None \ keys t" by blast assume "s \ t" then obtain x where neq: "lookup s x \ lookup t x" by (meson poly_mapping_eqI) have "x \ None" proof assume "x = None" with \None \ keys s\ and \None \ keys t\ have "x \ keys s" and "x \ keys t" by blast+ with neq show False by (simp add: in_keys_iff) qed then obtain y where x: "x = Some y" by blast have "restrict_indets_pp t \ restrict_indets_pp s" proof assume "restrict_indets_pp t = restrict_indets_pp s" hence "lookup (restrict_indets_pp t) y = lookup (restrict_indets_pp s) y" by (simp only:) hence "lookup s x = lookup t x" by (simp add: x lookup_restrict_indets_pp) with neq show False .. qed thus "keys (monomial (lookup p s) (restrict_indets_pp s)) \ keys (monomial (lookup p t) (restrict_indets_pp t)) = {}" by (simp add: subst_pp_extend_indets_subst) qed also have "\ = restrict_indets_pp ` keys p" by (auto split: if_split_asm) finally show ?thesis . qed lemma indets_restrict_indets_subset: "indets (restrict_indets p) \ the ` (indets p - {None})" proof fix x assume "x \ indets (restrict_indets p)" then obtain t where "t \ keys (restrict_indets p)" and "x \ keys t" by (rule in_indetsE) from this(1) keys_restrict_indets_subset have "t \ restrict_indets_pp ` keys p" .. then obtain s where "s \ keys p" and "t = restrict_indets_pp s" .. from \x \ keys t\ this(2) have "x \ the ` (keys s - {None})" by (simp only: keys_restrict_indets_pp) also from \s \ keys p\ have "\ \ the ` (indets p - {None})" by (intro image_mono Diff_mono keys_subset_indets subset_refl) finally show "x \ the ` (indets p - {None})" . qed lemma poly_deg_restrict_indets_le: "poly_deg (restrict_indets p) \ poly_deg p" proof (rule poly_deg_leI) fix t assume "t \ keys (restrict_indets p)" hence "t \ restrict_indets_pp ` keys p" using keys_restrict_indets_subset .. then obtain s where "s \ keys p" and "t = restrict_indets_pp s" .. from this(2) have "deg_pm t + lookup s None = deg_pm s" by (simp only: deg_pm_restrict_indets_pp) also from \s \ keys p\ have "\ \ poly_deg p" by (rule poly_deg_max_keys) finally show "deg_pm t \ poly_deg p" by simp qed lemma shows restrict_indets_zero [simp]: "restrict_indets 0 = 0" and restrict_indets_one [simp]: "restrict_indets 1 = 1" and restrict_indets_monomial: "restrict_indets (monomial c t) = monomial c (restrict_indets_pp t)" and restrict_indets_plus: "restrict_indets (p + q) = restrict_indets p + restrict_indets q" and restrict_indets_uminus: "restrict_indets (- r) = - restrict_indets (r::_ \\<^sub>0 _::comm_ring_1)" and restrict_indets_minus: "restrict_indets (r - r') = restrict_indets r - restrict_indets r'" and restrict_indets_times: "restrict_indets (p * q) = restrict_indets p * restrict_indets q" and restrict_indets_power: "restrict_indets (p ^ n) = restrict_indets p ^ n" and restrict_indets_sum: "restrict_indets (sum f A) = (\a\A. restrict_indets (f a))" and restrict_indets_prod: "restrict_indets (prod f A) = (\a\A. restrict_indets (f a))" by (simp_all add: restrict_indets_def poly_subst_monomial poly_subst_plus poly_subst_uminus poly_subst_minus poly_subst_times poly_subst_power poly_subst_sum poly_subst_prod subst_pp_restrict_indets_subst punit.monom_mult_monomial) lemma restrict_extend_indets [simp]: "restrict_indets (extend_indets p) = p" unfolding extend_indets_def restrict_indets_def poly_subst_poly_subst by (rule poly_subst_id) (simp add: extend_indets_subst_def restrict_indets_subst_def poly_subst_monomial subst_pp_single) lemma extend_restrict_indets: assumes "None \ indets p" shows "extend_indets (restrict_indets p) = p" unfolding extend_indets_def restrict_indets_def poly_subst_poly_subst proof (rule poly_subst_id) fix x assume "x \ indets p" with assms have "x \ None" by meson then obtain y where x: "x = Some y" by blast thus "poly_subst extend_indets_subst (monomial 1 (restrict_indets_subst x)) = monomial 1 (Poly_Mapping.single x 1)" by (simp add: extend_indets_subst_def restrict_indets_subst_def poly_subst_monomial subst_pp_single) qed lemma restrict_indets_dehomogenize [simp]: "restrict_indets (dehomogenize None p) = restrict_indets p" proof - have eq: "poly_subst (\x. (monomial 1 (restrict_indets_subst x))) (dehomo_subst None y) = monomial 1 (restrict_indets_subst y)" for y::"'x option" by (auto simp: restrict_indets_subst_def dehomo_subst_def poly_subst_monomial subst_pp_single) show ?thesis by (simp only: dehomogenize_def restrict_indets_def poly_subst_poly_subst eq) qed corollary restrict_indets_comp_dehomogenize: "restrict_indets \ dehomogenize None = restrict_indets" by (rule ext) (simp only: o_def restrict_indets_dehomogenize) corollary extend_restrict_indets_eq_dehomogenize: "extend_indets (restrict_indets p) = dehomogenize None p" proof - have "extend_indets (restrict_indets p) = extend_indets (restrict_indets (dehomogenize None p))" by simp also have "\ = dehomogenize None p" proof (intro extend_restrict_indets notI) assume "None \ indets (dehomogenize None p)" hence "None \ indets p - {None}" using indets_dehomogenize .. thus False by simp qed finally show ?thesis . qed corollary extend_indets_comp_restrict_indets: "extend_indets \ restrict_indets = dehomogenize None" by (rule ext) (simp only: o_def extend_restrict_indets_eq_dehomogenize) lemma restrict_homogenize_extend_indets [simp]: "restrict_indets (homogenize None (extend_indets p)) = p" proof - have "restrict_indets (homogenize None (extend_indets p)) = restrict_indets (dehomogenize None (homogenize None (extend_indets p)))" by (simp only: restrict_indets_dehomogenize) also have "\ = restrict_indets (dehomogenize None (extend_indets p))" by (simp only: dehomogenize_homogenize) also have "\ = p" by simp finally show ?thesis . qed lemma dehomogenize_extend_indets [simp]: "dehomogenize None (extend_indets p) = extend_indets p" by (simp add: indets_extend_indets) lemma restrict_indets_ideal: "restrict_indets ` ideal F = ideal (restrict_indets ` F)" using restrict_indets_plus restrict_indets_times proof (rule image_ideal_eq_surj) from restrict_extend_indets show "surj restrict_indets" by (rule surjI) qed lemma ideal_restrict_indets: "ideal G = ideal (homogenize None ` extend_indets ` F) \ ideal (restrict_indets ` G) = ideal F" by (simp flip: restrict_indets_ideal) (simp add: restrict_indets_ideal image_image) lemma extend_indets_ideal: "extend_indets ` ideal F = ideal (extend_indets ` F) \ P[- {None}]" proof - have "extend_indets ` ideal F = extend_indets ` restrict_indets ` ideal (extend_indets ` F)" by (simp add: restrict_indets_ideal image_image) also have "\ = ideal (extend_indets ` F) \ P[- {None}]" by (simp add: extend_restrict_indets_eq_dehomogenize dehomogenize_ideal image_image) finally show ?thesis . qed corollary extend_indets_ideal_subset: "extend_indets ` ideal F \ ideal (extend_indets ` F)" by (simp add: extend_indets_ideal) subsection \Canonical Isomorphisms between \P[X,Y]\ and \P[X][Y]\: \focus\ and \flatten\\ definition focus :: "'x set \ (('x \\<^sub>0 nat) \\<^sub>0 'a) \ (('x \\<^sub>0 nat) \\<^sub>0 ('x \\<^sub>0 nat) \\<^sub>0 'a::comm_monoid_add)" where "focus X p = (\t\keys p. monomial (monomial (lookup p t) (except t X)) (except t (- X)))" definition flatten :: "('a \\<^sub>0 'a \\<^sub>0 'b) \ ('a::comm_powerprod \\<^sub>0 'b::semiring_1)" where "flatten p = (\t\keys p. punit.monom_mult 1 t (lookup p t))" lemma focus_superset: assumes "finite A" and "keys p \ A" shows "focus X p = (\t\A. monomial (monomial (lookup p t) (except t X)) (except t (- X)))" unfolding focus_def using assms by (rule sum.mono_neutral_left) (simp add: in_keys_iff) lemma keys_focus: "keys (focus X p) = (\t. except t (- X)) ` keys p" proof have "keys (focus X p) \ (\t\keys p. keys (monomial (monomial (lookup p t) (except t X)) (except t (- X))))" unfolding focus_def by (rule keys_sum_subset) also have "\ \ (\t\keys p. {except t (- X)})" by (intro UN_mono subset_refl) simp also have "\ = (\t. except t (- X)) ` keys p" by (rule UNION_singleton_eq_range) finally show "keys (focus X p) \ (\t. except t (- X)) ` keys p" . next { fix s assume "s \ keys p" have "lookup (focus X p) (except s (- X)) = (\t\keys p. monomial (lookup p t) (except t X) when except t (- X) = except s (- X))" (is "_ = ?p") by (simp only: focus_def lookup_sum lookup_single) also have "\ \ 0" proof have "lookup ?p (except s X) = (\t\keys p. lookup p t when except t X = except s X \ except t (- X) = except s (- X))" by (simp add: lookup_sum lookup_single when_def if_distrib if_distribR) (metis (no_types, hide_lams) lookup_single_eq lookup_single_not_eq lookup_zero) also have "\ = (\t\{s}. lookup p t)" proof (intro sum.mono_neutral_cong_right ballI) fix t assume "t \ keys p - {s}" hence "t \ s" by simp hence "except t X + except t (- X) \ except s X + except s (- X)" by (simp flip: except_decomp) thus "(lookup p t when except t X = except s X \ except t (- X) = except s (- X)) = 0" by (auto simp: when_def) next from \s \ keys p\ show "{s} \ keys p" by simp qed simp_all also from \s \ keys p\ have "\ \ 0" by (simp add: in_keys_iff) finally have "except s X \ keys ?p" by (simp add: in_keys_iff) moreover assume "?p = 0" ultimately show False by simp qed finally have "except s (- X) \ keys (focus X p)" by (simp add: in_keys_iff) } thus "(\t. except t (- X)) ` keys p \ keys (focus X p)" by blast qed lemma keys_coeffs_focus_subset: assumes "c \ range (lookup (focus X p))" shows "keys c \ (\t. except t X) ` keys p" proof - from assms obtain s where "c = lookup (focus X p) s" .. hence "keys c = keys (lookup (focus X p) s)" by (simp only:) also have "\ \ (\t\keys p. keys (lookup (monomial (monomial (lookup p t) (except t X)) (except t (- X))) s))" unfolding focus_def lookup_sum by (rule keys_sum_subset) also from subset_refl have "\ \ (\t\keys p. {except t X})" by (rule UN_mono) (simp add: lookup_single when_def) also have "\ = (\t. except t X) ` keys p" by (rule UNION_singleton_eq_range) finally show ?thesis . qed lemma focus_in_Polys': assumes "p \ P[Y]" shows "focus X p \ P[Y \ X]" proof (intro PolysI subsetI) fix t assume "t \ keys (focus X p)" then obtain s where "s \ keys p" and t: "t = except s (- X)" unfolding keys_focus .. note this(1) also from assms have "keys p \ .[Y]" by (rule PolysD) finally have "keys s \ Y" by (rule PPsD) hence "keys t \ Y \ X" by (simp add: t keys_except le_infI1) thus "t \ .[Y \ X]" by (rule PPsI) qed corollary focus_in_Polys: "focus X p \ P[X]" proof - have "p \ P[UNIV]" by simp hence "focus X p \ P[UNIV \ X]" by (rule focus_in_Polys') thus ?thesis by simp qed lemma focus_coeffs_subset_Polys': assumes "p \ P[Y]" shows "range (lookup (focus X p)) \ P[Y - X]" proof (intro subsetI PolysI) fix c t assume "c \ range (lookup (focus X p))" hence "keys c \ (\t. except t X) ` keys p" by (rule keys_coeffs_focus_subset) moreover assume "t \ keys c" ultimately have "t \ (\t. except t X) ` keys p" .. then obtain s where "s \ keys p" and t: "t = except s X" .. note this(1) also from assms have "keys p \ .[Y]" by (rule PolysD) finally have "keys s \ Y" by (rule PPsD) hence "keys t \ Y - X" by (simp add: t keys_except Diff_mono) thus "t \ .[Y - X]" by (rule PPsI) qed corollary focus_coeffs_subset_Polys: "range (lookup (focus X p)) \ P[- X]" proof - have "p \ P[UNIV]" by simp hence "range (lookup (focus X p)) \ P[UNIV - X]" by (rule focus_coeffs_subset_Polys') thus ?thesis by (simp only: Compl_eq_Diff_UNIV) qed corollary lookup_focus_in_Polys: "lookup (focus X p) t \ P[- X]" using focus_coeffs_subset_Polys by blast lemma focus_zero [simp]: "focus X 0 = 0" by (simp add: focus_def) lemma focus_eq_zero_iff [iff]: "focus X p = 0 \ p = 0" by (simp only: keys_focus flip: keys_eq_empty_iff) simp lemma focus_one [simp]: "focus X 1 = 1" by (simp add: focus_def) lemma focus_monomial: "focus X (monomial c t) = monomial (monomial c (except t X)) (except t (- X))" by (simp add: focus_def) lemma focus_uminus [simp]: "focus X (- p) = - focus X p" by (simp add: focus_def keys_uminus single_uminus sum_negf) lemma focus_plus: "focus X (p + q) = focus X p + focus X q" proof - have "finite (keys p \ keys q)" by simp moreover have "keys (p + q) \ keys p \ keys q" by (rule Poly_Mapping.keys_add) ultimately show ?thesis by (simp add: focus_superset[where A="keys p \ keys q"] lookup_add single_add sum.distrib) qed lemma focus_minus: "focus X (p - q) = focus X p - focus X (q::_ \\<^sub>0 _::ab_group_add)" by (simp only: diff_conv_add_uminus focus_plus focus_uminus) lemma focus_times: "focus X (p * q) = focus X p * focus X q" proof - have eq: "focus X (monomial c s * q) = focus X (monomial c s) * focus X q" for c s proof - have "focus X (monomial c s * q) = focus X (punit.monom_mult c s q)" by (simp only: times_monomial_left) also have "\ = (\t\(+) s ` keys q. monomial (monomial (lookup (punit.monom_mult c s q) t) (except t X)) (except t (- X)))" by (rule focus_superset) (simp_all add: punit.keys_monom_mult_subset[simplified]) also have "\ = (\t\keys q. ((\t. monomial (monomial (lookup (punit.monom_mult c s q) t) (except t X)) (except t (- X))) \ ((+) s)) t)" by (rule sum.reindex) simp also have "\ = monomial (monomial c (except s X)) (except s (- X)) * (\t\keys q. monomial (monomial (lookup q t) (except t X)) (except t (- X)))" by (simp add: o_def punit.lookup_monom_mult except_plus times_monomial_monomial sum_distrib_left) also have "\ = focus X (monomial c s) * focus X q" by (simp only: focus_monomial focus_def[where p=q]) finally show ?thesis . qed show ?thesis by (induct p rule: poly_mapping_plus_induct) (simp_all add: ring_distribs focus_plus eq) qed lemma focus_sum: "focus X (sum f I) = (\i\I. focus X (f i))" by (induct I rule: infinite_finite_induct) (simp_all add: focus_plus) lemma focus_prod: "focus X (prod f I) = (\i\I. focus X (f i))" by (induct I rule: infinite_finite_induct) (simp_all add: focus_times) lemma focus_power [simp]: "focus X (f ^ m) = focus X f ^ m" by (induct m) (simp_all add: focus_times) lemma focus_Polys: assumes "p \ P[X]" shows "focus X p = (\t\keys p. monomial (monomial (lookup p t) 0) t)" unfolding focus_def proof (rule sum.cong) fix t assume "t \ keys p" also from assms have "\ \ .[X]" by (rule PolysD) finally have "keys t \ X" by (rule PPsD) hence "except t X = 0" and "except t (- X) = t" by (rule except_eq_zeroI, auto simp: except_id_iff) thus "monomial (monomial (lookup p t) (except t X)) (except t (- X)) = monomial (monomial (lookup p t) 0) t" by simp qed (fact refl) corollary lookup_focus_Polys: "p \ P[X] \ lookup (focus X p) t = monomial (lookup p t) 0" by (simp add: focus_Polys lookup_sum lookup_single when_def in_keys_iff) lemma focus_Polys_Compl: assumes "p \ P[- X]" shows "focus X p = monomial p 0" proof - have "focus X p = (\t\keys p. monomial (monomial (lookup p t) t) 0)" unfolding focus_def proof (rule sum.cong) fix t assume "t \ keys p" also from assms have "\ \ .[- X]" by (rule PolysD) finally have "keys t \ - X" by (rule PPsD) hence "except t (- X) = 0" and "except t X = t" by (rule except_eq_zeroI, auto simp: except_id_iff) thus "monomial (monomial (lookup p t) (except t X)) (except t (- X)) = monomial (monomial (lookup p t) t) 0" by simp qed (fact refl) also have "\ = monomial (\t\keys p. monomial (lookup p t) t) 0" by (simp only: monomial_sum) also have "\ = monomial p 0" by (simp only: poly_mapping_sum_monomials) finally show ?thesis . qed corollary focus_empty [simp]: "focus {} p = monomial p 0" by (rule focus_Polys_Compl) simp lemma focus_Int: assumes "p \ P[Y]" shows "focus (X \ Y) p = focus X p" unfolding focus_def using refl proof (rule sum.cong) fix t assume "t \ keys p" also from assms have "\ \ .[Y]" by (rule PolysD) finally have "keys t \ Y" by (rule PPsD) hence "keys t \ X \ Y" by blast hence "except t (X \ Y) = except t X + except t Y" by (rule except_Int) also from \keys t \ Y\ have "except t Y = 0" by (rule except_eq_zeroI) finally have eq: "except t (X \ Y) = except t X" by simp have "except t (- (X \ Y)) = except (except t (- Y)) (- X)" by (simp add: except_except Un_commute) also from \keys t \ Y\ have "except t (- Y) = t" by (auto simp: except_id_iff) finally show "monomial (monomial (lookup p t) (except t (X \ Y))) (except t (- (X \ Y))) = monomial (monomial (lookup p t) (except t X)) (except t (- X))" by (simp only: eq) qed lemma range_focusD: assumes "p \ range (focus X)" shows "p \ P[X]" and "range (lookup p) \ P[- X]" and "lookup p t \ P[- X]" using assms by (auto intro: focus_in_Polys lookup_focus_in_Polys) lemma range_focusI: assumes "p \ P[X]" and "lookup p ` keys (p::_ \\<^sub>0 _ \\<^sub>0 _::semiring_1) \ P[- X]" shows "p \ range (focus X)" using assms proof (induct p rule: poly_mapping_plus_induct_Polys) case 0 show ?case by simp next case (plus p c t) from plus.hyps(3) have 1: "keys (monomial c t) = {t}" by simp also from plus.hyps(4) have "\ \ keys p = {}" by simp finally have "keys (monomial c t + p) = keys (monomial c t) \ keys p" by (rule keys_add[symmetric]) hence 2: "keys (monomial c t + p) = insert t (keys p)" by (simp only: 1 flip: insert_is_Un) from \t \ .[X]\ have "keys t \ X" by (rule PPsD) hence eq1: "except t X = 0" and eq2: "except t (- X) = t" by (rule except_eq_zeroI, auto simp: except_id_iff) from plus.hyps(3, 4) plus.prems have "c \ P[- X]" and "lookup p ` keys p \ P[- X]" by (simp_all add: 2 lookup_add lookup_single in_keys_iff) (smt add.commute add.right_neutral image_cong plus.hyps(4) when_simps(2)) from this(2) have "p \ range (focus X)" by (rule plus.hyps) then obtain q where p: "p = focus X q" .. moreover from \c \ P[- X]\ have "monomial c t = focus X (monomial 1 t * c)" by (simp add: focus_times focus_monomial eq1 eq2 focus_Polys_Compl times_monomial_monomial) ultimately have "monomial c t + p = focus X (monomial 1 t * c + q)" by (simp only: focus_plus) thus ?case by (rule range_eqI) qed lemma inj_focus: "inj ((focus X) :: (('x \\<^sub>0 nat) \\<^sub>0 'a::ab_group_add) \ _)" proof (rule injI) fix p q :: "('x \\<^sub>0 nat) \\<^sub>0 'a" assume "focus X p = focus X q" hence "focus X (p - q) = 0" by (simp add: focus_minus) thus "p = q" by simp qed lemma flatten_superset: assumes "finite A" and "keys p \ A" shows "flatten p = (\t\A. punit.monom_mult 1 t (lookup p t))" unfolding flatten_def using assms by (rule sum.mono_neutral_left) (simp add: in_keys_iff) lemma keys_flatten_subset: "keys (flatten p) \ (\t\keys p. (+) t ` keys (lookup p t))" proof - have "keys (flatten p) \ (\t\keys p. keys (punit.monom_mult 1 t (lookup p t)))" unfolding flatten_def by (rule keys_sum_subset) also from subset_refl have "\ \ (\t\keys p. (+) t ` keys (lookup p t))" by (rule UN_mono) (rule punit.keys_monom_mult_subset[simplified]) finally show ?thesis . qed lemma flatten_in_Polys: assumes "p \ P[X]" and "lookup p ` keys p \ P[Y]" shows "flatten p \ P[X \ Y]" proof (intro PolysI subsetI) fix t assume "t \ keys (flatten p)" also have "\ \ (\t\keys p. (+) t ` keys (lookup p t))" by (rule keys_flatten_subset) finally obtain s where "s \ keys p" and "t \ (+) s ` keys (lookup p s)" .. from this(2) obtain s' where "s' \ keys (lookup p s)" and t: "t = s + s'" .. from assms(1) have "keys p \ .[X]" by (rule PolysD) with \s \ keys p\ have "s \ .[X]" .. also have "\ \ .[X \ Y]" by (rule PPs_mono) simp finally have 1: "s \ .[X \ Y]" . from \s \ keys p\ have "lookup p s \ lookup p ` keys p" by (rule imageI) also have "\ \ P[Y]" by fact finally have "keys (lookup p s) \ .[Y]" by (rule PolysD) with \s' \ _\ have "s' \ .[Y]" .. also have "\ \ .[X \ Y]" by (rule PPs_mono) simp finally have "s' \ .[X \ Y]" . with 1 show "t \ .[X \ Y]" unfolding t by (rule PPs_closed_plus) qed lemma flatten_zero [simp]: "flatten 0 = 0" by (simp add: flatten_def) lemma flatten_one [simp]: "flatten 1 = 1" by (simp add: flatten_def) lemma flatten_monomial: "flatten (monomial c t) = punit.monom_mult 1 t c" by (simp add: flatten_def) lemma flatten_uminus [simp]: "flatten (- p) = - flatten (p::_ \\<^sub>0 _ \\<^sub>0 _::ring)" by (simp add: flatten_def keys_uminus punit.monom_mult_uminus_right sum_negf) lemma flatten_plus: "flatten (p + q) = flatten p + flatten q" proof - have "finite (keys p \ keys q)" by simp moreover have "keys (p + q) \ keys p \ keys q" by (rule Poly_Mapping.keys_add) ultimately show ?thesis by (simp add: flatten_superset[where A="keys p \ keys q"] punit.monom_mult_dist_right lookup_add sum.distrib) qed lemma flatten_minus: "flatten (p - q) = flatten p - flatten (q::_ \\<^sub>0 _ \\<^sub>0 _::ring)" by (simp only: diff_conv_add_uminus flatten_plus flatten_uminus) lemma flatten_times: "flatten (p * q) = flatten p * flatten (q::_ \\<^sub>0 _ \\<^sub>0 'b::comm_semiring_1)" proof - have eq: "flatten (monomial c s * q) = flatten (monomial c s) * flatten q" for c s proof - have eq: "monomial 1 (t + s) = monomial 1 s * monomial (1::'b) t" for t by (simp add: times_monomial_monomial add.commute) have "flatten (monomial c s * q) = flatten (punit.monom_mult c s q)" by (simp only: times_monomial_left) also have "\ = (\t\(+) s ` keys q. punit.monom_mult 1 t (lookup (punit.monom_mult c s q) t))" by (rule flatten_superset) (simp_all add: punit.keys_monom_mult_subset[simplified]) also have "\ = (\t\keys q. ((\t. punit.monom_mult 1 t (lookup (punit.monom_mult c s q) t)) \ (+) s) t)" by (rule sum.reindex) simp thm times_monomial_left also have "\ = punit.monom_mult 1 s c * (\t\keys q. punit.monom_mult 1 t (lookup q t))" by (simp add: o_def punit.lookup_monom_mult sum_distrib_left) (simp add: algebra_simps eq flip: times_monomial_left) also have "\ = flatten (monomial c s) * flatten q" by (simp only: flatten_monomial flatten_def[where p=q]) finally show ?thesis . qed show ?thesis by (induct p rule: poly_mapping_plus_induct) (simp_all add: ring_distribs flatten_plus eq) qed lemma flatten_monom_mult: "flatten (punit.monom_mult c t p) = punit.monom_mult 1 t (c * flatten (p::_ \\<^sub>0 _ \\<^sub>0 'b::comm_semiring_1))" by (simp only: flatten_times flatten_monomial mult.assoc flip: times_monomial_left) lemma flatten_sum: "flatten (sum f I) = (\i\I. flatten (f i))" by (induct I rule: infinite_finite_induct) (simp_all add: flatten_plus) lemma flatten_prod: "flatten (prod f I) = (\i\I. flatten (f i :: _ \\<^sub>0 _::comm_semiring_1))" by (induct I rule: infinite_finite_induct) (simp_all add: flatten_times) lemma flatten_power [simp]: "flatten (f ^ m) = flatten (f:: _ \\<^sub>0 _::comm_semiring_1) ^ m" by (induct m) (simp_all add: flatten_times) lemma surj_flatten: "surj flatten" proof (rule surjI) fix p show "flatten (monomial p 0) = p" by (simp add: flatten_monomial) qed lemma flatten_focus [simp]: "flatten (focus X p) = p" by (induct p rule: poly_mapping_plus_induct) (simp_all add: focus_plus flatten_plus focus_monomial flatten_monomial punit.monom_mult_monomial add.commute flip: except_decomp) lemma focus_flatten: assumes "p \ P[X]" and "lookup p ` keys p \ P[- X]" shows "focus X (flatten p) = p" proof - from assms have "p \ range (focus X)" by (rule range_focusI) then obtain q where "p = focus X q" .. thus ?thesis by simp qed lemma image_focus_ideal: "focus X ` ideal F = ideal (focus X ` F) \ range (focus X)" proof from focus_plus focus_times have "focus X ` ideal F \ ideal (focus X ` F)" by (rule image_ideal_subset) moreover from subset_UNIV have "focus X ` ideal F \ range (focus X)" by (rule image_mono) ultimately show "focus X ` ideal F \ ideal (focus X ` F) \ range (focus X)" by blast next show "ideal (focus X ` F) \ range (focus X) \ focus X ` ideal F" proof fix p assume "p \ ideal (focus X ` F) \ range (focus X)" hence "p \ ideal (focus X ` F)" and "p \ range (focus X)" by simp_all from this(1) obtain F0 q where "F0 \ focus X ` F" and p: "p = (\f'\F0. q f' * f')" by (rule ideal.spanE) from this(1) obtain F' where "F' \ F" and F0: "F0 = focus X ` F'" by (rule subset_imageE) from inj_focus subset_UNIV have "inj_on (focus X) F'" by (rule inj_on_subset) from \p \ range _\ obtain p' where "p = focus X p'" .. hence "p = focus X (flatten p)" by simp also from \inj_on _ F'\ have "\ = focus X (\f'\F'. flatten (q (focus X f')) * f')" by (simp add: p F0 sum.reindex flatten_sum flatten_times) finally have "p = focus X (\f'\F'. flatten (q (focus X f')) * f')" . moreover have "(\f'\F'. flatten (q (focus X f')) * f') \ ideal F" proof show "(\f'\F'. flatten (q (focus X f')) * f') \ ideal F'" by (rule ideal.sum_in_spanI) next from \F' \ F\ show "ideal F' \ ideal F" by (rule ideal.span_mono) qed ultimately show "p \ focus X ` ideal F" by (rule image_eqI) qed qed lemma image_flatten_ideal: "flatten ` ideal F = ideal (flatten ` F)" using flatten_plus flatten_times surj_flatten by (rule image_ideal_eq_surj) lemma poly_eval_focus: "poly_eval a (focus X p) = poly_subst (\x. if x \ X then a x else monomial 1 (Poly_Mapping.single x 1)) p" proof - let ?b = "\x. if x \ X then a x else monomial 1 (Poly_Mapping.single x 1)" have *: "lookup (punit.monom_mult (monomial (lookup p t) (except t X)) 0 (subst_pp (\x. monomial (a x) 0) (except t (- X)))) 0 = punit.monom_mult (lookup p t) 0 (subst_pp ?b t)" for t proof - have 1: "subst_pp ?b (except t X) = monomial 1 (except t X)" by (rule subst_pp_id) (simp add: keys_except) from refl have 2: "subst_pp ?b (except t (- X)) = subst_pp a (except t (-X))" by (rule subst_pp_cong) (simp add: keys_except) have "lookup (punit.monom_mult (monomial (lookup p t) (except t X)) 0 (subst_pp (\x. monomial (a x) 0) (except t (- X)))) 0 = punit.monom_mult (lookup p t) (except t X) (subst_pp a (except t (- X)))" by (simp add: lookup_times_zero subst_pp_def lookup_prod_zero lookup_power_zero flip: times_monomial_left) also have "\ = punit.monom_mult (lookup p t) 0 (monomial 1 (except t X) * subst_pp a (except t (- X)))" by (simp add: times_monomial_monomial flip: times_monomial_left mult.assoc) also have "\ = punit.monom_mult (lookup p t) 0 (subst_pp ?b (except t X + except t (- X)))" by (simp only: subst_pp_plus 1 2) also have "\ = punit.monom_mult (lookup p t) 0 (subst_pp ?b t)" by (simp flip: except_decomp) finally show ?thesis . qed show ?thesis by (simp add: poly_eval_def focus_def poly_subst_sum lookup_sum poly_subst_monomial * flip: poly_subst_def) qed corollary poly_eval_poly_eval_focus: "poly_eval a (poly_eval b (focus X p)) = poly_eval (\x::'x. if x \ X then poly_eval a (b x) else a x) p" proof - have eq: "monomial (lookup (poly_subst (\y. monomial (a y) (0::'x \\<^sub>0 nat)) q) 0) 0 = poly_subst (\y. monomial (a y) (0::'x \\<^sub>0 nat)) q" for q by (intro poly_deg_zero_imp_monomial poly_deg_poly_subst_eq_zeroI) simp show ?thesis unfolding poly_eval_focus by (simp add: poly_eval_def poly_subst_poly_subst if_distrib poly_subst_monomial subst_pp_single eq cong: if_cong) qed lemma indets_poly_eval_focus_subset: "indets (poly_eval a (focus X p)) \ \ (indets ` a ` X) \ (indets p - X)" proof fix x assume "x \ indets (poly_eval a (focus X p))" also have "\ = indets (poly_subst (\x. if x \ X then a x else monomial 1 (Poly_Mapping.single x 1)) p)" (is "_ = indets (poly_subst ?f _)") by (simp only: poly_eval_focus) finally obtain y where "y \ indets p" and "x \ indets (?f y)" by (rule in_indets_poly_substE) from this(2) have "(x \ X \ x = y) \ (y \ X \ x \ indets (a y))" by (simp add: indets_monomial split: if_split_asm) thus "x \ \ (indets ` a ` X) \ (indets p - X)" proof (elim disjE conjE) assume "x \ X" and "x = y" with \y \ indets p\ have "x \ indets p - X" by simp thus ?thesis .. next assume "y \ X" and "x \ indets (a y)" hence "x \ \ (indets ` a ` X)" by blast thus ?thesis .. qed qed lemma lookup_poly_eval_focus: "lookup (poly_eval (\x. monomial (a x) 0) (focus X p)) t = poly_eval a (lookup (focus (- X) p) t)" proof - let ?f = "\x. if x \ X then monomial (a x) 0 else monomial 1 (Poly_Mapping.single x 1)" have eq: "subst_pp ?f s = monomial (\x\keys s \ X. a x ^ lookup s x) (except s X)" for s proof - have "subst_pp ?f s = (\x\(keys s \ X) \ (keys s - X). ?f x ^ lookup s x)" unfolding subst_pp_def by (rule prod.cong) blast+ also have "\ = (\x\keys s \ X. ?f x ^ lookup s x) * (\x\keys s - X. ?f x ^ lookup s x)" by (rule prod.union_disjoint) auto also have "\ = monomial (\x\keys s \ X. a x ^ lookup s x) (\x\keys s - X. Poly_Mapping.single x (lookup s x))" by (simp add: monomial_power_map_scale times_monomial_monomial flip: punit.monomial_prod_sum) also have "(\x\keys s - X. Poly_Mapping.single x (lookup s x)) = except s X" by (metis (mono_tags, lifting) DiffD2 keys_except lookup_except_eq_idI poly_mapping_sum_monomials sum.cong) finally show ?thesis . qed show ?thesis by (simp add: poly_eval_focus poly_subst_def lookup_sum eq flip: punit.map_scale_eq_monom_mult) (simp add: focus_def lookup_sum poly_eval_sum lookup_single when_distrib poly_eval_monomial keys_except lookup_except_when) qed lemma keys_poly_eval_focus_subset: "keys (poly_eval (\x. monomial (a x) 0) (focus X p)) \ (\t. except t X) ` keys p" proof fix t assume "t \ keys (poly_eval (\x. monomial (a x) 0) (focus X p))" hence "lookup (poly_eval (\x. monomial (a x) 0) (focus X p)) t \ 0" by (simp add: in_keys_iff) hence "poly_eval a (lookup (focus (- X) p) t) \ 0" by (simp add: lookup_poly_eval_focus) hence "t \ keys (focus (- X) p)" by (auto simp flip: lookup_not_eq_zero_eq_in_keys) thus "t \ (\t. except t X) ` keys p" by (simp add: keys_focus) qed lemma poly_eval_focus_in_Polys: assumes "p \ P[X]" shows "poly_eval (\x. monomial (a x) 0) (focus Y p) \ P[X - Y]" proof (rule PolysI_alt) have "indets (poly_eval (\x. monomial (a x) 0) (focus Y p)) \ \ (indets ` (\x. monomial (a x) 0) ` Y) \ (indets p - Y)" by (fact indets_poly_eval_focus_subset) also have "\ = indets p - Y" by simp also from assms have "\ \ X - Y" by (auto dest: PolysD) finally show "indets (poly_eval (\x. monomial (a x) 0) (focus Y p)) \ X - Y" . qed lemma image_poly_eval_focus_ideal: "poly_eval (\x. monomial (a x) 0) ` focus X ` ideal F = ideal (poly_eval (\x. monomial (a x) 0) ` focus X ` F) \ (P[- X]::(('x \\<^sub>0 nat) \\<^sub>0 'a::comm_ring_1) set)" proof - let ?h = "\f. poly_eval (\x. monomial (a x) 0) (focus X f)" have h_id: "?h p = p" if "p \ P[- X]" for p proof - from that have "focus X p \ P[- X \ X]" by (rule focus_in_Polys') also have "\ = P[{}]" by simp finally obtain c where eq: "focus X p = monomial c 0" unfolding Polys_empty .. hence "flatten (focus X p) = flatten (monomial c 0)" by (rule arg_cong) hence "c = p" by (simp add: flatten_monomial) thus "?h p = p" by (simp add: eq poly_eval_monomial) qed have rng: "range ?h = P[- X]" proof (intro subset_antisym subsetI, elim rangeE) fix b f assume b: "b = ?h f" have "?h f \ P[UNIV - X]" by (rule poly_eval_focus_in_Polys) simp thus "b \ P[- X]" by (simp add: b Compl_eq_Diff_UNIV) next fix p :: "('x \\<^sub>0 nat) \\<^sub>0 'a" assume "p \ P[- X]" hence "?h p = p" by (rule h_id) hence "p = ?h p" by (rule sym) thus "p \ range ?h" by (rule range_eqI) qed have "poly_eval (\x. monomial (a x) 0) ` focus X ` ideal F = ?h ` ideal F" by (fact image_image) also have "\ = ideal (?h ` F) \ range ?h" proof (rule image_ideal_eq_Int) fix p have "?h p \ range ?h" by (fact rangeI) also have "\ = P[- X]" by fact finally show "?h (?h p) = ?h p" by (rule h_id) qed (simp_all only: focus_plus poly_eval_plus focus_times poly_eval_times) also have "\ = ideal (poly_eval (\x. monomial (a x) 0) ` focus X ` F) \ P[- X]" by (simp only: image_image rng) finally show ?thesis . qed subsection \Locale @{term pm_powerprod}\ lemma varnum_eq_zero_iff: "varnum X t = 0 \ t \ .[X]" by (auto simp: varnum_def PPs_def) lemma dgrad_set_varnum: "dgrad_set (varnum X) 0 = .[X]" by (simp add: dgrad_set_def PPs_def varnum_eq_zero_iff) context ordered_powerprod begin abbreviation "lcf \ punit.lc" abbreviation "tcf \ punit.tc" abbreviation "lpp \ punit.lt" abbreviation "tpp \ punit.tt" end (* ordered_powerprod *) locale pm_powerprod = ordered_powerprod ord ord_strict for ord::"('x::{countable,linorder} \\<^sub>0 nat) \ ('x \\<^sub>0 nat) \ bool" (infixl "\" 50) and ord_strict (infixl "\" 50) begin sublocale gd_powerprod .. lemma PPs_closed_lpp: assumes "p \ P[X]" shows "lpp p \ .[X]" proof (cases "p = 0") case True thus ?thesis by (simp add: zero_in_PPs) next case False hence "lpp p \ keys p" by (rule punit.lt_in_keys) also from assms have "\ \ .[X]" by (rule PolysD) finally show ?thesis . qed lemma PPs_closed_tpp: assumes "p \ P[X]" shows "tpp p \ .[X]" proof (cases "p = 0") case True thus ?thesis by (simp add: zero_in_PPs) next case False hence "tpp p \ keys p" by (rule punit.tt_in_keys) also from assms have "\ \ .[X]" by (rule PolysD) finally show ?thesis . qed corollary PPs_closed_image_lpp: "F \ P[X] \ lpp ` F \ .[X]" by (auto intro: PPs_closed_lpp) corollary PPs_closed_image_tpp: "F \ P[X] \ tpp ` F \ .[X]" by (auto intro: PPs_closed_tpp) lemma hom_component_lpp: assumes "p \ 0" shows "hom_component p (deg_pm (lpp p)) \ 0" (is "?p \ 0") and "lpp (hom_component p (deg_pm (lpp p))) = lpp p" proof - from assms have "lpp p \ keys p" by (rule punit.lt_in_keys) hence *: "lpp p \ keys ?p" by (simp add: keys_hom_component) thus "?p \ 0" by auto from * show "lpp ?p = lpp p" proof (rule punit.lt_eqI_keys) fix t assume "t \ keys ?p" hence "t \ keys p" by (simp add: keys_hom_component) thus "t \ lpp p" by (rule punit.lt_max_keys) qed qed definition is_hom_ord :: "'x \ bool" where "is_hom_ord x \ (\s t. deg_pm s = deg_pm t \ (s \ t \ except s {x} \ except t {x}))" lemma is_hom_ordD: "is_hom_ord x \ deg_pm s = deg_pm t \ s \ t \ except s {x} \ except t {x}" by (simp add: is_hom_ord_def) lemma dgrad_p_set_varnum: "punit.dgrad_p_set (varnum X) 0 = P[X]" by (simp add: punit.dgrad_p_set_def dgrad_set_varnum Polys_def) end text \We must create a copy of @{locale pm_powerprod} to avoid infinite chains of interpretations.\ instantiation option :: (linorder) linorder begin fun less_eq_option :: "'a option \ 'a option \ bool" where "less_eq_option None _ = True" | "less_eq_option (Some x) None = False" | "less_eq_option (Some x) (Some y) = (x \ y)" definition less_option :: "'a option \ 'a option \ bool" where "less_option x y \ x \ y \ \ y \ x" instance proof fix x :: "'a option" show "x \ x" using less_eq_option.elims(3) by fastforce qed (auto simp: less_option_def elim!: less_eq_option.elims) end locale extended_ord_pm_powerprod = pm_powerprod begin definition extended_ord :: "('a option \\<^sub>0 nat) \ ('a option \\<^sub>0 nat) \ bool" where "extended_ord s t \ (restrict_indets_pp s \ restrict_indets_pp t \ (restrict_indets_pp s = restrict_indets_pp t \ lookup s None \ lookup t None))" definition extended_ord_strict :: "('a option \\<^sub>0 nat) \ ('a option \\<^sub>0 nat) \ bool" where "extended_ord_strict s t \ (restrict_indets_pp s \ restrict_indets_pp t \ (restrict_indets_pp s = restrict_indets_pp t \ lookup s None < lookup t None))" sublocale extended_ord: pm_powerprod extended_ord extended_ord_strict proof - have 1: "s = t" if "lookup s None = lookup t None" and "restrict_indets_pp s = restrict_indets_pp t" for s t :: "'a option \\<^sub>0 nat" proof (rule poly_mapping_eqI) fix y show "lookup s y = lookup t y" proof (cases y) case None with that(1) show ?thesis by simp next case y: (Some z) have "lookup s y = lookup (restrict_indets_pp s) z" by (simp only: lookup_restrict_indets_pp y) also have "\ = lookup (restrict_indets_pp t) z" by (simp only: that(2)) also have "\ = lookup t y" by (simp only: lookup_restrict_indets_pp y) finally show ?thesis . qed qed have 2: "0 \ t" if "t \ 0" for t::"'a \\<^sub>0 nat" using that zero_min by (rule ordered_powerprod_lin.dual_order.not_eq_order_implies_strict) show "pm_powerprod extended_ord extended_ord_strict" by standard (auto simp: extended_ord_def extended_ord_strict_def restrict_indets_pp_plus lookup_add 1 dest: plus_monotone_strict 2) qed lemma extended_ord_is_hom_ord: "extended_ord.is_hom_ord None" by (auto simp add: extended_ord_def lookup_restrict_indets_pp lookup_except extended_ord.is_hom_ord_def simp flip: deg_pm_restrict_indets_pp) end end (* theory *) diff --git a/thys/Transition_Systems_and_Automata/Automata/DBA/DBA_Combine.thy b/thys/Transition_Systems_and_Automata/Automata/DBA/DBA_Combine.thy --- a/thys/Transition_Systems_and_Automata/Automata/DBA/DBA_Combine.thy +++ b/thys/Transition_Systems_and_Automata/Automata/DBA/DBA_Combine.thy @@ -1,303 +1,304 @@ section \Deterministic Büchi Automata Combinations\ theory DBA_Combine imports "DBA" "DGBA" begin definition dbgail :: "('label, 'state) dba list \ ('label, 'state list) dgba" where "dbgail AA \ dgba - (INTER (set AA) dba.alphabet) + (\ (dba.alphabet ` set AA)) (map dba.initial AA) (\ a pp. map2 (\ A p. dba.transition A a p) AA pp) (map (\ k pp. dba.accepting (AA ! k) (pp ! k)) [0 ..< length AA])" lemma dbgail_trace_smap: assumes "length pp = length AA" "k < length AA" shows "smap (\ pp. pp ! k) (dgba.trace (dbgail AA) w pp) = dba.trace (AA ! k) w (pp ! k)" using assms unfolding dbgail_def by (coinduction arbitrary: w pp) (force) lemma dbgail_nodes_length: assumes "pp \ DGBA.nodes (dbgail AA)" shows "length pp = length AA" using assms unfolding dbgail_def by induct auto lemma dbgail_nodes[intro]: assumes "pp \ DGBA.nodes (dbgail AA)" "k < length pp" shows "pp ! k \ DBA.nodes (AA ! k)" using assms unfolding dbgail_def by induct auto lemma dbgail_nodes_finite[intro]: assumes "list_all (finite \ DBA.nodes) AA" shows "finite (DGBA.nodes (dbgail AA))" proof (rule finite_subset) show "DGBA.nodes (dbgail AA) \ listset (map DBA.nodes AA)" by (force simp: listset_member list_all2_conv_all_nth dbgail_nodes_length) have "finite (listset (map DBA.nodes AA)) \ list_all finite (map DBA.nodes AA)" by (rule listset_finite) (auto simp: list_all_iff) then show "finite (listset (map DBA.nodes AA))" using assms by (simp add: list.pred_map) qed lemma dbgail_nodes_card: assumes "list_all (finite \ DBA.nodes) AA" shows "card (DGBA.nodes (dbgail AA)) \ prod_list (map (card \ DBA.nodes) AA)" proof - have "card (DGBA.nodes (dbgail AA)) \ card (listset (map DBA.nodes AA))" proof (rule card_mono) have "finite (listset (map DBA.nodes AA)) \ list_all finite (map DBA.nodes AA)" by (rule listset_finite) (auto simp: list_all_iff) then show "finite (listset (map DBA.nodes AA))" using assms by (simp add: list.pred_map) show "DGBA.nodes (dbgail AA) \ listset (map DBA.nodes AA)" by (force simp: listset_member list_all2_conv_all_nth dbgail_nodes_length) qed also have "\ = prod_list (map (card \ DBA.nodes) AA)" by simp finally show ?thesis by this qed - lemma dbgail_language[simp]: "DGBA.language (dbgail AA) = INTER (set AA) DBA.language" + lemma dbgail_language[simp]: "DGBA.language (dbgail AA) = \ (DBA.language ` set AA)" proof safe fix w A assume 1: "w \ DGBA.language (dbgail AA)" "A \ set AA" obtain 2: "dgba.run (dbgail AA) w (dgba.initial (dbgail AA))" "gen infs (dgba.accepting (dbgail AA)) (dgba.trace (dbgail AA) w (dgba.initial (dbgail AA)))" using 1(1) by rule obtain k where 3: "A = AA ! k" "k < length AA" using 1(2) unfolding in_set_conv_nth by auto have 4: "(\ pp. dba.accepting A (pp ! k)) \ set (dgba.accepting (dbgail AA))" using 3 unfolding dbgail_def by auto show "w \ DBA.language A" proof show "dba.run A w (dba.initial A)" using 1(2) 2(1) unfolding DBA.run_alt_def DGBA.run_alt_def dbgail_def by auto have "True \ infs (\ pp. dba.accepting A (pp ! k)) (dgba.trace (dbgail AA) w (map dba.initial AA))" using 2(2) 4 unfolding dbgail_def by auto also have "\ \ infs (dba.accepting A) (smap (\ pp. pp ! k) (dgba.trace (dbgail AA) w (map dba.initial AA)))" by (simp add: comp_def) also have "smap (\ pp. pp ! k) (dgba.trace (dbgail AA) w (map dba.initial AA)) = dba.trace (AA ! k) w (map dba.initial AA ! k)" using 3(2) by (fastforce intro: dbgail_trace_smap) also have "\ = dba.trace A w (dba.initial A)" using 3 by auto finally show "infs (dba.accepting A) (dba.trace A w (dba.initial A))" by simp qed next fix w - assume 1: "w \ INTER (set AA) DBA.language" + assume 1: "w \ \ (DBA.language ` set AA)" have 2: "dba.run A w (dba.initial A)" "infs (dba.accepting A) (dba.trace A w (dba.initial A))" if "A \ set AA" for A using 1 that by auto show "w \ DGBA.language (dbgail AA)" proof (intro DGBA.language ballI gen) show "dgba.run (dbgail AA) w (dgba.initial (dbgail AA))" using 2(1) unfolding DBA.run_alt_def DGBA.run_alt_def dbgail_def by auto next fix P assume 3: "P \ set (dgba.accepting (dbgail AA))" obtain k where 4: "P = (\ pp. dba.accepting (AA ! k) (pp ! k))" "k < length AA" using 3 unfolding dbgail_def by auto have "True \ infs (dba.accepting (AA ! k)) (dba.trace (AA ! k) w (map dba.initial AA ! k))" using 2(2) 4(2) by auto also have "dba.trace (AA ! k) w (map dba.initial AA ! k) = smap (\ pp. pp ! k) (dgba.trace (dbgail AA) w (map dba.initial AA))" using 4(2) by (fastforce intro: dbgail_trace_smap[symmetric]) also have "infs (dba.accepting (AA ! k)) \ \ infs P (dgba.trace (dbgail AA) w (map dba.initial AA))" unfolding 4(1) by (simp add: comp_def) also have "map dba.initial AA = dgba.initial (dbgail AA)" unfolding dbgail_def by simp finally show "infs P (dgba.trace (dbgail AA) w (dgba.initial (dbgail AA)))" by simp qed qed definition dbail :: "('label, 'state) dba list \ ('label, 'state list degen) dba" where "dbail = dgbad \ dbgail" lemma dbail_nodes_finite[intro]: assumes "list_all (finite \ DBA.nodes) AA" shows "finite (DBA.nodes (dbail AA))" using dbgail_nodes_finite assms unfolding dbail_def by auto lemma dbail_nodes_card: assumes"list_all (finite \ DBA.nodes) AA" shows "card (DBA.nodes (dbail AA)) \ max 1 (length AA) * prod_list (map (card \ DBA.nodes) AA)" proof - have "card (DBA.nodes (dbail AA)) \ max 1 (length (dgba.accepting (dbgail AA))) * card (DGBA.nodes (dbgail AA))" unfolding dbail_def using dgbad_nodes_card by simp also have "length (dgba.accepting (dbgail AA)) = length AA" unfolding dbgail_def by simp also have "card (DGBA.nodes (dbgail AA)) \ prod_list (map (card \ DBA.nodes) AA)" using dbgail_nodes_card assms by this finally show ?thesis by simp qed - lemma dbail_language[simp]: "DBA.language (dbail AA) = INTER (set AA) DBA.language" - unfolding dbail_def using dgbad_language dbgail_language by auto + lemma dbail_language [simp]: + "DBA.language (dbail AA) = \ (DBA.language ` set AA)" + by (simp add: dbail_def) definition dbau :: "('label, 'state\<^sub>1) dba \ ('label, 'state\<^sub>2) dba \ ('label, 'state\<^sub>1 \ 'state\<^sub>2) dba" where "dbau A B \ dba (dba.alphabet A \ dba.alphabet B) (dba.initial A, dba.initial B) (\ a (p, q). (dba.transition A a p, dba.transition B a q)) (\ (p, q). dba.accepting A p \ dba.accepting B q)" (* TODO: can these be extracted as more general theorems about sscan? *) lemma dbau_fst[iff]: "infs (P \ fst) (dba.trace (dbau A B) w (p, q)) \ infs P (dba.trace A w p)" proof - let ?t = "dba.trace (dbau A B) w (p, q)" have "infs (P \ fst) ?t \ infs P (smap fst ?t)" by (simp add: comp_def) also have "smap fst ?t = dba.trace A w p" unfolding dbau_def by (coinduction arbitrary: w p q) (auto) finally show ?thesis by this qed lemma dbau_snd[iff]: "infs (P \ snd) (dba.trace (dbau A B) w (p, q)) \ infs P (dba.trace B w q)" proof - let ?t = "dba.trace (dbau A B) w (p, q)" have "infs (P \ snd) ?t \ infs P (smap snd ?t)" by (simp add: comp_def) also have "smap snd ?t = dba.trace B w q" unfolding dbau_def by (coinduction arbitrary: w p q) (auto) finally show ?thesis by this qed lemma dbau_nodes_fst[intro]: assumes "dba.alphabet A = dba.alphabet B" shows "fst ` DBA.nodes (dbau A B) \ DBA.nodes A" proof (rule subsetI, erule imageE) fix pq p assume "pq \ DBA.nodes (dbau A B)" "p = fst pq" then show "p \ DBA.nodes A" using assms unfolding dbau_def by (induct arbitrary: p) (auto) qed lemma dbau_nodes_snd[intro]: assumes "dba.alphabet A = dba.alphabet B" shows "snd ` DBA.nodes (dbau A B) \ DBA.nodes B" proof (rule subsetI, erule imageE) fix pq q assume "pq \ DBA.nodes (dbau A B)" "q = snd pq" then show "q \ DBA.nodes B" using assms unfolding dbau_def by (induct arbitrary: q) (auto) qed lemma dbau_nodes_finite[intro]: assumes "dba.alphabet A = dba.alphabet B" assumes "finite (DBA.nodes A)" "finite (DBA.nodes B)" shows "finite (DBA.nodes (dbau A B))" proof (rule finite_subset) show "DBA.nodes (dbau A B) \ DBA.nodes A \ DBA.nodes B" using dbau_nodes_fst[OF assms(1)] dbau_nodes_snd[OF assms(1)] unfolding image_subset_iff by force show "finite (DBA.nodes A \ DBA.nodes B)" using assms(2, 3) by simp qed lemma dbau_nodes_card[intro]: assumes "dba.alphabet A = dba.alphabet B" assumes "finite (DBA.nodes A)" "finite (DBA.nodes B)" shows "card (DBA.nodes (dbau A B)) \ card (DBA.nodes A) * card (DBA.nodes B)" proof - have "card (DBA.nodes (dbau A B)) \ card (DBA.nodes A \ DBA.nodes B)" proof (rule card_mono) show "finite (DBA.nodes A \ DBA.nodes B)" using assms(2, 3) by simp show "DBA.nodes (dbau A B) \ DBA.nodes A \ DBA.nodes B" using dbau_nodes_fst[OF assms(1)] dbau_nodes_snd[OF assms(1)] unfolding image_subset_iff by force qed also have "\ = card (DBA.nodes A) * card (DBA.nodes B)" using card_cartesian_product by this finally show ?thesis by this qed lemma dbau_language[simp]: assumes "dba.alphabet A = dba.alphabet B" shows "DBA.language (dbau A B) = DBA.language A \ DBA.language B" proof - have 1: "dba.alphabet (dbau A B) = dba.alphabet A \ dba.alphabet B" unfolding dbau_def by simp have 2: "dba.initial (dbau A B) = (dba.initial A, dba.initial B)" unfolding dbau_def by simp have 3: "dba.accepting (dbau A B) = (\ pq. (dba.accepting A \ fst) pq \ (dba.accepting B \ snd) pq)" unfolding dbau_def by auto have 4: "infs (dba.accepting (dbau A B)) (DBA.trace (dbau A B) w (p, q)) \ infs (dba.accepting A) (DBA.trace A w p) \ infs (dba.accepting B) (DBA.trace B w q)" for w p q unfolding 3 by blast show ?thesis using assms unfolding DBA.language_def DBA.run_alt_def 1 2 4 by auto qed definition dbaul :: "('label, 'state) dba list \ ('label, 'state list) dba" where "dbaul AA \ dba - (UNION (set AA) dba.alphabet) + (\ (dba.alphabet ` set AA)) (map dba.initial AA) (\ a pp. map2 (\ A p. dba.transition A a p) AA pp) (\ pp. \ k < length AA. dba.accepting (AA ! k) (pp ! k))" lemma dbaul_trace_smap: assumes "length pp = length AA" "k < length AA" shows "smap (\ pp. pp ! k) (dba.trace (dbaul AA) w pp) = dba.trace (AA ! k) w (pp ! k)" using assms unfolding dbaul_def by (coinduction arbitrary: w pp) (force) lemma dbaul_nodes_length: assumes "pp \ DBA.nodes (dbaul AA)" shows "length pp = length AA" using assms unfolding dbaul_def by induct auto lemma dbaul_nodes[intro]: - assumes "INTER (set AA) dba.alphabet = UNION (set AA) dba.alphabet" + assumes "\ (dba.alphabet ` set AA) = \ (dba.alphabet ` set AA)" assumes "pp \ DBA.nodes (dbaul AA)" "k < length pp" shows "pp ! k \ DBA.nodes (AA ! k)" using assms(2, 3, 1) unfolding dbaul_def by induct force+ lemma dbaul_nodes_finite[intro]: - assumes "INTER (set AA) dba.alphabet = UNION (set AA) dba.alphabet" + assumes "\ (dba.alphabet ` set AA) = \ (dba.alphabet ` set AA)" assumes "list_all (finite \ DBA.nodes) AA" shows "finite (DBA.nodes (dbaul AA))" proof (rule finite_subset) show "DBA.nodes (dbaul AA) \ listset (map DBA.nodes AA)" using assms(1) by (force simp: listset_member list_all2_conv_all_nth dbaul_nodes_length) have "finite (listset (map DBA.nodes AA)) \ list_all finite (map DBA.nodes AA)" by (rule listset_finite) (auto simp: list_all_iff) then show "finite (listset (map DBA.nodes AA))" using assms(2) by (simp add: list.pred_map) qed lemma dbaul_nodes_card: - assumes "INTER (set AA) dba.alphabet = UNION (set AA) dba.alphabet" + assumes "\ (dba.alphabet ` set AA) = \ (dba.alphabet ` set AA)" assumes "list_all (finite \ DBA.nodes) AA" shows "card (DBA.nodes (dbaul AA)) \ prod_list (map (card \ DBA.nodes) AA)" proof - have "card (DBA.nodes (dbaul AA)) \ card (listset (map DBA.nodes AA))" proof (rule card_mono) have "finite (listset (map DBA.nodes AA)) \ list_all finite (map DBA.nodes AA)" by (rule listset_finite) (auto simp: list_all_iff) then show "finite (listset (map DBA.nodes AA))" using assms(2) by (simp add: list.pred_map) show "DBA.nodes (dbaul AA) \ listset (map DBA.nodes AA)" using assms(1) by (force simp: listset_member list_all2_conv_all_nth dbaul_nodes_length) qed also have "\ = prod_list (map (card \ DBA.nodes) AA)" by simp finally show ?thesis by this qed lemma dbaul_language[simp]: - assumes "INTER (set AA) dba.alphabet = UNION (set AA) dba.alphabet" - shows "DBA.language (dbaul AA) = UNION (set AA) DBA.language" + assumes "\ (dba.alphabet ` set AA) = \ (dba.alphabet ` set AA)" + shows "DBA.language (dbaul AA) = \ (DBA.language ` set AA)" proof safe fix w assume 1: "w \ DBA.language (dbaul AA)" obtain 2: "dba.run (dbaul AA) w (dba.initial (dbaul AA))" "infs (dba.accepting (dbaul AA)) (dba.trace (dbaul AA) w (dba.initial (dbaul AA)))" using 1 by rule obtain k where 3: "k < length AA" "infs (\ pp. dba.accepting (AA ! k) (pp ! k)) (dba.trace (dbaul AA) w (map dba.initial AA))" using 2(2) unfolding dbaul_def by auto - show "w \ UNION (set AA) DBA.language" + show "w \ \ (DBA.language ` set AA)" proof (intro UN_I DBA.language) show "AA ! k \ set AA" using 3(1) by simp show "dba.run (AA ! k) w (dba.initial (AA ! k))" using assms 2(1) 3(1) unfolding DBA.run_alt_def DGBA.run_alt_def dbaul_def by force have "True \ infs (\ pp. dba.accepting (AA ! k) (pp ! k)) (dba.trace (dbaul AA) w (map dba.initial AA))" using 3(2) by auto also have "\ \ infs (dba.accepting (AA ! k)) (smap (\ pp. pp ! k) (dba.trace (dbaul AA) w (map dba.initial AA)))" by (simp add: comp_def) also have "smap (\ pp. pp ! k) (dba.trace (dbaul AA) w (map dba.initial AA)) = dba.trace (AA ! k) w (map dba.initial AA ! k)" using 3(1) by (fastforce intro: dbaul_trace_smap) also have "\ = dba.trace (AA ! k) w (dba.initial (AA ! k))" using 3 by auto finally show "infs (dba.accepting (AA ! k)) (dba.trace (AA ! k) w (dba.initial (AA ! k)))" by simp qed next fix A w assume 1: "A \ set AA" "w \ DBA.language A" obtain 2: "dba.run A w (dba.initial A)" "infs (dba.accepting A) (dba.trace A w (dba.initial A))" using 1(2) by rule obtain k where 3: "A = AA ! k" "k < length AA" using 1(1) unfolding in_set_conv_nth by auto show "w \ DBA.language (dbaul AA)" proof show "dba.run (dbaul AA) w (dba.initial (dbaul AA))" using 1(1) 2(1) unfolding DBA.run_alt_def DGBA.run_alt_def dbaul_def by auto have "True \ infs (dba.accepting (AA ! k)) (dba.trace (AA ! k) w (map dba.initial AA ! k))" using 2(2) 3 by auto also have "dba.trace (AA ! k) w (map dba.initial AA ! k) = smap (\ pp. pp ! k) (dba.trace (dbaul AA) w (map dba.initial AA))" using 3(2) by (fastforce intro: dbaul_trace_smap[symmetric]) also have "infs (dba.accepting (AA ! k)) \ \ infs (\ pp. dba.accepting (AA ! k) (pp ! k)) (dba.trace (dbaul AA) w (map dba.initial AA))" by (simp add: comp_def) finally show "infs (dba.accepting (dbaul AA)) (dba.trace (dbaul AA) w (dba.initial (dbaul AA)))" using 3(2) unfolding dbaul_def by auto qed qed end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/Automata/DCA/DCA_Combine.thy b/thys/Transition_Systems_and_Automata/Automata/DCA/DCA_Combine.thy --- a/thys/Transition_Systems_and_Automata/Automata/DCA/DCA_Combine.thy +++ b/thys/Transition_Systems_and_Automata/Automata/DCA/DCA_Combine.thy @@ -1,298 +1,298 @@ section \Deterministic Co-Büchi Automata Combinations\ theory DCA_Combine imports "DCA" "DGCA" begin definition dcai :: "('label, 'state\<^sub>1) dca \ ('label, 'state\<^sub>2) dca \ ('label, 'state\<^sub>1 \ 'state\<^sub>2) dca" where "dcai A B \ dca (dca.alphabet A \ dca.alphabet B) (dca.initial A, dca.initial B) (\ a (p, q). (dca.transition A a p, dca.transition B a q)) (\ (p, q). dca.rejecting A p \ dca.rejecting B q)" lemma dcai_fst[iff]: "infs (P \ fst) (dca.trace (dcai A B) w (p, q)) \ infs P (dca.trace A w p)" proof - let ?t = "dca.trace (dcai A B) w (p, q)" have "infs (P \ fst) ?t \ infs P (smap fst ?t)" by (simp add: comp_def) also have "smap fst ?t = dca.trace A w p" unfolding dcai_def by (coinduction arbitrary: w p q) (auto) finally show ?thesis by this qed lemma dcai_snd[iff]: "infs (P \ snd) (dca.trace (dcai A B) w (p, q)) \ infs P (dca.trace B w q)" proof - let ?t = "dca.trace (dcai A B) w (p, q)" have "infs (P \ snd) ?t \ infs P (smap snd ?t)" by (simp add: comp_def) also have "smap snd ?t = dca.trace B w q" unfolding dcai_def by (coinduction arbitrary: w p q) (auto) finally show ?thesis by this qed lemma dcai_nodes_fst[intro]: "fst ` DCA.nodes (dcai A B) \ DCA.nodes A" proof (rule subsetI, erule imageE) fix pq p assume "pq \ DCA.nodes (dcai A B)" "p = fst pq" then show "p \ DCA.nodes A" unfolding dcai_def by (induct arbitrary: p) (auto) qed lemma dcai_nodes_snd[intro]: "snd ` DCA.nodes (dcai A B) \ DCA.nodes B" proof (rule subsetI, erule imageE) fix pq q assume "pq \ DCA.nodes (dcai A B)" "q = snd pq" then show "q \ DCA.nodes B" unfolding dcai_def by (induct arbitrary: q) (auto) qed lemma dcai_nodes_finite[intro]: assumes "finite (DCA.nodes A)" "finite (DCA.nodes B)" shows "finite (DCA.nodes (dcai A B))" proof (rule finite_subset) show "DCA.nodes (dcai A B) \ DCA.nodes A \ DCA.nodes B" using dcai_nodes_fst dcai_nodes_snd unfolding image_subset_iff by force show "finite (DCA.nodes A \ DCA.nodes B)" using assms by simp qed lemma dcai_nodes_card[intro]: assumes "finite (DCA.nodes A)" "finite (DCA.nodes B)" shows "card (DCA.nodes (dcai A B)) \ card (DCA.nodes A) * card (DCA.nodes B)" proof - have "card (DCA.nodes (dcai A B)) \ card (DCA.nodes A \ DCA.nodes B)" proof (rule card_mono) show "finite (DCA.nodes A \ DCA.nodes B)" using assms by simp show "DCA.nodes (dcai A B) \ DCA.nodes A \ DCA.nodes B" using dcai_nodes_fst dcai_nodes_snd unfolding image_subset_iff by force qed also have "\ = card (DCA.nodes A) * card (DCA.nodes B)" using card_cartesian_product by this finally show ?thesis by this qed lemma dcai_language[simp]: "DCA.language (dcai A B) = DCA.language A \ DCA.language B" proof - have 1: "dca.alphabet (dcai A B) = dca.alphabet A \ dca.alphabet B" unfolding dcai_def by simp have 2: "dca.initial (dcai A B) = (dca.initial A, dca.initial B)" unfolding dcai_def by simp have 3: "dca.rejecting (dcai A B) = (\ pq. (dca.rejecting A \ fst) pq \ (dca.rejecting B \ snd) pq)" unfolding dcai_def by auto have 4: "infs (dca.rejecting (dcai A B)) (DCA.trace (dcai A B) w (p, q)) \ infs (dca.rejecting A) (DCA.trace A w p) \ infs (dca.rejecting B) (DCA.trace B w q)" for w p q unfolding 3 by blast show ?thesis unfolding DCA.language_def DCA.run_alt_def 1 2 4 by auto qed definition dcail :: "('label, 'state) dca list \ ('label, 'state list) dca" where "dcail AA \ dca - (INTER (set AA) dca.alphabet) + (\ (dca.alphabet ` set AA)) (map dca.initial AA) (\ a pp. map2 (\ A p. dca.transition A a p) AA pp) (\ pp. \ k < length AA. dca.rejecting (AA ! k) (pp ! k))" lemma dcail_trace_smap: assumes "length pp = length AA" "k < length AA" shows "smap (\ pp. pp ! k) (dca.trace (dcail AA) w pp) = dca.trace (AA ! k) w (pp ! k)" using assms unfolding dcail_def by (coinduction arbitrary: w pp) (force) lemma dcail_nodes_length: assumes "pp \ DCA.nodes (dcail AA)" shows "length pp = length AA" using assms unfolding dcail_def by induct auto lemma dcail_nodes[intro]: assumes "pp \ DCA.nodes (dcail AA)" "k < length pp" shows "pp ! k \ DCA.nodes (AA ! k)" using assms unfolding dcail_def by induct auto lemma dcail_finite[intro]: assumes "list_all (finite \ DCA.nodes) AA" shows "finite (DCA.nodes (dcail AA))" proof (rule finite_subset) show "DCA.nodes (dcail AA) \ listset (map DCA.nodes AA)" by (force simp: listset_member list_all2_conv_all_nth dcail_nodes_length) have "finite (listset (map DCA.nodes AA)) \ list_all finite (map DCA.nodes AA)" by (rule listset_finite) (auto simp: list_all_iff) then show "finite (listset (map DCA.nodes AA))" using assms by (simp add: list.pred_map) qed lemma dcail_nodes_card: assumes "list_all (finite \ DCA.nodes) AA" shows "card (DCA.nodes (dcail AA)) \ prod_list (map (card \ DCA.nodes) AA)" proof - have "card (DCA.nodes (dcail AA)) \ card (listset (map DCA.nodes AA))" proof (rule card_mono) have "finite (listset (map DCA.nodes AA)) \ list_all finite (map DCA.nodes AA)" by (rule listset_finite) (auto simp: list_all_iff) then show "finite (listset (map DCA.nodes AA))" using assms by (simp add: list.pred_map) show "DCA.nodes (dcail AA) \ listset (map DCA.nodes AA)" by (force simp: listset_member list_all2_conv_all_nth dcail_nodes_length) qed also have "\ = prod_list (map (card \ DCA.nodes) AA)" by simp finally show ?thesis by this qed - lemma dcail_language[simp]: "DCA.language (dcail AA) = INTER (set AA) DCA.language" + lemma dcail_language[simp]: "DCA.language (dcail AA) = \ (DCA.language ` set AA)" proof safe fix A w assume 1: "w \ DCA.language (dcail AA)" "A \ set AA" obtain 2: "dca.run (dcail AA) w (dca.initial (dcail AA))" "fins (dca.rejecting (dcail AA)) (dca.trace (dcail AA) w (dca.initial (dcail AA)))" using 1(1) by rule obtain k where 3: "A = AA ! k" "k < length AA" using 1(2) unfolding in_set_conv_nth by auto have 4: "fins (\ pp. dca.rejecting A (pp ! k)) (dca.trace (dcail AA) w (map dca.initial AA))" using 2(2) 3 unfolding dcail_def by auto show "w \ DCA.language A" proof show "dca.run A w (dca.initial A)" using 1(2) 2(1) unfolding DCA.run_alt_def dcail_def by auto have "True \ fins (\ pp. dca.rejecting A (pp ! k)) (dca.trace (dcail AA) w (map dca.initial AA))" using 4 by simp also have "\ \ fins (dca.rejecting A) (smap (\ pp. pp ! k) (dca.trace (dcail AA) w (map dca.initial AA)))" by (simp add: comp_def) also have "smap (\ pp. pp ! k) (dca.trace (dcail AA) w (map dca.initial AA)) = dca.trace (AA ! k) w (map dca.initial AA ! k)" using 3(2) by (fastforce intro: dcail_trace_smap) also have "\ = dca.trace A w (dca.initial A)" using 3 by auto finally show "fins (dca.rejecting A) (DCA.trace A w (dca.initial A))" by simp qed next fix w - assume 1: "w \ INTER (set AA) DCA.language" + assume 1: "w \ \ (DCA.language `set AA)" have 2: "dca.run A w (dca.initial A)" "fins (dca.rejecting A) (dca.trace A w (dca.initial A))" if "A \ set AA" for A using 1 that by auto have 3: "fins (\ pp. dca.rejecting (AA ! k) (pp ! k)) (dca.trace (dcail AA) w (map dca.initial AA))" if "k < length AA" for k proof - have "True \ fins (dca.rejecting (AA ! k)) (dca.trace (AA ! k) w (map dca.initial AA ! k))" using 2(2) that by auto also have "dca.trace (AA ! k) w (map dca.initial AA ! k) = smap (\ pp. pp ! k) (dca.trace (dcail AA) w (map dca.initial AA))" using that by (fastforce intro: dcail_trace_smap[symmetric]) also have "infs (dca.rejecting (AA ! k)) \ \ infs (\ pp. dca.rejecting (AA ! k) (pp ! k)) (dca.trace (dcail AA) w (map dca.initial AA))" by (simp add: comp_def) finally show ?thesis by simp qed show "w \ DCA.language (dcail AA)" proof show "dca.run (dcail AA) w (dca.initial (dcail AA))" using 2(1) unfolding DCA.run_alt_def dcail_def by auto show "fins (dca.rejecting (dcail AA)) (dca.trace (dcail AA) w (dca.initial (dcail AA)))" using 3 unfolding dcail_def by auto qed qed definition dcgaul :: "('label, 'state) dca list \ ('label, 'state list) dgca" where "dcgaul AA \ dgca - (UNION (set AA) dca.alphabet) + (\ (dca.alphabet ` set AA)) (map dca.initial AA) (\ a pp. map2 (\ A p. dca.transition A a p) AA pp) (map (\ k pp. dca.rejecting (AA ! k) (pp ! k)) [0 ..< length AA])" lemma dcgaul_trace_smap: assumes "length pp = length AA" "k < length AA" shows "smap (\ pp. pp ! k) (dgca.trace (dcgaul AA) w pp) = dca.trace (AA ! k) w (pp ! k)" using assms unfolding dcgaul_def by (coinduction arbitrary: w pp) (force) lemma dcgaul_nodes_length: assumes "pp \ DGCA.nodes (dcgaul AA)" shows "length pp = length AA" using assms unfolding dcgaul_def by induct auto lemma dcgaul_nodes[intro]: - assumes "INTER (set AA) dca.alphabet = UNION (set AA) dca.alphabet" + assumes "\ (dca.alphabet ` set AA) = \ (dca.alphabet ` set AA)" assumes "pp \ DGCA.nodes (dcgaul AA)" "k < length pp" shows "pp ! k \ DCA.nodes (AA ! k)" using assms(2, 3, 1) unfolding dcgaul_def by induct force+ lemma dcgaul_nodes_finite[intro]: - assumes "INTER (set AA) dca.alphabet = UNION (set AA) dca.alphabet" + assumes "\ (dca.alphabet ` set AA) = \ (dca.alphabet ` set AA)" assumes "list_all (finite \ DCA.nodes) AA" shows "finite (DGCA.nodes (dcgaul AA))" proof (rule finite_subset) show "DGCA.nodes (dcgaul AA) \ listset (map DCA.nodes AA)" using assms(1) by (force simp: listset_member list_all2_conv_all_nth dcgaul_nodes_length) have "finite (listset (map DCA.nodes AA)) \ list_all finite (map DCA.nodes AA)" by (rule listset_finite) (auto simp: list_all_iff) then show "finite (listset (map DCA.nodes AA))" using assms(2) by (simp add: list.pred_map) qed lemma dcgaul_nodes_card: - assumes "INTER (set AA) dca.alphabet = UNION (set AA) dca.alphabet" + assumes "\ (dca.alphabet ` set AA) = \ (dca.alphabet ` set AA)" assumes "list_all (finite \ DCA.nodes) AA" shows "card (DGCA.nodes (dcgaul AA)) \ prod_list (map (card \ DCA.nodes) AA)" proof - have "card (DGCA.nodes (dcgaul AA)) \ card (listset (map DCA.nodes AA))" proof (rule card_mono) have "finite (listset (map DCA.nodes AA)) \ list_all finite (map DCA.nodes AA)" by (rule listset_finite) (auto simp: list_all_iff) then show "finite (listset (map DCA.nodes AA))" using assms(2) by (simp add: list.pred_map) show "DGCA.nodes (dcgaul AA) \ listset (map DCA.nodes AA)" using assms(1) by (force simp: listset_member list_all2_conv_all_nth dcgaul_nodes_length) qed also have "\ = prod_list (map (card \ DCA.nodes) AA)" by simp finally show ?thesis by this qed lemma dcgaul_language[simp]: - assumes "INTER (set AA) dca.alphabet = UNION (set AA) dca.alphabet" - shows "DGCA.language (dcgaul AA) = UNION (set AA) DCA.language" + assumes "\ (dca.alphabet ` set AA) = \ (dca.alphabet ` set AA)" + shows "DGCA.language (dcgaul AA) = \ (DCA.language ` set AA)" proof safe fix w assume 1: "w \ DGCA.language (dcgaul AA)" obtain k where 2: "dgca.run (dcgaul AA) w (dgca.initial (dcgaul AA))" "k < length AA" "fins (\ pp. dca.rejecting (AA ! k) (pp ! k)) (dgca.trace (dcgaul AA) w (dgca.initial (dcgaul AA)))" using 1 unfolding dcgaul_def by force - show "w \ UNION (set AA) DCA.language" + show "w \ \ (DCA.language ` set AA)" proof (intro UN_I DCA.language) show "AA ! k \ set AA" using 2(2) by simp show "dca.run (AA ! k) w (dca.initial (AA ! k))" using assms 2(1, 2) unfolding DCA.run_alt_def DGCA.run_alt_def dcgaul_def by force have "True \ fins (\ pp. dca.rejecting (AA ! k) (pp ! k)) (dgca.trace (dcgaul AA) w (map dca.initial AA))" using 2(3) unfolding dcgaul_def by auto also have "\ \ fins (dca.rejecting (AA ! k)) (smap (\ pp. pp ! k) (dgca.trace (dcgaul AA) w (map dca.initial AA)))" by (simp add: comp_def) also have "smap (\ pp. pp ! k) (dgca.trace (dcgaul AA) w (map dca.initial AA)) = dca.trace (AA ! k) w (map dca.initial AA ! k)" using 2(2) by (fastforce intro: dcgaul_trace_smap) also have "\ = dca.trace (AA ! k) w (dca.initial (AA ! k))" using 2(2) by auto finally show "fins (dca.rejecting (AA ! k)) (dca.trace (AA ! k) w (dca.initial (AA ! k)))" by simp qed next fix A w assume 1: "A \ set AA" "w \ DCA.language A" obtain 2: "dca.run A w (dca.initial A)" "fins (dca.rejecting A) (dca.trace A w (dca.initial A))" using 1(2) by rule obtain k where 3: "A = AA ! k" "k < length AA" using 1(1) unfolding in_set_conv_nth by auto show "w \ DGCA.language (dcgaul AA)" proof (intro DGCA.language bexI cogen) show "dgca.run (dcgaul AA) w (dgca.initial (dcgaul AA))" using 1(1) 2(1) unfolding DCA.run_alt_def DGCA.run_alt_def dcgaul_def by auto have "True \ fins (dca.rejecting (AA ! k)) (dca.trace (AA ! k) w (map dca.initial AA ! k))" using 2(2) 3 by auto also have "dca.trace (AA ! k) w (map dca.initial AA ! k) = smap (\ pp. pp ! k) (dgca.trace (dcgaul AA) w (map dca.initial AA))" using 3(2) by (fastforce intro: dcgaul_trace_smap[symmetric]) also have "fins (dca.rejecting (AA ! k)) \ \ fins (\ pp. dca.rejecting (AA ! k) (pp ! k)) (dgca.trace (dcgaul AA) w (map dca.initial AA))" by (simp add: comp_def) also have "map dca.initial AA = dgca.initial (dcgaul AA)" unfolding dcgaul_def by simp finally show "fins (\ pp. dca.rejecting (AA ! k) (pp ! k)) (dgca.trace (dcgaul AA) w (dgca.initial (dcgaul AA)))" by simp show "(\ pp. dca.rejecting (AA ! k) (pp ! k)) \ set (dgca.rejecting (dcgaul AA))" unfolding dcgaul_def using 3(2) by simp qed qed definition dcaul :: "('label, 'state) dca list \ ('label, 'state list degen) dca" where "dcaul = dgcad \ dcgaul" lemma dcaul_nodes_finite[intro]: - assumes "INTER (set AA) dca.alphabet = UNION (set AA) dca.alphabet" + assumes "\ (dca.alphabet ` set AA) = \ (dca.alphabet ` set AA)" assumes "list_all (finite \ DCA.nodes) AA" shows "finite (DCA.nodes (dcaul AA))" using dcgaul_nodes_finite assms unfolding dcaul_def by auto lemma dcaul_nodes_card: - assumes "INTER (set AA) dca.alphabet = UNION (set AA) dca.alphabet" + assumes "\ (dca.alphabet ` set AA) = \ (dca.alphabet ` set AA)" assumes "list_all (finite \ DCA.nodes) AA" shows "card (DCA.nodes (dcaul AA)) \ max 1 (length AA) * prod_list (map (card \ DCA.nodes) AA)" proof - have "card (DCA.nodes (dcaul AA)) \ max 1 (length (dgca.rejecting (dcgaul AA))) * card (DGCA.nodes (dcgaul AA))" unfolding dcaul_def using dgcad_nodes_card by simp also have "length (dgca.rejecting (dcgaul AA)) = length AA" unfolding dcgaul_def by simp also have "card (DGCA.nodes (dcgaul AA)) \ prod_list (map (card \ DCA.nodes) AA)" using dcgaul_nodes_card assms by this finally show ?thesis by simp qed lemma dcaul_language[simp]: - assumes "INTER (set AA) dca.alphabet = UNION (set AA) dca.alphabet" - shows "DCA.language (dcaul AA) = UNION (set AA) DCA.language" + assumes "\ (dca.alphabet ` set AA) = \ (dca.alphabet ` set AA)" + shows "DCA.language (dcaul AA) = \ (DCA.language ` set AA)" unfolding dcaul_def using dgcad_language dcgaul_language[OF assms] by auto end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/Automata/DRA/DRA_Combine.thy b/thys/Transition_Systems_and_Automata/Automata/DRA/DRA_Combine.thy --- a/thys/Transition_Systems_and_Automata/Automata/DRA/DRA_Combine.thy +++ b/thys/Transition_Systems_and_Automata/Automata/DRA/DRA_Combine.thy @@ -1,223 +1,223 @@ section \Deterministic Rabin Automata Combinations\ theory DRA_Combine imports "DRA" "../DBA/DBA" "../DCA/DCA" begin definition from_dba :: "('label, 'state) dba \ ('label, 'state) dra" where "from_dba A \ dra (dba.alphabet A) (dba.initial A) (dba.transition A) [(dba.accepting A, bot)]" lemma from_dba_language[simp]: "DRA.language (from_dba A) = DBA.language A" unfolding DBA.language_def DRA.language_def from_dba_def DBA.run_def DRA.run_def by (auto 0 3) definition from_dca :: "('label, 'state) dca \ ('label, 'state) dra" where "from_dca A \ dra (dca.alphabet A) (dca.initial A) (dca.transition A) [(top, dca.rejecting A)]" lemma from_dca_language[simp]: "DRA.language (from_dca A) = DCA.language A" unfolding DCA.language_def DRA.language_def from_dca_def DCA.run_def DRA.run_def by (auto 0 3) definition dbcrai :: "('label, 'state\<^sub>1) dba \ ('label, 'state\<^sub>2) dca \ ('label, 'state\<^sub>1 \ 'state\<^sub>2) dra" where "dbcrai A B \ dra (dba.alphabet A \ dca.alphabet B) (dba.initial A, dca.initial B) (\ a (p, q). (dba.transition A a p, dca.transition B a q)) [(dba.accepting A \ fst, dca.rejecting B \ snd)]" lemma dbcrai_fst[iff]: "infs (P \ fst) (dra.trace (dbcrai A B) w (p, q)) \ infs P (dba.trace A w p)" proof - let ?t = "dra.trace (dbcrai A B) w (p, q)" have "infs (P \ fst) ?t \ infs P (smap fst ?t)" by (simp add: comp_def) also have "smap fst ?t = dba.trace A w p" unfolding dbcrai_def by (coinduction arbitrary: w p q) (auto) finally show ?thesis by this qed lemma dbcrai_snd[iff]: "infs (P \ snd) (dra.trace (dbcrai A B) w (p, q)) \ infs P (dca.trace B w q)" proof - let ?t = "dra.trace (dbcrai A B) w (p, q)" have "infs (P \ snd) ?t \ infs P (smap snd ?t)" by (simp add: comp_def) also have "smap snd ?t = dca.trace B w q" unfolding dbcrai_def by (coinduction arbitrary: w p q) (auto) finally show ?thesis by this qed lemma dbcrai_nodes_fst[intro]: "fst ` DRA.nodes (dbcrai A B) \ DBA.nodes A" proof (rule subsetI, erule imageE) fix pq p assume "pq \ DRA.nodes (dbcrai A B)" "p = fst pq" then show "p \ DBA.nodes A" unfolding dbcrai_def by (induct arbitrary: p) (auto) qed lemma dbcrai_nodes_snd[intro]: "snd ` DRA.nodes (dbcrai A B) \ DCA.nodes B" proof (rule subsetI, erule imageE) fix pq q assume "pq \ DRA.nodes (dbcrai A B)" "q = snd pq" then show "q \ DCA.nodes B" unfolding dbcrai_def by (induct arbitrary: q) (auto) qed lemma dbcrai_nodes_finite[intro]: assumes "finite (DBA.nodes A)" "finite (DCA.nodes B)" shows "finite (DRA.nodes (dbcrai A B))" proof (rule finite_subset) show "DRA.nodes (dbcrai A B) \ DBA.nodes A \ DCA.nodes B" using dbcrai_nodes_fst dbcrai_nodes_snd unfolding image_subset_iff by force show "finite (DBA.nodes A \ DCA.nodes B)" using assms by simp qed lemma dbcrai_nodes_card[intro]: assumes "finite (DBA.nodes A)" "finite (DCA.nodes B)" shows "card (DRA.nodes (dbcrai A B)) \ card (DBA.nodes A) * card (DCA.nodes B)" proof - have "card (DRA.nodes (dbcrai A B)) \ card (DBA.nodes A \ DCA.nodes B)" proof (rule card_mono) show "finite (DBA.nodes A \ DCA.nodes B)" using assms by simp show "DRA.nodes (dbcrai A B) \ DBA.nodes A \ DCA.nodes B" using dbcrai_nodes_fst dbcrai_nodes_snd unfolding image_subset_iff by force qed also have "\ = card (DBA.nodes A) * card (DCA.nodes B)" using card_cartesian_product by this finally show ?thesis by this qed lemma dbcrai_language[simp]: "DRA.language (dbcrai A B) = DBA.language A \ DCA.language B" proof - have 1: "dra.alphabet (dbcrai A B) = dba.alphabet A \ dca.alphabet B" unfolding dbcrai_def by simp have 2: "dra.initial (dbcrai A B) = (dba.initial A, dca.initial B)" unfolding dbcrai_def by simp have 3: "dra.accepting (dbcrai A B) = [(dba.accepting A \ fst, dca.rejecting B \ snd)]" unfolding dbcrai_def by simp have 4: "cogen rabin (dra.accepting (dbcrai A B)) (DRA.trace (dbcrai A B) w (p, q)) \ infs (dba.accepting A) (DBA.trace A w p) \ fins (rejecting B) (DCA.trace B w q)" for w p q unfolding cogen_def rabin_def 3 by simp show ?thesis unfolding DRA.language_def DBA.language_def DCA.language_def unfolding DRA.run_alt_def DBA.run_alt_def DCA.run_alt_def unfolding 1 2 4 by auto qed abbreviation (input) "get k pp \ pp ! k" definition draul :: "('label, 'state) dra list \ ('label, 'state list) dra" where "draul AA \ dra - (UNION (set AA) dra.alphabet) + (\ (dra.alphabet ` set AA)) (map dra.initial AA) (\ a pp. map2 (\ A p. dra.transition A a p) AA pp) (do { k \ [0 ..< length AA]; (f, g) \ dra.accepting (AA ! k); [(f \ get k, g \ get k)] })" lemma draul_get: assumes "length pp = length AA" "k < length AA" shows "infs (P \ get k) (dra.trace (draul AA) w pp) \ infs P (dra.trace (AA ! k) w (pp ! k))" proof - have "infs (P \ get k) (dra.trace (draul AA) w pp) \ infs P (smap (get k) (dra.trace (draul AA) w pp))" by (simp add: comp_def) also have "smap (get k) (dra.trace (draul AA) w pp) = dra.trace (AA ! k) w (pp ! k)" using assms unfolding draul_def by (coinduction arbitrary: w pp) (force) finally show ?thesis by this qed lemma draul_nodes_length: assumes "pp \ DRA.nodes (draul AA)" shows "length pp = length AA" using assms unfolding draul_def by induct auto lemma draul_nodes[intro]: - assumes "INTER (set AA) dra.alphabet = UNION (set AA) dra.alphabet" + assumes "\ (dra.alphabet ` set AA) = \ (dra.alphabet ` set AA)" assumes "pp \ DRA.nodes (draul AA)" "k < length pp" shows "pp ! k \ DRA.nodes (AA ! k)" using assms(2, 3, 1) unfolding draul_def by induct force+ lemma draul_nodes_finite[intro]: - assumes "INTER (set AA) dra.alphabet = UNION (set AA) dra.alphabet" + assumes "\ (dra.alphabet ` set AA) = \ (dra.alphabet ` set AA)" assumes "list_all (finite \ DRA.nodes) AA" shows "finite (DRA.nodes (draul AA))" proof (rule finite_subset) show "DRA.nodes (draul AA) \ listset (map DRA.nodes AA)" using assms(1) by (force simp: listset_member list_all2_conv_all_nth draul_nodes_length) have "finite (listset (map DRA.nodes AA)) \ list_all finite (map DRA.nodes AA)" by (rule listset_finite) (auto simp: list_all_iff) then show "finite (listset (map DRA.nodes AA))" using assms(2) by (simp add: list.pred_map) qed lemma draul_nodes_card: - assumes "INTER (set AA) dra.alphabet = UNION (set AA) dra.alphabet" + assumes "\ (dra.alphabet ` set AA) = \ (dra.alphabet ` set AA)" assumes "list_all (finite \ DRA.nodes) AA" shows "card (DRA.nodes (draul AA)) \ prod_list (map (card \ DRA.nodes) AA)" proof - have "card (DRA.nodes (draul AA)) \ card (listset (map DRA.nodes AA))" proof (rule card_mono) have "finite (listset (map DRA.nodes AA)) \ list_all finite (map DRA.nodes AA)" by (rule listset_finite) (auto simp: list_all_iff) then show "finite (listset (map DRA.nodes AA))" using assms(2) by (simp add: list.pred_map) show "DRA.nodes (draul AA) \ listset (map DRA.nodes AA)" using assms(1) by (force simp: listset_member list_all2_conv_all_nth draul_nodes_length) qed also have "\ = prod_list (map (card \ DRA.nodes) AA)" by simp finally show ?thesis by this qed lemma draul_language[simp]: - assumes "INTER (set AA) dra.alphabet = UNION (set AA) dra.alphabet" - shows "DRA.language (draul AA) = UNION (set AA) DRA.language" + assumes "\ (dra.alphabet ` set AA) = \ (dra.alphabet ` set AA)" + shows "DRA.language (draul AA) = \ (DRA.language ` set AA)" proof safe fix w assume 1: "w \ DRA.language (draul AA)" obtain 2: "dra.run (draul AA) w (dra.initial (draul AA))" "cogen rabin (dra.accepting (draul AA)) (dra.trace (draul AA) w (dra.initial (draul AA)))" using 1 by rule obtain I F where 3: "(I, F) \ set (dra.accepting (draul AA))" "infs I (dra.trace (draul AA) w (dra.initial (draul AA)))" "fins F (dra.trace (draul AA) w (dra.initial (draul AA)))" using 2(2) by blast obtain k P Q where 4: "k < length AA" "I = P \ get k" "F = Q \ get k" "(P, Q) \ set (dra.accepting (AA ! k))" using 3(1) unfolding draul_def List.bind_def by (auto simp: comp_def) - show "w \ UNION (set AA) DRA.language" + show "w \ \ (DRA.language ` set AA)" proof (intro UN_I DRA.language cogen rabin) show "AA ! k \ set AA" using 4(1) by auto show "DRA.run (AA ! k) w (dra.initial (AA ! k))" using assms 2(1) 4(1) unfolding DRA.run_alt_def draul_def by force show "(P, Q) \ set (dra.accepting (AA ! k))" using 4(4) by this show "(P, Q) = (P, Q)" by rule have "True \ infs (P \ get k) (dra.trace (draul AA) w (map dra.initial AA))" using 3(2) unfolding draul_def 4(2) by simp also have "\ \ infs P (dra.trace (AA ! k) w (map dra.initial AA ! k))" using 4(1) by (force intro!: draul_get) also have "map dra.initial AA ! k = dra.initial (AA ! k)" using 4(1) by simp finally show "infs P (dra.trace (AA ! k) w (dra.initial (AA ! k)))" by simp have "False \ infs (Q \ get k) (dra.trace (draul AA) w (map dra.initial AA))" using 3(3) unfolding draul_def 4(3) by simp also have "\ \ infs Q (dra.trace (AA ! k) w (map dra.initial AA ! k))" using 4(1) by (force intro!: draul_get) also have "map dra.initial AA ! k = dra.initial (AA ! k)" using 4(1) by simp finally show "fins Q (dra.trace (AA ! k) w (dra.initial (AA ! k)))" by simp qed next fix A w assume 1: "A \ set AA" "w \ DRA.language A" obtain 2: "dra.run A w (dra.initial A)" "cogen rabin (dra.accepting A) (dra.trace A w (dra.initial A))" using 1(2) by rule obtain I F where 3: "(I, F) \ set (dra.accepting A)" "infs I (dra.trace A w (dra.initial A))" "fins F (dra.trace A w (dra.initial A))" using 2(2) by blast obtain k where 4: "A = AA ! k" "k < length AA" using 1(1) unfolding in_set_conv_nth by auto show "w \ DRA.language (draul AA)" proof (intro DRA.language cogen rabin) show "dra.run (draul AA) w (dra.initial (draul AA))" using 1(1) 2(1) unfolding DRA.run_alt_def draul_def by auto show "(I \ get k, F \ get k) \ set (dra.accepting (draul AA))" unfolding draul_def List.bind_def using 3(1) 4 by (force simp: comp_def) show "(I \ get k, F \ get k) = (I \ get k, F \ get k)" by rule have "infs (I \ get k) (dra.trace (draul AA) w (dra.initial (draul AA))) \ infs (I \ get k) (dra.trace (draul AA) w (map dra.initial AA))" unfolding draul_def by simp also have "\ \ infs I (dra.trace (AA ! k) w (map dra.initial AA ! k))" using 4(2) by (force intro!: draul_get) also have "\ \ True" using 3(2) 4 by simp finally show "infs (I \ get k) (dra.trace (draul AA) w (dra.initial (draul AA)))" by simp have "infs (F \ get k) (dra.trace (draul AA) w (dra.initial (draul AA))) \ infs (F \ get k) (dra.trace (draul AA) w (map dra.initial AA))" unfolding draul_def by simp also have "\ \ infs F (dra.trace (AA ! k) w (map dra.initial AA ! k))" using 4(2) by (force intro!: draul_get) also have "\ \ False" using 3(3) 4 by simp finally show "fins (F \ get k) (dra.trace (draul AA) w (dra.initial (draul AA)))" by simp qed qed end \ No newline at end of file diff --git a/thys/UTP/toolkit/Countable_Set_Extra.thy b/thys/UTP/toolkit/Countable_Set_Extra.thy --- a/thys/UTP/toolkit/Countable_Set_Extra.thy +++ b/thys/UTP/toolkit/Countable_Set_Extra.thy @@ -1,365 +1,365 @@ (******************************************************************************) (* Project: Isabelle/UTP Toolkit *) (* File: Countable_Set_Extra.thy *) (* Authors: Simon Foster and Frank Zeyda *) (* Emails: simon.foster@york.ac.uk and frank.zeyda@york.ac.uk *) (******************************************************************************) section \ Countable Sets: Extra functions and properties \ theory Countable_Set_Extra imports "HOL-Library.Countable_Set_Type" Sequence FSet_Extra begin subsection \ Extra syntax \ notation cempty ("{}\<^sub>c") notation cin (infix "\\<^sub>c" 50) notation cUn (infixl "\\<^sub>c" 65) notation cInt (infixl "\\<^sub>c" 70) notation cDiff (infixl "-\<^sub>c" 65) notation cUnion ("\\<^sub>c_" [900] 900) notation cimage (infixr "`\<^sub>c" 90) abbreviation csubseteq :: "'a cset \ 'a cset \ bool" ("(_/ \\<^sub>c _)" [51, 51] 50) where "A \\<^sub>c B \ A \ B" abbreviation csubset :: "'a cset \ 'a cset \ bool" ("(_/ \\<^sub>c _)" [51, 51] 50) where "A \\<^sub>c B \ A < B" subsection \ Countable set functions \ setup_lifting type_definition_cset lift_definition cnin :: "'a \ 'a cset \ bool" (infix "\\<^sub>c" 50) is "(\)" . definition cBall :: "'a cset \ ('a \ bool) \ bool" where "cBall A P = (\x. x \\<^sub>c A \ P x)" definition cBex :: "'a cset \ ('a \ bool) \ bool" where "cBex A P = (\x. x \\<^sub>c A \ P x)" declare cBall_def [mono,simp] declare cBex_def [mono,simp] syntax "_cBall" :: "pttrn => 'a cset => bool => bool" ("(3\ _\\<^sub>c_./ _)" [0, 0, 10] 10) "_cBex" :: "pttrn => 'a cset => bool => bool" ("(3\ _\\<^sub>c_./ _)" [0, 0, 10] 10) translations "\ x\\<^sub>cA. P" == "CONST cBall A (%x. P)" "\ x\\<^sub>cA. P" == "CONST cBex A (%x. P)" definition cset_Collect :: "('a \ bool) \ 'a cset" where "cset_Collect = (acset o Collect)" lift_definition cset_Coll :: "'a cset \ ('a \ bool) \ 'a cset" is "\ A P. {x \ A. P x}" by (auto) lemma cset_Coll_equiv: "cset_Coll A P = cset_Collect (\ x. x \\<^sub>c A \ P x)" by (simp add:cset_Collect_def cset_Coll_def cin_def) declare cset_Collect_def [simp] syntax "_cColl" :: "pttrn => bool => 'a cset" ("(1{_./ _}\<^sub>c)") translations "{x . P}\<^sub>c" \ "(CONST cset_Collect) (\ x . P)" syntax (xsymbols) "_cCollect" :: "pttrn => 'a cset => bool => 'a cset" ("(1{_ \\<^sub>c/ _./ _}\<^sub>c)") translations "{x \\<^sub>c A. P}\<^sub>c" => "CONST cset_Coll A (\ x. P)" lemma cset_CollectI: "P (a :: 'a::countable) \ a \\<^sub>c {x. P x}\<^sub>c" by (simp add: cin_def) lemma cset_CollI: "\ a \\<^sub>c A; P a \ \ a \\<^sub>c {x \\<^sub>c A. P x}\<^sub>c" by (simp add: cin.rep_eq cset_Coll.rep_eq) lemma cset_CollectD: "(a :: 'a::countable) \\<^sub>c {x. P x}\<^sub>c \ P a" by (simp add: cin_def) lemma cset_Collect_cong: "(\x. P x = Q x) ==> {x. P x}\<^sub>c = {x. Q x}\<^sub>c" by simp text \ Avoid eta-contraction for robust pretty-printing. \ print_translation \ [Syntax_Trans.preserve_binder_abs_tr' @{const_syntax cset_Collect} @{syntax_const "_cColl"}] \ lift_definition cset_set :: "'a list \ 'a cset" is set using countable_finite by blast lemma countable_finite_power: "countable(A) \ countable {B. B \ A \ finite(B)}" by (metis Collect_conj_eq Int_commute countable_Collect_finite_subset) -lift_definition cINTER :: "'a cset \ ('a \ 'b cset) \ 'b cset" is -"\ A f. if (A = {}) then {} else INTER A f" - by (auto) +lift_definition cInter :: "'a cset cset \ 'a cset" ("\\<^sub>c_" [900] 900) + is "\A. if A = {} then {} else \ A" + using countable_INT [of _ _ id] by auto -definition cInter :: "'a cset cset \ 'a cset" ("\\<^sub>c_" [900] 900) where -"\\<^sub>c A = cINTER A id" +abbreviation (input) cINTER :: "'a cset \ ('a \ 'b cset) \ 'b cset" + where "cINTER A f \ cInter (cimage f A)" lift_definition cfinite :: "'a cset \ bool" is finite . lift_definition cInfinite :: "'a cset \ bool" is infinite . lift_definition clist :: "'a::linorder cset \ 'a list" is sorted_list_of_set . lift_definition ccard :: "'a cset \ nat" is card . lift_definition cPow :: "'a cset \ 'a cset cset" is "\ A. {B. B \\<^sub>c A \ cfinite(B)}" proof - fix A have "{B :: 'a cset. B \\<^sub>c A \ cfinite B} = acset ` {B :: 'a set. B \ rcset A \ finite B}" apply (auto simp add: cfinite.rep_eq cin_def less_eq_cset_def countable_finite) using image_iff apply fastforce done moreover have "countable {B :: 'a set. B \ rcset A \ finite B}" by (auto intro: countable_finite_power) ultimately show "countable {B. B \\<^sub>c A \ cfinite B}" by simp qed definition CCollect :: "('a \ bool option) \ 'a cset option" where "CCollect p = (if (None \ range p) then Some (cset_Collect (the \ p)) else None)" definition cset_mapM :: "'a option cset \ 'a cset option" where "cset_mapM A = (if (None \\<^sub>c A) then None else Some (the `\<^sub>c A))" lemma cset_mapM_Some_image [simp]: "cset_mapM (cimage Some A) = Some A" apply (auto simp add: cset_mapM_def) apply (metis cimage_cinsert cinsertI1 option.sel set_cinsert) done definition CCollect_ext :: "('a \ 'b option) \ ('a \ bool option) \ 'b cset option" where "CCollect_ext f p = do { xs \ CCollect p; cset_mapM (f `\<^sub>c xs) }" lemma the_Some_image [simp]: "the ` Some ` xs = xs" by (auto simp add:image_iff) lemma CCollect_ext_Some [simp]: "CCollect_ext Some xs = CCollect xs" apply (case_tac "CCollect xs") apply (auto simp add:CCollect_ext_def) done lift_definition list_of_cset :: "'a :: linorder cset \ 'a list" is sorted_list_of_set . lift_definition fset_cset :: "'a fset \ 'a cset" is id using uncountable_infinite by auto definition cset_count :: "'a cset \ 'a \ nat" where "cset_count A = (if (finite (rcset A)) then (SOME f::'a\nat. inj_on f (rcset A)) else (SOME f::'a\nat. bij_betw f (rcset A) UNIV))" lemma cset_count_inj_seq: "inj_on (cset_count A) (rcset A)" proof (cases "finite (rcset A)") case True note fin = this obtain count :: "'a \ nat" where count_inj: "inj_on count (rcset A)" by (metis countable_def mem_Collect_eq rcset) with fin show ?thesis by (metis (poly_guards_query) cset_count_def someI_ex) next case False note inf = this obtain count :: "'a \ nat" where count_bij: "bij_betw count (rcset A) UNIV" by (metis countableE_infinite inf mem_Collect_eq rcset) with inf have "bij_betw (cset_count A) (rcset A) UNIV" by (metis (poly_guards_query) cset_count_def someI_ex) thus ?thesis by (metis bij_betw_imp_inj_on) qed lemma cset_count_infinite_bij: assumes "infinite (rcset A)" shows "bij_betw (cset_count A) (rcset A) UNIV" proof - from assms obtain count :: "'a \ nat" where count_bij: "bij_betw count (rcset A) UNIV" by (metis countableE_infinite mem_Collect_eq rcset) with assms show ?thesis by (metis (poly_guards_query) cset_count_def someI_ex) qed definition cset_seq :: "'a cset \ (nat \ 'a)" where "cset_seq A i = (if (i \ range (cset_count A) \ inv_into (rcset A) (cset_count A) i \\<^sub>c A) then Some (inv_into (rcset A) (cset_count A) i) else None)" lemma cset_seq_ran: "ran (cset_seq A) = rcset(A)" apply (auto simp add: ran_def cset_seq_def cin.rep_eq) apply (metis cset_count_inj_seq inv_into_f_f rangeI) done lemma cset_seq_inj: "inj cset_seq" proof (rule injI) fix A B :: "'a cset" assume "cset_seq A = cset_seq B" thus "A = B" by (metis cset_seq_ran rcset_inverse) qed lift_definition cset2seq :: "'a cset \ 'a seq" is "(\ A i. if (i \ cset_count A ` rcset A) then inv_into (rcset A) (cset_count A) i else (SOME x. x \\<^sub>c A))" . lemma range_cset2seq: "A \ {}\<^sub>c \ range (Rep_seq (cset2seq A)) = rcset A" by (force intro: someI2 simp add: cset2seq.rep_eq cset_count_inj_seq bot_cset.rep_eq cin.rep_eq) lemma infinite_cset_count_surj: "infinite (rcset A) \ surj (cset_count A)" using bij_betw_imp_surj cset_count_infinite_bij by auto lemma cset2seq_inj: "inj_on cset2seq {A. A \ {}\<^sub>c}" apply (rule inj_onI) apply (simp) apply (metis range_cset2seq rcset_inject) done lift_definition nat_seq2set :: "nat seq \ nat set" is "\ f. prod_encode ` {(x, f x) | x. True}" . lemma inj_nat_seq2set: "inj nat_seq2set" proof (rule injI, transfer) fix f g assume "prod_encode ` {(x, f x) |x. True} = prod_encode ` {(x, g x) |x. True}" hence "{(x, f x) |x. True} = {(x, g x) |x. True}" by (simp add: inj_image_eq_iff[OF inj_prod_encode]) thus "f = g" by (auto simp add: set_eq_iff) qed lift_definition bit_seq_of_nat_set :: "nat set \ bool seq" is "\ A i. i \ A" . lemma bit_seq_of_nat_set_inj: "inj bit_seq_of_nat_set" apply (rule injI) apply transfer apply (auto simp add: fun_eq_iff) done lemma bit_seq_of_nat_cset_bij: "bij bit_seq_of_nat_set" apply (rule bijI) apply (fact bit_seq_of_nat_set_inj) apply transfer apply (rule surjI) apply auto done text \ This function is a partial injection from countable sets of natural sets to natural sets. When used with the Schroeder-Bernstein theorem, it can be used to conjure a total bijection between these two types. \ definition nat_set_cset_collapse :: "nat set cset \ nat set" where "nat_set_cset_collapse = inv bit_seq_of_nat_set \ seq_inj \ cset2seq \ (\ A. (bit_seq_of_nat_set `\<^sub>c A))" lemma nat_set_cset_collapse_inj: "inj_on nat_set_cset_collapse {A. A \ {}\<^sub>c}" proof - have "(`\<^sub>c) bit_seq_of_nat_set ` {A. A \ {}\<^sub>c} \ {A. A \ {}\<^sub>c}" by (auto simp add:cimage.rep_eq) thus ?thesis apply (simp add: nat_set_cset_collapse_def) apply (rule comp_inj_on) apply (meson bit_seq_of_nat_set_inj cset.inj_map injD inj_onI) apply (rule comp_inj_on) apply (metis cset2seq_inj subset_inj_on) apply (rule comp_inj_on) apply (rule subset_inj_on) apply (rule seq_inj) apply (simp) apply (meson UNIV_I bij_imp_bij_inv bij_is_inj bit_seq_of_nat_cset_bij subsetI subset_inj_on) done qed lemma inj_csingle: "inj csingle" by (auto intro: injI simp add: cinsert_def bot_cset.rep_eq) lemma range_csingle: "range csingle \ {A. A \ {}\<^sub>c}" by (auto) lift_definition csets :: "'a set \ 'a cset set" is "\ A. {B. B \ A \ countable B}" by auto lemma csets_finite: "finite A \ finite (csets A)" by (auto simp add: csets_def) lemma csets_infinite: "infinite A \ infinite (csets A)" by (auto simp add: csets_def, metis csets.abs_eq csets.rep_eq finite_countable_subset finite_imageI) lemma csets_UNIV: "csets (UNIV :: 'a set) = (UNIV :: 'a cset set)" by (auto simp add: csets_def, metis image_iff rcset rcset_inverse) lemma infinite_nempty_cset: assumes "infinite (UNIV :: 'a set)" shows "infinite ({A. A \ {}\<^sub>c} :: 'a cset set)" proof - have "infinite (UNIV :: 'a cset set)" by (metis assms csets_UNIV csets_infinite) hence "infinite ((UNIV :: 'a cset set) - {{}\<^sub>c})" by (rule infinite_remove) thus ?thesis by (auto) qed lemma nat_set_cset_partial_bij: obtains f :: "nat set cset \ nat set" where "bij_betw f {A. A \ {}\<^sub>c} UNIV" using Schroeder_Bernstein[OF nat_set_cset_collapse_inj, of UNIV csingle, simplified, OF inj_csingle range_csingle] by (auto) lemma nat_set_cset_bij: obtains f :: "nat set cset \ nat set" where "bij f" proof - obtain g :: "nat set cset \ nat set" where "bij_betw g {A. A \ {}\<^sub>c} UNIV" using nat_set_cset_partial_bij by blast moreover obtain h :: "nat set cset \ nat set cset" where "bij_betw h UNIV {A. A \ {}\<^sub>c}" proof - have "infinite (UNIV :: nat set cset set)" by (metis Finite_Set.finite_set csets_UNIV csets_infinite infinite_UNIV_char_0) then obtain h' :: "nat set cset \ nat set cset" where "bij_betw h' UNIV (UNIV - {{}\<^sub>c})" using infinite_imp_bij_betw[of "UNIV :: nat set cset set" "{}\<^sub>c"] by auto moreover have "(UNIV :: nat set cset set) - {{}\<^sub>c} = {A. A \ {}\<^sub>c}" by (auto) ultimately show ?thesis using that by (auto) qed ultimately have "bij (g \ h)" using bij_betw_trans by blast with that show ?thesis by (auto) qed definition "nat_set_cset_bij = (SOME f :: nat set cset \ nat set. bij f)" lemma bij_nat_set_cset_bij: "bij nat_set_cset_bij" by (metis nat_set_cset_bij nat_set_cset_bij_def someI_ex) lemma inj_on_image_csets: "inj_on f A \ inj_on ((`\<^sub>c) f) (csets A)" by (fastforce simp add: inj_on_def cimage_def cin_def csets_def) lemma image_csets_surj: "\ inj_on f A; f ` A = B \ \ (`\<^sub>c) f ` csets A = csets B" apply (auto simp add: cimage_def csets_def image_mono map_fun_def) apply (simp add: image_comp) apply (auto simp add: image_Collect) apply (erule subset_imageE) using countable_image_inj_on subset_inj_on by blast lemma bij_betw_image_csets: "bij_betw f A B \ bij_betw ((`\<^sub>c) f) (csets A) (csets B)" by (simp add: bij_betw_def inj_on_image_csets image_csets_surj) end \ No newline at end of file diff --git a/thys/UTP/utp/utp_healthy.thy b/thys/UTP/utp/utp_healthy.thy --- a/thys/UTP/utp/utp_healthy.thy +++ b/thys/UTP/utp/utp_healthy.thy @@ -1,318 +1,318 @@ section \ Healthiness Conditions \ theory utp_healthy imports utp_pred_laws begin subsection \ Main Definitions \ text \ We collect closure laws for healthiness conditions in the following theorem attribute. \ named_theorems closure type_synonym '\ health = "'\ upred \ '\ upred" text \ A predicate $P$ is healthy, under healthiness function $H$, if $P$ is a fixed-point of $H$. \ definition Healthy :: "'\ upred \ '\ health \ bool" (infix "is" 30) where "P is H \ (H P = P)" lemma Healthy_def': "P is H \ (H P = P)" unfolding Healthy_def by auto lemma Healthy_if: "P is H \ (H P = P)" unfolding Healthy_def by auto lemma Healthy_intro: "H(P) = P \ P is H" by (simp add: Healthy_def) declare Healthy_def' [upred_defs] abbreviation Healthy_carrier :: "'\ health \ '\ upred set" ("\_\\<^sub>H") where "\H\\<^sub>H \ {P. P is H}" lemma Healthy_carrier_image: "A \ \\\\<^sub>H \ \ ` A = A" by (auto simp add: image_def, (metis Healthy_if mem_Collect_eq subsetCE)+) lemma Healthy_carrier_Collect: "A \ \H\\<^sub>H \ A = {H(P) | P. P \ A}" by (simp add: Healthy_carrier_image Setcompr_eq_image) lemma Healthy_func: "\ F \ \\\<^sub>1\\<^sub>H \ \\\<^sub>2\\<^sub>H; P is \\<^sub>1 \ \ \\<^sub>2(F(P)) = F(P)" using Healthy_if by blast lemma Healthy_comp: "\ P is \\<^sub>1; P is \\<^sub>2 \ \ P is \\<^sub>1 \ \\<^sub>2" by (simp add: Healthy_def) lemma Healthy_apply_closed: assumes "F \ \H\\<^sub>H \ \H\\<^sub>H" "P is H" shows "F(P) is H" using assms(1) assms(2) by auto lemma Healthy_set_image_member: "\ P \ F ` A; \ x. F x is H \ \ P is H" by blast lemma Healthy_case_prod [closure]: "\ \ x y. P x y is H \ \ case_prod P v is H" by (simp add: prod.case_eq_if) lemma Healthy_SUPREMUM: - "A \ \H\\<^sub>H \ SUPREMUM A H = \ A" + "A \ \H\\<^sub>H \ Sup (H ` A) = \ A" by (drule Healthy_carrier_image, presburger) lemma Healthy_INFIMUM: - "A \ \H\\<^sub>H \ INFIMUM A H = \ A" + "A \ \H\\<^sub>H \ Inf (H ` A) = \ A" by (drule Healthy_carrier_image, presburger) lemma Healthy_nu [closure]: assumes "mono F" "F \ \id\\<^sub>H \ \H\\<^sub>H" shows "\ F is H" by (metis (mono_tags) Healthy_def Healthy_func assms eq_id_iff lfp_unfold) lemma Healthy_mu [closure]: assumes "mono F" "F \ \id\\<^sub>H \ \H\\<^sub>H" shows "\ F is H" by (metis (mono_tags) Healthy_def Healthy_func assms eq_id_iff gfp_unfold) lemma Healthy_subset_member: "\ A \ \H\\<^sub>H; P \ A \ \ H(P) = P" by (meson Ball_Collect Healthy_if) lemma is_Healthy_subset_member: "\ A \ \H\\<^sub>H; P \ A \ \ P is H" by blast subsection \ Properties of Healthiness Conditions \ definition Idempotent :: "'\ health \ bool" where "Idempotent(H) \ (\ P. H(H(P)) = H(P))" abbreviation Monotonic :: "'\ health \ bool" where "Monotonic(H) \ mono H" definition IMH :: "'\ health \ bool" where "IMH(H) \ Idempotent(H) \ Monotonic(H)" definition Antitone :: "'\ health \ bool" where "Antitone(H) \ (\ P Q. Q \ P \ (H(P) \ H(Q)))" definition Conjunctive :: "'\ health \ bool" where "Conjunctive(H) \ (\ Q. \ P. H(P) = (P \ Q))" definition FunctionalConjunctive :: "'\ health \ bool" where "FunctionalConjunctive(H) \ (\ F. \ P. H(P) = (P \ F(P)) \ Monotonic(F))" definition WeakConjunctive :: "'\ health \ bool" where "WeakConjunctive(H) \ (\ P. \ Q. H(P) = (P \ Q))" definition Disjunctuous :: "'\ health \ bool" where [upred_defs]: "Disjunctuous H = (\ P Q. H(P \ Q) = (H(P) \ H(Q)))" definition Continuous :: "'\ health \ bool" where [upred_defs]: "Continuous H = (\ A. A \ {} \ H (\ A) = \ (H ` A))" lemma Healthy_Idempotent [closure]: "Idempotent H \ H(P) is H" by (simp add: Healthy_def Idempotent_def) lemma Healthy_range: "Idempotent H \ range H = \H\\<^sub>H" by (auto simp add: image_def Healthy_if Healthy_Idempotent, metis Healthy_if) lemma Idempotent_id [simp]: "Idempotent id" by (simp add: Idempotent_def) lemma Idempotent_comp [intro]: "\ Idempotent f; Idempotent g; f \ g = g \ f \ \ Idempotent (f \ g)" by (auto simp add: Idempotent_def comp_def, metis) lemma Idempotent_image: "Idempotent f \ f ` f ` A = f ` A" by (metis (mono_tags, lifting) Idempotent_def image_cong image_image) lemma Monotonic_id [simp]: "Monotonic id" by (simp add: monoI) lemma Monotonic_id' [closure]: "mono (\ X. X)" by (simp add: monoI) lemma Monotonic_const [closure]: "Monotonic (\ x. c)" by (simp add: mono_def) lemma Monotonic_comp [intro]: "\ Monotonic f; Monotonic g \ \ Monotonic (f \ g)" by (simp add: mono_def) lemma Monotonic_inf [closure]: assumes "Monotonic P" "Monotonic Q" shows "Monotonic (\ X. P(X) \ Q(X))" using assms by (simp add: mono_def, rel_auto) lemma Monotonic_cond [closure]: assumes "Monotonic P" "Monotonic Q" shows "Monotonic (\ X. P(X) \ b \ Q(X))" by (simp add: assms cond_monotonic) lemma Conjuctive_Idempotent: "Conjunctive(H) \ Idempotent(H)" by (auto simp add: Conjunctive_def Idempotent_def) lemma Conjunctive_Monotonic: "Conjunctive(H) \ Monotonic(H)" unfolding Conjunctive_def mono_def using dual_order.trans by fastforce lemma Conjunctive_conj: assumes "Conjunctive(HC)" shows "HC(P \ Q) = (HC(P) \ Q)" using assms unfolding Conjunctive_def by (metis utp_pred_laws.inf.assoc utp_pred_laws.inf.commute) lemma Conjunctive_distr_conj: assumes "Conjunctive(HC)" shows "HC(P \ Q) = (HC(P) \ HC(Q))" using assms unfolding Conjunctive_def by (metis Conjunctive_conj assms utp_pred_laws.inf.assoc utp_pred_laws.inf_right_idem) lemma Conjunctive_distr_disj: assumes "Conjunctive(HC)" shows "HC(P \ Q) = (HC(P) \ HC(Q))" using assms unfolding Conjunctive_def using utp_pred_laws.inf_sup_distrib2 by fastforce lemma Conjunctive_distr_cond: assumes "Conjunctive(HC)" shows "HC(P \ b \ Q) = (HC(P) \ b \ HC(Q))" using assms unfolding Conjunctive_def by (metis cond_conj_distr utp_pred_laws.inf_commute) lemma FunctionalConjunctive_Monotonic: "FunctionalConjunctive(H) \ Monotonic(H)" unfolding FunctionalConjunctive_def by (metis mono_def utp_pred_laws.inf_mono) lemma WeakConjunctive_Refinement: assumes "WeakConjunctive(HC)" shows "P \ HC(P)" using assms unfolding WeakConjunctive_def by (metis utp_pred_laws.inf.cobounded1) lemma WeakCojunctive_Healthy_Refinement: assumes "WeakConjunctive(HC)" and "P is HC" shows "HC(P) \ P" using assms unfolding WeakConjunctive_def Healthy_def by simp lemma WeakConjunctive_implies_WeakConjunctive: "Conjunctive(H) \ WeakConjunctive(H)" unfolding WeakConjunctive_def Conjunctive_def by pred_auto declare Conjunctive_def [upred_defs] declare mono_def [upred_defs] lemma Disjunctuous_Monotonic: "Disjunctuous H \ Monotonic H" by (metis Disjunctuous_def mono_def semilattice_sup_class.le_iff_sup) lemma ContinuousD [dest]: "\ Continuous H; A \ {} \ \ H (\ A) = (\ P\A. H(P))" by (simp add: Continuous_def) lemma Continuous_Disjunctous: "Continuous H \ Disjunctuous H" apply (auto simp add: Continuous_def Disjunctuous_def) apply (rename_tac P Q) apply (drule_tac x="{P,Q}" in spec) apply (simp) done lemma Continuous_Monotonic [closure]: "Continuous H \ Monotonic H" by (simp add: Continuous_Disjunctous Disjunctuous_Monotonic) lemma Continuous_comp [intro]: "\ Continuous f; Continuous g \ \ Continuous (f \ g)" by (simp add: Continuous_def) lemma Continuous_const [closure]: "Continuous (\ X. P)" by pred_auto lemma Continuous_cond [closure]: assumes "Continuous F" "Continuous G" shows "Continuous (\ X. F(X) \ b \ G(X))" using assms by (pred_auto) text \ Closure laws derived from continuity \ lemma Sup_Continuous_closed [closure]: "\ Continuous H; \ i. i \ A \ P(i) is H; A \ {} \ \ (\ i\A. P(i)) is H" by (drule ContinuousD[of H "P ` A"], simp add: UINF_mem_UNIV[THEN sym] UINF_as_Sup[THEN sym]) (metis (no_types, lifting) Healthy_def' SUP_cong image_image) lemma UINF_mem_Continuous_closed [closure]: "\ Continuous H; \ i. i \ A \ P(i) is H; A \ {} \ \ (\ i\A \ P(i)) is H" by (simp add: Sup_Continuous_closed UINF_as_Sup_collect) lemma UINF_mem_Continuous_closed_pair [closure]: assumes "Continuous H" "\ i j. (i, j) \ A \ P i j is H" "A \ {}" shows "(\ (i,j)\A \ P i j) is H" proof - have "(\ (i,j)\A \ P i j) = (\ x\A \ P (fst x) (snd x))" by (rel_auto) also have "... is H" by (metis (mono_tags) UINF_mem_Continuous_closed assms(1) assms(2) assms(3) prod.collapse) finally show ?thesis . qed lemma UINF_mem_Continuous_closed_triple [closure]: assumes "Continuous H" "\ i j k. (i, j, k) \ A \ P i j k is H" "A \ {}" shows "(\ (i,j,k)\A \ P i j k) is H" proof - have "(\ (i,j,k)\A \ P i j k) = (\ x\A \ P (fst x) (fst (snd x)) (snd (snd x)))" by (rel_auto) also have "... is H" by (metis (mono_tags) UINF_mem_Continuous_closed assms(1) assms(2) assms(3) prod.collapse) finally show ?thesis . qed lemma UINF_mem_Continuous_closed_quad [closure]: assumes "Continuous H" "\ i j k l. (i, j, k, l) \ A \ P i j k l is H" "A \ {}" shows "(\ (i,j,k,l)\A \ P i j k l) is H" proof - have "(\ (i,j,k,l)\A \ P i j k l) = (\ x\A \ P (fst x) (fst (snd x)) (fst (snd (snd x))) (snd (snd (snd x))))" by (rel_auto) also have "... is H" by (metis (mono_tags) UINF_mem_Continuous_closed assms(1) assms(2) assms(3) prod.collapse) finally show ?thesis . qed lemma UINF_mem_Continuous_closed_quint [closure]: assumes "Continuous H" "\ i j k l m. (i, j, k, l, m) \ A \ P i j k l m is H" "A \ {}" shows "(\ (i,j,k,l,m)\A \ P i j k l m) is H" proof - have "(\ (i,j,k,l,m)\A \ P i j k l m) = (\ x\A \ P (fst x) (fst (snd x)) (fst (snd (snd x))) (fst (snd (snd (snd x)))) (snd (snd (snd (snd x)))))" by (rel_auto) also have "... is H" by (metis (mono_tags) UINF_mem_Continuous_closed assms(1) assms(2) assms(3) prod.collapse) finally show ?thesis . qed lemma UINF_ind_closed [closure]: assumes "Continuous H" "\ i. P i = true" "\ i. Q i is H" shows "UINF P Q is H" proof - from assms(2) have "UINF P Q = (\ i \ Q i)" by (rel_auto) also have "... is H" using UINF_mem_Continuous_closed[of H UNIV P] by (simp add: Sup_Continuous_closed UINF_as_Sup_collect' assms) finally show ?thesis . qed text \ All continuous functions are also Scott-continuous \ lemma sup_continuous_Continuous [closure]: "Continuous F \ sup_continuous F" by (simp add: Continuous_def sup_continuous_def) lemma USUP_healthy: "A \ \H\\<^sub>H \ (\ P\A \ F(P)) = (\ P\A \ F(H(P)))" by (rule USUP_cong, simp add: Healthy_subset_member) lemma UINF_healthy: "A \ \H\\<^sub>H \ (\ P\A \ F(P)) = (\ P\A \ F(H(P)))" by (rule UINF_cong, simp add: Healthy_subset_member) end \ No newline at end of file diff --git a/thys/UTP/utp/utp_rel_laws.thy b/thys/UTP/utp/utp_rel_laws.thy --- a/thys/UTP/utp/utp_rel_laws.thy +++ b/thys/UTP/utp/utp_rel_laws.thy @@ -1,997 +1,997 @@ section \ Relational Calculus Laws \ theory utp_rel_laws imports utp_rel utp_recursion begin subsection \ Conditional Laws \ lemma comp_cond_left_distr: "((P \ b \\<^sub>r Q) ;; R) = ((P ;; R) \ b \\<^sub>r (Q ;; R))" by (rel_auto) lemma cond_seq_left_distr: "out\ \ b \ ((P \ b \ Q) ;; R) = ((P ;; R) \ b \ (Q ;; R))" by (rel_auto) lemma cond_seq_right_distr: "in\ \ b \ (P ;; (Q \ b \ R)) = ((P ;; Q) \ b \ (P ;; R))" by (rel_auto) text \ Alternative expression of conditional using assumptions and choice \ lemma rcond_rassume_expand: "P \ b \\<^sub>r Q = ([b]\<^sup>\ ;; P) \ ([(\ b)]\<^sup>\ ;; Q)" by (rel_auto) subsection \ Precondition and Postcondition Laws \ theorem precond_equiv: "P = (P ;; true) \ (out\ \ P)" by (rel_auto) theorem postcond_equiv: "P = (true ;; P) \ (in\ \ P)" by (rel_auto) lemma precond_right_unit: "out\ \ p \ (p ;; true) = p" by (metis precond_equiv) lemma postcond_left_unit: "in\ \ p \ (true ;; p) = p" by (metis postcond_equiv) theorem precond_left_zero: assumes "out\ \ p" "p \ false" shows "(true ;; p) = true" using assms by (rel_auto) theorem feasibile_iff_true_right_zero: "P ;; true = true \ `\ out\ \ P`" by (rel_auto) subsection \ Sequential Composition Laws \ lemma seqr_assoc: "(P ;; Q) ;; R = P ;; (Q ;; R)" by (rel_auto) lemma seqr_left_unit [simp]: "II ;; P = P" by (rel_auto) lemma seqr_right_unit [simp]: "P ;; II = P" by (rel_auto) lemma seqr_left_zero [simp]: "false ;; P = false" by pred_auto lemma seqr_right_zero [simp]: "P ;; false = false" by pred_auto lemma impl_seqr_mono: "\ `P \ Q`; `R \ S` \ \ `(P ;; R) \ (Q ;; S)`" by (pred_blast) lemma seqr_mono: "\ P\<^sub>1 \ P\<^sub>2; Q\<^sub>1 \ Q\<^sub>2 \ \ (P\<^sub>1 ;; Q\<^sub>1) \ (P\<^sub>2 ;; Q\<^sub>2)" by (rel_blast) lemma seqr_monotonic: "\ mono P; mono Q \ \ mono (\ X. P X ;; Q X)" by (simp add: mono_def, rel_blast) lemma Monotonic_seqr_tail [closure]: assumes "Monotonic F" shows "Monotonic (\ X. P ;; F(X))" by (simp add: assms monoD monoI seqr_mono) lemma seqr_exists_left: "((\ $x \ P) ;; Q) = (\ $x \ (P ;; Q))" by (rel_auto) lemma seqr_exists_right: "(P ;; (\ $x\ \ Q)) = (\ $x\ \ (P ;; Q))" by (rel_auto) lemma seqr_or_distl: "((P \ Q) ;; R) = ((P ;; R) \ (Q ;; R))" by (rel_auto) lemma seqr_or_distr: "(P ;; (Q \ R)) = ((P ;; Q) \ (P ;; R))" by (rel_auto) lemma seqr_inf_distl: "((P \ Q) ;; R) = ((P ;; R) \ (Q ;; R))" by (rel_auto) lemma seqr_inf_distr: "(P ;; (Q \ R)) = ((P ;; Q) \ (P ;; R))" by (rel_auto) lemma seqr_and_distr_ufunc: "ufunctional P \ (P ;; (Q \ R)) = ((P ;; Q) \ (P ;; R))" by (rel_auto) lemma seqr_and_distl_uinj: "uinj R \ ((P \ Q) ;; R) = ((P ;; R) \ (Q ;; R))" by (rel_auto) lemma seqr_unfold: "(P ;; Q) = (\<^bold>\ v \ P\\v\/$\<^bold>v\\ \ Q\\v\/$\<^bold>v\)" by (rel_auto) lemma seqr_middle: assumes "vwb_lens x" shows "(P ;; Q) = (\<^bold>\ v \ P\\v\/$x\\ ;; Q\\v\/$x\)" using assms by (rel_auto', metis vwb_lens_wb wb_lens.source_stability) lemma seqr_left_one_point: assumes "vwb_lens x" shows "((P \ $x\ =\<^sub>u \v\) ;; Q) = (P\\v\/$x\\ ;; Q\\v\/$x\)" using assms by (rel_auto, metis vwb_lens_wb wb_lens.get_put) lemma seqr_right_one_point: assumes "vwb_lens x" shows "(P ;; ($x =\<^sub>u \v\ \ Q)) = (P\\v\/$x\\ ;; Q\\v\/$x\)" using assms by (rel_auto, metis vwb_lens_wb wb_lens.get_put) lemma seqr_left_one_point_true: assumes "vwb_lens x" shows "((P \ $x\) ;; Q) = (P\true/$x\\ ;; Q\true/$x\)" by (metis assms seqr_left_one_point true_alt_def upred_eq_true) lemma seqr_left_one_point_false: assumes "vwb_lens x" shows "((P \ \$x\) ;; Q) = (P\false/$x\\ ;; Q\false/$x\)" by (metis assms false_alt_def seqr_left_one_point upred_eq_false) lemma seqr_right_one_point_true: assumes "vwb_lens x" shows "(P ;; ($x \ Q)) = (P\true/$x\\ ;; Q\true/$x\)" by (metis assms seqr_right_one_point true_alt_def upred_eq_true) lemma seqr_right_one_point_false: assumes "vwb_lens x" shows "(P ;; (\$x \ Q)) = (P\false/$x\\ ;; Q\false/$x\)" by (metis assms false_alt_def seqr_right_one_point upred_eq_false) lemma seqr_insert_ident_left: assumes "vwb_lens x" "$x\ \ P" "$x \ Q" shows "(($x\ =\<^sub>u $x \ P) ;; Q) = (P ;; Q)" using assms by (rel_simp, meson vwb_lens_wb wb_lens_weak weak_lens.put_get) lemma seqr_insert_ident_right: assumes "vwb_lens x" "$x\ \ P" "$x \ Q" shows "(P ;; ($x\ =\<^sub>u $x \ Q)) = (P ;; Q)" using assms by (rel_simp, metis (no_types, hide_lams) vwb_lens_def wb_lens_def weak_lens.put_get) lemma seq_var_ident_lift: assumes "vwb_lens x" "$x\ \ P" "$x \ Q" shows "(($x\ =\<^sub>u $x \ P) ;; ($x\ =\<^sub>u $x \ Q)) = ($x\ =\<^sub>u $x \ (P ;; Q))" using assms by (rel_auto', metis (no_types, lifting) vwb_lens_wb wb_lens_weak weak_lens.put_get) lemma seqr_bool_split: assumes "vwb_lens x" shows "P ;; Q = (P\true/$x\\ ;; Q\true/$x\ \ P\false/$x\\ ;; Q\false/$x\)" using assms by (subst seqr_middle[of x], simp_all) lemma cond_inter_var_split: assumes "vwb_lens x" shows "(P \ $x\ \ Q) ;; R = (P\true/$x\\ ;; R\true/$x\ \ Q\false/$x\\ ;; R\false/$x\)" proof - have "(P \ $x\ \ Q) ;; R = (($x\ \ P) ;; R \ (\ $x\ \ Q) ;; R)" by (simp add: cond_def seqr_or_distl) also have "... = ((P \ $x\) ;; R \ (Q \ \$x\) ;; R)" by (rel_auto) also have "... = (P\true/$x\\ ;; R\true/$x\ \ Q\false/$x\\ ;; R\false/$x\)" by (simp add: seqr_left_one_point_true seqr_left_one_point_false assms) finally show ?thesis . qed theorem seqr_pre_transfer: "in\ \ q \ ((P \ q) ;; R) = (P ;; (q\<^sup>- \ R))" by (rel_auto) theorem seqr_pre_transfer': "((P \ \q\\<^sub>>) ;; R) = (P ;; (\q\\<^sub>< \ R))" by (rel_auto) theorem seqr_post_out: "in\ \ r \ (P ;; (Q \ r)) = ((P ;; Q) \ r)" by (rel_blast) lemma seqr_post_var_out: fixes x :: "(bool \ '\)" shows "(P ;; (Q \ $x\)) = ((P ;; Q) \ $x\)" by (rel_auto) theorem seqr_post_transfer: "out\ \ q \ (P ;; (q \ R)) = ((P \ q\<^sup>-) ;; R)" by (rel_auto) lemma seqr_pre_out: "out\ \ p \ ((p \ Q) ;; R) = (p \ (Q ;; R))" by (rel_blast) lemma seqr_pre_var_out: fixes x :: "(bool \ '\)" shows "(($x \ P) ;; Q) = ($x \ (P ;; Q))" by (rel_auto) lemma seqr_true_lemma: "(P = (\ ((\ P) ;; true))) = (P = (P ;; true))" by (rel_auto) lemma seqr_to_conj: "\ out\ \ P; in\ \ Q \ \ (P ;; Q) = (P \ Q)" by (metis postcond_left_unit seqr_pre_out utp_pred_laws.inf_top.right_neutral) lemma shEx_lift_seq_1 [uquant_lift]: "((\<^bold>\ x \ P x) ;; Q) = (\<^bold>\ x \ (P x ;; Q))" by rel_auto lemma shEx_mem_lift_seq_1 [uquant_lift]: assumes "out\ \ A" shows "((\<^bold>\ x \ A \ P x) ;; Q) = (\<^bold>\ x \ A \ (P x ;; Q))" using assms by rel_blast lemma shEx_lift_seq_2 [uquant_lift]: "(P ;; (\<^bold>\ x \ Q x)) = (\<^bold>\ x \ (P ;; Q x))" by rel_auto lemma shEx_mem_lift_seq_2 [uquant_lift]: assumes "in\ \ A" shows "(P ;; (\<^bold>\ x \ A \ Q x)) = (\<^bold>\ x \ A \ (P ;; Q x))" using assms by rel_blast subsection \ Iterated Sequential Composition Laws \ lemma iter_seqr_nil [simp]: "(;; i : [] \ P(i)) = II" by (simp add: seqr_iter_def) lemma iter_seqr_cons [simp]: "(;; i : (x # xs) \ P(i)) = P(x) ;; (;; i : xs \ P(i))" by (simp add: seqr_iter_def) subsection \ Quantale Laws \ lemma seq_Sup_distl: "P ;; (\ A) = (\ Q\A. P ;; Q)" by (transfer, auto) lemma seq_Sup_distr: "(\ A) ;; Q = (\ P\A. P ;; Q)" by (transfer, auto) lemma seq_UINF_distl: "P ;; (\ Q\A \ F(Q)) = (\ Q\A \ P ;; F(Q))" by (simp add: UINF_as_Sup_collect seq_Sup_distl) lemma seq_UINF_distl': "P ;; (\ Q \ F(Q)) = (\ Q \ P ;; F(Q))" by (metis UINF_mem_UNIV seq_UINF_distl) lemma seq_UINF_distr: "(\ P\A \ F(P)) ;; Q = (\ P\A \ F(P) ;; Q)" by (simp add: UINF_as_Sup_collect seq_Sup_distr) lemma seq_UINF_distr': "(\ P \ F(P)) ;; Q = (\ P \ F(P) ;; Q)" by (metis UINF_mem_UNIV seq_UINF_distr) lemma seq_SUP_distl: "P ;; (\i\A. Q(i)) = (\i\A. P ;; Q(i))" by (metis image_image seq_Sup_distl) lemma seq_SUP_distr: "(\i\A. P(i)) ;; Q = (\i\A. P(i) ;; Q)" by (simp add: seq_Sup_distr) subsection \ Skip Laws \ lemma cond_skip: "out\ \ b \ (b \ II) = (II \ b\<^sup>-)" by (rel_auto) lemma pre_skip_post: "(\b\\<^sub>< \ II) = (II \ \b\\<^sub>>)" by (rel_auto) lemma skip_var: fixes x :: "(bool \ '\)" shows "($x \ II) = (II \ $x\)" by (rel_auto) lemma skip_r_unfold: "vwb_lens x \ II = ($x\ =\<^sub>u $x \ II\\<^sub>\x)" by (rel_simp, metis mwb_lens.put_put vwb_lens_mwb vwb_lens_wb wb_lens.get_put) lemma skip_r_alpha_eq: "II = ($\<^bold>v\ =\<^sub>u $\<^bold>v)" by (rel_auto) lemma skip_ra_unfold: "II\<^bsub>x;y\<^esub> = ($x\ =\<^sub>u $x \ II\<^bsub>y\<^esub>)" by (rel_auto) lemma skip_res_as_ra: "\ vwb_lens y; x +\<^sub>L y \\<^sub>L 1\<^sub>L; x \ y \ \ II\\<^sub>\x = II\<^bsub>y\<^esub>" apply (rel_auto) apply (metis (no_types, lifting) lens_indep_def) apply (metis vwb_lens.put_eq) done subsection \ Assignment Laws \ lemma assigns_subst [usubst]: "\\\\<^sub>s \ \\\\<^sub>a = \\ \ \\\<^sub>a" by (rel_auto) lemma assigns_r_comp: "(\\\\<^sub>a ;; P) = (\\\\<^sub>s \ P)" by (rel_auto) lemma assigns_r_feasible: "(\\\\<^sub>a ;; true) = true" by (rel_auto) lemma assign_subst [usubst]: "\ mwb_lens x; mwb_lens y \ \ [$x \\<^sub>s \u\\<^sub><] \ (y := v) = (x, y) := (u, [x \\<^sub>s u] \ v)" by (rel_auto) lemma assign_vacuous_skip: assumes "vwb_lens x" shows "(x := &x) = II" using assms by rel_auto text \ The following law shows the case for the above law when $x$ is only mainly-well behaved. We require that the state is one of those in which $x$ is well defined using and assumption. \ lemma assign_vacuous_assume: assumes "mwb_lens x" shows "[(&\<^bold>v \\<^sub>u \\\<^bsub>x\<^esub>\)]\<^sup>\ ;; (x := &x) = [(&\<^bold>v \\<^sub>u \\\<^bsub>x\<^esub>\)]\<^sup>\" using assms by rel_auto lemma assign_simultaneous: assumes "vwb_lens y" "x \ y" shows "(x,y) := (e, &y) = (x := e)" by (simp add: assms usubst_upd_comm usubst_upd_var_id) lemma assigns_idem: "mwb_lens x \ (x,x) := (u,v) = (x := v)" by (simp add: usubst) lemma assigns_comp: "(\f\\<^sub>a ;; \g\\<^sub>a) = \g \ f\\<^sub>a" by (simp add: assigns_r_comp usubst) lemma assigns_cond: "(\f\\<^sub>a \ b \\<^sub>r \g\\<^sub>a) = \f \ b \\<^sub>s g\\<^sub>a" by (rel_auto) lemma assigns_r_conv: "bij f \ \f\\<^sub>a\<^sup>- = \inv f\\<^sub>a" by (rel_auto, simp_all add: bij_is_inj bij_is_surj surj_f_inv_f) lemma assign_pred_transfer: fixes x :: "('a \ '\)" assumes "$x \ b" "out\ \ b" shows "(b \ x := v) = (x := v \ b\<^sup>-)" using assms by (rel_blast) lemma assign_r_comp: "x := u ;; P = P\\u\\<^sub>" by (simp add: assigns_r_comp usubst alpha) lemma assign_test: "mwb_lens x \ (x := \u\ ;; x := \v\) = (x := \v\)" by (simp add: assigns_comp usubst) lemma assign_twice: "\ mwb_lens x; x \ f \ \ (x := e ;; x := f) = (x := f)" by (simp add: assigns_comp usubst unrest) lemma assign_commute: assumes "x \ y" "x \ f" "y \ e" shows "(x := e ;; y := f) = (y := f ;; x := e)" using assms by (rel_simp, simp_all add: lens_indep_comm) lemma assign_cond: fixes x :: "('a \ '\)" assumes "out\ \ b" shows "(x := e ;; (P \ b \ Q)) = ((x := e ;; P) \ (b\\e\\<^sub>) \ (x := e ;; Q))" by (rel_auto) lemma assign_rcond: fixes x :: "('a \ '\)" shows "(x := e ;; (P \ b \\<^sub>r Q)) = ((x := e ;; P) \ (b\e/x\) \\<^sub>r (x := e ;; Q))" by (rel_auto) lemma assign_r_alt_def: fixes x :: "('a \ '\)" shows "x := v = II\\v\\<^sub>" by (rel_auto) lemma assigns_r_ufunc: "ufunctional \f\\<^sub>a" by (rel_auto) lemma assigns_r_uinj: "inj f \ uinj \f\\<^sub>a" by (rel_simp, simp add: inj_eq) lemma assigns_r_swap_uinj: "\ vwb_lens x; vwb_lens y; x \ y \ \ uinj ((x,y) := (&y,&x))" by (metis assigns_r_uinj pr_var_def swap_usubst_inj) lemma assign_unfold: "vwb_lens x \ (x := v) = ($x\ =\<^sub>u \v\\<^sub>< \ II\\<^sub>\x)" apply (rel_auto, auto simp add: comp_def) using vwb_lens.put_eq by fastforce subsection \ Non-deterministic Assignment Laws \ lemma nd_assign_comp: "x \ y \ x := * ;; y := * = x,y := *" apply (rel_auto) using lens_indep_comm by fastforce+ lemma nd_assign_assign: "\ vwb_lens x; x \ e \ \ x := * ;; x := e = x := e" by (rel_auto) subsection \ Converse Laws \ lemma convr_invol [simp]: "p\<^sup>-\<^sup>- = p" by pred_auto lemma lit_convr [simp]: "\v\\<^sup>- = \v\" by pred_auto lemma uivar_convr [simp]: fixes x :: "('a \ '\)" shows "($x)\<^sup>- = $x\" by pred_auto lemma uovar_convr [simp]: fixes x :: "('a \ '\)" shows "($x\)\<^sup>- = $x" by pred_auto lemma uop_convr [simp]: "(uop f u)\<^sup>- = uop f (u\<^sup>-)" by (pred_auto) lemma bop_convr [simp]: "(bop f u v)\<^sup>- = bop f (u\<^sup>-) (v\<^sup>-)" by (pred_auto) lemma eq_convr [simp]: "(p =\<^sub>u q)\<^sup>- = (p\<^sup>- =\<^sub>u q\<^sup>-)" by (pred_auto) lemma not_convr [simp]: "(\ p)\<^sup>- = (\ p\<^sup>-)" by (pred_auto) lemma disj_convr [simp]: "(p \ q)\<^sup>- = (q\<^sup>- \ p\<^sup>-)" by (pred_auto) lemma conj_convr [simp]: "(p \ q)\<^sup>- = (q\<^sup>- \ p\<^sup>-)" by (pred_auto) lemma seqr_convr [simp]: "(p ;; q)\<^sup>- = (q\<^sup>- ;; p\<^sup>-)" by (rel_auto) lemma pre_convr [simp]: "\p\\<^sub><\<^sup>- = \p\\<^sub>>" by (rel_auto) lemma post_convr [simp]: "\p\\<^sub>>\<^sup>- = \p\\<^sub><" by (rel_auto) subsection \ Assertion and Assumption Laws \ declare sublens_def [lens_defs del] lemma assume_false: "[false]\<^sup>\ = false" by (rel_auto) lemma assume_true: "[true]\<^sup>\ = II" by (rel_auto) lemma assume_seq: "[b]\<^sup>\ ;; [c]\<^sup>\ = [(b \ c)]\<^sup>\" by (rel_auto) lemma assert_false: "{false}\<^sub>\ = true" by (rel_auto) lemma assert_true: "{true}\<^sub>\ = II" by (rel_auto) lemma assert_seq: "{b}\<^sub>\ ;; {c}\<^sub>\ = {(b \ c)}\<^sub>\" by (rel_auto) subsection \ Frame and Antiframe Laws \ named_theorems frame lemma frame_all [frame]: "\:[P] = P" by (rel_auto) lemma frame_none [frame]: "\:[P] = (P \ II)" by (rel_auto) lemma frame_commute: assumes "$y \ P" "$y\ \ P" "$x \ Q" "$x\ \ Q" "x \ y" shows "x:[P] ;; y:[Q] = y:[Q] ;; x:[P]" apply (insert assms) apply (rel_auto) apply (rename_tac s s' s\<^sub>0) apply (subgoal_tac "(s \\<^sub>L s' on y) \\<^sub>L s\<^sub>0 on x = s\<^sub>0 \\<^sub>L s' on y") apply (metis lens_indep_get lens_indep_sym lens_override_def) apply (simp add: lens_indep.lens_put_comm lens_override_def) apply (rename_tac s s' s\<^sub>0) apply (subgoal_tac "put\<^bsub>y\<^esub> (put\<^bsub>x\<^esub> s (get\<^bsub>x\<^esub> (put\<^bsub>x\<^esub> s\<^sub>0 (get\<^bsub>x\<^esub> s')))) (get\<^bsub>y\<^esub> (put\<^bsub>y\<^esub> s (get\<^bsub>y\<^esub> s\<^sub>0))) = put\<^bsub>x\<^esub> s\<^sub>0 (get\<^bsub>x\<^esub> s')") apply (metis lens_indep_get lens_indep_sym) apply (metis lens_indep.lens_put_comm) done lemma frame_contract_RID: assumes "vwb_lens x" "P is RID(x)" "x \ y" shows "(x;y):[P] = y:[P]" proof - from assms(1,3) have "(x;y):[RID(x)(P)] = y:[RID(x)(P)]" apply (rel_auto) apply (simp add: lens_indep.lens_put_comm) apply (metis (no_types) vwb_lens_wb wb_lens.get_put) done thus ?thesis by (simp add: Healthy_if assms) qed lemma frame_miracle [simp]: "x:[false] = false" by (rel_auto) lemma frame_skip [simp]: "vwb_lens x \ x:[II] = II" by (rel_auto) lemma frame_assign_in [frame]: "\ vwb_lens a; x \\<^sub>L a \ \ a:[x := v] = x := v" by (rel_auto, simp_all add: lens_get_put_quasi_commute lens_put_of_quotient) lemma frame_conj_true [frame]: "\ {$x,$x\} \ P; vwb_lens x \ \ (P \ x:[true]) = x:[P]" by (rel_auto) lemma frame_is_assign [frame]: "vwb_lens x \ x:[$x\ =\<^sub>u \v\\<^sub><] = x := v" by (rel_auto) lemma frame_seq [frame]: "\ vwb_lens x; {$x,$x\} \ P; {$x,$x\} \ Q \ \ x:[P ;; Q] = x:[P] ;; x:[Q]" apply (rel_auto) apply (metis mwb_lens.put_put vwb_lens_mwb vwb_lens_wb wb_lens_def weak_lens.put_get) apply (metis mwb_lens.put_put vwb_lens_mwb) done lemma frame_to_antiframe [frame]: "\ x \ y; x +\<^sub>L y = 1\<^sub>L \ \ x:[P] = y:\P\" by (rel_auto, metis lens_indep_def, metis lens_indep_def surj_pair) lemma rel_frext_miracle [frame]: "a:[false]\<^sup>+ = false" by (rel_auto) lemma rel_frext_skip [frame]: "vwb_lens a \ a:[II]\<^sup>+ = II" by (rel_auto) lemma rel_frext_seq [frame]: "vwb_lens a \ a:[P ;; Q]\<^sup>+ = (a:[P]\<^sup>+ ;; a:[Q]\<^sup>+)" apply (rel_auto) apply (rename_tac s s' s\<^sub>0) apply (rule_tac x="put\<^bsub>a\<^esub> s s\<^sub>0" in exI) apply (auto) apply (metis mwb_lens.put_put vwb_lens_mwb) done lemma rel_frext_assigns [frame]: "vwb_lens a \ a:[\\\\<^sub>a]\<^sup>+ = \\ \\<^sub>s a\\<^sub>a" by (rel_auto) lemma rel_frext_rcond [frame]: "a:[P \ b \\<^sub>r Q]\<^sup>+ = (a:[P]\<^sup>+ \ b \\<^sub>p a \\<^sub>r a:[Q]\<^sup>+)" by (rel_auto) lemma rel_frext_commute: "x \ y \ x:[P]\<^sup>+ ;; y:[Q]\<^sup>+ = y:[Q]\<^sup>+ ;; x:[P]\<^sup>+" apply (rel_auto) apply (rename_tac a c b) apply (subgoal_tac "\b a. get\<^bsub>y\<^esub> (put\<^bsub>x\<^esub> b a) = get\<^bsub>y\<^esub> b") apply (metis (no_types, hide_lams) lens_indep_comm lens_indep_get) apply (simp add: lens_indep.lens_put_irr2) apply (subgoal_tac "\b c. get\<^bsub>x\<^esub> (put\<^bsub>y\<^esub> b c) = get\<^bsub>x\<^esub> b") apply (subgoal_tac "\b a. get\<^bsub>y\<^esub> (put\<^bsub>x\<^esub> b a) = get\<^bsub>y\<^esub> b") apply (metis (mono_tags, lifting) lens_indep_comm) apply (simp_all add: lens_indep.lens_put_irr2) done lemma antiframe_disj [frame]: "(x:\P\ \ x:\Q\) = x:\P \ Q\" by (rel_auto) lemma antiframe_seq [frame]: "\ vwb_lens x; $x\ \ P; $x \ Q \ \ (x:\P\ ;; x:\Q\) = x:\P ;; Q\" apply (rel_auto) apply (metis vwb_lens_wb wb_lens_def weak_lens.put_get) apply (metis vwb_lens_wb wb_lens.put_twice wb_lens_def weak_lens.put_get) done lemma nameset_skip: "vwb_lens x \ (ns x \ II) = II\<^bsub>x\<^esub>" by (rel_auto, meson vwb_lens_wb wb_lens.get_put) lemma nameset_skip_ra: "vwb_lens x \ (ns x \ II\<^bsub>x\<^esub>) = II\<^bsub>x\<^esub>" by (rel_auto) declare sublens_def [lens_defs] subsection \ While Loop Laws \ theorem while_unfold: "while b do P od = ((P ;; while b do P od) \ b \\<^sub>r II)" proof - have m:"mono (\X. (P ;; X) \ b \\<^sub>r II)" by (auto intro: monoI seqr_mono cond_mono) have "(while b do P od) = (\ X \ (P ;; X) \ b \\<^sub>r II)" by (simp add: while_top_def) also have "... = ((P ;; (\ X \ (P ;; X) \ b \\<^sub>r II)) \ b \\<^sub>r II)" by (subst lfp_unfold, simp_all add: m) also have "... = ((P ;; while b do P od) \ b \\<^sub>r II)" by (simp add: while_top_def) finally show ?thesis . qed theorem while_false: "while false do P od = II" by (subst while_unfold, rel_auto) theorem while_true: "while true do P od = false" apply (simp add: while_top_def alpha) apply (rule antisym) apply (simp_all) apply (rule lfp_lowerbound) apply (rel_auto) done theorem while_bot_unfold: "while\<^sub>\ b do P od = ((P ;; while\<^sub>\ b do P od) \ b \\<^sub>r II)" proof - have m:"mono (\X. (P ;; X) \ b \\<^sub>r II)" by (auto intro: monoI seqr_mono cond_mono) have "(while\<^sub>\ b do P od) = (\ X \ (P ;; X) \ b \\<^sub>r II)" by (simp add: while_bot_def) also have "... = ((P ;; (\ X \ (P ;; X) \ b \\<^sub>r II)) \ b \\<^sub>r II)" by (subst gfp_unfold, simp_all add: m) also have "... = ((P ;; while\<^sub>\ b do P od) \ b \\<^sub>r II)" by (simp add: while_bot_def) finally show ?thesis . qed theorem while_bot_false: "while\<^sub>\ false do P od = II" by (simp add: while_bot_def mu_const alpha) theorem while_bot_true: "while\<^sub>\ true do P od = (\ X \ P ;; X)" by (simp add: while_bot_def alpha) text \ An infinite loop with a feasible body corresponds to a program error (non-termination). \ theorem while_infinite: "P ;; true\<^sub>h = true \ while\<^sub>\ true do P od = true" apply (simp add: while_bot_true) apply (rule antisym) apply (simp) apply (rule gfp_upperbound) apply (simp) done subsection \ Algebraic Properties \ interpretation upred_semiring: semiring_1 where times = seqr and one = skip_r and zero = false\<^sub>h and plus = Lattices.sup by (unfold_locales, (rel_auto)+) declare upred_semiring.power_Suc [simp del] text \ We introduce the power syntax derived from semirings \ abbreviation upower :: "'\ hrel \ nat \ '\ hrel" (infixr "\<^bold>^" 80) where "upower P n \ upred_semiring.power P n" translations "P \<^bold>^ i" <= "CONST power.power II op ;; P i" "P \<^bold>^ i" <= "(CONST power.power II op ;; P) i" text \ Set up transfer tactic for powers \ lemma upower_rep_eq: "\P \<^bold>^ i\\<^sub>e = (\ b. b \ ({p. \P\\<^sub>e p} ^^ i))" proof (induct i arbitrary: P) case 0 then show ?case by (auto, rel_auto) next case (Suc i) show ?case by (simp add: Suc seqr.rep_eq relpow_commute upred_semiring.power_Suc) qed lemma upower_rep_eq_alt: "\power.power \id\\<^sub>a (;;) P i\\<^sub>e = (\b. b \ ({p. \P\\<^sub>e p} ^^ i))" by (metis skip_r_def upower_rep_eq) update_uexpr_rep_eq_thms lemma Sup_power_expand: fixes P :: "nat \ 'a::complete_lattice" shows "P(0) \ (\i. P(i+1)) = (\i. P(i))" proof - have "UNIV = insert (0::nat) {1..}" by auto moreover have "(\i. P(i)) = \ (P ` UNIV)" by (blast) - moreover have "\ (P ` insert 0 {1..}) = P(0) \ SUPREMUM {1..} P" + moreover have "\ (P ` insert 0 {1..}) = P(0) \ \ (P ` {1..})" by (simp) - moreover have "SUPREMUM {1..} P = (\i. P(i+1))" + moreover have "\ (P ` {1..}) = (\i. P(i+1))" by (simp add: atLeast_Suc_greaterThan greaterThan_0) ultimately show ?thesis by (simp only:) qed lemma Sup_upto_Suc: "(\i\{0..Suc n}. P \<^bold>^ i) = (\i\{0..n}. P \<^bold>^ i) \ P \<^bold>^ Suc n" proof - have "(\i\{0..Suc n}. P \<^bold>^ i) = (\i\insert (Suc n) {0..n}. P \<^bold>^ i)" by (simp add: atLeast0_atMost_Suc) also have "... = P \<^bold>^ Suc n \ (\i\{0..n}. P \<^bold>^ i)" by (simp) finally show ?thesis by (simp add: Lattices.sup_commute) qed text \ The following two proofs are adapted from the AFP entry \href{https://www.isa-afp.org/entries/Kleene_Algebra.shtml}{Kleene Algebra}. See also~\cite{Armstrong2012,Armstrong2015}. \ lemma upower_inductl: "Q \ ((P ;; Q) \ R) \ Q \ P \<^bold>^ n ;; R" proof (induct n) case 0 then show ?case by (auto) next case (Suc n) then show ?case by (auto simp add: upred_semiring.power_Suc, metis (no_types, hide_lams) dual_order.trans order_refl seqr_assoc seqr_mono) qed lemma upower_inductr: assumes "Q \ R \ (Q ;; P)" shows "Q \ R ;; (P \<^bold>^ n)" using assms proof (induct n) case 0 then show ?case by auto next case (Suc n) have "R ;; P \<^bold>^ Suc n = (R ;; P \<^bold>^ n) ;; P" by (metis seqr_assoc upred_semiring.power_Suc2) also have "Q ;; P \ ..." by (meson Suc.hyps assms eq_iff seqr_mono) also have "Q \ ..." using assms by auto finally show ?case . qed lemma SUP_atLeastAtMost_first: fixes P :: "nat \ 'a::complete_lattice" assumes "m \ n" shows "(\i\{m..n}. P(i)) = P(m) \ (\i\{Suc m..n}. P(i))" by (metis SUP_insert assms atLeastAtMost_insertL) lemma upower_seqr_iter: "P \<^bold>^ n = (;; Q : replicate n P \ Q)" by (induct n, simp_all add: upred_semiring.power_Suc) lemma assigns_power: "\f\\<^sub>a \<^bold>^ n = \f ^^ n\\<^sub>a" by (induct n, rel_auto+) subsection \ Kleene Star \ definition ustar :: "'\ hrel \ '\ hrel" ("_\<^sup>\" [999] 999) where "P\<^sup>\ = (\i\{0..} \ P\<^bold>^i)" lemma ustar_rep_eq: "\P\<^sup>\\\<^sub>e = (\b. b \ ({p. \P\\<^sub>e p}\<^sup>*))" by (simp add: ustar_def, rel_auto, simp_all add: relpow_imp_rtrancl rtrancl_imp_relpow) update_uexpr_rep_eq_thms subsection \ Kleene Plus \ purge_notation trancl ("(_\<^sup>+)" [1000] 999) definition uplus :: "'\ hrel \ '\ hrel" ("_\<^sup>+" [999] 999) where [upred_defs]: "P\<^sup>+ = P ;; P\<^sup>\" lemma uplus_power_def: "P\<^sup>+ = (\ i \ P \<^bold>^ (Suc i))" by (simp add: uplus_def ustar_def seq_UINF_distl' UINF_atLeast_Suc upred_semiring.power_Suc) subsection \ Omega \ definition uomega :: "'\ hrel \ '\ hrel" ("_\<^sup>\" [999] 999) where "P\<^sup>\ = (\ X \ P ;; X)" subsection \ Relation Algebra Laws \ theorem RA1: "(P ;; (Q ;; R)) = ((P ;; Q) ;; R)" by (simp add: seqr_assoc) theorem RA2: "(P ;; II) = P" "(II ;; P) = P" by simp_all theorem RA3: "P\<^sup>-\<^sup>- = P" by simp theorem RA4: "(P ;; Q)\<^sup>- = (Q\<^sup>- ;; P\<^sup>-)" by simp theorem RA5: "(P \ Q)\<^sup>- = (P\<^sup>- \ Q\<^sup>-)" by (rel_auto) theorem RA6: "((P \ Q) ;; R) = (P;;R \ Q;;R)" using seqr_or_distl by blast theorem RA7: "((P\<^sup>- ;; (\(P ;; Q))) \ (\Q)) = (\Q)" by (rel_auto) subsection \ Kleene Algebra Laws \ lemma ustar_alt_def: "P\<^sup>\ = (\ i \ P \<^bold>^ i)" by (simp add: ustar_def) theorem ustar_sub_unfoldl: "P\<^sup>\ \ II \ (P;;P\<^sup>\)" by (rel_simp, simp add: rtrancl_into_trancl2 trancl_into_rtrancl) theorem ustar_inductl: assumes "Q \ R" "Q \ P ;; Q" shows "Q \ P\<^sup>\ ;; R" proof - have "P\<^sup>\ ;; R = (\ i. P \<^bold>^ i ;; R)" by (simp add: ustar_def UINF_as_Sup_collect' seq_SUP_distr) also have "Q \ ..." by (simp add: SUP_least assms upower_inductl) finally show ?thesis . qed theorem ustar_inductr: assumes "Q \ R" "Q \ Q ;; P" shows "Q \ R ;; P\<^sup>\" proof - have "R ;; P\<^sup>\ = (\ i. R ;; P \<^bold>^ i)" by (simp add: ustar_def UINF_as_Sup_collect' seq_SUP_distl) also have "Q \ ..." by (simp add: SUP_least assms upower_inductr) finally show ?thesis . qed lemma ustar_refines_nu: "(\ X \ (P ;; X) \ II) \ P\<^sup>\" by (metis (no_types, lifting) lfp_greatest semilattice_sup_class.le_sup_iff semilattice_sup_class.sup_idem upred_semiring.mult_2_right upred_semiring.one_add_one ustar_inductl) lemma ustar_as_nu: "P\<^sup>\ = (\ X \ (P ;; X) \ II)" proof (rule antisym) show "(\ X \ (P ;; X) \ II) \ P\<^sup>\" by (simp add: ustar_refines_nu) show "P\<^sup>\ \ (\ X \ (P ;; X) \ II)" by (metis lfp_lowerbound upred_semiring.add_commute ustar_sub_unfoldl) qed lemma ustar_unfoldl: "P\<^sup>\ = II \ (P ;; P\<^sup>\)" apply (simp add: ustar_as_nu) apply (subst lfp_unfold) apply (rule monoI) apply (rel_auto)+ done text \ While loop can be expressed using Kleene star \ lemma while_star_form: "while b do P od = (P \ b \\<^sub>r II)\<^sup>\ ;; [(\b)]\<^sup>\" proof - have 1: "Continuous (\X. P ;; X \ b \\<^sub>r II)" by (rel_auto) have "while b do P od = (\i. ((\X. P ;; X \ b \\<^sub>r II) ^^ i) false)" by (simp add: "1" false_upred_def sup_continuous_Continuous sup_continuous_lfp while_top_def) also have "... = ((\X. P ;; X \ b \\<^sub>r II) ^^ 0) false \ (\i. ((\X. P ;; X \ b \\<^sub>r II) ^^ (i+1)) false)" by (subst Sup_power_expand, simp) also have "... = (\i. ((\X. P ;; X \ b \\<^sub>r II) ^^ (i+1)) false)" by (simp) also have "... = (\i. (P \ b \\<^sub>r II)\<^bold>^i ;; (false \ b \\<^sub>r II))" proof (rule SUP_cong, simp_all) fix i show "P ;; ((\X. P ;; X \ b \\<^sub>r II) ^^ i) false \ b \\<^sub>r II = (P \ b \\<^sub>r II) \<^bold>^ i ;; (false \ b \\<^sub>r II)" proof (induct i) case 0 then show ?case by simp next case (Suc i) then show ?case by (simp add: upred_semiring.power_Suc) (metis (no_types, lifting) RA1 comp_cond_left_distr cond_L6 upred_semiring.mult.left_neutral) qed qed also have "... = (\i\{0..} \ (P \ b \\<^sub>r II)\<^bold>^i ;; [(\b)]\<^sup>\)" by (rel_auto) also have "... = (P \ b \\<^sub>r II)\<^sup>\ ;; [(\b)]\<^sup>\" by (metis seq_UINF_distr ustar_def) finally show ?thesis . qed subsection \ Omega Algebra Laws \ lemma uomega_induct: "P ;; P\<^sup>\ \ P\<^sup>\" by (simp add: uomega_def, metis eq_refl gfp_unfold monoI seqr_mono) subsection \ Refinement Laws \ lemma skip_r_refine: "(p \ p) \ II" by pred_blast lemma conj_refine_left: "(Q \ P) \ R \ P \ (Q \ R)" by (rel_auto) lemma pre_weak_rel: assumes "`Pre \ I`" and "(I \ Post) \ P" shows "(Pre \ Post) \ P" using assms by(rel_auto) lemma cond_refine_rel: assumes "S \ (\b\\<^sub>< \ P)" "S \ (\\b\\<^sub>< \ Q)" shows "S \ P \ b \\<^sub>r Q" by (metis aext_not assms(1) assms(2) cond_def lift_rcond_def utp_pred_laws.le_sup_iff) lemma seq_refine_pred: assumes "(\b\\<^sub>< \ \s\\<^sub>>) \ P" and "(\s\\<^sub>< \ \c\\<^sub>>) \ Q" shows "(\b\\<^sub>< \ \c\\<^sub>>) \ (P ;; Q)" using assms by rel_auto lemma seq_refine_unrest: assumes "out\ \ b" "in\ \ c" assumes "(b \ \s\\<^sub>>) \ P" and "(\s\\<^sub>< \ c) \ Q" shows "(b \ c) \ (P ;; Q)" using assms by rel_blast subsection \ Domain and Range Laws \ lemma Dom_conv_Ran: "Dom(P\<^sup>-) = Ran(P)" by (rel_auto) lemma Ran_conv_Dom: "Ran(P\<^sup>-) = Dom(P)" by (rel_auto) lemma Dom_skip: "Dom(II) = true" by (rel_auto) lemma Dom_assigns: "Dom(\\\\<^sub>a) = true" by (rel_auto) lemma Dom_miracle: "Dom(false) = false" by (rel_auto) lemma Dom_assume: "Dom([b]\<^sup>\) = b" by (rel_auto) lemma Dom_seq: "Dom(P ;; Q) = Dom(P ;; [Dom(Q)]\<^sup>\)" by (rel_auto) lemma Dom_disj: "Dom(P \ Q) = (Dom(P) \ Dom(Q))" by (rel_auto) lemma Dom_inf: "Dom(P \ Q) = (Dom(P) \ Dom(Q))" by (rel_auto) lemma assume_Dom: "[Dom(P)]\<^sup>\ ;; P = P" by (rel_auto) end \ No newline at end of file