diff --git a/thys/Collections/GenCF/Intf/Intf_Comp.thy b/thys/Collections/GenCF/Intf/Intf_Comp.thy --- a/thys/Collections/GenCF/Intf/Intf_Comp.thy +++ b/thys/Collections/GenCF/Intf/Intf_Comp.thy @@ -1,875 +1,857 @@ section \\isaheader{Orderings By Comparison Operator}\ theory Intf_Comp imports Automatic_Refinement.Automatic_Refinement begin subsection \Basic Definitions\ datatype comp_res = LESS | EQUAL | GREATER consts i_comp_res :: interface abbreviation "comp_res_rel \ Id :: (comp_res \ _) set" lemmas [autoref_rel_intf] = REL_INTFI[of comp_res_rel i_comp_res] definition "comp2le cmp a b \ case cmp a b of LESS \ True | EQUAL \ True | GREATER \ False" definition "comp2lt cmp a b \ case cmp a b of LESS \ True | EQUAL \ False | GREATER \ False" definition "comp2eq cmp a b \ case cmp a b of LESS \ False | EQUAL \ True | GREATER \ False" locale linorder_on = fixes D :: "'a set" fixes cmp :: "'a \ 'a \ comp_res" assumes lt_eq: "\x\D; y\D\ \ cmp x y = LESS \ (cmp y x = GREATER)" assumes refl[simp, intro!]: "x\D \ cmp x x = EQUAL" assumes trans[trans]: "\ x\D; y\D; z\D; cmp x y = LESS; cmp y z = LESS\ \ cmp x z = LESS" "\ x\D; y\D; z\D; cmp x y = LESS; cmp y z = EQUAL\ \ cmp x z = LESS" "\ x\D; y\D; z\D; cmp x y = EQUAL; cmp y z = LESS\ \ cmp x z = LESS" "\ x\D; y\D; z\D; cmp x y = EQUAL; cmp y z = EQUAL\ \ cmp x z = EQUAL" begin abbreviation "le \ comp2le cmp" abbreviation "lt \ comp2lt cmp" lemma eq_sym: "\x\D; y\D\ \ cmp x y = EQUAL \ cmp y x = EQUAL" apply (cases "cmp y x") using lt_eq lt_eq[symmetric] by auto end abbreviation "linorder \ linorder_on UNIV" lemma linorder_to_class: assumes "linorder cmp" assumes [simp]: "\x y. cmp x y = EQUAL \ x=y" shows "class.linorder (comp2le cmp) (comp2lt cmp)" proof - interpret linorder_on UNIV cmp by fact show ?thesis apply (unfold_locales) unfolding comp2le_def comp2lt_def apply (auto split: comp_res.split comp_res.split_asm) using lt_eq apply simp using lt_eq apply simp using lt_eq[symmetric] apply simp apply (drule (1) trans[rotated 3], simp_all) [] apply (drule (1) trans[rotated 3], simp_all) [] apply (drule (1) trans[rotated 3], simp_all) [] apply (drule (1) trans[rotated 3], simp_all) [] using lt_eq apply simp using lt_eq apply simp using lt_eq[symmetric] apply simp done qed definition "dflt_cmp le lt a b \ if lt a b then LESS else if le a b then EQUAL else GREATER" lemma (in linorder) class_to_linorder: "linorder (dflt_cmp (\) (<))" apply (unfold_locales) unfolding dflt_cmp_def by (auto split: if_split_asm) lemma restrict_linorder: "\linorder_on D cmp ; D'\D\ \ linorder_on D' cmp" apply (rule linorder_on.intro) apply (drule (1) rev_subsetD)+ apply (erule (2) linorder_on.lt_eq) apply (drule (1) rev_subsetD)+ apply (erule (1) linorder_on.refl) apply (drule (1) rev_subsetD)+ apply (erule (5) linorder_on.trans) apply (drule (1) rev_subsetD)+ apply (erule (5) linorder_on.trans) apply (drule (1) rev_subsetD)+ apply (erule (5) linorder_on.trans) apply (drule (1) rev_subsetD)+ apply (erule (5) linorder_on.trans) done subsection \Operations on Linear Orderings\ text \Map with injective function\ definition cmp_img where "cmp_img f cmp a b \ cmp (f a) (f b)" lemma img_linorder[intro?]: assumes LO: "linorder_on (f`D) cmp" shows "linorder_on D (cmp_img f cmp)" apply unfold_locales unfolding cmp_img_def apply (rule linorder_on.lt_eq[OF LO], auto) [] apply (rule linorder_on.refl[OF LO], auto) [] apply (erule (1) linorder_on.trans[OF LO, rotated -2], auto) [] apply (erule (1) linorder_on.trans[OF LO, rotated -2], auto) [] apply (erule (1) linorder_on.trans[OF LO, rotated -2], auto) [] apply (erule (1) linorder_on.trans[OF LO, rotated -2], auto) [] done text \Combine\ definition "cmp_combine D1 cmp1 D2 cmp2 a b \ if a\D1 \ b\D1 then cmp1 a b else if a\D1 \ b\D2 then LESS else if a\D2 \ b\D1 then GREATER else cmp2 a b " (* TODO: Move *) lemma UnE': assumes "x\A\B" obtains "x\A" | "x\A" "x\B" using assms by blast lemma combine_linorder[intro?]: assumes "linorder_on D1 cmp1" assumes "linorder_on D2 cmp2" assumes "D = D1\D2" shows "linorder_on D (cmp_combine D1 cmp1 D2 cmp2)" apply unfold_locales unfolding cmp_combine_def using assms apply - apply (simp only:) apply (elim UnE) apply (auto dest: linorder_on.lt_eq) [4] apply (simp only:) apply (elim UnE) apply (auto dest: linorder_on.refl) [2] apply (simp only:) apply (elim UnE') apply simp_all [8] apply (erule (5) linorder_on.trans) apply (erule (5) linorder_on.trans) apply (simp only:) apply (elim UnE') apply simp_all [8] apply (erule (5) linorder_on.trans) apply (erule (5) linorder_on.trans) apply (simp only:) apply (elim UnE') apply simp_all [8] apply (erule (5) linorder_on.trans) apply (erule (5) linorder_on.trans) apply (simp only:) apply (elim UnE') apply simp_all [8] apply (erule (5) linorder_on.trans) apply (erule (5) linorder_on.trans) done subsection \Universal Linear Ordering\ text \With Zorn's Lemma, we get a universal linear (even wf) ordering\ definition "univ_order_rel \ (SOME r. well_order_on UNIV r)" definition "univ_cmp x y \ if x=y then EQUAL else if (x,y)\univ_order_rel then LESS else GREATER" lemma univ_wo: "well_order_on UNIV univ_order_rel" unfolding univ_order_rel_def using well_order_on[of UNIV] .. lemma univ_linorder[intro?]: "linorder univ_cmp" apply unfold_locales unfolding univ_cmp_def apply (auto split: if_split_asm) using univ_wo apply - unfolding well_order_on_def linear_order_on_def partial_order_on_def preorder_on_def apply (auto simp add: antisym_def) [] apply (unfold total_on_def, fast) [] apply (unfold trans_def, fast) [] apply (auto simp add: antisym_def) [] done text \Extend any linear order to a universal order\ definition "cmp_extend D cmp \ cmp_combine D cmp UNIV univ_cmp" lemma extend_linorder[intro?]: "linorder_on D cmp \ linorder (cmp_extend D cmp)" unfolding cmp_extend_def apply rule apply assumption apply rule by simp subsubsection \Lexicographic Order on Lists\ fun cmp_lex where "cmp_lex cmp [] [] = EQUAL" | "cmp_lex cmp [] _ = LESS" | "cmp_lex cmp _ [] = GREATER" | "cmp_lex cmp (a#l) (b#m) = ( case cmp a b of LESS \ LESS | EQUAL \ cmp_lex cmp l m | GREATER \ GREATER)" primrec cmp_lex' where "cmp_lex' cmp [] m = (case m of [] \ EQUAL | _ \ LESS)" | "cmp_lex' cmp (a#l) m = (case m of [] \ GREATER | (b#m) \ (case cmp a b of LESS \ LESS | EQUAL \ cmp_lex' cmp l m | GREATER \ GREATER ))" lemma cmp_lex_alt: "cmp_lex cmp l m = cmp_lex' cmp l m" apply (induct l arbitrary: m) apply (auto split: comp_res.split list.split) done lemma (in linorder_on) lex_linorder[intro?]: "linorder_on (lists D) (cmp_lex cmp)" proof fix l m assume "l\lists D" "m\lists D" thus "(cmp_lex cmp l m = LESS) = (cmp_lex cmp m l = GREATER)" apply (induct cmp\cmp l m rule: cmp_lex.induct) apply (auto split: comp_res.split simp: lt_eq) apply (auto simp: lt_eq[symmetric]) done next fix x assume "x\lists D" thus "cmp_lex cmp x x = EQUAL" by (induct x) auto next fix x y z assume M: "x\lists D" "y\lists D" "z\lists D" { assume "cmp_lex cmp x y = LESS" "cmp_lex cmp y z = LESS" thus "cmp_lex cmp x z = LESS" using M apply (induct cmp\cmp x y arbitrary: z rule: cmp_lex.induct) apply (auto split: comp_res.split_asm comp_res.split) apply (case_tac z, auto) [] apply (case_tac z, auto split: comp_res.split_asm comp_res.split, (drule (4) trans, simp)+ ) [] apply (case_tac z, auto split: comp_res.split_asm comp_res.split, (drule (4) trans, simp)+ ) [] done } { assume "cmp_lex cmp x y = LESS" "cmp_lex cmp y z = EQUAL" thus "cmp_lex cmp x z = LESS" using M apply (induct cmp\cmp x y arbitrary: z rule: cmp_lex.induct) apply (auto split: comp_res.split_asm comp_res.split) apply (case_tac z, auto) [] apply (case_tac z, auto split: comp_res.split_asm comp_res.split, (drule (4) trans, simp)+ ) [] apply (case_tac z, auto split: comp_res.split_asm comp_res.split, (drule (4) trans, simp)+ ) [] done } { assume "cmp_lex cmp x y = EQUAL" "cmp_lex cmp y z = LESS" thus "cmp_lex cmp x z = LESS" using M apply (induct cmp\cmp x y arbitrary: z rule: cmp_lex.induct) apply (auto split: comp_res.split_asm comp_res.split) apply (case_tac z, auto split: comp_res.split_asm comp_res.split, (drule (4) trans, simp)+ ) [] done } { assume "cmp_lex cmp x y = EQUAL" "cmp_lex cmp y z = EQUAL" thus "cmp_lex cmp x z = EQUAL" using M apply (induct cmp\cmp x y arbitrary: z rule: cmp_lex.induct) apply (auto split: comp_res.split_asm comp_res.split) apply (case_tac z) apply (auto split: comp_res.split_asm comp_res.split) apply (drule (4) trans, simp)+ done } qed subsubsection \Lexicographic Order on Pairs\ fun cmp_prod where "cmp_prod cmp1 cmp2 (a1,a2) (b1,b2) = ( case cmp1 a1 b1 of LESS \ LESS | EQUAL \ cmp2 a2 b2 | GREATER \ GREATER)" lemma cmp_prod_alt: "cmp_prod = (\cmp1 cmp2 (a1,a2) (b1,b2). ( case cmp1 a1 b1 of LESS \ LESS | EQUAL \ cmp2 a2 b2 | GREATER \ GREATER))" by (auto intro!: ext) lemma prod_linorder[intro?]: assumes A: "linorder_on A cmp1" assumes B: "linorder_on B cmp2" shows "linorder_on (A\B) (cmp_prod cmp1 cmp2)" proof - interpret A: linorder_on A cmp1 + B: linorder_on B cmp2 by fact+ show ?thesis apply unfold_locales apply (auto split: comp_res.split comp_res.split_asm, simp_all add: A.lt_eq B.lt_eq, simp_all add: A.lt_eq[symmetric] ) [] apply (auto split: comp_res.split comp_res.split_asm) [] apply (auto split: comp_res.split comp_res.split_asm) [] apply (drule (4) A.trans B.trans, simp)+ apply (auto split: comp_res.split comp_res.split_asm) [] apply (drule (4) A.trans B.trans, simp)+ apply (auto split: comp_res.split comp_res.split_asm) [] apply (drule (4) A.trans B.trans, simp)+ apply (auto split: comp_res.split comp_res.split_asm) [] apply (drule (4) A.trans B.trans, simp)+ done qed subsection \Universal Ordering for Sets that is Effective for Finite Sets\ subsubsection \Sorted Lists of Sets\ text \Some more results about sorted lists of finite sets\ lemma set_to_map_set_is_map_of: "distinct (map fst l) \ set_to_map (set l) = map_of l" apply (induct l) apply (auto simp: set_to_map_insert) done context linorder begin - lemma sorted_list_of_set_eq_nil[simp]: - assumes "finite A" - shows "sorted_list_of_set A = [] \ A={}" - using assms - apply (induct rule: finite_induct) - apply simp - apply simp - done - lemma sorted_list_of_set_eq_nil2[simp]: assumes "finite A" shows "[] = sorted_list_of_set A \ A={}" using assms by (auto dest: sym) lemma set_insort[simp]: "set (insort x l) = insert x (set l)" by (induct l) auto lemma sorted_list_of_set_inj_aux: fixes A B :: "'a set" assumes "finite A" assumes "finite B" assumes "sorted_list_of_set A = sorted_list_of_set B" shows "A=B" using assms proof - from \finite B\ have "B = set (sorted_list_of_set B)" by simp also from assms have "\ = set (sorted_list_of_set (A))" by simp also from \finite A\ have "set (sorted_list_of_set (A)) = A" by simp finally show ?thesis by simp qed lemma sorted_list_of_set_inj: "inj_on sorted_list_of_set (Collect finite)" apply (rule inj_onI) using sorted_list_of_set_inj_aux by blast - lemma the_sorted_list_of_set: - assumes "distinct l" - assumes "sorted l" - shows "sorted_list_of_set (set l) = l" - using assms - by (simp - add: sorted_list_of_set_sort_remdups distinct_remdups_id sorted_sort_id) - - definition "sorted_list_of_map m \ map (\k. (k, the (m k))) (sorted_list_of_set (dom m))" lemma the_sorted_list_of_map: assumes "distinct (map fst l)" assumes "sorted (map fst l)" shows "sorted_list_of_map (map_of l) = l" proof - have "dom (map_of l) = set (map fst l)" by (induct l) force+ hence "sorted_list_of_set (dom (map_of l)) = map fst l" - using the_sorted_list_of_set[OF assms] by simp + using sorted_list_of_set.idem_if_sorted_distinct[OF assms(2,1)] by simp hence "sorted_list_of_map (map_of l) = map (\k. (k, the (map_of l k))) (map fst l)" unfolding sorted_list_of_map_def by simp also have "\ = l" using \distinct (map fst l)\ proof (induct l) case Nil thus ?case by simp next case (Cons a l) hence 1: "distinct (map fst l)" and 2: "fst a\fst`set l" and 3: "map (\k. (k, the (map_of l k))) (map fst l) = l" by simp_all from 2 have [simp]: "\(\x\set l. fst x = fst a)" by (auto simp: image_iff) show ?case apply simp apply (subst (3) 3[symmetric]) apply simp done qed finally show ?thesis . qed lemma map_of_sorted_list_of_map[simp]: assumes FIN: "finite (dom m)" shows "map_of (sorted_list_of_map m) = m" unfolding sorted_list_of_map_def proof - have "set (sorted_list_of_set (dom m)) = dom m" and DIST: "distinct (sorted_list_of_set (dom m))" by (simp_all add: FIN) have [simp]: "(fst \ (\k. (k, the (m k)))) = id" by auto have [simp]: "(\k. (k, the (m k))) ` dom m = map_to_set m" by (auto simp: map_to_set_def) show "map_of (map (\k. (k, the (m k))) (sorted_list_of_set (dom m))) = m" apply (subst set_to_map_set_is_map_of[symmetric]) apply (simp add: DIST) apply (subst set_map) apply (simp add: FIN map_to_set_inverse) done qed lemma sorted_list_of_map_inj_aux: fixes A B :: "'a\'b" assumes [simp]: "finite (dom A)" assumes [simp]: "finite (dom B)" assumes E: "sorted_list_of_map A = sorted_list_of_map B" shows "A=B" using assms proof - have "A = map_of (sorted_list_of_map A)" by simp also note E also have "map_of (sorted_list_of_map B) = B" by simp finally show ?thesis . qed lemma sorted_list_of_map_inj: "inj_on sorted_list_of_map (Collect (finite o dom))" apply (rule inj_onI) using sorted_list_of_map_inj_aux by auto end definition "cmp_set cmp \ cmp_extend (Collect finite) ( cmp_img (linorder.sorted_list_of_set (comp2le cmp)) (cmp_lex cmp) )" thm img_linorder lemma set_ord_linear[intro?]: "linorder cmp \ linorder (cmp_set cmp)" unfolding cmp_set_def apply rule apply rule apply (rule restrict_linorder) apply (erule linorder_on.lex_linorder) apply simp done definition "cmp_map cmpk cmpv \ cmp_extend (Collect (finite o dom)) ( cmp_img (linorder.sorted_list_of_map (comp2le cmpk)) (cmp_lex (cmp_prod cmpk cmpv)) ) " lemma map_to_set_inj[intro!]: "inj map_to_set" apply (rule inj_onI) unfolding map_to_set_def apply (rule ext) apply (case_tac "x xa") apply (case_tac [!] "y xa") apply force+ done corollary map_to_set_inj'[intro!]: "inj_on map_to_set S" by (metis map_to_set_inj subset_UNIV subset_inj_on) lemma map_ord_linear[intro?]: assumes A: "linorder cmpk" assumes B: "linorder cmpv" shows "linorder (cmp_map cmpk cmpv)" proof - interpret lk: linorder_on UNIV cmpk by fact interpret lv: linorder_on UNIV cmpv by fact show ?thesis unfolding cmp_map_def apply rule apply rule apply (rule restrict_linorder) apply (rule linorder_on.lex_linorder) apply (rule) apply fact apply fact apply simp done qed locale eq_linorder_on = linorder_on + assumes cmp_imp_equal: "\x\D; y\D\ \ cmp x y = EQUAL \ x = y" begin lemma cmp_eq[simp]: "\x\D; y\D\ \ cmp x y = EQUAL \ x = y" by (auto simp: cmp_imp_equal) end abbreviation "eq_linorder \ eq_linorder_on UNIV" lemma dflt_cmp_2inv[simp]: "dflt_cmp (comp2le cmp) (comp2lt cmp) = cmp" unfolding dflt_cmp_def[abs_def] comp2le_def[abs_def] comp2lt_def[abs_def] apply (auto split: comp_res.splits intro!: ext) done lemma (in linorder) dflt_cmp_inv2[simp]: shows "(comp2le (dflt_cmp (\) (<)))= (\)" "(comp2lt (dflt_cmp (\) (<)))= (<)" proof - show "(comp2lt (dflt_cmp (\) (<)))= (<)" unfolding dflt_cmp_def[abs_def] comp2le_def[abs_def] comp2lt_def[abs_def] apply (auto split: comp_res.splits intro!: ext) done show "(comp2le (dflt_cmp (\) (<))) = (\)" unfolding dflt_cmp_def[abs_def] comp2le_def[abs_def] comp2lt_def[abs_def] apply (auto split: comp_res.splits intro!: ext) done qed lemma eq_linorder_class_conv: "eq_linorder cmp \ class.linorder (comp2le cmp) (comp2lt cmp)" proof assume "eq_linorder cmp" then interpret eq_linorder_on UNIV cmp . have "linorder cmp" by unfold_locales show "class.linorder (comp2le cmp) (comp2lt cmp)" apply (rule linorder_to_class) apply fact by simp next assume "class.linorder (comp2le cmp) (comp2lt cmp)" then interpret linorder "comp2le cmp" "comp2lt cmp" . from class_to_linorder interpret linorder_on UNIV cmp by simp show "eq_linorder cmp" proof fix x y assume "cmp x y = EQUAL" hence "comp2le cmp x y" "\comp2lt cmp x y" by (auto simp: comp2le_def comp2lt_def) thus "x=y" by simp qed qed lemma (in linorder) class_to_eq_linorder: "eq_linorder (dflt_cmp (\) (<))" proof - interpret linorder_on UNIV "dflt_cmp (\) (<)" by (rule class_to_linorder) show ?thesis apply unfold_locales apply (auto simp: dflt_cmp_def split: if_split_asm) done qed lemma eq_linorder_comp2eq_eq: assumes "eq_linorder cmp" shows "comp2eq cmp = (=)" proof - interpret eq_linorder_on UNIV cmp by fact show ?thesis apply (intro ext) unfolding comp2eq_def apply (auto split: comp_res.split dest: refl) done qed lemma restrict_eq_linorder: assumes "eq_linorder_on D cmp" assumes S: "D'\D" shows "eq_linorder_on D' cmp" proof - interpret eq_linorder_on D cmp by fact show ?thesis apply (rule eq_linorder_on.intro) apply (rule restrict_linorder[where D=D]) apply unfold_locales [] apply fact apply unfold_locales using S apply - apply (drule (1) rev_subsetD)+ apply auto done qed lemma combine_eq_linorder[intro?]: assumes A: "eq_linorder_on D1 cmp1" assumes B: "eq_linorder_on D2 cmp2" assumes EQ: "D=D1\D2" shows "eq_linorder_on D (cmp_combine D1 cmp1 D2 cmp2)" proof - interpret A: eq_linorder_on D1 cmp1 by fact interpret B: eq_linorder_on D2 cmp2 by fact interpret linorder_on "(D1 \ D2)" "(cmp_combine D1 cmp1 D2 cmp2)" apply rule apply unfold_locales by simp show ?thesis apply (simp only: EQ) apply unfold_locales unfolding cmp_combine_def by (auto split: if_split_asm) qed lemma img_eq_linorder[intro?]: assumes A: "eq_linorder_on (f`D) cmp" assumes INJ: "inj_on f D" shows "eq_linorder_on D (cmp_img f cmp)" proof - interpret eq_linorder_on "f`D" cmp by fact interpret L: linorder_on "(D)" "(cmp_img f cmp)" apply rule apply unfold_locales done show ?thesis apply unfold_locales unfolding cmp_img_def using INJ apply (auto dest: inj_onD) done qed lemma univ_eq_linorder[intro?]: shows "eq_linorder univ_cmp" apply (rule eq_linorder_on.intro) apply rule apply unfold_locales unfolding univ_cmp_def apply (auto split: if_split_asm) done lemma extend_eq_linorder[intro?]: assumes "eq_linorder_on D cmp" shows "eq_linorder (cmp_extend D cmp)" proof - interpret eq_linorder_on D cmp by fact show ?thesis unfolding cmp_extend_def apply (rule) apply fact apply rule by simp qed lemma lex_eq_linorder[intro?]: assumes "eq_linorder_on D cmp" shows "eq_linorder_on (lists D) (cmp_lex cmp)" proof - interpret eq_linorder_on D cmp by fact show ?thesis apply (rule eq_linorder_on.intro) apply rule apply unfold_locales subgoal for l m apply (induct cmp\cmp l m rule: cmp_lex.induct) apply (auto split: comp_res.splits) done done qed lemma prod_eq_linorder[intro?]: assumes "eq_linorder_on D1 cmp1" assumes "eq_linorder_on D2 cmp2" shows "eq_linorder_on (D1\D2) (cmp_prod cmp1 cmp2)" proof - interpret A: eq_linorder_on D1 cmp1 by fact interpret B: eq_linorder_on D2 cmp2 by fact show ?thesis apply (rule eq_linorder_on.intro) apply rule apply unfold_locales apply (auto split: comp_res.splits) done qed lemma set_ord_eq_linorder[intro?]: "eq_linorder cmp \ eq_linorder (cmp_set cmp)" unfolding cmp_set_def apply rule apply rule apply (rule restrict_eq_linorder) apply rule apply assumption apply simp apply (rule linorder.sorted_list_of_set_inj) apply (subst (asm) eq_linorder_class_conv) . lemma map_ord_eq_linorder[intro?]: "\eq_linorder cmpk; eq_linorder cmpv\ \ eq_linorder (cmp_map cmpk cmpv)" unfolding cmp_map_def apply rule apply rule apply (rule restrict_eq_linorder) apply rule apply rule apply assumption apply assumption apply simp apply (rule linorder.sorted_list_of_map_inj) apply (subst (asm) eq_linorder_class_conv) . definition cmp_unit :: "unit \ unit \ comp_res" where [simp]: "cmp_unit u v \ EQUAL" lemma cmp_unit_eq_linorder: "eq_linorder cmp_unit" by unfold_locales simp_all subsection \Parametricity\ lemma param_cmp_extend[param]: assumes "(cmp,cmp')\R \ R \ Id" assumes "Range R \ D" shows "(cmp,cmp_extend D cmp') \ R \ R \ Id" unfolding cmp_extend_def cmp_combine_def[abs_def] using assms apply clarsimp by (blast dest!: fun_relD) lemma param_cmp_img[param]: "(cmp_img,cmp_img) \ (Ra\Rb) \ (Rb\Rb\Rc) \ Ra \ Ra \ Rc" unfolding cmp_img_def[abs_def] by parametricity lemma param_comp_res[param]: "(LESS,LESS)\Id" "(EQUAL,EQUAL)\Id" "(GREATER,GREATER)\Id" "(case_comp_res,case_comp_res)\Ra\Ra\Ra\Id\Ra" by (auto split: comp_res.split) term cmp_lex lemma param_cmp_lex[param]: "(cmp_lex,cmp_lex)\(Ra\Rb\Id)\\Ra\list_rel\\Rb\list_rel\Id" unfolding cmp_lex_alt[abs_def] cmp_lex'_def by (parametricity) term cmp_prod lemma param_cmp_prod[param]: "(cmp_prod,cmp_prod)\ (Ra\Rb\Id)\(Rc\Rd\Id)\\Ra,Rc\prod_rel\\Rb,Rd\prod_rel\Id" unfolding cmp_prod_alt by (parametricity) lemma param_cmp_unit[param]: "(cmp_unit,cmp_unit)\Id\Id\Id" by auto lemma param_comp2eq[param]: "(comp2eq,comp2eq)\(R\R\Id)\R\R\Id" unfolding comp2eq_def[abs_def] by (parametricity) lemma cmp_combine_paramD: assumes "(cmp,cmp_combine D1 cmp1 D2 cmp2)\R\R\Id" assumes "Range R \ D1" shows "(cmp,cmp1)\R\R\Id" using assms unfolding cmp_combine_def[abs_def] apply (intro fun_relI) apply (drule_tac x=a in fun_relD, assumption) apply (drule_tac x=aa in fun_relD, assumption) apply (drule RangeI, drule (1) rev_subsetD) apply (drule RangeI, drule (1) rev_subsetD) apply simp done lemma cmp_extend_paramD: assumes "(cmp,cmp_extend D cmp')\R\R\Id" assumes "Range R \ D" shows "(cmp,cmp')\R\R\Id" using assms unfolding cmp_extend_def apply (rule cmp_combine_paramD) done subsection \Tuning of Generated Implementation\ lemma [autoref_post_simps]: "comp2eq (dflt_cmp (\) ((<)::_::linorder\_)) = (=)" by (simp add: class_to_eq_linorder eq_linorder_comp2eq_eq) end diff --git a/thys/Collections/Iterator/SetIterator.thy b/thys/Collections/Iterator/SetIterator.thy --- a/thys/Collections/Iterator/SetIterator.thy +++ b/thys/Collections/Iterator/SetIterator.thy @@ -1,795 +1,797 @@ (* Title: Iterators over Finite Sets Author: Thomas Tuerk Maintainer: Thomas Tuerk *) section \\isaheader{Iterators over Finite Sets}\ theory SetIterator imports Automatic_Refinement.Misc Automatic_Refinement.Foldi (*"../../Refine_Monadic/Refine_Monadic"*) begin text \When reasoning about finite sets, it is often handy to be able to iterate over the elements of the set. If the finite set is represented by a list, the @{term fold} operation can be used. For general finite sets, the order of elements is not fixed though. Therefore, nondeterministic iterators are introduced in this theory.\ subsection \General Iterators\ type_synonym ('x,'\) set_iterator = "('\\bool) \ ('x\'\\'\) \ '\ \ '\" locale set_iterator_genord = fixes iti::"('x,'\) set_iterator" and S0::"'x set" and R::"'x \ 'x \ bool" assumes foldli_transform: "\l0. distinct l0 \ S0 = set l0 \ sorted_wrt R l0 \ iti = foldli l0" begin text \Let's prove some trivial lemmata to show that the formalisation agrees with the view of iterators described above.\ lemma set_iterator_weaken_R : "(\x y. \x \ S0; y \ S0; R x y\ \ R' x y) \ set_iterator_genord iti S0 R'" by (metis set_iterator_genord.intro foldli_transform sorted_wrt_mono_rel) lemma finite_S0 : shows "finite S0" by (metis finite_set foldli_transform) lemma iti_stop_rule_cond : assumes not_c: "\(c \)" shows "iti c f \ = \" proof - from foldli_transform obtain l0 where l0_props: "iti c = foldli l0 c" by blast with foldli_not_cond [of c \ l0 f, OF not_c] show ?thesis by simp qed lemma iti_stop_rule_emp : assumes S0_emp: "S0 = {}" shows "iti c f \ = \" proof - from foldli_transform obtain l0 where l0_props: "S0 = set l0" "iti c = foldli l0 c" by blast with foldli.simps(1) [of c f \] S0_emp show ?thesis by simp qed text \Reducing everything to folding is cumbersome. Let's define a few high-level inference rules.\ lemma iteratei_rule_P: assumes "I S0 \0" "\S \ x. \ c \; x \ S; I S \; S \ S0; \y\S - {x}. R x y; \y\S0 - S. R y x\ \ I (S - {x}) (f x \)" "\\. I {} \ \ P \" "\\ S. \ S \ S0; S \ {}; \ c \; I S \; \x\S. \y\S0-S. R y x \ \ P \" shows "P (iti c f \0)" proof - { fix P I \0 f assume I: "I S0 \0" and R: "\S \ x. \c \; x \ S; I S \; S \ S0; \y\S-{x}. R x y\ \ I (S - {x}) (f x \)" and C1: "I {} (iti c f \0) \ P" and C2:"\S. \S \ S0; S \ {}; \ c (iti c f \0); I S (iti c f \0)\ \ P" from foldli_transform obtain l0 where l0_props: "distinct l0" "S0 = set l0" "sorted_wrt R l0" "iti c = foldli l0 c" by auto from I R have "I {} (iti c f \0) \ (\S. S \ S0 \ S \ {} \ \ (c (iti c f \0)) \ I S (iti c f \0))" unfolding l0_props using l0_props(1,3) proof (induct l0 arbitrary: \0) case Nil thus ?case by simp next case (Cons x l0) note ind_hyp = Cons(1) note I_x_l0 = Cons(2) note step = Cons(3) from Cons(4) have dist_l0: "distinct l0" and x_nin_l0: "x \ set l0" by simp_all from Cons(5) have R_l0: "\y\set l0. R x y" and sort_l0: "sorted_wrt R l0" by simp_all show ?case proof (cases "c \0") case False with I_x_l0 show ?thesis apply (rule_tac disjI2) apply (rule_tac exI[where x="set (x # l0)"]) apply (simp) done next case True note c_\0 = this from step[of \0 x "set (x # l0)"] I_x_l0 R_l0 c_\0 x_nin_l0 have step': "I (set l0) (f x \0)" by (simp_all add: Ball_def) from ind_hyp [OF step' step dist_l0 sort_l0] have "I {} (foldli l0 c f (f x \0)) \ (\S. S \ set l0 \ S \ {} \ \ c (foldli l0 c f (f x \0)) \ I S (foldli l0 c f (f x \0)))" by (fastforce) thus ?thesis by (simp add: c_\0 subset_iff) metis qed qed with C1 C2 have "P" by blast } note aux = this from assms show ?thesis apply (rule_tac aux [of "\S \. I S \ \ (\x\S. \y\S0-S. R y x)" \0 f]) apply auto done qed text \Instead of removing elements one by one from the invariant, adding them is sometimes more natural.\ lemma iteratei_rule_insert_P: assumes pre : "I {} \0" "\S \ x. \ c \; x \ S0 - S; I S \; S \ S0; \y\(S0 - S) - {x}. R x y; \y\S. R y x\ \ I (insert x S) (f x \)" "\\. I S0 \ \ P \" "\\ S. \ S \ S0; S \ S0; \ (c \); I S \; \x\S0-S. \y\S. R y x \ \ P \" shows "P (iti c f \0)" proof - let ?I' = "\S \. I (S0 - S) \" have pre1: "!!\ S. \ S \ S0; S \ {}; \ (c \); ?I' S \; \x\S. \y\S0-S. R y x\ \ P \" proof - fix S \ assume AA: "S \ S0" "S \ {}" "\ (c \)" "?I' S \" "\x\S. \y\S0-S. R y x" with pre(4) [of "S0 - S"] show "P \" by auto qed have pre2 :"\x S \. \c \; x \ S; S \ S0; ?I' S \; \y\S - {x}. R x y; \y\S0-S. R y x \ \ ?I' (S - {x}) (f x \)" proof - fix x S \ assume AA : "c \" "x \ S" "S \ S0" "?I' S \" "\y\S - {x}. R x y" "\y\S0 - S. R y x" from AA(2) AA(3) have "S0 - (S - {x}) = insert x (S0 - S)" "S0 - (S0 - S) = S" by auto moreover note pre(2) [of \ x "S0 - S"] AA ultimately show "?I' (S - {x}) (f x \)" by auto qed show "P (iti c f \0)" apply (rule iteratei_rule_P [of ?I' \0 c f P]) apply (simp add: pre) apply (rule pre2) apply simp_all apply (simp add: pre(3)) apply (simp add: pre1) done qed text \Show that iti without interruption is related to fold\ lemma iti_fold: assumes lc_f: "comp_fun_commute f" shows "iti (\_. True) f \0 = Finite_Set.fold f \0 S0" proof (rule iteratei_rule_insert_P [where I = "\X \'. \' = Finite_Set.fold f \0 X"]) fix S \ x assume "x \ S0 - S" "S \ S0" and \_eq: "\ = Finite_Set.fold f \0 S" from finite_S0 \S \ S0\ have fin_S: "finite S" by (metis finite_subset) from \x \ S0 - S\ have x_nin_S: "x \ S" by simp - note fold_eq = comp_fun_commute.fold_insert [OF lc_f fin_S x_nin_S] - show "f x \ = Finite_Set.fold f \0 (insert x S)" + interpret comp_fun_commute: comp_fun_commute f + by (fact lc_f) + note fold_eq = comp_fun_commute.fold_insert [OF fin_S x_nin_S] + show "f x \ = Finite_Set.fold f \0 (insert x S)" by (simp add: fold_eq \_eq) qed simp_all end subsection \Iterators over Maps\ type_synonym ('k,'v,'\) map_iterator = "('k\'v,'\) set_iterator" text \Iterator over the key-value pairs of a finite map are called iterators over maps.\ abbreviation "map_iterator_genord it m R \ set_iterator_genord it (map_to_set m) R" subsection \Unordered Iterators\ text \Often one does not care about the order in which the elements are processed. Therefore, the selection function can be set to not impose any further restrictings. This leads to considerably simpler theorems.\ definition "set_iterator it S0 \ set_iterator_genord it S0 (\_ _. True)" abbreviation "map_iterator it m \ set_iterator it (map_to_set m)" lemma set_iterator_intro : "set_iterator_genord it S0 R \ set_iterator it S0" unfolding set_iterator_def apply (rule set_iterator_genord.set_iterator_weaken_R [where R = R]) apply simp_all done lemma set_iterator_no_cond_rule_P: "\ set_iterator it S0; I S0 \0; !!S \ x. \ x \ S; I S \; S \ S0 \ \ I (S - {x}) (f x \); !!\. I {} \ \ P \ \ \ P (it (\_. True) f \0)" unfolding set_iterator_def using set_iterator_genord.iteratei_rule_P [of it S0 "\_ _. True" I \0 "\_. True" f P] by simp lemma set_iterator_no_cond_rule_insert_P: "\ set_iterator it S0; I {} \0; !!S \ x. \ x \ S0 - S; I S \; S \ S0 \ \ I (insert x S) (f x \); !!\. I S0 \ \ P \ \ \ P (it (\_. True) f \0)" unfolding set_iterator_def using set_iterator_genord.iteratei_rule_insert_P [of it S0 "\_ _. True" I \0 "\_. True" f P] by simp lemma set_iterator_rule_P: "\ set_iterator it S0; I S0 \0; !!S \ x. \ c \; x \ S; I S \; S \ S0 \ \ I (S - {x}) (f x \); !!\. I {} \ \ P \; !!\ S. S \ S0 \ S \ {} \ \ c \ \ I S \ \ P \ \ \ P (it c f \0)" unfolding set_iterator_def using set_iterator_genord.iteratei_rule_P [of it S0 "\_ _. True" I \0 c f P] by simp lemma set_iterator_rule_insert_P: "\ set_iterator it S0; I {} \0; !!S \ x. \ c \; x \ S0 - S; I S \; S \ S0 \ \ I (insert x S) (f x \); !!\. I S0 \ \ P \; !!\ S. S \ S0 \ S \ S0 \ \ c \ \ I S \ \ P \ \ \ P (it c f \0)" unfolding set_iterator_def using set_iterator_genord.iteratei_rule_insert_P [of it S0 "\_ _. True" I \0 c f P] by simp text\The following rules is adapted for maps. Instead of a set of key-value pairs the invariant now only sees the keys.\ lemma map_iterator_genord_rule_P: assumes "map_iterator_genord it m R" and I0: "I (dom m) \0" and IP: "!!k v it \. \ c \; k \ it; m k = Some v; it \ dom m; I it \; \k' v'. k' \ it-{k} \ m k' = Some v' \ R (k, v) (k', v'); \k' v'. k' \ it \ m k' = Some v' \ R (k', v') (k, v)\ \ I (it - {k}) (f (k, v) \)" and IF: "!!\. I {} \ \ P \" and II: "!!\ it. \ it \ dom m; it \ {}; \ c \; I it \; \k v k' v'. k \ it \ m k = Some v \ k' \ it \ m k' = Some v' \ R (k, v) (k', v') \ \ P \" shows "P (it c f \0)" proof (rule set_iterator_genord.iteratei_rule_P [of it "map_to_set m" R "\S \. I (fst ` S) \" \0 c f P]) show "map_iterator_genord it m R" by fact next show "I (fst ` map_to_set m) \0" using I0 by (simp add: map_to_set_dom[symmetric]) next fix \ assume "I (fst ` {}) \" with IF show "P \" by simp next fix \ S assume "S \ map_to_set m" "S \ {}" "\ c \" "I (fst ` S) \" and R_prop: "\x\S. \y\map_to_set m - S. R y x" let ?S' = "fst ` S" show "P \" proof (rule II [where it = ?S']) from \S \ map_to_set m\ show "?S' \ dom m" unfolding map_to_set_dom by auto next from \S \ {}\ show "?S' \ {}" by auto next show "\ (c \)" by fact next show "I (fst ` S) \" by fact next show "\k v k' v'. k \ fst ` S \ m k = Some v \ k' \ fst ` S \ m k' = Some v' \ R (k, v) (k', v')" proof (intro allI impI, elim conjE ) fix k v k' v' assume pre_k: "k \ fst ` S" "m k = Some v" assume pre_k': "k' \ fst ` S" "m k' = Some v'" from \S \ map_to_set m\ pre_k' have kv'_in: "(k', v') \ S" unfolding map_to_set_def by auto from \S \ map_to_set m\ pre_k have kv_in: "(k, v) \ map_to_set m - S" unfolding map_to_set_def by (auto simp add: image_iff) from R_prop kv_in kv'_in show "R (k, v) (k',v')" by simp qed qed next fix \ S kv assume "S \ map_to_set m" "kv \ S" "c \" and I_S': "I (fst ` S) \" and R_S: "\kv'\S - {kv}. R kv kv'" and R_not_S: "\kv'\map_to_set m - S. R kv' kv" let ?S' = "fst ` S" obtain k v where kv_eq[simp]: "kv = (k, v)" by (rule prod.exhaust) have "I (fst ` S - {k}) (f (k, v) \)" proof (rule IP) show "c \" by fact next from \kv \ S\ show "k \ ?S'" by (auto simp add: image_iff Bex_def) next from \kv \ S\ \S \ map_to_set m\ have "kv \ map_to_set m" by auto thus m_k_eq: "m k = Some v" unfolding map_to_set_def by simp next from \S \ map_to_set m\ show S'_subset: "?S' \ dom m" unfolding map_to_set_dom by auto next show "I (fst ` S) \" by fact next from \S \ map_to_set m\ \kv \ S\ have S_simp: "{(k', v'). k' \ (fst ` S) - {k} \ m k' = Some v'} = S - {kv}" unfolding map_to_set_def subset_iff apply (auto simp add: image_iff Bex_def) apply (metis option.inject) done from R_S[unfolded S_simp[symmetric]] R_not_S show "\k' v'. k' \ fst ` S - {k} \ m k' = Some v' \ R (k, v) (k', v') " by simp next from \S \ map_to_set m\ R_not_S show "\k' v'. k' \ fst ` S \ m k' = Some v' \ R (k', v') (k, v)" apply (simp add: Ball_def map_to_set_def subset_iff image_iff) apply metis done qed moreover from \S \ map_to_set m\ \kv \ S\ have "fst ` (S - {kv}) = fst ` S - {k}" apply (simp add: set_eq_iff image_iff Bex_def map_to_set_def subset_iff) apply (metis option.inject) done ultimately show "I (fst ` (S - {kv})) (f kv \)" by simp qed lemma map_iterator_genord_rule_insert_P: assumes "map_iterator_genord it m R" and I0: "I {} \0" and IP: "!!k v it \. \ c \; k \ dom m - it; m k = Some v; it \ dom m; I it \; \k' v'. k' \ (dom m - it) - {k} \ m k' = Some v' \ R (k, v) (k', v'); \k' v'. k' \ it \ m k' = Some v' \ R (k', v') (k, v)\ \ I (insert k it) (f (k, v) \)" and IF: "!!\. I (dom m) \ \ P \" and II: "!!\ it. \ it \ dom m; it \ dom m; \ c \; I it \; \k v k' v'. k \ it \ m k = Some v \ k' \ it \ m k' = Some v' \ R (k, v) (k', v') \ \ P \" shows "P (it c f \0)" proof (rule map_iterator_genord_rule_P [of it m R "\S \. I (dom m - S) \"]) show "map_iterator_genord it m R" by fact next show "I (dom m - dom m) \0" using I0 by simp next fix \ assume "I (dom m - {}) \" with IF show "P \" by simp next fix \ it assume assms: "it \ dom m" "it \ {}" "\ c \" "I (dom m - it) \" "\k v k' v'. k \ it \ m k = Some v \ k' \ it \ m k' = Some v' \ R (k, v) (k', v')" from assms have "dom m - it \ dom m" by auto with II[of "dom m - it" \] assms show "P \" apply (simp add: subset_iff dom_def) apply (metis option.simps(2)) done next fix k v it \ assume assms: "c \" "k \ it" "m k = Some v" "it \ dom m" "I (dom m - it) \" "\k' v'. k' \ it - {k} \ m k' = Some v' \ R (k, v) (k', v')" "\k' v'. k' \ it \ m k' = Some v' \ R (k', v') (k, v)" hence "insert k (dom m - it) = (dom m - (it - {k}))" "dom m - (dom m - it) = it" by auto with assms IP[of \ k "dom m - it" v] show "I (dom m - (it - {k})) (f (k, v) \)" by (simp_all add: subset_iff) qed lemma map_iterator_rule_P: assumes "map_iterator it m" and I0: "I (dom m) \0" and IP: "!!k v it \. \ c \; k \ it; m k = Some v; it \ dom m; I it \ \ \ I (it - {k}) (f (k, v) \)" and IF: "!!\. I {} \ \ P \" and II: "!!\ it. \ it \ dom m; it \ {}; \ c \; I it \ \ \ P \" shows "P (it c f \0)" using assms map_iterator_genord_rule_P[of it m "\_ _. True" I \0 c f P] unfolding set_iterator_def by simp lemma map_iterator_rule_insert_P: assumes "map_iterator it m" and I0: "I {} \0" and IP: "!!k v it \. \ c \; k \ dom m - it; m k = Some v; it \ dom m; I it \ \ \ I (insert k it) (f (k, v) \)" and IF: "!!\. I (dom m) \ \ P \" and II: "!!\ it. \ it \ dom m; it \ dom m; \ c \; I it \ \ \ P \" shows "P (it c f \0)" using assms map_iterator_genord_rule_insert_P[of it m "\_ _. True" I \0 c f P] unfolding set_iterator_def by simp lemma map_iterator_no_cond_rule_P: assumes "map_iterator it m" and I0: "I (dom m) \0" and IP: "!!k v it \. \ k \ it; m k = Some v; it \ dom m; I it \ \ \ I (it - {k}) (f (k, v) \)" and IF: "!!\. I {} \ \ P \" shows "P (it (\_. True) f \0)" using assms map_iterator_genord_rule_P[of it m "\_ _. True" I \0 "\_. True" f P] unfolding set_iterator_def by simp lemma map_iterator_no_cond_rule_insert_P: assumes "map_iterator it m" and I0: "I {} \0" and IP: "!!k v it \. \ k \ dom m - it; m k = Some v; it \ dom m; I it \ \ \ I (insert k it) (f (k, v) \)" and IF: "!!\. I (dom m) \ \ P \" shows "P (it (\_. True) f \0)" using assms map_iterator_genord_rule_insert_P[of it m "\_ _. True" I \0 "\_. True" f P] unfolding set_iterator_def by simp subsection \Ordered Iterators\ text \Selecting according to a linear order is another case that is interesting. Ordered iterators over maps, i.\,e.\ iterators over key-value pairs, use an order on the keys.\ context linorder begin definition "set_iterator_linord it S0 \ set_iterator_genord it S0 (\)" definition "set_iterator_rev_linord it S0 \ set_iterator_genord it S0 (\)" definition "set_iterator_map_linord it S0 \ set_iterator_genord it S0 (\(k,_) (k',_). k\k')" definition "set_iterator_map_rev_linord it S0 \ set_iterator_genord it S0 (\(k,_) (k',_). k\k')" abbreviation "map_iterator_linord it m \ set_iterator_map_linord it (map_to_set m)" abbreviation "map_iterator_rev_linord it m \ set_iterator_map_rev_linord it (map_to_set m)" lemma set_iterator_linord_rule_P: "\ set_iterator_linord it S0; I S0 \0; !!S \ x. \ c \; x \ S; I S \; S \ S0; \x'. x' \ S0-S \ x' \ x; \x'. x' \ S \ x \ x'\ \ I (S - {x}) (f x \); !!\. I {} \ \ P \; !!\ S. S \ S0 \ S \ {} \ (\x x'. \x \ S; x' \ S0-S\ \ x' \ x) \ \ c \ \ I S \ \ P \ \ \ P (it c f \0)" unfolding set_iterator_linord_def apply (rule set_iterator_genord.iteratei_rule_P [of it S0 "(\)" I \0 c f P]) apply (simp_all add: Ball_def) apply (metis order_refl) done lemma set_iterator_linord_rule_insert_P: "\ set_iterator_linord it S0; I {} \0; !!S \ x. \ c \; x \ S0 - S; I S \; S \ S0; \x'. x' \ S \ x' \ x; \x'. x' \ S0 - S \ x \ x'\ \ I (insert x S) (f x \); !!\. I S0 \ \ P \; !!\ S. S \ S0 \ S \ S0 \ (\x x'. \x \ S0-S; x' \ S\ \ x' \ x) \ \ c \ \ I S \ \ P \ \ \ P (it c f \0)" unfolding set_iterator_linord_def apply (rule set_iterator_genord.iteratei_rule_insert_P [of it S0 "(\)" I \0 c f P]) apply (simp_all add: Ball_def) apply (metis order_refl) done lemma set_iterator_rev_linord_rule_P: "\ set_iterator_rev_linord it S0; I S0 \0; !!S \ x. \ c \; x \ S; I S \; S \ S0; \x'. x' \ S0-S \ x \ x'; \x'. x' \ S \ x' \ x\ \ I (S - {x}) (f x \); !!\. I {} \ \ P \; !!\ S. S \ S0 \ S \ {} \ (\x x'. \x \ S; x' \ S0-S\ \ x \ x') \ \ c \ \ I S \ \ P \ \ \ P (it c f \0)" unfolding set_iterator_rev_linord_def apply (rule set_iterator_genord.iteratei_rule_P [of it S0 "(\)" I \0 c f P]) apply (simp_all add: Ball_def) apply (metis order_refl) done lemma set_iterator_rev_linord_rule_insert_P: "\ set_iterator_rev_linord it S0; I {} \0; !!S \ x. \ c \; x \ S0 - S; I S \; S \ S0; \x'. x' \ S \ x \ x'; \x'. x' \ S0 - S \ x' \ x\ \ I (insert x S) (f x \); !!\. I S0 \ \ P \; !!\ S. S \ S0 \ S \ S0 \ (\x x'. \x \ S0-S; x' \ S\ \ x \ x') \ \ c \ \ I S \ \ P \ \ \ P (it c f \0)" unfolding set_iterator_rev_linord_def apply (rule set_iterator_genord.iteratei_rule_insert_P [of it S0 "(\)" I \0 c f P]) apply (simp_all add: Ball_def) apply (metis order_refl) done lemma set_iterator_map_linord_rule_P: "\ set_iterator_map_linord it S0; I S0 \0; !!S \ k v. \ c \; (k, v) \ S; I S \; S \ S0; \k' v'. (k', v') \ S0-S \ k' \ k; \k' v'. (k', v') \ S \ k \ k'\ \ I (S - {(k,v)}) (f (k,v) \); !!\. I {} \ \ P \; !!\ S. S \ S0 \ S \ {} \ (\k v k' v'. \(k, v) \ S0-S; (k', v') \ S\ \ k \ k') \ \ c \ \ I S \ \ P \ \ \ P (it c f \0)" unfolding set_iterator_map_linord_def apply (rule set_iterator_genord.iteratei_rule_P [of it S0 "(\(k,_) (k',_). k \ k')" I \0 c f P]) apply simp_all apply (auto simp add: Ball_def) apply (metis order_refl) apply metis done lemma set_iterator_map_linord_rule_insert_P: "\ set_iterator_map_linord it S0; I {} \0; !!S \ k v. \ c \; (k, v) \ S0 - S; I S \; S \ S0; \k' v'. (k', v') \ S \ k' \ k; \k' v'. (k',v') \ S0 - S \ k \ k'\ \ I (insert (k,v) S) (f (k,v) \); !!\. I S0 \ \ P \; !!\ S. S \ S0 \ S \ S0 \ (\k v k' v'. \(k, v) \ S; (k', v') \ S0-S\ \ k \ k') \ \ c \ \ I S \ \ P \ \ \ P (it c f \0)" unfolding set_iterator_map_linord_def apply (rule set_iterator_genord.iteratei_rule_insert_P [of it S0 "(\(k,_) (k',_). k \ k')" I \0 c f P]) apply simp_all apply (auto simp add: Ball_def) apply (metis order_refl) apply metis done lemma set_iterator_map_rev_linord_rule_P: "\ set_iterator_map_rev_linord it S0; I S0 \0; !!S \ k v. \ c \; (k, v) \ S; I S \; S \ S0; \k' v'. (k', v') \ S0-S \ k \ k'; \k' v'. (k', v') \ S \ k' \ k\ \ I (S - {(k,v)}) (f (k,v) \); !!\. I {} \ \ P \; !!\ S. S \ S0 \ S \ {} \ (\k v k' v'. \(k, v) \ S0-S; (k', v') \ S\ \ k' \ k) \ \ c \ \ I S \ \ P \ \ \ P (it c f \0)" unfolding set_iterator_map_rev_linord_def apply (rule set_iterator_genord.iteratei_rule_P [of it S0 "(\(k,_) (k',_). k \ k')" I \0 c f P]) apply simp_all apply (auto simp add: Ball_def) apply (metis order_refl) apply metis done lemma set_iterator_map_rev_linord_rule_insert_P: "\ set_iterator_map_rev_linord it S0; I {} \0; !!S \ k v. \ c \; (k, v) \ S0 - S; I S \; S \ S0; \k' v'. (k', v') \ S \ k \ k'; \k' v'. (k',v') \ S0 - S \ k' \ k\ \ I (insert (k,v) S) (f (k,v) \); !!\. I S0 \ \ P \; !!\ S. S \ S0 \ S \ S0 \ (\k v k' v'. \(k, v) \ S; (k', v') \ S0-S\ \ k' \ k) \ \ c \ \ I S \ \ P \ \ \ P (it c f \0)" unfolding set_iterator_map_rev_linord_def apply (rule set_iterator_genord.iteratei_rule_insert_P [of it S0 "(\(k,_) (k',_). k \ k')" I \0 c f P]) apply simp_all apply (auto simp add: Ball_def) apply (metis order_refl) apply metis done lemma map_iterator_linord_rule_P: assumes "map_iterator_linord it m" and I0: "I (dom m) \0" and IP: "!!k v it \. \ c \; k \ it; m k = Some v; it \ dom m; I it \; \k'. k' \ it \ k \ k'; \k'. k' \ (dom m)-it \ k' \ k\ \ I (it - {k}) (f (k, v) \)" and IF: "!!\. I {} \ \ P \" and II: "!!\ it. \ it \ dom m; it \ {}; \ c \; I it \; \k k'. \k \ (dom m)-it; k' \ it\ \ k \ k'\ \ P \" shows "P (it c f \0)" using assms unfolding set_iterator_map_linord_def by (rule map_iterator_genord_rule_P) auto lemma map_iterator_linord_rule_insert_P: assumes "map_iterator_linord it m" and I0: "I {} \0" and IP: "!!k v it \. \ c \; k \ dom m - it; m k = Some v; it \ dom m; I it \; \k'. k' \ (dom m - it) \ k \ k'; \k'. k' \ it \ k' \ k \ \ I (insert k it) (f (k, v) \)" and IF: "!!\. I (dom m) \ \ P \" and II: "!!\ it. \ it \ dom m; it \ dom m; \ c \; I it \; \k k'. \k \ it; k' \ (dom m)-it\ \ k \ k'\ \ P \" shows "P (it c f \0)" using assms unfolding set_iterator_map_linord_def by (rule map_iterator_genord_rule_insert_P) auto lemma map_iterator_rev_linord_rule_P: assumes "map_iterator_rev_linord it m" and I0: "I (dom m) \0" and IP: "!!k v it \. \ c \; k \ it; m k = Some v; it \ dom m; I it \; \k'. k' \ it \ k' \ k; \k'. k' \ (dom m)-it \ k \ k'\ \ I (it - {k}) (f (k, v) \)" and IF: "!!\. I {} \ \ P \" and II: "!!\ it. \ it \ dom m; it \ {}; \ c \; I it \; \k k'. \k \ (dom m)-it; k' \ it\ \ k' \ k\ \ P \" shows "P (it c f \0)" using assms unfolding set_iterator_map_rev_linord_def by (rule map_iterator_genord_rule_P) auto lemma map_iterator_rev_linord_rule_insert_P: assumes "map_iterator_rev_linord it m" and I0: "I {} \0" and IP: "!!k v it \. \ c \; k \ dom m - it; m k = Some v; it \ dom m; I it \; \k'. k' \ (dom m - it) \ k' \ k; \k'. k' \ it \ k \ k'\ \ I (insert k it) (f (k, v) \)" and IF: "!!\. I (dom m) \ \ P \" and II: "!!\ it. \ it \ dom m; it \ dom m; \ c \; I it \; \k k'. \k \ it; k' \ (dom m)-it\ \ k' \ k\ \ P \" shows "P (it c f \0)" using assms unfolding set_iterator_map_rev_linord_def by (rule map_iterator_genord_rule_insert_P) auto end subsection \Conversions to foldli\ lemma set_iterator_genord_foldli_conv : "set_iterator_genord iti S R \ (\l0. distinct l0 \ S = set l0 \ sorted_wrt R l0 \ iti = foldli l0)" unfolding set_iterator_genord_def by simp lemma set_iterator_genord_I [intro] : "\distinct l0; S = set l0; sorted_wrt R l0; iti = foldli l0\ \ set_iterator_genord iti S R" unfolding set_iterator_genord_foldli_conv by blast lemma set_iterator_foldli_conv : "set_iterator iti S \ (\l0. distinct l0 \ S = set l0 \ iti = foldli l0)" unfolding set_iterator_def set_iterator_genord_def by simp lemma set_iterator_I [intro] : "\distinct l0; S = set l0; iti = foldli l0\ \ set_iterator iti S" unfolding set_iterator_foldli_conv by blast context linorder begin lemma set_iterator_linord_foldli_conv : "set_iterator_linord iti S \ (\l0. distinct l0 \ S = set l0 \ sorted l0 \ iti = foldli l0)" unfolding set_iterator_linord_def set_iterator_genord_def by simp lemma set_iterator_linord_I [intro] : "\distinct l0; S = set l0; sorted l0; iti = foldli l0\ \ set_iterator_linord iti S" unfolding set_iterator_linord_foldli_conv by blast lemma set_iterator_rev_linord_foldli_conv : "set_iterator_rev_linord iti S \ (\l0. distinct l0 \ S = set l0 \ sorted (rev l0) \ iti = foldli l0)" unfolding set_iterator_rev_linord_def set_iterator_genord_def by simp lemma set_iterator_rev_linord_I [intro] : "\distinct l0; S = set l0; sorted (rev l0); iti = foldli l0\ \ set_iterator_rev_linord iti S" unfolding set_iterator_rev_linord_foldli_conv by blast end lemma map_iterator_genord_foldli_conv : "map_iterator_genord iti m R \ (\(l0::('k \ 'v) list). distinct (map fst l0) \ m = map_of l0 \ sorted_wrt R l0 \ iti = foldli l0)" proof - { fix l0 :: "('k \ 'v) list" assume dist: "distinct l0" have "(map_to_set m = set l0) \ (distinct (map fst l0) \ m = map_of l0)" proof (cases "distinct (map fst l0)") case True thus ?thesis by (metis map_of_map_to_set) next case False note not_dist_fst = this with dist have "~(inj_on fst (set l0))" by (simp add: distinct_map) hence "set l0 \ map_to_set m" by (rule_tac notI) (simp add: map_to_set_def inj_on_def) with not_dist_fst show ?thesis by simp qed } thus ?thesis unfolding set_iterator_genord_def distinct_map by metis qed lemma map_iterator_genord_I [intro] : "\distinct (map fst l0); m = map_of l0; sorted_wrt R l0; iti = foldli l0\ \ map_iterator_genord iti m R" unfolding map_iterator_genord_foldli_conv by blast lemma map_iterator_foldli_conv : "map_iterator iti m \ (\l0. distinct (map fst l0) \ m = map_of l0 \ iti = foldli l0)" unfolding set_iterator_def map_iterator_genord_foldli_conv by simp lemma map_iterator_I [intro] : "\distinct (map fst l0); m = map_of l0; iti = foldli l0\ \ map_iterator iti m" unfolding map_iterator_foldli_conv by blast context linorder begin lemma sorted_wrt_keys_map_fst: "sorted_wrt (\(k,_) (k',_). R k k') l = sorted_wrt R (map fst l)" by (induct l) auto lemma map_iterator_linord_foldli_conv : "map_iterator_linord iti m \ (\l0. distinct (map fst l0) \ m = map_of l0 \ sorted (map fst l0) \ iti = foldli l0)" unfolding set_iterator_map_linord_def map_iterator_genord_foldli_conv by (simp add: sorted_wrt_keys_map_fst) lemma map_iterator_linord_I [intro] : "\distinct (map fst l0); m = map_of l0; sorted (map fst l0); iti = foldli l0\ \ map_iterator_linord iti m" unfolding map_iterator_linord_foldli_conv by blast lemma map_iterator_rev_linord_foldli_conv : "map_iterator_rev_linord iti m \ (\l0. distinct (map fst l0) \ m = map_of l0 \ sorted (rev (map fst l0)) \ iti = foldli l0)" unfolding set_iterator_map_rev_linord_def map_iterator_genord_foldli_conv by (simp add: sorted_wrt_keys_map_fst) lemma map_iterator_rev_linord_I [intro] : "\distinct (map fst l0); m = map_of l0; sorted (rev (map fst l0)); iti = foldli l0\ \ map_iterator_rev_linord iti m" unfolding map_iterator_rev_linord_foldli_conv by blast end end diff --git a/thys/Containers/Set_Impl.thy b/thys/Containers/Set_Impl.thy --- a/thys/Containers/Set_Impl.thy +++ b/thys/Containers/Set_Impl.thy @@ -1,2069 +1,2071 @@ (* Title: Containers/Set_Impl.thy Author: Andreas Lochbihler, KIT René Thiemann, UIBK *) theory Set_Impl imports Collection_Enum DList_Set RBT_Set2 Closure_Set Containers_Generator Complex_Main begin section \Different implementations of sets\ subsection \Auxiliary functions\ text \A simple quicksort implementation\ context ord begin function (sequential) quicksort_acc :: "'a list \ 'a list \ 'a list" and quicksort_part :: "'a list \ 'a \ 'a list \ 'a list \ 'a list \ 'a list \ 'a list" where "quicksort_acc ac [] = ac" | "quicksort_acc ac [x] = x # ac" | "quicksort_acc ac (x # xs) = quicksort_part ac x [] [] [] xs" | "quicksort_part ac x lts eqs gts [] = quicksort_acc (eqs @ x # quicksort_acc ac gts) lts" | "quicksort_part ac x lts eqs gts (z # zs) = (if z > x then quicksort_part ac x lts eqs (z # gts) zs else if z < x then quicksort_part ac x (z # lts) eqs gts zs else quicksort_part ac x lts (z # eqs) gts zs)" by pat_completeness simp_all lemma length_quicksort_accp: "quicksort_acc_quicksort_part_dom (Inl (ac, xs)) \ length (quicksort_acc ac xs) = length ac + length xs" and length_quicksort_partp: "quicksort_acc_quicksort_part_dom (Inr (ac, x, lts, eqs, gts, zs)) \ length (quicksort_part ac x lts eqs gts zs) = length ac + 1 + length lts + length eqs + length gts + length zs" apply(induct rule: quicksort_acc_quicksort_part.pinduct) apply(simp_all add: quicksort_acc.psimps quicksort_part.psimps) done termination apply(relation "measure (case_sum (\(_, xs). 2 * length xs ^ 2) (\(_, _, lts, eqs, gts, zs). 2 * (length lts + length eqs + length gts + length zs) ^ 2 + length zs + 1))") apply(simp_all add: power2_eq_square add_mult_distrib add_mult_distrib2 length_quicksort_accp) done definition quicksort :: "'a list \ 'a list" where "quicksort = quicksort_acc []" lemma set_quicksort_acc [simp]: "set (quicksort_acc ac xs) = set ac \ set xs" and set_quicksort_part [simp]: "set (quicksort_part ac x lts eqs gts zs) = set ac \ {x} \ set lts \ set eqs \ set gts \ set zs" by(induct ac xs and ac x lts eqs gts zs rule: quicksort_acc_quicksort_part.induct)(auto split: if_split_asm) lemma set_quicksort [simp]: "set (quicksort xs) = set xs" by(simp add: quicksort_def) lemma distinct_quicksort_acc: "distinct (quicksort_acc ac xs) = distinct (ac @ xs)" and distinct_quicksort_part: "distinct (quicksort_part ac x lts eqs gts zs) = distinct (ac @ [x] @ lts @ eqs @ gts @ zs)" by(induct ac xs and ac x lts eqs gts zs rule: quicksort_acc_quicksort_part.induct) auto lemma distinct_quicksort [simp]: "distinct (quicksort xs) = distinct xs" by(simp add: quicksort_def distinct_quicksort_acc) end lemmas [code] = ord.quicksort_acc.simps quicksort_acc.simps ord.quicksort_part.simps quicksort_part.simps ord.quicksort_def quicksort_def context linorder begin lemma sorted_quicksort_acc: "\ sorted ac; \x \ set xs. \a \ set ac. x < a \ \ sorted (quicksort_acc ac xs)" and sorted_quicksort_part: "\ sorted ac; \y \ set lts \ {x} \ set eqs \ set gts \ set zs. \a \ set ac. y < a; \y \ set lts. y < x; \y \ set eqs. y = x; \y \ set gts. y > x \ \ sorted (quicksort_part ac x lts eqs gts zs)" proof(induction ac xs and ac x lts eqs gts zs rule: quicksort_acc_quicksort_part.induct) case 1 thus ?case by simp next case 2 thus ?case by(auto) next case 3 thus ?case by simp next case (4 ac x lts eqs gts) note ac_greater = \\y\set lts \ {x} \ set eqs \ set gts \ set []. \a\set ac. y < a\ have "sorted eqs" "set eqs \ {x}" using \\y\set eqs. y = x\ by(induct eqs)(simp_all) moreover have "\y \ set ac \ set gts. x \ y" using \\a\set gts. x < a\ ac_greater by auto moreover have "sorted (quicksort_acc ac gts)" using \sorted ac\ ac_greater by(auto intro: "4.IH") ultimately have "sorted (eqs @ x # quicksort_acc ac gts)" by(auto simp add: sorted_append) moreover have "\y\set lts. \a\set (eqs @ x # quicksort_acc ac gts). y < a" using \\y\set lts. y < x\ ac_greater \\a\set gts. x < a\ \\y\set eqs. y = x\ by fastforce ultimately show ?case by(simp add: "4.IH") next case 5 thus ?case by(simp add: not_less order_eq_iff) qed lemma sorted_quicksort [simp]: "sorted (quicksort xs)" by(simp add: quicksort_def sorted_quicksort_acc) lemma insort_key_append1: "\y \ set ys. f x < f y \ insort_key f x (xs @ ys) = insort_key f x xs @ ys" proof(induct xs) case Nil thus ?case by(cases ys) auto qed simp lemma insort_key_append2: "\y \ set xs. f x > f y \ insort_key f x (xs @ ys) = xs @ insort_key f x ys" by(induct xs) auto lemma sort_key_append: "\x\set xs. \y\set ys. f x < f y \ sort_key f (xs @ ys) = sort_key f xs @ sort_key f ys" by(induct xs)(simp_all add: insort_key_append1) definition single_list :: "'a \ 'a list" where "single_list a = [a]" lemma to_single_list: "x # xs = single_list x @ xs" by(simp add: single_list_def) lemma sort_snoc: "sort (xs @ [x]) = insort x (sort xs)" by(induct xs)(simp_all add: insort_left_comm) lemma sort_append_swap: "sort (xs @ ys) = sort (ys @ xs)" by(induct xs arbitrary: ys rule: rev_induct)(simp_all add: sort_snoc[symmetric]) lemma sort_append_swap2: "sort (xs @ ys @ zs) = sort (ys @ xs @ zs)" by(induct xs)(simp_all, subst (1 2) sort_append_swap, simp) lemma sort_Cons_append_swap: "sort (x # xs) = sort (xs @ [x])" by(subst sort_append_swap) simp lemma sort_append_Cons_swap: "sort (ys @ x # xs) = sort (ys @ xs @ [x])" apply(induct ys) apply(simp only: append.simps sort_Cons_append_swap) apply simp done lemma quicksort_acc_conv_sort: "quicksort_acc ac xs = sort xs @ ac" and quicksort_part_conv_sort: "\ \y \ set lts. y < x; \y \ set eqs. y = x; \y \ set gts. y > x \ \ quicksort_part ac x lts eqs gts zs = sort (lts @ eqs @ gts @ x # zs) @ ac" proof(induct ac xs and ac x lts eqs gts zs rule: quicksort_acc_quicksort_part.induct) case 1 thus ?case by simp next case 2 thus ?case by simp next case 3 thus ?case by simp next case (4 ac x lts eqs gts) note eqs = \\y\set eqs. y = x\ { fix eqs assume "\y\set eqs. y = x" hence "insort x eqs = x # eqs" by(induct eqs) simp_all } note [simp] = this from eqs have [simp]: "sort eqs = eqs" by(induct eqs) simp_all from eqs have [simp]: "eqs @ [x] = x # eqs" by(induct eqs) simp_all show ?case using 4 apply(subst sort_key_append) apply(auto 4 3 dest: bspec)[1] apply(simp add: append_assoc[symmetric] sort_snoc del: append_assoc) apply(subst sort_key_append) apply(auto 4 3 simp add: insort_key_append1 dest: bspec) done next case (5 ac x lts eqs gts z zs) have "\ \ z < x; \ x < z \ \ z = x" by simp thus ?case using 5 apply(simp del: sort_key_simps) apply(safe, simp_all del: sort_key_simps add: to_single_list) apply(subst sort_append_swap) apply(fold append_assoc) apply(subst (2) sort_append_swap) apply(subst sort_append_swap2) apply(unfold append_assoc) apply(rule refl) apply(subst (1 5) append_assoc[symmetric]) apply(subst (1 2) sort_append_swap) apply(unfold append_assoc) apply(subst sort_append_swap2) apply(subst (1 2) sort_append_swap) apply(unfold append_assoc) apply(subst sort_append_swap2) apply(rule refl) apply(subst (2 6) append_assoc[symmetric]) apply(subst (2 5) append_assoc[symmetric]) apply(subst (1 2) sort_append_swap2) apply(subst (4) append_assoc) apply(subst (2) sort_append_swap2) apply simp done qed lemma quicksort_conv_sort: "quicksort xs = sort xs" by(simp add: quicksort_def quicksort_acc_conv_sort) lemma sort_remdups: "sort (remdups xs) = remdups (sort xs)" by(rule sorted_distinct_set_unique) simp_all end text \Removing duplicates from a sorted list\ context ord begin fun remdups_sorted :: "'a list \ 'a list" where "remdups_sorted [] = []" | "remdups_sorted [x] = [x]" | "remdups_sorted (x#y#xs) = (if x < y then x # remdups_sorted (y#xs) else remdups_sorted (y#xs))" end lemmas [code] = ord.remdups_sorted.simps context linorder begin lemma [simp]: assumes "sorted xs" shows sorted_remdups_sorted: "sorted (remdups_sorted xs)" and set_remdups_sorted: "set (remdups_sorted xs) = set xs" using assms by(induct xs rule: remdups_sorted.induct)(auto) lemma distinct_remdups_sorted [simp]: "sorted xs \ distinct (remdups_sorted xs)" by(induct xs rule: remdups_sorted.induct)(auto) lemma remdups_sorted_conv_remdups: "sorted xs \ remdups_sorted xs = remdups xs" by(induct xs rule: remdups_sorted.induct)(auto) end text \An specialised operation to convert a finite set into a sorted list\ definition csorted_list_of_set :: "'a :: ccompare set \ 'a list" where [code del]: "csorted_list_of_set A = (if ID CCOMPARE('a) = None \ \ finite A then undefined else linorder.sorted_list_of_set cless_eq A)" lemma csorted_list_of_set_set [simp]: "\ ID CCOMPARE('a :: ccompare) = Some c; linorder.sorted (le_of_comp c) xs; distinct xs \ \ linorder.sorted_list_of_set (le_of_comp c) (set xs) = xs" by(simp add: distinct_remdups_id linorder.sorted_list_of_set_sort_remdups[OF ID_ccompare] linorder.sorted_sort_id[OF ID_ccompare]) lemma csorted_list_of_set_split: fixes A :: "'a :: ccompare set" shows "P (csorted_list_of_set A) \ (\xs. ID CCOMPARE('a) \ None \ finite A \ A = set xs \ distinct xs \ linorder.sorted cless_eq xs \ P xs) \ (ID CCOMPARE('a) = None \ \ finite A \ P undefined)" by(auto simp add: csorted_list_of_set_def linorder.sorted_list_of_set[OF ID_ccompare]) code_identifier code_module Set \ (SML) Set_Impl | code_module Set_Impl \ (SML) Set_Impl subsection \Delete code equation with set as constructor\ lemma is_empty_unfold [code_unfold]: "set_eq A {} = Set.is_empty A" "set_eq {} A = Set.is_empty A" by(auto simp add: Set.is_empty_def set_eq_def) definition is_UNIV :: "'a set \ bool" where [code del]: "is_UNIV A \ A = UNIV" lemma is_UNIV_unfold [code_unfold]: "A = UNIV \ is_UNIV A" "UNIV = A \ is_UNIV A" "set_eq A UNIV \ is_UNIV A" "set_eq UNIV A \ is_UNIV A" by(auto simp add: is_UNIV_def set_eq_def) lemma [code_unfold del, symmetric, code_post del]: "x \ set xs \ List.member xs x" by(simp add: List.member_def) lemma [code_unfold del, symmetric, code_post del]: "finite \ Cardinality.finite'" by(simp) lemma [code_unfold del, symmetric, code_post del]: "card \ Cardinality.card'" by simp declare [[code drop: Set.empty Set.is_empty uminus_set_inst.uminus_set Set.member Set.insert Set.remove UNIV Set.filter image Set.subset_eq Ball Bex Set.union minus_set_inst.minus_set Set.inter card Set.bind the_elem Pow sum Gcd Lcm Product_Type.product Id_on Image trancl relcomp wf Min Inf_fin Max Sup_fin "Inf :: 'a set set \ 'a set" "Sup :: 'a set set \ 'a set" sorted_list_of_set List.map_project Sup_pred_inst.Sup_pred finite Cardinality.finite' card Cardinality.card' Inf_pred_inst.Inf_pred pred_of_set Cardinality.subset' Cardinality.eq_set Wellfounded.acc Bleast can_select "set_eq :: 'a set \ 'a set \ bool" irrefl bacc set_of_pred set_of_seq ]] declare Cardinality.finite'_def[code] Cardinality.card'_def[code] subsection \Set implementations\ definition Collect_set :: "('a \ bool) \ 'a set" where [simp]: "Collect_set = Collect" definition DList_set :: "'a :: ceq set_dlist \ 'a set" where "DList_set = Collect o DList_Set.member" definition RBT_set :: "'a :: ccompare set_rbt \ 'a set" where "RBT_set = Collect o RBT_Set2.member" definition Complement :: "'a set \ 'a set" where [simp]: "Complement A = - A" definition Set_Monad :: "'a list \ 'a set" where [simp]: "Set_Monad = set" code_datatype Collect_set DList_set RBT_set Set_Monad Complement lemma DList_set_empty [simp]: "DList_set DList_Set.empty = {}" by(simp add: DList_set_def) lemma RBT_set_empty [simp]: "RBT_set RBT_Set2.empty = {}" by(simp add: RBT_set_def) lemma RBT_set_conv_keys: "ID CCOMPARE('a :: ccompare) \ None \ RBT_set (t :: 'a set_rbt) = set (RBT_Set2.keys t)" by(clarsimp simp add: RBT_set_def member_conv_keys) subsection \Set operations\ text \ A collection of all the theorems about @{const Complement}. \ ML \ structure Set_Complement_Eqs = Named_Thms ( val name = @{binding set_complement_code} val description = "Code equations involving set complement" ) \ setup \Set_Complement_Eqs.setup\ text \Various fold operations over sets\ typedef ('a, 'b) comp_fun_commute = "{f :: 'a \ 'b \ 'b. comp_fun_commute f}" morphisms comp_fun_commute_apply Abs_comp_fun_commute by(rule exI[where x="\_. id"])(simp, unfold_locales, auto) setup_lifting type_definition_comp_fun_commute lemma comp_fun_commute_apply' [simp]: - "comp_fun_commute (comp_fun_commute_apply f)" -using comp_fun_commute_apply[of f] by simp + "comp_fun_commute_on UNIV (comp_fun_commute_apply f)" +using comp_fun_commute_apply[of f] by (simp add: comp_fun_commute_def') lift_definition set_fold_cfc :: "('a, 'b) comp_fun_commute \ 'b \ 'a set \ 'b" is "Finite_Set.fold" . declare [[code drop: set_fold_cfc]] lemma set_fold_cfc_code [code]: fixes xs :: "'a :: ceq list" and dxs :: "'a :: ceq set_dlist" and rbt :: "'b :: ccompare set_rbt" shows set_fold_cfc_Complement[set_complement_code]: "set_fold_cfc f''' b (Complement A) = Code.abort (STR ''set_fold_cfc not supported on Complement'') (\_. set_fold_cfc f''' b (Complement A))" and "set_fold_cfc f''' b (Collect_set P) = Code.abort (STR ''set_fold_cfc not supported on Collect_set'') (\_. set_fold_cfc f''' b (Collect_set P))" "set_fold_cfc f b (Set_Monad xs) = (case ID CEQ('a) of None \ Code.abort (STR ''set_fold_cfc Set_Monad: ceq = None'') (\_. set_fold_cfc f b (Set_Monad xs)) | Some eq \ List.fold (comp_fun_commute_apply f) (equal_base.list_remdups eq xs) b)" (is ?Set_Monad) "set_fold_cfc f' b (DList_set dxs) = (case ID CEQ('a) of None \ Code.abort (STR ''set_fold_cfc DList_set: ceq = None'') (\_. set_fold_cfc f' b (DList_set dxs)) | Some _ \ DList_Set.fold (comp_fun_commute_apply f') dxs b)" (is ?DList_set) "set_fold_cfc f'' b (RBT_set rbt) = (case ID CCOMPARE('b) of None \ Code.abort (STR ''set_fold_cfc RBT_set: ccompare = None'') (\_. set_fold_cfc f'' b (RBT_set rbt)) | Some _ \ RBT_Set2.fold (comp_fun_commute_apply f'') rbt b)" (is ?RBT_set) proof - + note fold_set_fold_remdups = comp_fun_commute_def' comp_fun_commute_on.fold_set_fold_remdups[OF _ subset_UNIV] show ?Set_Monad - by(auto split: option.split dest!: Collection_Eq.ID_ceq simp add: set_fold_cfc_def comp_fun_commute.fold_set_fold_remdups) - show ?DList_set - apply(auto split: option.split simp add: DList_set_def) - apply transfer - apply(auto dest: Collection_Eq.ID_ceq simp add: List.member_def[abs_def] comp_fun_commute.fold_set_fold_remdups distinct_remdups_id) + by(auto split: option.split dest!: Collection_Eq.ID_ceq simp add: set_fold_cfc_def fold_set_fold_remdups) + show ?DList_set + apply(auto split: option.splits simp add: DList_set_def) + apply transfer + apply(auto dest: Collection_Eq.ID_ceq simp add: List.member_def[abs_def] fold_set_fold_remdups distinct_remdups_id) done show ?RBT_set apply(auto split: option.split simp add: RBT_set_conv_keys fold_conv_fold_keys) apply transfer - apply(simp add: comp_fun_commute.fold_set_fold_remdups distinct_remdups_id linorder.distinct_keys[OF ID_ccompare] ord.is_rbt_rbt_sorted) + apply(simp add: fold_set_fold_remdups distinct_remdups_id linorder.distinct_keys[OF ID_ccompare] ord.is_rbt_rbt_sorted) done qed simp_all typedef ('a, 'b) comp_fun_idem = "{f :: 'a \ 'b \ 'b. comp_fun_idem f}" morphisms comp_fun_idem_apply Abs_comp_fun_idem by(rule exI[where x="\_. id"])(simp, unfold_locales, auto) setup_lifting type_definition_comp_fun_idem lemma comp_fun_idem_apply' [simp]: - "comp_fun_idem (comp_fun_idem_apply f)" -using comp_fun_idem_apply[of f] by simp + "comp_fun_idem_on UNIV (comp_fun_idem_apply f)" +using comp_fun_idem_apply[of f] by (simp add: comp_fun_idem_def') lift_definition set_fold_cfi :: "('a, 'b) comp_fun_idem \ 'b \ 'a set \ 'b" is "Finite_Set.fold" . declare [[code drop: set_fold_cfi]] lemma set_fold_cfi_code [code]: fixes xs :: "'a list" and dxs :: "'b :: ceq set_dlist" and rbt :: "'c :: ccompare set_rbt" shows "set_fold_cfi f b (Complement A) = Code.abort (STR ''set_fold_cfi not supported on Complement'') (\_. set_fold_cfi f b (Complement A))" "set_fold_cfi f b (Collect_set P) = Code.abort (STR ''set_fold_cfi not supported on Collect_set'') (\_. set_fold_cfi f b (Collect_set P))" "set_fold_cfi f b (Set_Monad xs) = List.fold (comp_fun_idem_apply f) xs b" (is ?Set_Monad) "set_fold_cfi f' b (DList_set dxs) = (case ID CEQ('b) of None \ Code.abort (STR ''set_fold_cfi DList_set: ceq = None'') (\_. set_fold_cfi f' b (DList_set dxs)) | Some _ \ DList_Set.fold (comp_fun_idem_apply f') dxs b)" (is ?DList_set) "set_fold_cfi f'' b (RBT_set rbt) = (case ID CCOMPARE('c) of None \ Code.abort (STR ''set_fold_cfi RBT_set: ccompare = None'') (\_. set_fold_cfi f'' b (RBT_set rbt)) | Some _ \ RBT_Set2.fold (comp_fun_idem_apply f'') rbt b)" (is ?RBT_set) proof - show ?Set_Monad - by(auto split: option.split dest!: Collection_Eq.ID_ceq simp add: set_fold_cfi_def comp_fun_idem.fold_set_fold) + by(auto split: option.split dest!: Collection_Eq.ID_ceq simp add: set_fold_cfi_def comp_fun_idem_def' comp_fun_idem_on.fold_set_fold[OF _ subset_UNIV]) show ?DList_set apply(auto split: option.split simp add: DList_set_def) apply transfer - apply(auto dest: Collection_Eq.ID_ceq simp add: List.member_def[abs_def] comp_fun_idem.fold_set_fold) + apply(auto dest: Collection_Eq.ID_ceq simp add: List.member_def[abs_def] comp_fun_idem_def' comp_fun_idem_on.fold_set_fold[OF _ subset_UNIV]) done show ?RBT_set apply(auto split: option.split simp add: RBT_set_conv_keys fold_conv_fold_keys) apply transfer - apply(simp add: comp_fun_idem.fold_set_fold) + apply(simp add: comp_fun_idem_def' comp_fun_idem_on.fold_set_fold[OF _ subset_UNIV]) done qed simp_all typedef 'a semilattice_set = "{f :: 'a \ 'a \ 'a. semilattice_set f}" morphisms semilattice_set_apply Abs_semilattice_set proof show "(\x y. if x = y then x else undefined) \ ?semilattice_set" unfolding mem_Collect_eq by(unfold_locales) simp_all qed setup_lifting type_definition_semilattice_set lemma semilattice_set_apply' [simp]: "semilattice_set (semilattice_set_apply f)" using semilattice_set_apply[of f] by simp lemma comp_fun_idem_semilattice_set_apply [simp]: - "comp_fun_idem (semilattice_set_apply f)" + "comp_fun_idem_on UNIV (semilattice_set_apply f)" proof - interpret semilattice_set "semilattice_set_apply f" by simp show ?thesis by(unfold_locales)(simp_all add: fun_eq_iff left_commute) qed lift_definition set_fold1 :: "'a semilattice_set \ 'a set \ 'a" is "semilattice_set.F" . lemma (in semilattice_set) F_set_conv_fold: "xs \ [] \ F (set xs) = Finite_Set.fold f (hd xs) (set (tl xs))" by(clarsimp simp add: neq_Nil_conv eq_fold) lemma set_fold1_code [code]: fixes rbt :: "'a :: {ccompare, lattice} set_rbt" and dxs :: "'b :: {ceq, lattice} set_dlist" shows set_fold1_Complement[set_complement_code]: "set_fold1 f (Complement A) = Code.abort (STR ''set_fold1: Complement'') (\_. set_fold1 f (Complement A))" and "set_fold1 f (Collect_set P) = Code.abort (STR ''set_fold1: Collect_set'') (\_. set_fold1 f (Collect_set P))" and "set_fold1 f (Set_Monad (x # xs)) = fold (semilattice_set_apply f) xs x" (is "?Set_Monad") and "set_fold1 f' (DList_set dxs) = (case ID CEQ('b) of None \ Code.abort (STR ''set_fold1 DList_set: ceq = None'') (\_. set_fold1 f' (DList_set dxs)) | Some _ \ if DList_Set.null dxs then Code.abort (STR ''set_fold1 DList_set: empty set'') (\_. set_fold1 f' (DList_set dxs)) else DList_Set.fold (semilattice_set_apply f') (DList_Set.tl dxs) (DList_Set.hd dxs))" (is "?DList_set") and "set_fold1 f'' (RBT_set rbt) = (case ID CCOMPARE('a) of None \ Code.abort (STR ''set_fold1 RBT_set: ccompare = None'') (\_. set_fold1 f'' (RBT_set rbt)) | Some _ \ if RBT_Set2.is_empty rbt then Code.abort (STR ''set_fold1 RBT_set: empty set'') (\_. set_fold1 f'' (RBT_set rbt)) else RBT_Set2.fold1 (semilattice_set_apply f'') rbt)" (is "?RBT_set") proof - + note fold_set_fold = comp_fun_idem_def' comp_fun_idem_on.fold_set_fold[OF _ subset_UNIV] show ?Set_Monad - by(simp add: set_fold1_def semilattice_set.eq_fold comp_fun_idem.fold_set_fold) + by(simp add: set_fold1_def semilattice_set.eq_fold fold_set_fold) show ?DList_set - by(simp add: set_fold1_def semilattice_set.F_set_conv_fold comp_fun_idem.fold_set_fold DList_set_def DList_Set.Collect_member split: option.split)(transfer, simp) + by(simp add: set_fold1_def semilattice_set.F_set_conv_fold fold_set_fold DList_set_def DList_Set.Collect_member split: option.split)(transfer, simp) show ?RBT_set - by(simp add: set_fold1_def semilattice_set.F_set_conv_fold comp_fun_idem.fold_set_fold RBT_set_def RBT_Set2.member_conv_keys RBT_Set2.fold1_conv_fold split: option.split) + by(simp add: set_fold1_def semilattice_set.F_set_conv_fold fold_set_fold RBT_set_def RBT_Set2.member_conv_keys RBT_Set2.fold1_conv_fold split: option.split) qed simp_all text \Implementation of set operations\ lemma Collect_code [code]: fixes P :: "'a :: cenum \ bool" shows "Collect P = (case ID CENUM('a) of None \ Collect_set P | Some (enum, _) \ Set_Monad (filter P enum))" by(auto split: option.split dest: in_cenum) lemma finite_code [code]: fixes dxs :: "'a :: ceq set_dlist" and rbt :: "'b :: ccompare set_rbt" and A :: "'c :: finite_UNIV set" and P :: "'c \ bool" shows "finite (DList_set dxs) = (case ID CEQ('a) of None \ Code.abort (STR ''finite DList_set: ceq = None'') (\_. finite (DList_set dxs)) | Some _ \ True)" "finite (RBT_set rbt) = (case ID CCOMPARE('b) of None \ Code.abort (STR ''finite RBT_set: ccompare = None'') (\_. finite (RBT_set rbt)) | Some _ \ True)" and finite_Complement [set_complement_code]: "finite (Complement A) \ (if of_phantom (finite_UNIV :: 'c finite_UNIV) then True else if finite A then False else Code.abort (STR ''finite Complement: infinite set'') (\_. finite (Complement A)))" and "finite (Set_Monad xs) = True" "finite (Collect_set P) \ of_phantom (finite_UNIV :: 'c finite_UNIV) \ Code.abort (STR ''finite Collect_set'') (\_. finite (Collect_set P))" by(auto simp add: DList_set_def RBT_set_def member_conv_keys card_gt_0_iff finite_UNIV split: option.split elim: finite_subset[rotated 1]) lemma card_code [code]: fixes dxs :: "'a :: ceq set_dlist" and xs :: "'a list" and rbt :: "'b :: ccompare set_rbt" and A :: "'c :: card_UNIV set" shows "card (DList_set dxs) = (case ID CEQ('a) of None \ Code.abort (STR ''card DList_set: ceq = None'') (\_. card (DList_set dxs)) | Some _ \ DList_Set.length dxs)" "card (RBT_set rbt) = (case ID CCOMPARE('b) of None \ Code.abort (STR ''card RBT_set: ccompare = None'') (\_. card (RBT_set rbt)) | Some _ \ length (RBT_Set2.keys rbt))" "card (Set_Monad xs) = (case ID CEQ('a) of None \ Code.abort (STR ''card Set_Monad: ceq = None'') (\_. card (Set_Monad xs)) | Some eq \ length (equal_base.list_remdups eq xs))" and card_Complement [set_complement_code]: "card (Complement A) = (let a = card A; s = CARD('c) in if s > 0 then s - a else if finite A then 0 else Code.abort (STR ''card Complement: infinite'') (\_. card (Complement A)))" by(auto simp add: RBT_set_def member_conv_keys distinct_card DList_set_def Let_def card_UNIV Compl_eq_Diff_UNIV card_Diff_subset_Int card_gt_0_iff finite_subset[of A UNIV] List.card_set dest: Collection_Eq.ID_ceq split: option.split) lemma is_UNIV_code [code]: fixes rbt :: "'a :: {cproper_interval, finite_UNIV} set_rbt" and A :: "'b :: card_UNIV set" shows "is_UNIV A \ (let a = CARD('b); b = card A in if a > 0 then a = b else if b > 0 then False else Code.abort (STR ''is_UNIV called on infinite type and set'') (\_. is_UNIV A))" (is ?generic) "is_UNIV (RBT_set rbt) = (case ID CCOMPARE('a) of None \ Code.abort (STR ''is_UNIV RBT_set: ccompare = None'') (\_. is_UNIV (RBT_set rbt)) | Some _ \ of_phantom (finite_UNIV :: 'a finite_UNIV) \ proper_intrvl.exhaustive_fusion cproper_interval rbt_keys_generator (RBT_Set2.init rbt))" (is ?rbt) proof - { fix c assume linorder: "ID CCOMPARE('a) = Some c" have "is_UNIV (RBT_set rbt) = (finite (UNIV :: 'a set) \ proper_intrvl.exhaustive cproper_interval (RBT_Set2.keys rbt))" (is "?lhs \ ?rhs") proof assume ?lhs have "finite (UNIV :: 'a set)" unfolding \?lhs\[unfolded is_UNIV_def, symmetric] using linorder by(simp add: finite_code) moreover hence "proper_intrvl.exhaustive cproper_interval (RBT_Set2.keys rbt)" using linorder \?lhs\ by(simp add: linorder_proper_interval.exhaustive_correct[OF ID_ccompare_interval[OF linorder]] sorted_RBT_Set_keys is_UNIV_def RBT_set_conv_keys) ultimately show ?rhs .. next assume ?rhs thus ?lhs using linorder by(auto simp add: linorder_proper_interval.exhaustive_correct[OF ID_ccompare_interval[OF linorder]] sorted_RBT_Set_keys is_UNIV_def RBT_set_conv_keys) qed } thus ?rbt by(auto simp add: finite_UNIV proper_intrvl.exhaustive_fusion_def unfoldr_rbt_keys_generator is_UNIV_def split: option.split) show ?generic by(auto simp add: Let_def is_UNIV_def dest: card_seteq[of UNIV A] dest!: card_ge_0_finite) qed lemma is_empty_code [code]: fixes dxs :: "'a :: ceq set_dlist" and rbt :: "'b :: ccompare set_rbt" and A :: "'c set" shows "Set.is_empty (Set_Monad xs) \ xs = []" "Set.is_empty (DList_set dxs) \ (case ID CEQ('a) of None \ Code.abort (STR ''is_empty DList_set: ceq = None'') (\_. Set.is_empty (DList_set dxs)) | Some _ \ DList_Set.null dxs)" (is ?DList_set) "Set.is_empty (RBT_set rbt) \ (case ID CCOMPARE('b) of None \ Code.abort (STR ''is_empty RBT_set: ccompare = None'') (\_. Set.is_empty (RBT_set rbt)) | Some _ \ RBT_Set2.is_empty rbt)" (is ?RBT_set) and is_empty_Complement [set_complement_code]: "Set.is_empty (Complement A) \ is_UNIV A" (is ?Complement) proof - show ?DList_set by(clarsimp simp add: DList_set_def Set.is_empty_def DList_Set.member_empty_empty split: option.split) show ?RBT_set by(clarsimp simp add: RBT_set_def Set.is_empty_def RBT_Set2.member_empty_empty[symmetric] fun_eq_iff simp del: RBT_Set2.member_empty_empty split: option.split) show ?Complement by(auto simp add: is_UNIV_def Set.is_empty_def) qed(simp_all add: Set.is_empty_def List.null_def) lemma Set_insert_code [code]: fixes dxs :: "'a :: ceq set_dlist" and rbt :: "'b :: ccompare set_rbt" shows "\x. Set.insert x (Collect_set A) = (case ID CEQ('a) of None \ Code.abort (STR ''insert Collect_set: ceq = None'') (\_. Set.insert x (Collect_set A)) | Some eq \ Collect_set (equal_base.fun_upd eq A x True))" "\x. Set.insert x (Set_Monad xs) = Set_Monad (x # xs)" "\x. Set.insert x (DList_set dxs) = (case ID CEQ('a) of None \ Code.abort (STR ''insert DList_set: ceq = None'') (\_. Set.insert x (DList_set dxs)) | Some _ \ DList_set (DList_Set.insert x dxs))" "\x. Set.insert x (RBT_set rbt) = (case ID CCOMPARE('b) of None \ Code.abort (STR ''insert RBT_set: ccompare = None'') (\_. Set.insert x (RBT_set rbt)) | Some _ \ RBT_set (RBT_Set2.insert x rbt))" and insert_Complement [set_complement_code]: "\x. Set.insert x (Complement X) = Complement (Set.remove x X)" by(auto split: option.split dest: equal.equal_eq[OF ID_ceq] simp add: DList_set_def DList_Set.member_insert RBT_set_def) lemma Set_member_code [code]: fixes xs :: "'a :: ceq list" shows "\x. x \ Collect_set A \ A x" "\x. x \ DList_set dxs \ DList_Set.member dxs x" "\x. x \ RBT_set rbt \ RBT_Set2.member rbt x" and mem_Complement [set_complement_code]: "\x. x \ Complement X \ x \ X" and "\x. x \ Set_Monad xs \ (case ID CEQ('a) of None \ Code.abort (STR ''member Set_Monad: ceq = None'') (\_. x \ Set_Monad xs) | Some eq \ equal_base.list_member eq xs x)" by(auto simp add: DList_set_def RBT_set_def List.member_def split: option.split dest!: Collection_Eq.ID_ceq) lemma Set_remove_code [code]: fixes rbt :: "'a :: ccompare set_rbt" and dxs :: "'b :: ceq set_dlist" shows "\x. Set.remove x (Collect_set A) = (case ID CEQ('b) of None \ Code.abort (STR ''remove Collect: ceq = None'') (\_. Set.remove x (Collect_set A)) | Some eq \ Collect_set (equal_base.fun_upd eq A x False))" "\x. Set.remove x (DList_set dxs) = (case ID CEQ('b) of None \ Code.abort (STR ''remove DList_set: ceq = None'') (\_. Set.remove x (DList_set dxs)) | Some _ \ DList_set (DList_Set.remove x dxs))" "\x. Set.remove x (RBT_set rbt) = (case ID CCOMPARE('a) of None \ Code.abort (STR ''remove RBT_set: ccompare = None'') (\_. Set.remove x (RBT_set rbt)) | Some _ \ RBT_set (RBT_Set2.remove x rbt))" and remove_Complement [set_complement_code]: "\x A. Set.remove x (Complement A) = Complement (Set.insert x A)" by(auto split: option.split if_split_asm dest: equal.equal_eq[OF ID_ceq] simp add: DList_set_def DList_Set.member_remove RBT_set_def) lemma Set_uminus_code [code, set_complement_code]: "- A = Complement A" "- (Collect_set P) = Collect_set (\x. \ P x)" "- (Complement B) = B" by auto text \ These equations represent complements as true complements. If you want that the complement operations returns an explicit enumeration of the elements, use the following set of equations which use @{class cenum}. \ lemma Set_uminus_cenum: fixes A :: "'a :: cenum set" shows "- A = (case ID CENUM('a) of None \ Complement A | Some (enum, _) \ Set_Monad (filter (\x. x \ A) enum))" and "- (Complement B) = B" by(auto split: option.split dest: ID_cEnum) lemma Set_minus_code [code]: fixes rbt1 rbt2 :: "'a :: ccompare set_rbt" shows "A - B = A \ (- B)" "RBT_set rbt1 - RBT_set rbt2 = (case ID CCOMPARE('a) of None \ Code.abort (STR ''minus RBT_set RBT_set: ccompare = None'') (\_. RBT_set rbt1 - RBT_set rbt2) | Some _ \ RBT_set (RBT_Set2.minus rbt1 rbt2))" by (auto simp: Set_member_code(3) split: option.splits) lemma Set_union_code [code]: fixes rbt1 rbt2 :: "'a :: ccompare set_rbt" and rbt :: "'b :: {ccompare, ceq} set_rbt" and dxs :: "'b set_dlist" and dxs1 dxs2 :: "'c :: ceq set_dlist" shows "RBT_set rbt1 \ RBT_set rbt2 = (case ID CCOMPARE('a) of None \ Code.abort (STR ''union RBT_set RBT_set: ccompare = None'') (\_. RBT_set rbt1 \ RBT_set rbt2) | Some _ \ RBT_set (RBT_Set2.union rbt1 rbt2))" (is ?RBT_set_RBT_set) "RBT_set rbt \ DList_set dxs = (case ID CCOMPARE('b) of None \ Code.abort (STR ''union RBT_set DList_set: ccompare = None'') (\_. RBT_set rbt \ DList_set dxs) | Some _ \ case ID CEQ('b) of None \ Code.abort (STR ''union RBT_set DList_set: ceq = None'') (\_. RBT_set rbt \ DList_set dxs) | Some _ \ RBT_set (DList_Set.fold RBT_Set2.insert dxs rbt))" (is ?RBT_set_DList_set) "DList_set dxs \ RBT_set rbt = (case ID CCOMPARE('b) of None \ Code.abort (STR ''union DList_set RBT_set: ccompare = None'') (\_. RBT_set rbt \ DList_set dxs) | Some _ \ case ID CEQ('b) of None \ Code.abort (STR ''union DList_set RBT_set: ceq = None'') (\_. RBT_set rbt \ DList_set dxs) | Some _ \ RBT_set (DList_Set.fold RBT_Set2.insert dxs rbt))" (is ?DList_set_RBT_set) "DList_set dxs1 \ DList_set dxs2 = (case ID CEQ('c) of None \ Code.abort (STR ''union DList_set DList_set: ceq = None'') (\_. DList_set dxs1 \ DList_set dxs2) | Some _ \ DList_set (DList_Set.union dxs1 dxs2))" (is ?DList_set_DList_set) "Set_Monad zs \ RBT_set rbt2 = (case ID CCOMPARE('a) of None \ Code.abort (STR ''union Set_Monad RBT_set: ccompare = None'') (\_. Set_Monad zs \ RBT_set rbt2) | Some _ \ RBT_set (fold RBT_Set2.insert zs rbt2))" (is ?Set_Monad_RBT_set) "RBT_set rbt1 \ Set_Monad zs = (case ID CCOMPARE('a) of None \ Code.abort (STR ''union RBT_set Set_Monad: ccompare = None'') (\_. RBT_set rbt1 \ Set_Monad zs) | Some _ \ RBT_set (fold RBT_Set2.insert zs rbt1))" (is ?RBT_set_Set_Monad) "Set_Monad ws \ DList_set dxs2 = (case ID CEQ('c) of None \ Code.abort (STR ''union Set_Monad DList_set: ceq = None'') (\_. Set_Monad ws \ DList_set dxs2) | Some _ \ DList_set (fold DList_Set.insert ws dxs2))" (is ?Set_Monad_DList_set) "DList_set dxs1 \ Set_Monad ws = (case ID CEQ('c) of None \ Code.abort (STR ''union DList_set Set_Monad: ceq = None'') (\_. DList_set dxs1 \ Set_Monad ws) | Some _ \ DList_set (fold DList_Set.insert ws dxs1))" (is ?DList_set_Set_Monad) "Set_Monad xs \ Set_Monad ys = Set_Monad (xs @ ys)" "Collect_set A \ B = Collect_set (\x. A x \ x \ B)" "B \ Collect_set A = Collect_set (\x. A x \ x \ B)" and Set_union_Complement [set_complement_code]: "Complement B \ B' = Complement (B \ - B')" "B' \ Complement B = Complement (- B' \ B)" proof - show ?RBT_set_RBT_set ?Set_Monad_RBT_set ?RBT_set_Set_Monad by(auto split: option.split simp add: RBT_set_def) show ?RBT_set_DList_set ?DList_set_RBT_set by(auto split: option.split simp add: RBT_set_def DList_set_def DList_Set.fold_def DList_Set.member_def List.member_def dest: equal.equal_eq[OF ID_ceq]) show ?DList_set_Set_Monad ?Set_Monad_DList_set by(auto split: option.split simp add: DList_set_def DList_Set.member_fold_insert) show ?DList_set_DList_set by(auto split: option.split simp add: DList_set_def DList_Set.member_union) qed(auto) lemma Set_inter_code [code]: fixes rbt1 rbt2 :: "'a :: ccompare set_rbt" and rbt :: "'b :: {ccompare, ceq} set_rbt" and dxs :: "'b set_dlist" and dxs1 dxs2 :: "'c :: ceq set_dlist" and xs1 xs2 :: "'c list" shows "Collect_set A'' \ J = Collect_set (\x. A'' x \ x \ J)" (is ?collect1) "J \ Collect_set A'' = Collect_set (\x. A'' x \ x \ J)" (is ?collect2) "Set_Monad xs'' \ I = Set_Monad (filter (\x. x \ I) xs'')" (is ?monad1) "I \ Set_Monad xs'' = Set_Monad (filter (\x. x \ I) xs'')" (is ?monad2) "DList_set dxs1 \ H = (case ID CEQ('c) of None \ Code.abort (STR ''inter DList_set1: ceq = None'') (\_. DList_set dxs1 \ H) | Some eq \ DList_set (DList_Set.filter (\x. x \ H) dxs1))" (is ?dlist1) "H \ DList_set dxs2 = (case ID CEQ('c) of None \ Code.abort (STR ''inter DList_set2: ceq = None'') (\_. H \ DList_set dxs2) | Some eq \ DList_set (DList_Set.filter (\x. x \ H) dxs2))" (is ?dlist2) "RBT_set rbt1 \ G = (case ID CCOMPARE('a) of None \ Code.abort (STR ''inter RBT_set1: ccompare = None'') (\_. RBT_set rbt1 \ G) | Some _ \ RBT_set (RBT_Set2.filter (\x. x \ G) rbt1))" (is ?rbt1) "G \ RBT_set rbt2 = (case ID CCOMPARE('a) of None \ Code.abort (STR ''inter RBT_set2: ccompare = None'') (\_. G \ RBT_set rbt2) | Some _ \ RBT_set (RBT_Set2.filter (\x. x \ G) rbt2))" (is ?rbt2) and Set_inter_Complement [set_complement_code]: "Complement B'' \ Complement B''' = Complement (B'' \ B''')" (is ?complement) and "Set_Monad xs \ RBT_set rbt1 = (case ID CCOMPARE('a) of None \ Code.abort (STR ''inter Set_Monad RBT_set: ccompare = None'') (\_. RBT_set rbt1 \ Set_Monad xs) | Some _ \ RBT_set (RBT_Set2.inter_list rbt1 xs))" (is ?monad_rbt) "Set_Monad xs' \ DList_set dxs2 = (case ID CEQ('c) of None \ Code.abort (STR ''inter Set_Monad DList_set: ceq = None'') (\_. Set_Monad xs' \ DList_set dxs2) | Some eq \ DList_set (DList_Set.filter (equal_base.list_member eq xs') dxs2))" (is ?monad_dlist) "Set_Monad xs1 \ Set_Monad xs2 = (case ID CEQ('c) of None \ Code.abort (STR ''inter Set_Monad Set_Monad: ceq = None'') (\_. Set_Monad xs1 \ Set_Monad xs2) | Some eq \ Set_Monad (filter (equal_base.list_member eq xs2) xs1))" (is ?monad) "DList_set dxs \ RBT_set rbt = (case ID CCOMPARE('b) of None \ Code.abort (STR ''inter DList_set RBT_set: ccompare = None'') (\_. DList_set dxs \ RBT_set rbt) | Some _ \ case ID CEQ('b) of None \ Code.abort (STR ''inter DList_set RBT_set: ceq = None'') (\_. DList_set dxs \ RBT_set rbt) | Some _ \ RBT_set (RBT_Set2.inter_list rbt (list_of_dlist dxs)))" (is ?dlist_rbt) "DList_set dxs1 \ DList_set dxs2 = (case ID CEQ('c) of None \ Code.abort (STR ''inter DList_set DList_set: ceq = None'') (\_. DList_set dxs1 \ DList_set dxs2) | Some _ \ DList_set (DList_Set.filter (DList_Set.member dxs2) dxs1))" (is ?dlist) "DList_set dxs1 \ Set_Monad xs' = (case ID CEQ('c) of None \ Code.abort (STR ''inter DList_set Set_Monad: ceq = None'') (\_. DList_set dxs1 \ Set_Monad xs') | Some eq \ DList_set (DList_Set.filter (equal_base.list_member eq xs') dxs1))" (is ?dlist_monad) "RBT_set rbt1 \ RBT_set rbt2 = (case ID CCOMPARE('a) of None \ Code.abort (STR ''inter RBT_set RBT_set: ccompare = None'') (\_. RBT_set rbt1 \ RBT_set rbt2) | Some _ \ RBT_set (RBT_Set2.inter rbt1 rbt2))" (is ?rbt_rbt) "RBT_set rbt \ DList_set dxs = (case ID CCOMPARE('b) of None \ Code.abort (STR ''inter RBT_set DList_set: ccompare = None'') (\_. RBT_set rbt \ DList_set dxs) | Some _ \ case ID CEQ('b) of None \ Code.abort (STR ''inter RBT_set DList_set: ceq = None'') (\_. RBT_set rbt \ DList_set dxs) | Some _ \ RBT_set (RBT_Set2.inter_list rbt (list_of_dlist dxs)))" (is ?rbt_dlist) "RBT_set rbt1 \ Set_Monad xs = (case ID CCOMPARE('a) of None \ Code.abort (STR ''inter RBT_set Set_Monad: ccompare = None'') (\_. RBT_set rbt1 \ Set_Monad xs) | Some _ \ RBT_set (RBT_Set2.inter_list rbt1 xs))" (is ?rbt_monad) proof - show ?rbt_rbt ?rbt1 ?rbt2 ?rbt_dlist ?rbt_monad ?dlist_rbt ?monad_rbt by(auto simp add: RBT_set_def DList_set_def DList_Set.member_def List.member_def dest: equal.equal_eq[OF ID_ceq] split: option.split) show ?dlist ?dlist1 ?dlist2 ?dlist_monad ?monad_dlist ?monad ?monad1 ?monad2 ?collect1 ?collect2 ?complement by(auto simp add: DList_set_def List.member_def dest!: Collection_Eq.ID_ceq split: option.splits) qed lemma Set_bind_code [code]: fixes dxs :: "'a :: ceq set_dlist" and rbt :: "'b :: ccompare set_rbt" shows "Set.bind (Set_Monad xs) f = fold ((\) \ f) xs (Set_Monad [])" (is ?Set_Monad) "Set.bind (DList_set dxs) f' = (case ID CEQ('a) of None \ Code.abort (STR ''bind DList_set: ceq = None'') (\_. Set.bind (DList_set dxs) f') | Some _ \ DList_Set.fold (union \ f') dxs {})" (is ?DList) "Set.bind (RBT_set rbt) f'' = (case ID CCOMPARE('b) of None \ Code.abort (STR ''bind RBT_set: ccompare = None'') (\_. Set.bind (RBT_set rbt) f'') | Some _ \ RBT_Set2.fold (union \ f'') rbt {})" (is ?RBT) proof - show ?Set_Monad by(simp add: set_bind_conv_fold) show ?DList by(auto simp add: DList_set_def DList_Set.member_def List.member_def List.member_def[abs_def] set_bind_conv_fold DList_Set.fold_def split: option.split dest: equal.equal_eq[OF ID_ceq] ID_ceq) show ?RBT by(clarsimp split: option.split simp add: RBT_set_def RBT_Set2.fold_conv_fold_keys RBT_Set2.member_conv_keys set_bind_conv_fold) qed lemma UNIV_code [code]: "UNIV = - {}" by(simp) lift_definition inf_sls :: "'a :: lattice semilattice_set" is "inf" by unfold_locales lemma Inf_fin_code [code]: "Inf_fin A = set_fold1 inf_sls A" by transfer(simp add: Inf_fin_def) lift_definition sup_sls :: "'a :: lattice semilattice_set" is "sup" by unfold_locales lemma Sup_fin_code [code]: "Sup_fin A = set_fold1 sup_sls A" by transfer(simp add: Sup_fin_def) lift_definition inf_cfi :: "('a :: lattice, 'a) comp_fun_idem" is "inf" by(rule comp_fun_idem_inf) lemma Inf_code: fixes A :: "'a :: complete_lattice set" shows "Inf A = (if finite A then set_fold_cfi inf_cfi top A else Code.abort (STR ''Inf: infinite'') (\_. Inf A))" by transfer(simp add: Inf_fold_inf) lift_definition sup_cfi :: "('a :: lattice, 'a) comp_fun_idem" is "sup" by(rule comp_fun_idem_sup) lemma Sup_code: fixes A :: "'a :: complete_lattice set" shows "Sup A = (if finite A then set_fold_cfi sup_cfi bot A else Code.abort (STR ''Sup: infinite'') (\_. Sup A))" by transfer(simp add: Sup_fold_sup) lemmas Inter_code [code] = Inf_code[where ?'a = "_ :: type set"] lemmas Union_code [code] = Sup_code[where ?'a = "_ :: type set"] lemmas Predicate_Inf_code [code] = Inf_code[where ?'a = "_ :: type Predicate.pred"] lemmas Predicate_Sup_code [code] = Sup_code[where ?'a = "_ :: type Predicate.pred"] lemmas Inf_fun_code [code] = Inf_code[where ?'a = "_ :: type \ _ :: complete_lattice"] lemmas Sup_fun_code [code] = Sup_code[where ?'a = "_ :: type \ _ :: complete_lattice"] lift_definition min_sls :: "'a :: linorder semilattice_set" is min by unfold_locales lemma Min_code [code]: "Min A = set_fold1 min_sls A" by transfer(simp add: Min_def) lift_definition max_sls :: "'a :: linorder semilattice_set" is max by unfold_locales lemma Max_code [code]: "Max A = set_fold1 max_sls A" by transfer(simp add: Max_def) text \ We do not implement @{term Ball}, @{term Bex}, and @{term sorted_list_of_set} for @{term Collect_set} using @{term cEnum}, because it should already have been converted to an explicit list of elements if that is possible. \ lemma Ball_code [code]: fixes rbt :: "'a :: ccompare set_rbt" and dxs :: "'b :: ceq set_dlist" shows "Ball (Set_Monad xs) P = list_all P xs" "Ball (DList_set dxs) P' = (case ID CEQ('b) of None \ Code.abort (STR ''Ball DList_set: ceq = None'') (\_. Ball (DList_set dxs) P') | Some _ \ DList_Set.dlist_all P' dxs)" "Ball (RBT_set rbt) P'' = (case ID CCOMPARE('a) of None \ Code.abort (STR ''Ball RBT_set: ccompare = None'') (\_. Ball (RBT_set rbt) P'') | Some _ \ RBT_Set2.all P'' rbt)" by(simp_all add: DList_set_def RBT_set_def list_all_iff dlist_all_conv_member RBT_Set2.all_conv_all_member split: option.splits) lemma Bex_code [code]: fixes rbt :: "'a :: ccompare set_rbt" and dxs :: "'b :: ceq set_dlist" shows "Bex (Set_Monad xs) P = list_ex P xs" "Bex (DList_set dxs) P' = (case ID CEQ('b) of None \ Code.abort (STR ''Bex DList_set: ceq = None'') (\_. Bex (DList_set dxs) P') | Some _ \ DList_Set.dlist_ex P' dxs)" "Bex (RBT_set rbt) P'' = (case ID CCOMPARE('a) of None \ Code.abort (STR ''Bex RBT_set: ccompare = None'') (\_. Bex (RBT_set rbt) P'') | Some _ \ RBT_Set2.ex P'' rbt)" by(simp_all add: DList_set_def RBT_set_def list_ex_iff dlist_ex_conv_member RBT_Set2.ex_conv_ex_member split: option.splits) lemma csorted_list_of_set_code [code]: fixes rbt :: "'a :: ccompare set_rbt" and dxs :: "'b :: {ccompare, ceq} set_dlist" and xs :: "'a :: ccompare list" shows "csorted_list_of_set (RBT_set rbt) = (case ID CCOMPARE('a) of None \ Code.abort (STR ''csorted_list_of_set RBT_set: ccompare = None'') (\_. csorted_list_of_set (RBT_set rbt)) | Some _ \ RBT_Set2.keys rbt)" "csorted_list_of_set (DList_set dxs) = (case ID CEQ('b) of None \ Code.abort (STR ''csorted_list_of_set DList_set: ceq = None'') (\_. csorted_list_of_set (DList_set dxs)) | Some _ \ case ID CCOMPARE('b) of None \ Code.abort (STR ''csorted_list_of_set DList_set: ccompare = None'') (\_. csorted_list_of_set (DList_set dxs)) | Some c \ ord.quicksort (lt_of_comp c) (list_of_dlist dxs))" "csorted_list_of_set (Set_Monad xs) = (case ID CCOMPARE('a) of None \ Code.abort (STR ''csorted_list_of_set Set_Monad: ccompare = None'') (\_. csorted_list_of_set (Set_Monad xs)) | Some c \ ord.remdups_sorted (lt_of_comp c) (ord.quicksort (lt_of_comp c) xs))" by(auto split: option.split simp add: RBT_set_def DList_set_def DList_Set.Collect_member member_conv_keys sorted_RBT_Set_keys linorder.sorted_list_of_set_sort_remdups[OF ID_ccompare] linorder.quicksort_conv_sort[OF ID_ccompare] distinct_remdups_id distinct_list_of_dlist linorder.remdups_sorted_conv_remdups[OF ID_ccompare] linorder.sorted_sort[OF ID_ccompare] linorder.sort_remdups[OF ID_ccompare] csorted_list_of_set_def) lemma cless_set_code [code]: fixes rbt rbt' :: "'a :: ccompare set_rbt" and rbt1 rbt2 :: "'b :: cproper_interval set_rbt" and A B :: "'a set" and A' B' :: "'b set" shows "cless_set A B \ (case ID CCOMPARE('a) of None \ Code.abort (STR ''cless_set: ccompare = None'') (\_. cless_set A B) | Some c \ if finite A \ finite B then ord.lexordp (\x y. lt_of_comp c y x) (csorted_list_of_set A) (csorted_list_of_set B) else Code.abort (STR ''cless_set: infinite set'') (\_. cless_set A B))" (is "?fin_fin") and cless_set_Complement2 [set_complement_code]: "cless_set A' (Complement B') \ (case ID CCOMPARE('b) of None \ Code.abort (STR ''cless_set Complement2: ccompare = None'') (\_. cless_set A' (Complement B')) | Some c \ if finite A' \ finite B' then finite (UNIV :: 'b set) \ proper_intrvl.set_less_aux_Compl (lt_of_comp c) cproper_interval None (csorted_list_of_set A') (csorted_list_of_set B') else Code.abort (STR ''cless_set Complement2: infinite set'') (\_. cless_set A' (Complement B')))" (is "?fin_Compl_fin") and cless_set_Complement1 [set_complement_code]: "cless_set (Complement A') B' \ (case ID CCOMPARE('b) of None \ Code.abort (STR ''cless_set Complement1: ccompare = None'') (\_. cless_set (Complement A') B') | Some c \ if finite A' \ finite B' then finite (UNIV :: 'b set) \ proper_intrvl.Compl_set_less_aux (lt_of_comp c) cproper_interval None (csorted_list_of_set A') (csorted_list_of_set B') else Code.abort (STR ''cless_set Complement1: infinite set'') (\_. cless_set (Complement A') B'))" (is "?Compl_fin_fin") and cless_set_Complement12 [set_complement_code]: "cless_set (Complement A) (Complement B) \ (case ID CCOMPARE('a) of None \ Code.abort (STR ''cless_set Complement Complement: ccompare = None'') (\_. cless_set (Complement A) (Complement B)) | Some _ \ cless B A)" (is ?Compl_Compl) and "cless_set (RBT_set rbt) (RBT_set rbt') \ (case ID CCOMPARE('a) of None \ Code.abort (STR ''cless_set RBT_set RBT_set: ccompare = None'') (\_. cless_set (RBT_set rbt) (RBT_set rbt')) | Some c \ ord.lexord_fusion (\x y. lt_of_comp c y x) rbt_keys_generator rbt_keys_generator (RBT_Set2.init rbt) (RBT_Set2.init rbt'))" (is ?rbt_rbt) and cless_set_rbt_Complement2 [set_complement_code]: "cless_set (RBT_set rbt1) (Complement (RBT_set rbt2)) \ (case ID CCOMPARE('b) of None \ Code.abort (STR ''cless_set RBT_set (Complement RBT_set): ccompare = None'') (\_. cless_set (RBT_set rbt1) (Complement (RBT_set rbt2))) | Some c \ finite (UNIV :: 'b set) \ proper_intrvl.set_less_aux_Compl_fusion (lt_of_comp c) cproper_interval rbt_keys_generator rbt_keys_generator None (RBT_Set2.init rbt1) (RBT_Set2.init rbt2))" (is ?rbt_Compl) and cless_set_rbt_Complement1 [set_complement_code]: "cless_set (Complement (RBT_set rbt1)) (RBT_set rbt2) \ (case ID CCOMPARE('b) of None \ Code.abort (STR ''cless_set (Complement RBT_set) RBT_set: ccompare = None'') (\_. cless_set (Complement (RBT_set rbt1)) (RBT_set rbt2)) | Some c \ finite (UNIV :: 'b set) \ proper_intrvl.Compl_set_less_aux_fusion (lt_of_comp c) cproper_interval rbt_keys_generator rbt_keys_generator None (RBT_Set2.init rbt1) (RBT_Set2.init rbt2))" (is ?Compl_rbt) proof - note [split] = option.split csorted_list_of_set_split and [simp] = le_of_comp_of_ords_linorder[OF ID_ccompare] lt_of_comp_of_ords finite_subset[OF subset_UNIV] ccompare_set_def ID_Some ord.lexord_fusion_def proper_intrvl.Compl_set_less_aux_fusion_def proper_intrvl.set_less_aux_Compl_fusion_def unfoldr_rbt_keys_generator RBT_set_def sorted_RBT_Set_keys member_conv_keys linorder.set_less_finite_iff[OF ID_ccompare] linorder.set_less_aux_code[OF ID_ccompare, symmetric] linorder.Compl_set_less_Compl[OF ID_ccompare] linorder.infinite_set_less_Complement[OF ID_ccompare] linorder.infinite_Complement_set_less[OF ID_ccompare] linorder_proper_interval.set_less_aux_Compl2_conv_set_less_aux_Compl[OF ID_ccompare_interval, symmetric] linorder_proper_interval.Compl1_set_less_aux_conv_Compl_set_less_aux[OF ID_ccompare_interval, symmetric] show ?Compl_Compl by simp show ?rbt_rbt by auto show ?rbt_Compl by(cases "finite (UNIV :: 'b set)") auto show ?Compl_rbt by(cases "finite (UNIV :: 'b set)") auto show ?fin_fin by auto show ?fin_Compl_fin by(cases "finite (UNIV :: 'b set)", auto) show ?Compl_fin_fin by(cases "finite (UNIV :: 'b set)") auto qed lemma le_of_comp_set_less_eq: "le_of_comp (comp_of_ords (ord.set_less_eq le) (ord.set_less le)) = ord.set_less_eq le" by (rule le_of_comp_of_ords_gen, simp add: ord.set_less_def) lemma cless_eq_set_code [code]: fixes rbt rbt' :: "'a :: ccompare set_rbt" and rbt1 rbt2 :: "'b :: cproper_interval set_rbt" and A B :: "'a set" and A' B' :: "'b set" shows "cless_eq_set A B \ (case ID CCOMPARE('a) of None \ Code.abort (STR ''cless_eq_set: ccompare = None'') (\_. cless_eq_set A B) | Some c \ if finite A \ finite B then ord.lexordp_eq (\x y. lt_of_comp c y x) (csorted_list_of_set A) (csorted_list_of_set B) else Code.abort (STR ''cless_eq_set: infinite set'') (\_. cless_eq_set A B))" (is "?fin_fin") and cless_eq_set_Complement2 [set_complement_code]: "cless_eq_set A' (Complement B') \ (case ID CCOMPARE('b) of None \ Code.abort (STR ''cless_eq_set Complement2: ccompare = None'') (\_. cless_eq_set A' (Complement B')) | Some c \ if finite A' \ finite B' then finite (UNIV :: 'b set) \ proper_intrvl.set_less_eq_aux_Compl (lt_of_comp c) cproper_interval None (csorted_list_of_set A') (csorted_list_of_set B') else Code.abort (STR ''cless_eq_set Complement2: infinite set'') (\_. cless_eq_set A' (Complement B')))" (is "?fin_Compl_fin") and cless_eq_set_Complement1 [set_complement_code]: "cless_eq_set (Complement A') B' \ (case ID CCOMPARE('b) of None \ Code.abort (STR ''cless_eq_set Complement1: ccompare = None'') (\_. cless_eq_set (Complement A') B') | Some c \ if finite A' \ finite B' then finite (UNIV :: 'b set) \ proper_intrvl.Compl_set_less_eq_aux (lt_of_comp c) cproper_interval None (csorted_list_of_set A') (csorted_list_of_set B') else Code.abort (STR ''cless_eq_set Complement1: infinite set'') (\_. cless_eq_set (Complement A') B'))" (is "?Compl_fin_fin") and cless_eq_set_Complement12 [set_complement_code]: "cless_eq_set (Complement A) (Complement B) \ (case ID CCOMPARE('a) of None \ Code.abort (STR ''cless_eq_set Complement Complement: ccompare = None'') (\_. cless_eq (Complement A) (Complement B)) | Some c \ cless_eq_set B A)" (is ?Compl_Compl) "cless_eq_set (RBT_set rbt) (RBT_set rbt') \ (case ID CCOMPARE('a) of None \ Code.abort (STR ''cless_eq_set RBT_set RBT_set: ccompare = None'') (\_. cless_eq_set (RBT_set rbt) (RBT_set rbt')) | Some c \ ord.lexord_eq_fusion (\x y. lt_of_comp c y x) rbt_keys_generator rbt_keys_generator (RBT_Set2.init rbt) (RBT_Set2.init rbt'))" (is ?rbt_rbt) and cless_eq_set_rbt_Complement2 [set_complement_code]: "cless_eq_set (RBT_set rbt1) (Complement (RBT_set rbt2)) \ (case ID CCOMPARE('b) of None \ Code.abort (STR ''cless_eq_set RBT_set (Complement RBT_set): ccompare = None'') (\_. cless_eq_set (RBT_set rbt1) (Complement (RBT_set rbt2))) | Some c \ finite (UNIV :: 'b set) \ proper_intrvl.set_less_eq_aux_Compl_fusion (lt_of_comp c) cproper_interval rbt_keys_generator rbt_keys_generator None (RBT_Set2.init rbt1) (RBT_Set2.init rbt2))" (is ?rbt_Compl) and cless_eq_set_rbt_Complement1 [set_complement_code]: "cless_eq_set (Complement (RBT_set rbt1)) (RBT_set rbt2) \ (case ID CCOMPARE('b) of None \ Code.abort (STR ''cless_eq_set (Complement RBT_set) RBT_set: ccompare = None'') (\_. cless_eq_set (Complement (RBT_set rbt1)) (RBT_set rbt2)) | Some c \ finite (UNIV :: 'b set) \ proper_intrvl.Compl_set_less_eq_aux_fusion (lt_of_comp c) cproper_interval rbt_keys_generator rbt_keys_generator None (RBT_Set2.init rbt1) (RBT_Set2.init rbt2))" (is ?Compl_rbt) proof - note [split] = option.split csorted_list_of_set_split and [simp] = le_of_comp_set_less_eq finite_subset[OF subset_UNIV] ccompare_set_def ID_Some ord.lexord_eq_fusion_def proper_intrvl.Compl_set_less_eq_aux_fusion_def proper_intrvl.set_less_eq_aux_Compl_fusion_def unfoldr_rbt_keys_generator RBT_set_def sorted_RBT_Set_keys member_conv_keys linorder.set_less_eq_finite_iff[OF ID_ccompare] linorder.set_less_eq_aux_code[OF ID_ccompare, symmetric] linorder.Compl_set_less_eq_Compl[OF ID_ccompare] linorder.infinite_set_less_eq_Complement[OF ID_ccompare] linorder.infinite_Complement_set_less_eq[OF ID_ccompare] linorder_proper_interval.set_less_eq_aux_Compl2_conv_set_less_eq_aux_Compl[OF ID_ccompare_interval, symmetric] linorder_proper_interval.Compl1_set_less_eq_aux_conv_Compl_set_less_eq_aux[OF ID_ccompare_interval, symmetric] show ?Compl_Compl by simp show ?rbt_rbt by auto show ?rbt_Compl by(cases "finite (UNIV :: 'b set)") auto show ?Compl_rbt by(cases "finite (UNIV :: 'b set)") auto show ?fin_fin by auto show ?fin_Compl_fin by (cases "finite (UNIV :: 'b set)", auto) show ?Compl_fin_fin by(cases "finite (UNIV :: 'b set)") auto qed lemma cproper_interval_set_Some_Some_code [code]: fixes rbt1 rbt2 :: "'a :: cproper_interval set_rbt" and A B :: "'a set" shows "cproper_interval (Some A) (Some B) \ (case ID CCOMPARE('a) of None \ Code.abort (STR ''cproper_interval: ccompare = None'') (\_. cproper_interval (Some A) (Some B)) | Some c \ finite (UNIV :: 'a set) \ proper_intrvl.proper_interval_set_aux (lt_of_comp c) cproper_interval (csorted_list_of_set A) (csorted_list_of_set B))" (is ?fin_fin) and cproper_interval_set_Some_Some_Complement [set_complement_code]: "cproper_interval (Some A) (Some (Complement B)) \ (case ID CCOMPARE('a) of None \ Code.abort (STR ''cproper_interval Complement2: ccompare = None'') (\_. cproper_interval (Some A) (Some (Complement B))) | Some c \ finite (UNIV :: 'a set) \ proper_intrvl.proper_interval_set_Compl_aux (lt_of_comp c) cproper_interval None 0 (csorted_list_of_set A) (csorted_list_of_set B))" (is ?fin_Compl_fin) and cproper_interval_set_Some_Complement_Some [set_complement_code]: "cproper_interval (Some (Complement A)) (Some B) \ (case ID CCOMPARE('a) of None \ Code.abort (STR ''cproper_interval Complement1: ccompare = None'') (\_. cproper_interval (Some (Complement A)) (Some B)) | Some c \ finite (UNIV :: 'a set) \ proper_intrvl.proper_interval_Compl_set_aux (lt_of_comp c) cproper_interval None (csorted_list_of_set A) (csorted_list_of_set B))" (is ?Compl_fin_fin) and cproper_interval_set_Some_Complement_Some_Complement [set_complement_code]: "cproper_interval (Some (Complement A)) (Some (Complement B)) \ (case ID CCOMPARE('a) of None \ Code.abort (STR ''cproper_interval Complement Complement: ccompare = None'') (\_. cproper_interval (Some (Complement A)) (Some (Complement B))) | Some _ \ cproper_interval (Some B) (Some A))" (is ?Compl_Compl) "cproper_interval (Some (RBT_set rbt1)) (Some (RBT_set rbt2)) \ (case ID CCOMPARE('a) of None \ Code.abort (STR ''cproper_interval RBT_set RBT_set: ccompare = None'') (\_. cproper_interval (Some (RBT_set rbt1)) (Some (RBT_set rbt2))) | Some c \ finite (UNIV :: 'a set) \ proper_intrvl.proper_interval_set_aux_fusion (lt_of_comp c) cproper_interval rbt_keys_generator rbt_keys_generator (RBT_Set2.init rbt1) (RBT_Set2.init rbt2))" (is ?rbt_rbt) and cproper_interval_set_Some_rbt_Some_Complement [set_complement_code]: "cproper_interval (Some (RBT_set rbt1)) (Some (Complement (RBT_set rbt2))) \ (case ID CCOMPARE('a) of None \ Code.abort (STR ''cproper_interval RBT_set (Complement RBT_set): ccompare = None'') (\_. cproper_interval (Some (RBT_set rbt1)) (Some (Complement (RBT_set rbt2)))) | Some c \ finite (UNIV :: 'a set) \ proper_intrvl.proper_interval_set_Compl_aux_fusion (lt_of_comp c) cproper_interval rbt_keys_generator rbt_keys_generator None 0 (RBT_Set2.init rbt1) (RBT_Set2.init rbt2))" (is ?rbt_Compl_rbt) and cproper_interval_set_Some_Complement_Some_rbt [set_complement_code]: "cproper_interval (Some (Complement (RBT_set rbt1))) (Some (RBT_set rbt2)) \ (case ID CCOMPARE('a) of None \ Code.abort (STR ''cproper_interval (Complement RBT_set) RBT_set: ccompare = None'') (\_. cproper_interval (Some (Complement (RBT_set rbt1))) (Some (RBT_set rbt2))) | Some c \ finite (UNIV :: 'a set) \ proper_intrvl.proper_interval_Compl_set_aux_fusion (lt_of_comp c) cproper_interval rbt_keys_generator rbt_keys_generator None (RBT_Set2.init rbt1) (RBT_Set2.init rbt2))" (is ?Compl_rbt_rbt) proof - note [split] = option.split csorted_list_of_set_split and [simp] = lt_of_comp_of_ords finite_subset[OF subset_UNIV] ccompare_set_def ID_Some linorder.set_less_finite_iff[OF ID_ccompare] RBT_set_def sorted_RBT_Set_keys member_conv_keys linorder.distinct_entries[OF ID_ccompare] unfoldr_rbt_keys_generator proper_intrvl.proper_interval_set_aux_fusion_def proper_intrvl.proper_interval_set_Compl_aux_fusion_def proper_intrvl.proper_interval_Compl_set_aux_fusion_def linorder_proper_interval.proper_interval_set_aux[OF ID_ccompare_interval] linorder_proper_interval.proper_interval_set_Compl_aux[OF ID_ccompare_interval] linorder_proper_interval.proper_interval_Compl_set_aux[OF ID_ccompare_interval] and [cong] = conj_cong show ?Compl_Compl by(clarsimp simp add: Complement_cproper_interval_set_Complement simp del: cproper_interval_set_Some_Some) show ?rbt_rbt ?rbt_Compl_rbt ?Compl_rbt_rbt by auto show ?fin_fin ?fin_Compl_fin ?Compl_fin_fin by auto qed context ord begin fun sorted_list_subset :: "('a \ 'a \ bool) \ 'a list \ 'a list \ bool" where "sorted_list_subset eq [] ys = True" | "sorted_list_subset eq (x # xs) [] = False" | "sorted_list_subset eq (x # xs) (y # ys) \ (if eq x y then sorted_list_subset eq xs ys else x > y \ sorted_list_subset eq (x # xs) ys)" end context linorder begin lemma sorted_list_subset_correct: "\ sorted xs; distinct xs; sorted ys; distinct ys \ \ sorted_list_subset (=) xs ys \ set xs \ set ys" apply(induct "(=) :: 'a \ 'a \ bool" xs ys rule: sorted_list_subset.induct) apply(auto 6 2) using order_antisym apply auto done end context ord begin definition sorted_list_subset_fusion :: "('a \ 'a \ bool) \ ('a, 's1) generator \ ('a, 's2) generator \ 's1 \ 's2 \ bool" where "sorted_list_subset_fusion eq g1 g2 s1 s2 = sorted_list_subset eq (list.unfoldr g1 s1) (list.unfoldr g2 s2)" lemma sorted_list_subset_fusion_code: "sorted_list_subset_fusion eq g1 g2 s1 s2 = (if list.has_next g1 s1 then let (x, s1') = list.next g1 s1 in list.has_next g2 s2 \ ( let (y, s2') = list.next g2 s2 in if eq x y then sorted_list_subset_fusion eq g1 g2 s1' s2' else y < x \ sorted_list_subset_fusion eq g1 g2 s1 s2') else True)" unfolding sorted_list_subset_fusion_def by(subst (1 2 5) list.unfoldr.simps)(simp add: split_beta Let_def) end lemmas [code] = ord.sorted_list_subset_fusion_code text \ Define a new constant for the subset operation because @{theory "HOL-Library.Cardinality"} introduces @{const "Cardinality.subset'"} and rewrites @{const "subset"} to @{const "Cardinality.subset'"} based on the sort of the element type. \ definition subset_eq :: "'a set \ 'a set \ bool" where [simp, code del]: "subset_eq = (\)" lemma subseteq_code [code]: "(\) = subset_eq" by simp lemma subset'_code [code]: "Cardinality.subset' = subset_eq" by simp lemma subset_eq_code [folded subset_eq_def, code]: fixes A1 A2 :: "'a set" and rbt :: "'b :: ccompare set_rbt" and rbt1 rbt2 :: "'d :: {ccompare, ceq} set_rbt" and dxs :: "'c :: ceq set_dlist" and xs :: "'c list" shows "RBT_set rbt \ B \ (case ID CCOMPARE('b) of None \ Code.abort (STR ''subset RBT_set1: ccompare = None'') (\_. RBT_set rbt \ B) | Some _ \ list_all_fusion rbt_keys_generator (\x. x \ B) (RBT_Set2.init rbt))" (is ?rbt) "DList_set dxs \ C \ (case ID CEQ('c) of None \ Code.abort (STR ''subset DList_set1: ceq = None'') (\_. DList_set dxs \ C) | Some _ \ DList_Set.dlist_all (\x. x \ C) dxs)" (is ?dlist) "Set_Monad xs \ C \ list_all (\x. x \ C) xs" (is ?Set_Monad) and Collect_subset_eq_Complement [folded subset_eq_def, set_complement_code]: "Collect_set P \ Complement A \ A \ {x. \ P x}" (is ?Collect_set_Compl) and Complement_subset_eq_Complement [folded subset_eq_def, set_complement_code]: "Complement A1 \ Complement A2 \ A2 \ A1" (is ?Compl) and "RBT_set rbt1 \ RBT_set rbt2 \ (case ID CCOMPARE('d) of None \ Code.abort (STR ''subset RBT_set RBT_set: ccompare = None'') (\_. RBT_set rbt1 \ RBT_set rbt2) | Some c \ (case ID CEQ('d) of None \ ord.sorted_list_subset_fusion (lt_of_comp c) (\ x y. c x y = Eq) rbt_keys_generator rbt_keys_generator (RBT_Set2.init rbt1) (RBT_Set2.init rbt2) | Some eq \ ord.sorted_list_subset_fusion (lt_of_comp c) eq rbt_keys_generator rbt_keys_generator (RBT_Set2.init rbt1) (RBT_Set2.init rbt2)))" (is ?rbt_rbt) proof - show ?rbt_rbt by (auto simp add: comparator.eq[OF ID_ccompare'] RBT_set_def member_conv_keys unfoldr_rbt_keys_generator ord.sorted_list_subset_fusion_def linorder.sorted_list_subset_correct[OF ID_ccompare] sorted_RBT_Set_keys split: option.split dest!: ID_ceq[THEN equal.equal_eq] del: iffI) show ?rbt by(auto simp add: RBT_set_def member_conv_keys list_all_fusion_def unfoldr_rbt_keys_generator keys.rep_eq list_all_iff split: option.split) show ?dlist by(auto simp add: DList_set_def dlist_all_conv_member split: option.split) show ?Set_Monad by(auto simp add: list_all_iff split: option.split) show ?Collect_set_Compl ?Compl by auto qed hide_const (open) subset_eq hide_fact (open) subset_eq_def lemma eq_set_code [code]: "Cardinality.eq_set = set_eq" by(simp add: set_eq_def) lemma set_eq_code [code]: fixes rbt1 rbt2 :: "'b :: {ccompare, ceq} set_rbt" shows "set_eq A B \ A \ B \ B \ A" and set_eq_Complement_Complement [set_complement_code]: "set_eq (Complement A) (Complement B) = set_eq A B" and "set_eq (RBT_set rbt1) (RBT_set rbt2) = (case ID CCOMPARE('b) of None \ Code.abort (STR ''set_eq RBT_set RBT_set: ccompare = None'') (\_. set_eq (RBT_set rbt1) (RBT_set rbt2)) | Some c \ (case ID CEQ('b) of None \ list_all2_fusion (\ x y. c x y = Eq) rbt_keys_generator rbt_keys_generator (RBT_Set2.init rbt1) (RBT_Set2.init rbt2) | Some eq \ list_all2_fusion eq rbt_keys_generator rbt_keys_generator (RBT_Set2.init rbt1) (RBT_Set2.init rbt2)))" (is ?rbt_rbt) proof - show ?rbt_rbt by (auto 4 3 split: option.split simp add: comparator.eq[OF ID_ccompare'] sorted_RBT_Set_keys list_all2_fusion_def unfoldr_rbt_keys_generator RBT_set_conv_keys set_eq_def list.rel_eq dest!: ID_ceq[THEN equal.equal_eq] intro: linorder.sorted_distinct_set_unique[OF ID_ccompare]) qed(auto simp add: set_eq_def) lemma Set_project_code [code]: "Set.filter P A = A \ Collect_set P" by(auto simp add: Set.filter_def) lemma Set_image_code [code]: fixes dxs :: "'a :: ceq set_dlist" and rbt :: "'b :: ccompare set_rbt" shows "image f (Set_Monad xs) = Set_Monad (map f xs)" "image f (Collect_set A) = Code.abort (STR ''image Collect_set'') (\_. image f (Collect_set A))" and image_Complement_Complement [set_complement_code]: "image f (Complement (Complement B)) = image f B" and "image g (DList_set dxs) = (case ID CEQ('a) of None \ Code.abort (STR ''image DList_set: ceq = None'') (\_. image g (DList_set dxs)) | Some _ \ DList_Set.fold (insert \ g) dxs {})" (is ?dlist) "image h (RBT_set rbt) = (case ID CCOMPARE('b) of None \ Code.abort (STR ''image RBT_set: ccompare = None'') (\_. image h (RBT_set rbt)) | Some _ \ RBT_Set2.fold (insert \ h) rbt {})" (is ?rbt) proof - { fix xs have "fold (insert \ g) xs {} = g ` set xs" by(induct xs rule: rev_induct) simp_all } thus ?dlist by(simp add: DList_set_def DList_Set.fold_def DList_Set.Collect_member split: option.split) { fix xs have "fold (insert \ h) xs {} = h ` set xs" by(induct xs rule: rev_induct) simp_all } thus ?rbt by(auto simp add: RBT_set_def fold_conv_fold_keys member_conv_keys split: option.split) qed simp_all lemma the_elem_code [code]: fixes dxs :: "'a :: ceq set_dlist" and rbt :: "'b :: ccompare set_rbt" shows "the_elem (Set_Monad [x]) = x" "the_elem (DList_set dxs) = (case ID CEQ('a) of None \ Code.abort (STR ''the_elem DList_set: ceq = None'') (\_. the_elem (DList_set dxs)) | Some _ \ case list_of_dlist dxs of [x] \ x | _ \ Code.abort (STR ''the_elem DList_set: not unique'') (\_. the_elem (DList_set dxs)))" "the_elem (RBT_set rbt) = (case ID CCOMPARE('b) of None \ Code.abort (STR ''the_elem RBT_set: ccompare = None'') (\_. the_elem (RBT_set rbt)) | Some _ \ case RBT_Mapping2.impl_of rbt of RBT_Impl.Branch _ RBT_Impl.Empty x _ RBT_Impl.Empty \ x | _ \ Code.abort (STR ''the_elem RBT_set: not unique'') (\_. the_elem (RBT_set rbt)))" by(auto simp add: RBT_set_def DList_set_def DList_Set.Collect_member the_elem_def member_conv_keys split: option.split list.split rbt.split)(simp add: RBT_Set2.keys_def) lemma Pow_set_conv_fold: "Pow (set xs \ A) = fold (\x A. A \ insert x ` A) xs (Pow A)" by(induct xs rule: rev_induct)(auto simp add: Pow_insert) lemma Pow_code [code]: fixes dxs :: "'a :: ceq set_dlist" and rbt :: "'b :: ccompare set_rbt" shows "Pow A = Collect_set (\B. B \ A)" "Pow (Set_Monad xs) = fold (\x A. A \ insert x ` A) xs {{}}" "Pow (DList_set dxs) = (case ID CEQ('a) of None \ Code.abort (STR ''Pow DList_set: ceq = None'') (\_. Pow (DList_set dxs)) | Some _ \ DList_Set.fold (\x A. A \ insert x ` A) dxs {{}})" "Pow (RBT_set rbt) = (case ID CCOMPARE('b) of None \ Code.abort (STR ''Pow RBT_set: ccompare = None'') (\_. Pow (RBT_set rbt)) | Some _ \ RBT_Set2.fold (\x A. A \ insert x ` A) rbt {{}})" by(auto simp add: DList_set_def DList_Set.Collect_member DList_Set.fold_def RBT_set_def fold_conv_fold_keys member_conv_keys Pow_set_conv_fold[where A="{}", simplified] split: option.split) lemma fold_singleton: "Finite_Set.fold f x {y} = f y x" by(fastforce simp add: Finite_Set.fold_def intro: fold_graph.intros elim: fold_graph.cases) lift_definition sum_cfc :: "('a \ 'b :: comm_monoid_add) \ ('a, 'b) comp_fun_commute" is "\f :: 'a \ 'b. plus \ f" by(unfold_locales)(simp add: fun_eq_iff add.left_commute) lemma sum_code [code]: "sum f A = (if finite A then set_fold_cfc (sum_cfc f) 0 A else 0)" by transfer(simp add: sum.eq_fold) lemma product_code [code]: fixes dxs :: "'a :: ceq set_dlist" and dys :: "'b :: ceq set_dlist" and rbt1 :: "'c :: ccompare set_rbt" and rbt2 :: "'d :: ccompare set_rbt" shows "Product_Type.product A B = Collect_set (\(x, y). x \ A \ y \ B)" "Product_Type.product (Set_Monad xs) (Set_Monad ys) = Set_Monad (fold (\x. fold (\y rest. (x, y) # rest) ys) xs [])" (is ?Set_Monad) "Product_Type.product (DList_set dxs) B1 = (case ID CEQ('a) of None \ Code.abort (STR ''product DList_set1: ceq = None'') (\_. Product_Type.product (DList_set dxs) B1) | Some _ \ DList_Set.fold (\x rest. Pair x ` B1 \ rest) dxs {})" (is "?dlist1") "Product_Type.product A1 (DList_set dys) = (case ID CEQ('b) of None \ Code.abort (STR ''product DList_set2: ceq = None'') (\_. Product_Type.product A1 (DList_set dys)) | Some _ \ DList_Set.fold (\y rest. (\x. (x, y)) ` A1 \ rest) dys {})" (is "?dlist2") "Product_Type.product (DList_set dxs) (DList_set dys) = (case ID CEQ('a) of None \ Code.abort (STR ''product DList_set DList_set: ceq1 = None'') (\_. Product_Type.product (DList_set dxs) (DList_set dys)) | Some _ \ case ID CEQ('b) of None \ Code.abort (STR ''product DList_set DList_set: ceq2 = None'') (\_. Product_Type.product (DList_set dxs) (DList_set dys)) | Some _ \ DList_set (DList_Set.product dxs dys))" "Product_Type.product (RBT_set rbt1) B2 = (case ID CCOMPARE('c) of None \ Code.abort (STR ''product RBT_set: ccompare1 = None'') (\_. Product_Type.product (RBT_set rbt1) B2) | Some _ \ RBT_Set2.fold (\x rest. Pair x ` B2 \ rest) rbt1 {})" (is "?rbt1") "Product_Type.product A2 (RBT_set rbt2) = (case ID CCOMPARE('d) of None \ Code.abort (STR ''product RBT_set: ccompare2 = None'') (\_. Product_Type.product A2 (RBT_set rbt2)) | Some _ \ RBT_Set2.fold (\y rest. (\x. (x, y)) ` A2 \ rest) rbt2 {})" (is "?rbt2") "Product_Type.product (RBT_set rbt1) (RBT_set rbt2) = (case ID CCOMPARE('c) of None \ Code.abort (STR ''product RBT_set RBT_set: ccompare1 = None'') (\_. Product_Type.product (RBT_set rbt1) (RBT_set rbt2)) | Some _ \ case ID CCOMPARE('d) of None \ Code.abort (STR ''product RBT_set RBT_set: ccompare2 = None'') (\_. Product_Type.product (RBT_set rbt1) (RBT_set rbt2)) | Some _ \ RBT_set (RBT_Set2.product rbt1 rbt2))" proof - have [simp]: "\a zs. fold (\y. (#) (a, y)) ys zs = rev (map (Pair a) ys) @ zs" by(induct ys) simp_all have [simp]: "\zs. fold (\x. fold (\y rest. (x, y) # rest) ys) xs zs = rev (concat (map (\x. map (Pair x) ys) xs)) @ zs" by(induct xs) simp_all show ?Set_Monad by(auto simp add: Product_Type.product_def) { fix xs :: "'a list" have "fold (\x. (\) (Pair x ` B1)) xs {} = set xs \ B1" by(induct xs rule: rev_induct) auto } thus ?dlist1 by(simp add: Product_Type.product_def DList_set_def DList_Set.fold.rep_eq DList_Set.Collect_member split: option.split) { fix ys :: "'b list" have "fold (\y. (\) ((\x. (x, y)) ` A1)) ys {} = A1 \ set ys" by(induct ys rule: rev_induct) auto } thus ?dlist2 by(simp add: Product_Type.product_def DList_set_def DList_Set.fold.rep_eq DList_Set.Collect_member split: option.split) { fix xs :: "'c list" have "fold (\x. (\) (Pair x ` B2)) xs {} = set xs \ B2" by(induct xs rule: rev_induct) auto } thus ?rbt1 by(simp add: Product_Type.product_def RBT_set_def RBT_Set2.member_product RBT_Set2.member_conv_keys fold_conv_fold_keys split: option.split) { fix ys :: "'d list" have "fold (\y. (\) ((\x. (x, y)) ` A2)) ys {} = A2 \ set ys" by(induct ys rule: rev_induct) auto } thus ?rbt2 by(simp add: Product_Type.product_def RBT_set_def RBT_Set2.member_product RBT_Set2.member_conv_keys fold_conv_fold_keys split: option.split) qed(auto simp add: RBT_set_def DList_set_def Product_Type.product_def DList_Set.product_member RBT_Set2.member_product split: option.split) lemma Id_on_code [code]: fixes A :: "'a :: ceq set" and dxs :: "'a set_dlist" and P :: "'a \ bool" and rbt :: "'b :: ccompare set_rbt" shows "Id_on B = (\x. (x, x)) ` B" and Id_on_Complement [set_complement_code]: "Id_on (Complement A) = (case ID CEQ('a) of None \ Code.abort (STR ''Id_on Complement: ceq = None'') (\_. Id_on (Complement A)) | Some eq \ Collect_set (\(x, y). eq x y \ x \ A))" and "Id_on (Collect_set P) = (case ID CEQ('a) of None \ Code.abort (STR ''Id_on Collect_set: ceq = None'') (\_. Id_on (Collect_set P)) | Some eq \ Collect_set (\(x, y). eq x y \ P x))" "Id_on (DList_set dxs) = (case ID CEQ('a) of None \ Code.abort (STR ''Id_on DList_set: ceq = None'') (\_. Id_on (DList_set dxs)) | Some _ \ DList_set (DList_Set.Id_on dxs))" "Id_on (RBT_set rbt) = (case ID CCOMPARE('b) of None \ Code.abort (STR ''Id_on RBT_set: ccompare = None'') (\_. Id_on (RBT_set rbt)) | Some _ \ RBT_set (RBT_Set2.Id_on rbt))" by(auto simp add: DList_set_def RBT_set_def DList_Set.member_Id_on RBT_Set2.member_Id_on dest: equal.equal_eq[OF ID_ceq] split: option.split) lemma Image_code [code]: fixes dxs :: "('a :: ceq \ 'b :: ceq) set_dlist" and rbt :: "('c :: ccompare \ 'd :: ccompare) set_rbt" shows "X `` Y = snd ` Set.filter (\(x, y). x \ Y) X" (is ?generic) "Set_Monad rxs `` A = Set_Monad (fold (\(x, y) rest. if x \ A then y # rest else rest) rxs [])" (is ?Set_Monad) "DList_set dxs `` B = (case ID CEQ('a) of None \ Code.abort (STR ''Image DList_set: ceq1 = None'') (\_. DList_set dxs `` B) | Some _ \ case ID CEQ('b) of None \ Code.abort (STR ''Image DList_set: ceq2 = None'') (\_. DList_set dxs `` B) | Some _ \ DList_Set.fold (\(x, y) acc. if x \ B then insert y acc else acc) dxs {})" (is ?DList_set) "RBT_set rbt `` C = (case ID CCOMPARE('c) of None \ Code.abort (STR ''Image RBT_set: ccompare1 = None'') (\_. RBT_set rbt `` C) | Some _ \ case ID CCOMPARE('d) of None \ Code.abort (STR ''Image RBT_set: ccompare2 = None'') (\_. RBT_set rbt `` C) | Some _ \ RBT_Set2.fold (\(x, y) acc. if x \ C then insert y acc else acc) rbt {})" (is ?RBT_set) proof - show ?generic by(auto intro: rev_image_eqI) have "set (fold (\(x, y) rest. if x \ A then y # rest else rest) rxs []) = set rxs `` A" by(induct rxs rule: rev_induct)(auto split: if_split_asm) thus ?Set_Monad by(auto) { fix dxs :: "('a \ 'b) list" have "fold (\(x, y) acc. if x \ B then insert y acc else acc) dxs {} = set dxs `` B" by(induct dxs rule: rev_induct)(auto split: if_split_asm) } thus ?DList_set by(clarsimp simp add: DList_set_def Collect_member ceq_prod_def ID_Some DList_Set.fold.rep_eq split: option.split) { fix rbt :: "(('c \ 'd) \ unit) list" have "fold (\(a, _). case a of (x, y) \ \acc. if x \ C then insert y acc else acc) rbt {} = (fst ` set rbt) `` C" by(induct rbt rule: rev_induct)(auto simp add: split_beta split: if_split_asm) } thus ?RBT_set by(clarsimp simp add: RBT_set_def ccompare_prod_def ID_Some RBT_Set2.fold.rep_eq member_conv_keys RBT_Set2.keys.rep_eq RBT_Impl.fold_def RBT_Impl.keys_def split: option.split) qed lemma insert_relcomp: "insert (a, b) A O B = A O B \ {a} \ {c. (b, c) \ B}" by auto lemma trancl_code [code]: "trancl A = (if finite A then ntrancl (card A - 1) A else Code.abort (STR ''trancl: infinite set'') (\_. trancl A))" by (simp add: finite_trancl_ntranl) lemma set_relcomp_set: "set xs O set ys = fold (\(x, y). fold (\(y', z) A. if y = y' then insert (x, z) A else A) ys) xs {}" proof(induct xs rule: rev_induct) case Nil show ?case by simp next case (snoc x xs) note [[show_types]] { fix a :: 'a and b :: 'c and X :: "('a \ 'b) set" have "fold (\(y', z) A. if b = y' then insert (a, z) A else A) ys X = X \ {a} \ {c. (b, c) \ set ys}" by(induct ys arbitrary: X rule: rev_induct)(auto split: if_split_asm) } thus ?case using snoc by(cases x)(simp add: insert_relcomp) qed lemma If_not: "(if \ a then b else c) = (if a then c else b)" by auto lemma relcomp_code [code]: fixes rbt1 :: "('a :: ccompare \ 'b :: ccompare) set_rbt" and rbt2 :: "('b \ 'c :: ccompare) set_rbt" and rbt3 :: "('a \ 'd :: {ccompare, ceq}) set_rbt" and rbt4 :: "('d \ 'a) set_rbt" and rbt5 :: "('b \ 'a) set_rbt" and dxs1 :: "('d \ 'e :: ceq) set_dlist" and dxs2 :: "('e \ 'd) set_dlist" and dxs3 :: "('e \ 'f :: ceq) set_dlist" and dxs4 :: "('f \ 'g :: ceq) set_dlist" and xs1 :: "('h \ 'i :: ceq) list" and xs2 :: "('i \ 'j) list" and xs3 :: "('b \ 'h) list" and xs4 :: "('h \ 'b) list" and xs5 :: "('f \ 'h) list" and xs6 :: "('h \ 'f) list" shows "RBT_set rbt1 O RBT_set rbt2 = (case ID CCOMPARE('a) of None \ Code.abort (STR ''relcomp RBT_set RBT_set: ccompare1 = None'') (\_. RBT_set rbt1 O RBT_set rbt2) | Some _ \ case ID CCOMPARE('b) of None \ Code.abort (STR ''relcomp RBT_set RBT_set: ccompare2 = None'') (\_. RBT_set rbt1 O RBT_set rbt2) | Some c_b \ case ID CCOMPARE('c) of None \ Code.abort (STR ''relcomp RBT_set RBT_set: ccompare3 = None'') (\_. RBT_set rbt1 O RBT_set rbt2) | Some _ \ RBT_Set2.fold (\(x, y). RBT_Set2.fold (\(y', z) A. if c_b y y' \ Eq then A else insert (x, z) A) rbt2) rbt1 {})" (is ?rbt_rbt) "RBT_set rbt3 O DList_set dxs1 = (case ID CCOMPARE('a) of None \ Code.abort (STR ''relcomp RBT_set DList_set: ccompare1 = None'') (\_. RBT_set rbt3 O DList_set dxs1) | Some _ \ case ID CCOMPARE('d) of None \ Code.abort (STR ''relcomp RBT_set DList_set: ccompare2 = None'') (\_. RBT_set rbt3 O DList_set dxs1) | Some _ \ case ID CEQ('d) of None \ Code.abort (STR ''relcomp RBT_set DList_set: ceq2 = None'') (\_. RBT_set rbt3 O DList_set dxs1) | Some eq \ case ID CEQ('e) of None \ Code.abort (STR ''relcomp RBT_set DList_set: ceq3 = None'') (\_. RBT_set rbt3 O DList_set dxs1) | Some _ \ RBT_Set2.fold (\(x, y). DList_Set.fold (\(y', z) A. if eq y y' then insert (x, z) A else A) dxs1) rbt3 {})" (is ?rbt_dlist) "DList_set dxs2 O RBT_set rbt4 = (case ID CEQ('e) of None \ Code.abort (STR ''relcomp DList_set RBT_set: ceq1 = None'') (\_. DList_set dxs2 O RBT_set rbt4) | Some _ \ case ID CCOMPARE('d) of None \ Code.abort (STR ''relcomp DList_set RBT_set: ceq2 = None'') (\_. DList_set dxs2 O RBT_set rbt4) | Some _ \ case ID CEQ('d) of None \ Code.abort (STR ''relcomp DList_set RBT_set: ccompare2 = None'') (\_. DList_set dxs2 O RBT_set rbt4) | Some eq \ case ID CCOMPARE('a) of None \ Code.abort (STR ''relcomp DList_set RBT_set: ccompare3 = None'') (\_. DList_set dxs2 O RBT_set rbt4) | Some _ \ DList_Set.fold (\(x, y). RBT_Set2.fold (\(y', z) A. if eq y y' then insert (x, z) A else A) rbt4) dxs2 {})" (is ?dlist_rbt) "DList_set dxs3 O DList_set dxs4 = (case ID CEQ('e) of None \ Code.abort (STR ''relcomp DList_set DList_set: ceq1 = None'') (\_. DList_set dxs3 O DList_set dxs4) | Some _ \ case ID CEQ('f) of None \ Code.abort (STR ''relcomp DList_set DList_set: ceq2 = None'') (\_. DList_set dxs3 O DList_set dxs4) | Some eq \ case ID CEQ('g) of None \ Code.abort (STR ''relcomp DList_set DList_set: ceq3 = None'') (\_. DList_set dxs3 O DList_set dxs4) | Some _ \ DList_Set.fold (\(x, y). DList_Set.fold (\(y', z) A. if eq y y' then insert (x, z) A else A) dxs4) dxs3 {})" (is ?dlist_dlist) "Set_Monad xs1 O Set_Monad xs2 = (case ID CEQ('i) of None \ Code.abort (STR ''relcomp Set_Monad Set_Monad: ceq = None'') (\_. Set_Monad xs1 O Set_Monad xs2) | Some eq \ fold (\(x, y). fold (\(y', z) A. if eq y y' then insert (x, z) A else A) xs2) xs1 {})" (is ?monad_monad) "RBT_set rbt1 O Set_Monad xs3 = (case ID CCOMPARE('a) of None \ Code.abort (STR ''relcomp RBT_set Set_Monad: ccompare1 = None'') (\_. RBT_set rbt1 O Set_Monad xs3) | Some _ \ case ID CCOMPARE('b) of None \ Code.abort (STR ''relcomp RBT_set Set_Monad: ccompare2 = None'') (\_. RBT_set rbt1 O Set_Monad xs3) | Some c_b \ RBT_Set2.fold (\(x, y). fold (\(y', z) A. if c_b y y' \ Eq then A else insert (x, z) A) xs3) rbt1 {})" (is ?rbt_monad) "Set_Monad xs4 O RBT_set rbt5 = (case ID CCOMPARE('a) of None \ Code.abort (STR ''relcomp Set_Monad RBT_set: ccompare1 = None'') (\_. Set_Monad xs4 O RBT_set rbt5) | Some _ \ case ID CCOMPARE('b) of None \ Code.abort (STR ''relcomp Set_Monad RBT_set: ccompare2 = None'') (\_. Set_Monad xs4 O RBT_set rbt5) | Some c_b \ fold (\(x, y). RBT_Set2.fold (\(y', z) A. if c_b y y' \ Eq then A else insert (x, z) A) rbt5) xs4 {})" (is ?monad_rbt) "DList_set dxs3 O Set_Monad xs5 = (case ID CEQ('e) of None \ Code.abort (STR ''relcomp DList_set Set_Monad: ceq1 = None'') (\_. DList_set dxs3 O Set_Monad xs5) | Some _ \ case ID CEQ('f) of None \ Code.abort (STR ''relcomp DList_set Set_Monad: ceq2 = None'') (\_. DList_set dxs3 O Set_Monad xs5) | Some eq \ DList_Set.fold (\(x, y). fold (\(y', z) A. if eq y y' then insert (x, z) A else A) xs5) dxs3 {})" (is ?dlist_monad) "Set_Monad xs6 O DList_set dxs4 = (case ID CEQ('f) of None \ Code.abort (STR ''relcomp Set_Monad DList_set: ceq1 = None'') (\_. Set_Monad xs6 O DList_set dxs4) | Some eq \ case ID CEQ('g) of None \ Code.abort (STR ''relcomp Set_Monad DList_set: ceq2 = None'') (\_. Set_Monad xs6 O DList_set dxs4) | Some _ \ fold (\(x, y). DList_Set.fold (\(y', z) A. if eq y y' then insert (x, z) A else A) dxs4) xs6 {})" (is ?monad_dlist) proof - show ?rbt_rbt ?rbt_monad ?monad_rbt by(auto simp add: comparator.eq[OF ID_ccompare'] RBT_set_def ccompare_prod_def member_conv_keys ID_Some RBT_Set2.fold_conv_fold_keys' RBT_Set2.keys.rep_eq If_not set_relcomp_set split: option.split del: equalityI) show ?rbt_dlist ?dlist_rbt ?dlist_dlist ?monad_monad ?dlist_monad ?monad_dlist by(auto simp add: RBT_set_def DList_set_def member_conv_keys ID_Some ccompare_prod_def ceq_prod_def Collect_member RBT_Set2.fold_conv_fold_keys' RBT_Set2.keys.rep_eq DList_Set.fold.rep_eq set_relcomp_set dest: equal.equal_eq[OF ID_ceq] split: option.split del: equalityI) qed lemma irrefl_code [code]: fixes r :: "('a :: {ceq, ccompare} \ 'a) set" shows "irrefl r \ (case ID CEQ('a) of Some eq \ (\(x, y) \ r. \ eq x y) | None \ case ID CCOMPARE('a) of None \ Code.abort (STR ''irrefl: ceq = None & ccompare = None'') (\_. irrefl r) | Some c \ (\(x, y) \ r. c x y \ Eq))" apply(auto simp add: irrefl_distinct comparator.eq[OF ID_ccompare'] split: option.split dest!: ID_ceq[THEN equal.equal_eq]) done lemma wf_code [code]: fixes rbt :: "('a :: ccompare \ 'a) set_rbt" and dxs :: "('b :: ceq \ 'b) set_dlist" shows "wf (Set_Monad xs) = acyclic (Set_Monad xs)" "wf (RBT_set rbt) = (case ID CCOMPARE('a) of None \ Code.abort (STR ''wf RBT_set: ccompare = None'') (\_. wf (RBT_set rbt)) | Some _ \ acyclic (RBT_set rbt))" "wf (DList_set dxs) = (case ID CEQ('b) of None \ Code.abort (STR ''wf DList_set: ceq = None'') (\_. wf (DList_set dxs)) | Some _ \ acyclic (DList_set dxs))" by(auto simp add: wf_iff_acyclic_if_finite split: option.split del: iffI)(simp_all add: wf_iff_acyclic_if_finite finite_code ccompare_prod_def ceq_prod_def ID_Some) lemma bacc_code [code]: "bacc R 0 = - snd ` R" "bacc R (Suc n) = (let rec = bacc R n in rec \ - snd ` (Set.filter (\(y, x). y \ rec) R))" by(auto intro: rev_image_eqI simp add: Let_def) (* TODO: acc could also be computed for infinite universes if r is finite *) lemma acc_code [code]: fixes A :: "('a :: {finite, card_UNIV} \ 'a) set" shows "Wellfounded.acc A = bacc A (of_phantom (card_UNIV :: 'a card_UNIV))" by(simp add: card_UNIV acc_bacc_eq) lemma sorted_list_of_set_code [code]: fixes dxs :: "'a :: {linorder, ceq} set_dlist" and rbt :: "'b :: {linorder, ccompare} set_rbt" shows "sorted_list_of_set (Set_Monad xs) = sort (remdups xs)" "sorted_list_of_set (DList_set dxs) = (case ID CEQ('a) of None \ Code.abort (STR ''sorted_list_of_set DList_set: ceq = None'') (\_. sorted_list_of_set (DList_set dxs)) | Some _ \ sort (list_of_dlist dxs))" "sorted_list_of_set (RBT_set rbt) = (case ID CCOMPARE('b) of None \ Code.abort (STR ''sorted_list_of_set RBT_set: ccompare = None'') (\_. sorted_list_of_set (RBT_set rbt)) | Some _ \ sort (RBT_Set2.keys rbt))" \ \We must sort the keys because @{term ccompare}'s ordering need not coincide with @{term linorder}'s.\ by(auto simp add: DList_set_def RBT_set_def sorted_list_of_set_sort_remdups Collect_member distinct_remdups_id distinct_list_of_dlist member_conv_keys split: option.split) lemma map_project_set: "List.map_project f (set xs) = set (List.map_filter f xs)" by(auto simp add: List.map_project_def List.map_filter_def intro: rev_image_eqI) lemma map_project_simps: shows map_project_empty: "List.map_project f {} = {}" and map_project_insert: "List.map_project f (insert x A) = (case f x of None \ List.map_project f A | Some y \ insert y (List.map_project f A))" by(auto simp add: List.map_project_def split: option.split) lemma map_project_conv_fold: "List.map_project f (set xs) = fold (\x A. case f x of None \ A | Some y \ insert y A) xs {}" by(induct xs rule: rev_induct)(simp_all add: map_project_simps cong: option.case_cong) lemma map_project_code [code]: fixes dxs :: "'a :: ceq set_dlist" and rbt :: "'b :: ccompare set_rbt" shows "List.map_project f (Set_Monad xs) = Set_Monad (List.map_filter f xs)" "List.map_project g (DList_set dxs) = (case ID CEQ('a) of None \ Code.abort (STR ''map_project DList_set: ceq = None'') (\_. List.map_project g (DList_set dxs)) | Some _ \ DList_Set.fold (\x A. case g x of None \ A | Some y \ insert y A) dxs {})" (is ?dlist) "List.map_project h (RBT_set rbt) = (case ID CCOMPARE('b) of None \ Code.abort (STR ''map_project RBT_set: ccompare = None'') (\_. List.map_project h (RBT_set rbt)) | Some _ \ RBT_Set2.fold (\x A. case h x of None \ A | Some y \ insert y A) rbt {})" (is ?rbt) proof - show ?dlist ?rbt by(auto split: option.split simp add: RBT_set_def DList_set_def DList_Set.fold.rep_eq Collect_member map_project_conv_fold RBT_Set2.fold_conv_fold_keys member_conv_keys del: equalityI) qed(auto simp add: List.map_project_def List.map_filter_def intro: rev_image_eqI) lemma Bleast_code [code]: "Bleast A P = (if finite A then case filter P (sorted_list_of_set A) of [] \ abort_Bleast A P | x # xs \ x else abort_Bleast A P)" proof(cases "finite A") case True hence *: "A = set (sorted_list_of_set A)" by(simp add: sorted_list_of_set) show ?thesis using True by(subst (1 3) *)(unfold Bleast_code, simp add: sorted_sort_id) qed(simp add: abort_Bleast_def Bleast_def) lemma can_select_code [code]: fixes xs :: "'a :: ceq list" and dxs :: "'a :: ceq set_dlist" and rbt :: "'b :: ccompare set_rbt" shows "can_select P (Set_Monad xs) = (case ID CEQ('a) of None \ Code.abort (STR ''can_select Set_Monad: ceq = None'') (\_. can_select P (Set_Monad xs)) | Some eq \ case filter P xs of Nil \ False | x # xs \ list_all (eq x) xs)" (is ?Set_Monad) "can_select Q (DList_set dxs) = (case ID CEQ('a) of None \ Code.abort (STR ''can_select DList_set: ceq = None'') (\_. can_select Q (DList_set dxs)) | Some _ \ DList_Set.length (DList_Set.filter Q dxs) = 1)" (is ?dlist) "can_select R (RBT_set rbt) = (case ID CCOMPARE('b) of None \ Code.abort (STR ''can_select RBT_set: ccompare = None'') (\_. can_select R (RBT_set rbt)) | Some _ \ singleton_list_fusion (filter_generator R rbt_keys_generator) (RBT_Set2.init rbt))" (is ?rbt) proof - show ?Set_Monad apply(auto split: option.split list.split dest!: ID_ceq[THEN equal.equal_eq] dest: filter_eq_ConsD simp add: can_select_def filter_empty_conv list_all_iff) apply(drule filter_eq_ConsD, fastforce) apply(drule filter_eq_ConsD, clarsimp, blast) done show ?dlist by(clarsimp simp add: can_select_def card_eq_length[symmetric] Set_member_code card_eq_Suc_0_ex1 simp del: card_eq_length split: option.split) note [simp del] = distinct_keys show ?rbt using distinct_keys[of rbt] apply(auto simp add: can_select_def singleton_list_fusion_def unfoldr_filter_generator unfoldr_rbt_keys_generator Set_member_code member_conv_keys filter_empty_conv empty_filter_conv split: option.split list.split dest: filter_eq_ConsD) apply(drule filter_eq_ConsD, fastforce) apply(drule filter_eq_ConsD, fastforce simp add: empty_filter_conv) apply(drule filter_eq_ConsD) apply clarsimp apply(drule Cons_eq_filterD) apply clarify apply(simp (no_asm_use)) apply blast done qed lemma pred_of_set_code [code]: fixes dxs :: "'a :: ceq set_dlist" and rbt :: "'b :: ccompare set_rbt" shows "pred_of_set (Set_Monad xs) = fold (sup \ Predicate.single) xs bot" "pred_of_set (DList_set dxs) = (case ID CEQ('a) of None \ Code.abort (STR ''pred_of_set DList_set: ceq = None'') (\_. pred_of_set (DList_set dxs)) | Some _ \ DList_Set.fold (sup \ Predicate.single) dxs bot)" "pred_of_set (RBT_set rbt) = (case ID CCOMPARE('b) of None \ Code.abort (STR ''pred_of_set RBT_set: ccompare = None'') (\_. pred_of_set (RBT_set rbt)) | Some _ \ RBT_Set2.fold (sup \ Predicate.single) rbt bot)" by(auto simp add: pred_of_set_set_fold_sup fold_map DList_set_def RBT_set_def Collect_member member_conv_keys DList_Set.fold.rep_eq fold_conv_fold_keys split: option.split) text \ @{typ "'a Predicate.pred"} is implemented as a monad, so we keep the monad when converting to @{typ "'a set"}. For this case, @{term insert_monad} and @{term union_monad} avoid the unnecessary dictionary construction. \ definition insert_monad :: "'a \ 'a set \ 'a set" where [simp]: "insert_monad = insert" definition union_monad :: "'a set \ 'a set \ 'a set" where [simp]: "union_monad = (\)" lemma insert_monad_code [code]: "insert_monad x (Set_Monad xs) = Set_Monad (x # xs)" by simp lemma union_monad_code [code]: "union_monad (Set_Monad xs) (Set_Monad ys) = Set_Monad (xs @ ys)" by(simp) lemma set_of_pred_code [code]: "set_of_pred (Predicate.Seq f) = (case f () of seq.Empty \ Set_Monad [] | seq.Insert x P \ insert_monad x (set_of_pred P) | seq.Join P xq \ union_monad (set_of_pred P) (set_of_seq xq))" by(simp add: of_pred_code cong: seq.case_cong) lemma set_of_seq_code [code]: "set_of_seq seq.Empty = Set_Monad []" "set_of_seq (seq.Insert x P) = insert_monad x (set_of_pred P)" "set_of_seq (seq.Join P xq) = union_monad (set_of_pred P) (set_of_seq xq)" by(simp_all add: of_seq_code) hide_const (open) insert_monad union_monad subsection \Type class instantiations\ datatype set_impl = Set_IMPL declare set_impl.eq.simps [code del] set_impl.size [code del] set_impl.rec [code del] set_impl.case [code del] lemma [code]: fixes x :: set_impl shows "size x = 0" and "size_set_impl x = 0" by(case_tac [!] x) simp_all definition set_Choose :: set_impl where [simp]: "set_Choose = Set_IMPL" definition set_Collect :: set_impl where [simp]: "set_Collect = Set_IMPL" definition set_DList :: set_impl where [simp]: "set_DList = Set_IMPL" definition set_RBT :: set_impl where [simp]: "set_RBT = Set_IMPL" definition set_Monad :: set_impl where [simp]: "set_Monad = Set_IMPL" code_datatype set_Choose set_Collect set_DList set_RBT set_Monad definition set_empty_choose :: "'a set" where [simp]: "set_empty_choose = {}" lemma set_empty_choose_code [code]: "(set_empty_choose :: 'a :: {ceq, ccompare} set) = (case CCOMPARE('a) of Some _ \ RBT_set RBT_Set2.empty | None \ case CEQ('a) of None \ Set_Monad [] | Some _ \ DList_set (DList_Set.empty))" by(simp split: option.split) definition set_impl_choose2 :: "set_impl \ set_impl \ set_impl" where [simp]: "set_impl_choose2 = (\_ _. Set_IMPL)" lemma set_impl_choose2_code [code]: "set_impl_choose2 x y = set_Choose" "set_impl_choose2 set_Collect set_Collect = set_Collect" "set_impl_choose2 set_DList set_DList = set_DList" "set_impl_choose2 set_RBT set_RBT = set_RBT" "set_impl_choose2 set_Monad set_Monad = set_Monad" by(simp_all) definition set_empty :: "set_impl \ 'a set" where [simp]: "set_empty = (\_. {})" lemma set_empty_code [code]: "set_empty set_Collect = Collect_set (\_. False)" "set_empty set_DList = DList_set DList_Set.empty" "set_empty set_RBT = RBT_set RBT_Set2.empty" "set_empty set_Monad = Set_Monad []" "set_empty set_Choose = set_empty_choose" by(simp_all) class set_impl = fixes set_impl :: "('a, set_impl) phantom" syntax (input) "_SET_IMPL" :: "type => logic" ("(1SET'_IMPL/(1'(_')))") parse_translation \ let fun set_impl_tr [ty] = (Syntax.const @{syntax_const "_constrain"} $ Syntax.const @{const_syntax "set_impl"} $ (Syntax.const @{type_syntax phantom} $ ty $ Syntax.const @{type_syntax set_impl})) | set_impl_tr ts = raise TERM ("set_impl_tr", ts); in [(@{syntax_const "_SET_IMPL"}, K set_impl_tr)] end \ declare [[code drop: "{}"]] lemma empty_code [code, code_unfold]: "({} :: 'a :: set_impl set) = set_empty (of_phantom SET_IMPL('a))" by simp subsection \Generator for the @{class set_impl}-class\ text \ This generator registers itself at the derive-manager for the classes @{class set_impl}. Here, one can choose the desired implementation via the parameter. \begin{itemize} \item \texttt{instantiation type :: (type,\ldots,type) (rbt,dlist,collect,monad,choose, or arbitrary constant name) set-impl} \end{itemize} \ text \ This generator can be used for arbitrary types, not just datatypes. \ ML_file \set_impl_generator.ML\ derive (dlist) set_impl unit bool derive (rbt) set_impl nat derive (set_RBT) set_impl int (* shows usage of constant names *) derive (dlist) set_impl Enum.finite_1 Enum.finite_2 Enum.finite_3 derive (rbt) set_impl integer natural derive (rbt) set_impl char instantiation sum :: (set_impl, set_impl) set_impl begin definition "SET_IMPL('a + 'b) = Phantom('a + 'b) (set_impl_choose2 (of_phantom SET_IMPL('a)) (of_phantom SET_IMPL('b)))" instance .. end instantiation prod :: (set_impl, set_impl) set_impl begin definition "SET_IMPL('a * 'b) = Phantom('a * 'b) (set_impl_choose2 (of_phantom SET_IMPL('a)) (of_phantom SET_IMPL('b)))" instance .. end derive (choose) set_impl list derive (rbt) set_impl String.literal instantiation option :: (set_impl) set_impl begin definition "SET_IMPL('a option) = Phantom('a option) (of_phantom SET_IMPL('a))" instance .. end derive (monad) set_impl "fun" derive (choose) set_impl set instantiation phantom :: (type, set_impl) set_impl begin definition "SET_IMPL(('a, 'b) phantom) = Phantom (('a, 'b) phantom) (of_phantom SET_IMPL('b))" instance .. end text \ We enable automatic implementation selection for sets constructed by @{const set}, although they could be directly converted using @{const Set_Monad} in constant time. However, then it is more likely that the parameters of binary operators have different implementations, which can lead to less efficient execution. However, we test whether automatic selection picks @{const Set_Monad} anyway and take a short-cut. \ definition set_aux :: "set_impl \ 'a list \ 'a set" where [simp, code del]: "set_aux _ = set" lemma set_aux_code [code]: defines "conv \ foldl (\s (x :: 'a). insert x s)" shows "set_aux impl = conv (set_empty impl)" (is "?thesis1") "set_aux set_Choose = (case CCOMPARE('a :: {ccompare, ceq}) of Some _ \ conv (RBT_set RBT_Set2.empty) | None \ case CEQ('a) of None \ Set_Monad | Some _ \ conv (DList_set DList_Set.empty))" (is "?thesis2") "set_aux set_Monad = Set_Monad" proof - have "conv {} = set" by(rule ext)(induct_tac x rule: rev_induct, simp_all add: conv_def) thus ?thesis1 ?thesis2 by(simp_all split: option.split) qed simp lemma set_code [code]: fixes xs :: "'a :: set_impl list" shows "set xs = set_aux (of_phantom (ID SET_IMPL('a))) xs" by(simp) subsection \Pretty printing for sets\ text \ @{term code_post} marks contexts (as hypothesis) in which we use code\_post as a decision procedure rather than a pretty-printing engine. The intended use is to enable more rules when proving assumptions of rewrite rules. \ definition code_post :: bool where "code_post = True" lemma conj_code_post [code_post]: assumes code_post shows "True & x \ x" "False & x \ False" by simp_all text \ A flag to switch post-processing of sets on and off. Use \verb$declare pretty_sets[code_post del]$ to disable pretty printing of sets in value. \ definition code_post_set :: bool where pretty_sets [code_post, simp]: "code_post_set = True" definition collapse_RBT_set :: "'a set_rbt \ 'a :: ccompare set \ 'a set" where "collapse_RBT_set r M = set (RBT_Set2.keys r) \ M" lemma RBT_set_collapse_RBT_set [code_post]: fixes r :: "'a :: ccompare set_rbt" assumes "code_post \ is_ccompare TYPE('a)" and code_post_set shows "RBT_set r = collapse_RBT_set r {}" using assms by(clarsimp simp add: code_post_def is_ccompare_def RBT_set_def member_conv_keys collapse_RBT_set_def) lemma collapse_RBT_set_Branch [code_post]: "collapse_RBT_set (Mapping_RBT (Branch c l x v r)) M = collapse_RBT_set (Mapping_RBT l) (insert x (collapse_RBT_set (Mapping_RBT r) M))" unfolding collapse_RBT_set_def by(auto simp add: is_ccompare_def set_keys_Mapping_RBT) lemma collapse_RBT_set_Empty [code_post]: "collapse_RBT_set (Mapping_RBT rbt.Empty) M = M" by(auto simp add: collapse_RBT_set_def set_keys_Mapping_RBT) definition collapse_DList_set :: "'a :: ceq set_dlist \ 'a set" where "collapse_DList_set dxs = set (DList_Set.list_of_dlist dxs)" lemma DList_set_collapse_DList_set [code_post]: fixes dxs :: "'a :: ceq set_dlist" assumes "code_post \ is_ceq TYPE('a)" and code_post_set shows "DList_set dxs = collapse_DList_set dxs" using assms by(clarsimp simp add: code_post_def DList_set_def is_ceq_def collapse_DList_set_def Collect_member) lemma collapse_DList_set_empty [code_post]: "collapse_DList_set (Abs_dlist []) = {}" by(simp add: collapse_DList_set_def Abs_dlist_inverse) lemma collapse_DList_set_Cons [code_post]: "collapse_DList_set (Abs_dlist (x # xs)) = insert x (collapse_DList_set (Abs_dlist xs))" by(simp add: collapse_DList_set_def set_list_of_dlist_Abs_dlist) lemma Set_Monad_code_post [code_post]: assumes code_post_set shows "Set_Monad [] = {}" and "Set_Monad (x#xs) = insert x (Set_Monad xs)" by simp_all end diff --git a/thys/Core_DOM/standard/Core_DOM_Heap_WF.thy b/thys/Core_DOM/standard/Core_DOM_Heap_WF.thy --- a/thys/Core_DOM/standard/Core_DOM_Heap_WF.thy +++ b/thys/Core_DOM/standard/Core_DOM_Heap_WF.thy @@ -1,8045 +1,8045 @@ (*********************************************************************************** * Copyright (c) 2016-2018 The University of Sheffield, UK * * 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. * * 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 HOLDER 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. * * SPDX-License-Identifier: BSD-2-Clause ***********************************************************************************) section\Wellformedness\ text\In this theory, we discuss the wellformedness of the DOM. First, we define wellformedness and, second, we show for all functions for querying and modifying the DOM to what extend they preserve wellformendess.\ theory Core_DOM_Heap_WF imports "Core_DOM_Functions" begin locale l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_child_nodes_defs get_child_nodes get_child_nodes_locs + l_get_disconnected_nodes_defs get_disconnected_nodes get_disconnected_nodes_locs for get_child_nodes :: "(_::linorder) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" begin definition a_owner_document_valid :: "(_) heap \ bool" where "a_owner_document_valid h \ (\node_ptr \ fset (node_ptr_kinds h). ((\document_ptr. document_ptr |\| document_ptr_kinds h \ node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r) \ (\parent_ptr. parent_ptr |\| object_ptr_kinds h \ node_ptr \ set |h \ get_child_nodes parent_ptr|\<^sub>r)))" lemma a_owner_document_valid_code [code]: "a_owner_document_valid h \ node_ptr_kinds h |\| fset_of_list (concat (map (\parent. |h \ get_child_nodes parent|\<^sub>r) (sorted_list_of_fset (object_ptr_kinds h)) @ map (\parent. |h \ get_disconnected_nodes parent|\<^sub>r) (sorted_list_of_fset (document_ptr_kinds h)))) " apply(auto simp add: a_owner_document_valid_def l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_owner_document_valid_def)[1] proof - fix x assume 1: " \node_ptr\fset (node_ptr_kinds h). (\document_ptr. document_ptr |\| document_ptr_kinds h \ node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r) \ (\parent_ptr. parent_ptr |\| object_ptr_kinds h \ node_ptr \ set |h \ get_child_nodes parent_ptr|\<^sub>r)" assume 2: "x |\| node_ptr_kinds h" assume 3: "x |\| fset_of_list (concat (map (\parent. |h \ get_disconnected_nodes parent|\<^sub>r) (sorted_list_of_fset (document_ptr_kinds h))))" have "\(\document_ptr. document_ptr |\| document_ptr_kinds h \ x \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r)" using 1 2 3 by (smt UN_I fset_of_list_elem image_eqI notin_fset set_concat set_map sorted_list_of_fset_simps(1)) then have "(\parent_ptr. parent_ptr |\| object_ptr_kinds h \ x \ set |h \ get_child_nodes parent_ptr|\<^sub>r)" using 1 2 by auto then obtain parent_ptr where parent_ptr: "parent_ptr |\| object_ptr_kinds h \ x \ set |h \ get_child_nodes parent_ptr|\<^sub>r" by auto moreover have "parent_ptr \ set (sorted_list_of_fset (object_ptr_kinds h))" using parent_ptr by auto moreover have "|h \ get_child_nodes parent_ptr|\<^sub>r \ set (map (\parent. |h \ get_child_nodes parent|\<^sub>r) (sorted_list_of_fset (object_ptr_kinds h)))" using calculation(2) by auto ultimately show "x |\| fset_of_list (concat (map (\parent. |h \ get_child_nodes parent|\<^sub>r) (sorted_list_of_fset (object_ptr_kinds h))))" using fset_of_list_elem by fastforce next fix node_ptr assume 1: "node_ptr_kinds h |\| fset_of_list (concat (map (\parent. |h \ get_child_nodes parent|\<^sub>r) (sorted_list_of_fset (object_ptr_kinds h)))) |\| fset_of_list (concat (map (\parent. |h \ get_disconnected_nodes parent|\<^sub>r) (sorted_list_of_fset (document_ptr_kinds h))))" assume 2: "node_ptr |\| node_ptr_kinds h" assume 3: "\parent_ptr. parent_ptr |\| object_ptr_kinds h \ node_ptr \ set |h \ get_child_nodes parent_ptr|\<^sub>r" have "node_ptr \ set (concat (map (\parent. |h \ get_child_nodes parent|\<^sub>r) (sorted_list_of_fset (object_ptr_kinds h)))) \ node_ptr \ set (concat (map (\parent. |h \ get_disconnected_nodes parent|\<^sub>r) (sorted_list_of_fset (document_ptr_kinds h))))" using 1 2 by (meson fin_mono fset_of_list_elem funion_iff) then show "\document_ptr. document_ptr |\| document_ptr_kinds h \ node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r" using 3 by auto qed definition a_parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" where "a_parent_child_rel h = {(parent, child). parent |\| object_ptr_kinds h \ child \ cast ` set |h \ get_child_nodes parent|\<^sub>r}" lemma a_parent_child_rel_code [code]: "a_parent_child_rel h = set (concat (map (\parent. map (\child. (parent, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child)) |h \ get_child_nodes parent|\<^sub>r) (sorted_list_of_fset (object_ptr_kinds h))) )" by(auto simp add: a_parent_child_rel_def) definition a_acyclic_heap :: "(_) heap \ bool" where "a_acyclic_heap h = acyclic (a_parent_child_rel h)" definition a_all_ptrs_in_heap :: "(_) heap \ bool" where "a_all_ptrs_in_heap h \ (\ptr \ fset (object_ptr_kinds h). set |h \ get_child_nodes ptr|\<^sub>r \ fset (node_ptr_kinds h)) \ (\document_ptr \ fset (document_ptr_kinds h). set |h \ get_disconnected_nodes document_ptr|\<^sub>r \ fset (node_ptr_kinds h))" definition a_distinct_lists :: "(_) heap \ bool" where "a_distinct_lists h = distinct (concat ( (map (\ptr. |h \ get_child_nodes ptr|\<^sub>r) |h \ object_ptr_kinds_M|\<^sub>r) @ (map (\document_ptr. |h \ get_disconnected_nodes document_ptr|\<^sub>r) |h \ document_ptr_kinds_M|\<^sub>r) ))" definition a_heap_is_wellformed :: "(_) heap \ bool" where "a_heap_is_wellformed h \ a_acyclic_heap h \ a_all_ptrs_in_heap h \ a_distinct_lists h \ a_owner_document_valid h" end locale l_heap_is_wellformed_defs = fixes heap_is_wellformed :: "(_) heap \ bool" fixes parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" global_interpretation l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs defines heap_is_wellformed = "l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_heap_is_wellformed get_child_nodes get_disconnected_nodes" and parent_child_rel = "l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_parent_child_rel get_child_nodes" and acyclic_heap = a_acyclic_heap and all_ptrs_in_heap = a_all_ptrs_in_heap and distinct_lists = a_distinct_lists and owner_document_valid = a_owner_document_valid . locale l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs + l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs + l_heap_is_wellformed_defs heap_is_wellformed parent_child_rel + l_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs for known_ptr :: "(_::linorder) object_ptr \ bool" and type_wf :: "(_) heap \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" + assumes heap_is_wellformed_impl: "heap_is_wellformed = a_heap_is_wellformed" assumes parent_child_rel_impl: "parent_child_rel = a_parent_child_rel" begin lemmas heap_is_wellformed_def = heap_is_wellformed_impl[unfolded a_heap_is_wellformed_def] lemmas parent_child_rel_def = parent_child_rel_impl[unfolded a_parent_child_rel_def] lemmas acyclic_heap_def = a_acyclic_heap_def[folded parent_child_rel_impl] lemma parent_child_rel_node_ptr: "(parent, child) \ parent_child_rel h \ is_node_ptr_kind child" by(auto simp add: parent_child_rel_def) lemma parent_child_rel_child_nodes: assumes "known_ptr parent" and "h \ get_child_nodes parent \\<^sub>r children" and "child \ set children" shows "(parent, cast child) \ parent_child_rel h" using assms apply(auto simp add: parent_child_rel_def is_OK_returns_result_I )[1] using get_child_nodes_ptr_in_heap by blast lemma parent_child_rel_child_nodes2: assumes "known_ptr parent" and "h \ get_child_nodes parent \\<^sub>r children" and "child \ set children" and "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child = child_obj" shows "(parent, child_obj) \ parent_child_rel h" using assms parent_child_rel_child_nodes by blast lemma parent_child_rel_finite: "finite (parent_child_rel h)" proof - have "parent_child_rel h = (\ptr \ set |h \ object_ptr_kinds_M|\<^sub>r. (\child \ set |h \ get_child_nodes ptr|\<^sub>r. {(ptr, cast child)}))" by(auto simp add: parent_child_rel_def) moreover have "finite (\ptr \ set |h \ object_ptr_kinds_M|\<^sub>r. (\child \ set |h \ get_child_nodes ptr|\<^sub>r. {(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child)}))" by simp ultimately show ?thesis by simp qed lemma distinct_lists_no_parent: assumes "a_distinct_lists h" assumes "h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes" assumes "node_ptr \ set disc_nodes" shows "\(\parent_ptr. parent_ptr |\| object_ptr_kinds h \ node_ptr \ set |h \ get_child_nodes parent_ptr|\<^sub>r)" using assms apply(auto simp add: a_distinct_lists_def)[1] proof - fix parent_ptr :: "(_) object_ptr" assume a1: "parent_ptr |\| object_ptr_kinds h" assume a2: "(\x\fset (object_ptr_kinds h). set |h \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h). set |h \ get_disconnected_nodes x|\<^sub>r) = {}" assume a3: "h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes" assume a4: "node_ptr \ set disc_nodes" assume a5: "node_ptr \ set |h \ get_child_nodes parent_ptr|\<^sub>r" have f6: "parent_ptr \ fset (object_ptr_kinds h)" using a1 by auto have f7: "document_ptr \ fset (document_ptr_kinds h)" using a3 by (meson fmember.rep_eq get_disconnected_nodes_ptr_in_heap is_OK_returns_result_I) have "|h \ get_disconnected_nodes document_ptr|\<^sub>r = disc_nodes" using a3 by simp then show False using f7 f6 a5 a4 a2 by blast qed lemma distinct_lists_disconnected_nodes: assumes "a_distinct_lists h" and "h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes" shows "distinct disc_nodes" proof - have h1: "distinct (concat (map (\document_ptr. |h \ get_disconnected_nodes document_ptr|\<^sub>r) |h \ document_ptr_kinds_M|\<^sub>r))" using assms(1) by(simp add: a_distinct_lists_def) then show ?thesis using concat_map_all_distinct[OF h1] assms(2) is_OK_returns_result_I get_disconnected_nodes_ok by (metis (no_types, lifting) DocumentMonad.ptr_kinds_ptr_kinds_M l_get_disconnected_nodes.get_disconnected_nodes_ptr_in_heap l_get_disconnected_nodes_axioms select_result_I2) qed lemma distinct_lists_children: assumes "a_distinct_lists h" and "known_ptr ptr" and "h \ get_child_nodes ptr \\<^sub>r children" shows "distinct children" proof (cases "children = []", simp) assume "children \ []" have h1: "distinct (concat ((map (\ptr. |h \ get_child_nodes ptr|\<^sub>r) |h \ object_ptr_kinds_M|\<^sub>r)))" using assms(1) by(simp add: a_distinct_lists_def) show ?thesis using concat_map_all_distinct[OF h1] assms(2) assms(3) by (metis (no_types, lifting) ObjectMonad.ptr_kinds_ptr_kinds_M get_child_nodes_ptr_in_heap is_OK_returns_result_I select_result_I2) qed lemma heap_is_wellformed_children_in_heap: assumes "heap_is_wellformed h" assumes "h \ get_child_nodes ptr \\<^sub>r children" assumes "child \ set children" shows "child |\| node_ptr_kinds h" using assms apply(auto simp add: heap_is_wellformed_def a_all_ptrs_in_heap_def)[1] by (metis (no_types, lifting) finite_set_in is_OK_returns_result_I local.get_child_nodes_ptr_in_heap select_result_I2 subsetD) lemma heap_is_wellformed_one_parent: assumes "heap_is_wellformed h" assumes "h \ get_child_nodes ptr \\<^sub>r children" assumes "h \ get_child_nodes ptr' \\<^sub>r children'" assumes "set children \ set children' \ {}" shows "ptr = ptr'" using assms proof (auto simp add: heap_is_wellformed_def a_distinct_lists_def)[1] fix x :: "(_) node_ptr" assume a1: "ptr \ ptr'" assume a2: "h \ get_child_nodes ptr \\<^sub>r children" assume a3: "h \ get_child_nodes ptr' \\<^sub>r children'" assume a4: "distinct (concat (map (\ptr. |h \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h)))))" have f5: "|h \ get_child_nodes ptr|\<^sub>r = children" using a2 by simp have "|h \ get_child_nodes ptr'|\<^sub>r = children'" using a3 by (meson select_result_I2) then have "ptr \ set (sorted_list_of_set (fset (object_ptr_kinds h))) \ ptr' \ set (sorted_list_of_set (fset (object_ptr_kinds h))) \ set children \ set children' = {}" using f5 a4 a1 by (meson distinct_concat_map_E(1)) then show False using a3 a2 by (metis (no_types) assms(4) finite_fset fmember.rep_eq is_OK_returns_result_I local.get_child_nodes_ptr_in_heap set_sorted_list_of_set) qed lemma parent_child_rel_child: "h \ get_child_nodes ptr \\<^sub>r children \ child \ set children \ (ptr, cast child) \ parent_child_rel h" by (simp add: is_OK_returns_result_I get_child_nodes_ptr_in_heap parent_child_rel_def) lemma parent_child_rel_acyclic: "heap_is_wellformed h \ acyclic (parent_child_rel h)" by (simp add: acyclic_heap_def local.heap_is_wellformed_def) lemma heap_is_wellformed_disconnected_nodes_distinct: "heap_is_wellformed h \ h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes \ distinct disc_nodes" using distinct_lists_disconnected_nodes local.heap_is_wellformed_def by blast lemma parent_child_rel_parent_in_heap: "(parent, child_ptr) \ parent_child_rel h \ parent |\| object_ptr_kinds h" using local.parent_child_rel_def by blast lemma parent_child_rel_child_in_heap: "heap_is_wellformed h \ type_wf h \ known_ptr parent \ (parent, child_ptr) \ parent_child_rel h \ child_ptr |\| object_ptr_kinds h" apply(auto simp add: heap_is_wellformed_def parent_child_rel_def a_all_ptrs_in_heap_def)[1] using get_child_nodes_ok by (meson finite_set_in subsetD) lemma heap_is_wellformed_disc_nodes_in_heap: "heap_is_wellformed h \ h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes \ node \ set disc_nodes \ node |\| node_ptr_kinds h" by (metis (no_types, lifting) finite_set_in is_OK_returns_result_I local.a_all_ptrs_in_heap_def local.get_disconnected_nodes_ptr_in_heap local.heap_is_wellformed_def select_result_I2 subsetD) lemma heap_is_wellformed_one_disc_parent: "heap_is_wellformed h \ h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes \ h \ get_disconnected_nodes document_ptr' \\<^sub>r disc_nodes' \ set disc_nodes \ set disc_nodes' \ {} \ document_ptr = document_ptr'" using DocumentMonad.ptr_kinds_ptr_kinds_M concat_append distinct_append distinct_concat_map_E(1) is_OK_returns_result_I local.a_distinct_lists_def local.get_disconnected_nodes_ptr_in_heap local.heap_is_wellformed_def select_result_I2 proof - assume a1: "heap_is_wellformed h" assume a2: "h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes" assume a3: "h \ get_disconnected_nodes document_ptr' \\<^sub>r disc_nodes'" assume a4: "set disc_nodes \ set disc_nodes' \ {}" have f5: "|h \ get_disconnected_nodes document_ptr|\<^sub>r = disc_nodes" using a2 by (meson select_result_I2) have f6: "|h \ get_disconnected_nodes document_ptr'|\<^sub>r = disc_nodes'" using a3 by (meson select_result_I2) have "\nss nssa. \ distinct (concat (nss @ nssa)) \ distinct (concat nssa::(_) node_ptr list)" by (metis (no_types) concat_append distinct_append) then have "distinct (concat (map (\d. |h \ get_disconnected_nodes d|\<^sub>r) |h \ document_ptr_kinds_M|\<^sub>r))" using a1 local.a_distinct_lists_def local.heap_is_wellformed_def by blast then show ?thesis using f6 f5 a4 a3 a2 by (meson DocumentMonad.ptr_kinds_ptr_kinds_M distinct_concat_map_E(1) is_OK_returns_result_I local.get_disconnected_nodes_ptr_in_heap) qed lemma heap_is_wellformed_children_disc_nodes_different: "heap_is_wellformed h \ h \ get_child_nodes ptr \\<^sub>r children \ h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes \ set children \ set disc_nodes = {}" by (metis (no_types, hide_lams) disjoint_iff_not_equal distinct_lists_no_parent is_OK_returns_result_I local.get_child_nodes_ptr_in_heap local.heap_is_wellformed_def select_result_I2) lemma heap_is_wellformed_children_disc_nodes: "heap_is_wellformed h \ node_ptr |\| node_ptr_kinds h \ \(\parent \ fset (object_ptr_kinds h). node_ptr \ set |h \ get_child_nodes parent|\<^sub>r) \ (\document_ptr \ fset (document_ptr_kinds h). node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r)" apply(auto simp add: heap_is_wellformed_def a_distinct_lists_def a_owner_document_valid_def)[1] by (meson fmember.rep_eq) lemma heap_is_wellformed_children_distinct: "heap_is_wellformed h \ h \ get_child_nodes ptr \\<^sub>r children \ distinct children" by (metis (no_types, lifting) ObjectMonad.ptr_kinds_ptr_kinds_M concat_append distinct_append distinct_concat_map_E(2) is_OK_returns_result_I local.a_distinct_lists_def local.get_child_nodes_ptr_in_heap local.heap_is_wellformed_def select_result_I2) end locale l_heap_is_wellformed = l_type_wf + l_known_ptr + l_heap_is_wellformed_defs + l_get_child_nodes_defs + l_get_disconnected_nodes_defs + assumes heap_is_wellformed_children_in_heap: "heap_is_wellformed h \ h \ get_child_nodes ptr \\<^sub>r children \ child \ set children \ child |\| node_ptr_kinds h" assumes heap_is_wellformed_disc_nodes_in_heap: "heap_is_wellformed h \ h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes \ node \ set disc_nodes \ node |\| node_ptr_kinds h" assumes heap_is_wellformed_one_parent: "heap_is_wellformed h \ h \ get_child_nodes ptr \\<^sub>r children \ h \ get_child_nodes ptr' \\<^sub>r children' \ set children \ set children' \ {} \ ptr = ptr'" assumes heap_is_wellformed_one_disc_parent: "heap_is_wellformed h \ h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes \ h \ get_disconnected_nodes document_ptr' \\<^sub>r disc_nodes' \ set disc_nodes \ set disc_nodes' \ {} \ document_ptr = document_ptr'" assumes heap_is_wellformed_children_disc_nodes_different: "heap_is_wellformed h \ h \ get_child_nodes ptr \\<^sub>r children \ h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes \ set children \ set disc_nodes = {}" assumes heap_is_wellformed_disconnected_nodes_distinct: "heap_is_wellformed h \ h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes \ distinct disc_nodes" assumes heap_is_wellformed_children_distinct: "heap_is_wellformed h \ h \ get_child_nodes ptr \\<^sub>r children \ distinct children" assumes heap_is_wellformed_children_disc_nodes: "heap_is_wellformed h \ node_ptr |\| node_ptr_kinds h \ \(\parent \ fset (object_ptr_kinds h). node_ptr \ set |h \ get_child_nodes parent|\<^sub>r) \ (\document_ptr \ fset (document_ptr_kinds h). node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r)" assumes parent_child_rel_child: "h \ get_child_nodes ptr \\<^sub>r children \ child \ set children \ (ptr, cast child) \ parent_child_rel h" assumes parent_child_rel_finite: "heap_is_wellformed h \ finite (parent_child_rel h)" assumes parent_child_rel_acyclic: "heap_is_wellformed h \ acyclic (parent_child_rel h)" assumes parent_child_rel_node_ptr: "(parent, child_ptr) \ parent_child_rel h \ is_node_ptr_kind child_ptr" assumes parent_child_rel_parent_in_heap: "(parent, child_ptr) \ parent_child_rel h \ parent |\| object_ptr_kinds h" assumes parent_child_rel_child_in_heap: "heap_is_wellformed h \ type_wf h \ known_ptr parent \ (parent, child_ptr) \ parent_child_rel h \ child_ptr |\| object_ptr_kinds h" interpretation i_heap_is_wellformed?: l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed parent_child_rel apply(unfold_locales) by(auto simp add: heap_is_wellformed_def parent_child_rel_def) declare l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma heap_is_wellformed_is_l_heap_is_wellformed [instances]: "l_heap_is_wellformed type_wf known_ptr heap_is_wellformed parent_child_rel get_child_nodes get_disconnected_nodes" apply(auto simp add: l_heap_is_wellformed_def)[1] using heap_is_wellformed_children_in_heap apply blast using heap_is_wellformed_disc_nodes_in_heap apply blast using heap_is_wellformed_one_parent apply blast using heap_is_wellformed_one_disc_parent apply blast using heap_is_wellformed_children_disc_nodes_different apply blast using heap_is_wellformed_disconnected_nodes_distinct apply blast using heap_is_wellformed_children_distinct apply blast using heap_is_wellformed_children_disc_nodes apply blast using parent_child_rel_child apply (blast) using parent_child_rel_child apply(blast) using parent_child_rel_finite apply blast using parent_child_rel_acyclic apply blast using parent_child_rel_node_ptr apply blast using parent_child_rel_parent_in_heap apply blast using parent_child_rel_child_in_heap apply blast done subsection \get\_parent\ locale l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs get_parent get_parent_locs + l_heap_is_wellformed type_wf known_ptr heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs for known_ptr :: "(_::linorder) object_ptr \ bool" and type_wf :: "(_) heap \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and known_ptrs :: "(_) heap \ bool" and get_parent :: "(_) node_ptr \ ((_) heap, exception, (_) object_ptr option) prog" and get_parent_locs :: "((_) heap \ (_) heap \ bool) set" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma child_parent_dual: assumes heap_is_wellformed: "heap_is_wellformed h" assumes "h \ get_child_nodes ptr \\<^sub>r children" assumes "child \ set children" assumes "known_ptrs h" assumes type_wf: "type_wf h" shows "h \ get_parent child \\<^sub>r Some ptr" proof - obtain ptrs where ptrs: "h \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have h1: "ptr \ set ptrs" using get_child_nodes_ok assms(2) is_OK_returns_result_I by (metis (no_types, hide_lams) ObjectMonad.ptr_kinds_ptr_kinds_M \\thesis. (\ptrs. h \ object_ptr_kinds_M \\<^sub>r ptrs \ thesis) \ thesis\ get_child_nodes_ptr_in_heap returns_result_eq select_result_I2) let ?P = "(\ptr. get_child_nodes ptr \ (\children. return (child \ set children)))" let ?filter = "filter_M ?P ptrs" have "h \ ok ?filter" using ptrs type_wf using get_child_nodes_ok apply(auto intro!: filter_M_is_OK_I bind_is_OK_pure_I get_child_nodes_ok simp add: bind_pure_I)[1] using assms(4) local.known_ptrs_known_ptr by blast then obtain parent_ptrs where parent_ptrs: "h \ ?filter \\<^sub>r parent_ptrs" by auto have h5: "\!x. x \ set ptrs \ h \ Heap_Error_Monad.bind (get_child_nodes x) (\children. return (child \ set children)) \\<^sub>r True" apply(auto intro!: bind_pure_returns_result_I)[1] using heap_is_wellformed_one_parent proof - have "h \ (return (child \ set children)::((_) heap, exception, bool) prog) \\<^sub>r True" by (simp add: assms(3)) then show "\z. z \ set ptrs \ h \ Heap_Error_Monad.bind (get_child_nodes z) (\ns. return (child \ set ns)) \\<^sub>r True" by (metis (no_types) assms(2) bind_pure_returns_result_I2 h1 is_OK_returns_result_I local.get_child_nodes_pure select_result_I2) next fix x y assume 0: "x \ set ptrs" and 1: "h \ Heap_Error_Monad.bind (get_child_nodes x) (\children. return (child \ set children)) \\<^sub>r True" and 2: "y \ set ptrs" and 3: "h \ Heap_Error_Monad.bind (get_child_nodes y) (\children. return (child \ set children)) \\<^sub>r True" and 4: "(\h ptr children ptr' children'. heap_is_wellformed h \ h \ get_child_nodes ptr \\<^sub>r children \ h \ get_child_nodes ptr' \\<^sub>r children' \ set children \ set children' \ {} \ ptr = ptr')" then show "x = y" by (metis (no_types, lifting) bind_returns_result_E disjoint_iff_not_equal heap_is_wellformed return_returns_result) qed have "child |\| node_ptr_kinds h" using heap_is_wellformed_children_in_heap heap_is_wellformed assms(2) assms(3) by fast moreover have "parent_ptrs = [ptr]" apply(rule filter_M_ex1[OF parent_ptrs h1 h5]) using ptrs assms(2) assms(3) by(auto simp add: object_ptr_kinds_M_defs bind_pure_I intro!: bind_pure_returns_result_I) ultimately show ?thesis using ptrs parent_ptrs by(auto simp add: bind_pure_I get_parent_def elim!: bind_returns_result_E2 intro!: bind_pure_returns_result_I filter_M_pure_I) (*slow, ca 1min *) qed lemma parent_child_rel_parent: assumes "heap_is_wellformed h" and "h \ get_parent child_node \\<^sub>r Some parent" shows "(parent, cast child_node) \ parent_child_rel h" using assms parent_child_rel_child get_parent_child_dual by auto lemma heap_wellformed_induct [consumes 1, case_names step]: assumes "heap_is_wellformed h" and step: "\parent. (\children child. h \ get_child_nodes parent \\<^sub>r children \ child \ set children \ P (cast child)) \ P parent" shows "P ptr" proof - fix ptr have "wf ((parent_child_rel h)\)" by (simp add: assms(1) finite_acyclic_wf_converse parent_child_rel_acyclic parent_child_rel_finite) then show "?thesis" proof (induct rule: wf_induct_rule) case (less parent) then show ?case using assms parent_child_rel_child by (meson converse_iff) qed qed lemma heap_wellformed_induct2 [consumes 3, case_names not_in_heap empty_children step]: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and not_in_heap: "\parent. parent |\| object_ptr_kinds h \ P parent" and empty_children: "\parent. h \ get_child_nodes parent \\<^sub>r [] \ P parent" and step: "\parent children child. h \ get_child_nodes parent \\<^sub>r children \ child \ set children \ P (cast child) \ P parent" shows "P ptr" proof(insert assms(1), induct rule: heap_wellformed_induct) case (step parent) then show ?case proof(cases "parent |\| object_ptr_kinds h") case True then obtain children where children: "h \ get_child_nodes parent \\<^sub>r children" using get_child_nodes_ok assms(2) assms(3) by (meson is_OK_returns_result_E local.known_ptrs_known_ptr) then show ?thesis proof (cases "children = []") case True then show ?thesis using children empty_children by simp next case False then show ?thesis using assms(6) children last_in_set step.hyps by blast qed next case False then show ?thesis by (simp add: not_in_heap) qed qed lemma heap_wellformed_induct_rev [consumes 1, case_names step]: assumes "heap_is_wellformed h" and step: "\child. (\parent child_node. cast child_node = child \ h \ get_parent child_node \\<^sub>r Some parent \ P parent) \ P child" shows "P ptr" proof - fix ptr have "wf ((parent_child_rel h))" by (simp add: assms(1) local.parent_child_rel_acyclic local.parent_child_rel_finite wf_iff_acyclic_if_finite) then show "?thesis" proof (induct rule: wf_induct_rule) case (less child) show ?case using assms get_parent_child_dual by (metis less.hyps parent_child_rel_parent) qed qed end interpretation i_get_parent_wf?: l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs get_parent get_parent_locs heap_is_wellformed parent_child_rel get_disconnected_nodes using instances by(simp add: l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs get_parent get_parent_locs heap_is_wellformed parent_child_rel get_disconnected_nodes get_disconnected_nodes_locs + l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed parent_child_rel for known_ptr :: "(_::linorder) object_ptr \ bool" and type_wf :: "(_) heap \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and known_ptrs :: "(_) heap \ bool" and get_parent :: "(_) node_ptr \ ((_) heap, exception, (_) object_ptr option) prog" and get_parent_locs :: "((_) heap \ (_) heap \ bool) set" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma preserves_wellformedness_writes_needed: assumes heap_is_wellformed: "heap_is_wellformed h" and "h \ f \\<^sub>h h'" and "writes SW f h h'" and preserved_get_child_nodes: "\h h' w. w \ SW \ h \ w \\<^sub>h h' \ \object_ptr. \r \ get_child_nodes_locs object_ptr. r h h'" and preserved_get_disconnected_nodes: "\h h' w. w \ SW \ h \ w \\<^sub>h h' \ \document_ptr. \r \ get_disconnected_nodes_locs document_ptr. r h h'" and preserved_object_pointers: "\h h' w. w \ SW \ h \ w \\<^sub>h h' \ \object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'" shows "heap_is_wellformed h'" proof - have object_ptr_kinds_eq3: "object_ptr_kinds h = object_ptr_kinds h'" using assms(2) assms(3) object_ptr_kinds_preserved preserved_object_pointers by blast then have object_ptr_kinds_eq: "\ptrs. h \ object_ptr_kinds_M \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" unfolding object_ptr_kinds_M_defs by simp then have object_ptr_kinds_eq2: "|h \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" using select_result_eq by force then have node_ptr_kinds_eq2: "|h \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by auto then have node_ptr_kinds_eq3: "node_ptr_kinds h = node_ptr_kinds h'" by auto have document_ptr_kinds_eq2: "|h \ document_ptr_kinds_M|\<^sub>r = |h' \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_eq2 document_ptr_kinds_M_eq by auto then have document_ptr_kinds_eq3: "document_ptr_kinds h = document_ptr_kinds h'" by auto have children_eq: "\ptr children. h \ get_child_nodes ptr \\<^sub>r children = h' \ get_child_nodes ptr \\<^sub>r children" apply(rule reads_writes_preserved[OF get_child_nodes_reads assms(3) assms(2)]) using preserved_get_child_nodes by fast then have children_eq2: "\ptr. |h \ get_child_nodes ptr|\<^sub>r = |h' \ get_child_nodes ptr|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq: "\document_ptr disconnected_nodes. h \ get_disconnected_nodes document_ptr \\<^sub>r disconnected_nodes = h' \ get_disconnected_nodes document_ptr \\<^sub>r disconnected_nodes" apply(rule reads_writes_preserved[OF get_disconnected_nodes_reads assms(3) assms(2)]) using preserved_get_disconnected_nodes by fast then have disconnected_nodes_eq2: "\document_ptr. |h \ get_disconnected_nodes document_ptr|\<^sub>r = |h' \ get_disconnected_nodes document_ptr|\<^sub>r" using select_result_eq by force have get_parent_eq: "\ptr parent. h \ get_parent ptr \\<^sub>r parent = h' \ get_parent ptr \\<^sub>r parent" apply(rule reads_writes_preserved[OF get_parent_reads assms(3) assms(2)]) using preserved_get_child_nodes preserved_object_pointers unfolding get_parent_locs_def by fast have "a_acyclic_heap h" using heap_is_wellformed by (simp add: heap_is_wellformed_def) have "parent_child_rel h' \ parent_child_rel h" proof fix x assume "x \ parent_child_rel h'" then show "x \ parent_child_rel h" by(simp add: parent_child_rel_def children_eq2 object_ptr_kinds_eq3) qed then have "a_acyclic_heap h'" using \a_acyclic_heap h\ acyclic_heap_def acyclic_subset by blast moreover have "a_all_ptrs_in_heap h" using heap_is_wellformed by (simp add: heap_is_wellformed_def) then have "a_all_ptrs_in_heap h'" by (simp add: children_eq2 disconnected_nodes_eq2 document_ptr_kinds_eq3 l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_all_ptrs_in_heap_def node_ptr_kinds_eq3 object_ptr_kinds_eq3) moreover have h0: "a_distinct_lists h" using heap_is_wellformed by (simp add: heap_is_wellformed_def) have h1: "map (\ptr. |h \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h))) = map (\ptr. |h' \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h')))" by (simp add: children_eq2 object_ptr_kinds_eq3) have h2: "map (\document_ptr. |h \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h))) = map (\document_ptr. |h' \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h')))" using disconnected_nodes_eq document_ptr_kinds_eq2 select_result_eq by force have "a_distinct_lists h'" using h0 by(simp add: a_distinct_lists_def h1 h2) moreover have "a_owner_document_valid h" using heap_is_wellformed by (simp add: heap_is_wellformed_def) then have "a_owner_document_valid h'" by(auto simp add: a_owner_document_valid_def children_eq2 disconnected_nodes_eq2 object_ptr_kinds_eq3 node_ptr_kinds_eq3 document_ptr_kinds_eq3) ultimately show ?thesis by (simp add: heap_is_wellformed_def) qed end interpretation i_get_parent_wf2?: l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs get_parent get_parent_locs heap_is_wellformed parent_child_rel get_disconnected_nodes get_disconnected_nodes_locs using l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms by (simp add: l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_get_parent_wf = l_type_wf + l_known_ptrs + l_heap_is_wellformed_defs + l_get_child_nodes_defs + l_get_parent_defs + assumes child_parent_dual: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_child_nodes ptr \\<^sub>r children \ child \ set children \ h \ get_parent child \\<^sub>r Some ptr" assumes heap_wellformed_induct [consumes 1, case_names step]: "heap_is_wellformed h \ (\parent. (\children child. h \ get_child_nodes parent \\<^sub>r children \ child \ set children \ P (cast child)) \ P parent) \ P ptr" assumes heap_wellformed_induct_rev [consumes 1, case_names step]: "heap_is_wellformed h \ (\child. (\parent child_node. cast child_node = child \ h \ get_parent child_node \\<^sub>r Some parent \ P parent) \ P child) \ P ptr" assumes parent_child_rel_parent: "heap_is_wellformed h \ h \ get_parent child_node \\<^sub>r Some parent \ (parent, cast child_node) \ parent_child_rel h" lemma get_parent_wf_is_l_get_parent_wf [instances]: "l_get_parent_wf type_wf known_ptr known_ptrs heap_is_wellformed parent_child_rel get_child_nodes get_parent" using known_ptrs_is_l_known_ptrs apply(auto simp add: l_get_parent_wf_def l_get_parent_wf_axioms_def)[1] using child_parent_dual heap_wellformed_induct heap_wellformed_induct_rev parent_child_rel_parent by metis+ subsection \get\_disconnected\_nodes\ subsection \set\_disconnected\_nodes\ subsubsection \get\_disconnected\_nodes\ locale l_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_disconnected_nodes_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs + l_heap_is_wellformed type_wf known_ptr heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs for known_ptr :: "(_) object_ptr \ bool" and type_wf :: "(_) heap \ bool" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma remove_from_disconnected_nodes_removes: assumes "heap_is_wellformed h" assumes "h \ get_disconnected_nodes ptr \\<^sub>r disc_nodes" assumes "h \ set_disconnected_nodes ptr (remove1 node_ptr disc_nodes) \\<^sub>h h'" assumes "h' \ get_disconnected_nodes ptr \\<^sub>r disc_nodes'" shows "node_ptr \ set disc_nodes'" using assms by (metis distinct_remove1_removeAll heap_is_wellformed_disconnected_nodes_distinct set_disconnected_nodes_get_disconnected_nodes member_remove remove_code(1) returns_result_eq) end locale l_set_disconnected_nodes_get_disconnected_nodes_wf = l_heap_is_wellformed + l_set_disconnected_nodes_get_disconnected_nodes + assumes remove_from_disconnected_nodes_removes: "heap_is_wellformed h \ h \ get_disconnected_nodes ptr \\<^sub>r disc_nodes \ h \ set_disconnected_nodes ptr (remove1 node_ptr disc_nodes) \\<^sub>h h' \ h' \ get_disconnected_nodes ptr \\<^sub>r disc_nodes' \ node_ptr \ set disc_nodes'" interpretation i_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M?: l_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs heap_is_wellformed parent_child_rel get_child_nodes using instances by (simp add: l_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_disconnected_nodes_get_disconnected_nodes_wf_is_l_set_disconnected_nodes_get_disconnected_nodes_wf [instances]: "l_set_disconnected_nodes_get_disconnected_nodes_wf type_wf known_ptr heap_is_wellformed parent_child_rel get_child_nodes get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs" apply(auto simp add: l_set_disconnected_nodes_get_disconnected_nodes_wf_def l_set_disconnected_nodes_get_disconnected_nodes_wf_axioms_def instances)[1] using remove_from_disconnected_nodes_removes apply fast done subsection \get\_root\_node\ locale l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed type_wf known_ptr heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs + l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs get_parent get_parent_locs + l_get_parent_wf type_wf known_ptr known_ptrs heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs get_parent get_parent_locs + l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_ancestors get_ancestors_locs get_root_node get_root_node_locs for known_ptr :: "(_::linorder) object_ptr \ bool" and type_wf :: "(_) heap \ bool" and known_ptrs :: "(_) heap \ bool" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and get_parent :: "(_) node_ptr \ ((_) heap, exception, (_) object_ptr option) prog" and get_parent_locs :: "((_) heap \ (_) heap \ bool) set" and get_ancestors :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr list) prog" and get_ancestors_locs :: "((_) heap \ (_) heap \ bool) set" and get_root_node :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr) prog" and get_root_node_locs :: "((_) heap \ (_) heap \ bool) set" begin lemma get_ancestors_reads: assumes "heap_is_wellformed h" shows "reads get_ancestors_locs (get_ancestors node_ptr) h h'" proof (insert assms(1), induct rule: heap_wellformed_induct_rev) case (step child) then show ?case using [[simproc del: Product_Type.unit_eq]] get_parent_reads[unfolded reads_def] apply(simp (no_asm) add: get_ancestors_def) by(auto simp add: get_ancestors_locs_def reads_subset[OF return_reads] get_parent_reads_pointers intro!: reads_bind_pure reads_subset[OF check_in_heap_reads] reads_subset[OF get_parent_reads] reads_subset[OF get_child_nodes_reads] split: option.splits) qed lemma get_ancestors_ok: assumes "heap_is_wellformed h" and "ptr |\| object_ptr_kinds h" and "known_ptrs h" and type_wf: "type_wf h" shows "h \ ok (get_ancestors ptr)" proof (insert assms(1) assms(2), induct rule: heap_wellformed_induct_rev) case (step child) then show ?case using assms(3) assms(4) apply(simp (no_asm) add: get_ancestors_def) apply(simp add: assms(1) get_parent_parent_in_heap) by(auto intro!: bind_is_OK_pure_I bind_pure_I get_parent_ok split: option.splits) qed lemma get_root_node_ptr_in_heap: assumes "h \ ok (get_root_node ptr)" shows "ptr |\| object_ptr_kinds h" using assms unfolding get_root_node_def using get_ancestors_ptr_in_heap by auto lemma get_root_node_ok: assumes "heap_is_wellformed h" "known_ptrs h" "type_wf h" and "ptr |\| object_ptr_kinds h" shows "h \ ok (get_root_node ptr)" unfolding get_root_node_def using assms get_ancestors_ok by auto lemma get_ancestors_parent: assumes "heap_is_wellformed h" and "h \ get_parent child \\<^sub>r Some parent" shows "h \ get_ancestors (cast child) \\<^sub>r (cast child) # parent # ancestors \ h \ get_ancestors parent \\<^sub>r parent # ancestors" proof assume a1: "h \ get_ancestors (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child # parent # ancestors" then have "h \ Heap_Error_Monad.bind (check_in_heap (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child)) (\_. Heap_Error_Monad.bind (get_parent child) (\x. Heap_Error_Monad.bind (case x of None \ return [] | Some x \ get_ancestors x) (\ancestors. return (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child # ancestors)))) \\<^sub>r cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child # parent # ancestors" by(simp add: get_ancestors_def) then show "h \ get_ancestors parent \\<^sub>r parent # ancestors" using assms(2) apply(auto elim!: bind_returns_result_E2 split: option.splits)[1] using returns_result_eq by fastforce next assume "h \ get_ancestors parent \\<^sub>r parent # ancestors" then show "h \ get_ancestors (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child # parent # ancestors" using assms(2) apply(simp (no_asm) add: get_ancestors_def) apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1] by (metis (full_types) assms(2) check_in_heap_ptr_in_heap is_OK_returns_result_I local.get_parent_ptr_in_heap node_ptr_kinds_commutes old.unit.exhaust select_result_I) qed lemma get_ancestors_never_empty: assumes "heap_is_wellformed h" and "h \ get_ancestors child \\<^sub>r ancestors" shows "ancestors \ []" proof(insert assms(2), induct arbitrary: ancestors rule: heap_wellformed_induct_rev[OF assms(1)]) case (1 child) then show ?case proof (induct "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child") case None then show ?case apply(simp add: get_ancestors_def) by(auto elim!: bind_returns_result_E2 split: option.splits) next case (Some child_node) then obtain parent_opt where parent_opt: "h \ get_parent child_node \\<^sub>r parent_opt" apply(simp add: get_ancestors_def) by(auto elim!: bind_returns_result_E2 split: option.splits) with Some show ?case proof(induct parent_opt) case None then show ?case apply(simp add: get_ancestors_def) by(auto elim!: bind_returns_result_E2 split: option.splits) next case (Some option) then show ?case apply(simp add: get_ancestors_def) by(auto elim!: bind_returns_result_E2 split: option.splits) qed qed qed lemma get_ancestors_subset: assumes "heap_is_wellformed h" and "h \ get_ancestors ptr \\<^sub>r ancestors" and "ancestor \ set ancestors" and "h \ get_ancestors ancestor \\<^sub>r ancestor_ancestors" and type_wf: "type_wf h" and known_ptrs: "known_ptrs h" shows "set ancestor_ancestors \ set ancestors" proof (insert assms(1) assms(2) assms(3), induct ptr arbitrary: ancestors rule: heap_wellformed_induct_rev) case (step child) have "child |\| object_ptr_kinds h" using get_ancestors_ptr_in_heap step(2) by auto (* then have "h \ check_in_heap child \\<^sub>r ()" using returns_result_select_result by force *) show ?case proof (induct "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child") case None then have "ancestors = [child]" using step(2) step(3) by(auto simp add: get_ancestors_def elim!: bind_returns_result_E2) show ?case using step(2) step(3) apply(auto simp add: \ancestors = [child]\)[1] using assms(4) returns_result_eq by fastforce next case (Some child_node) note s1 = Some obtain parent_opt where parent_opt: "h \ get_parent child_node \\<^sub>r parent_opt" using \child |\| object_ptr_kinds h\ assms(1) Some[symmetric] get_parent_ok[OF type_wf known_ptrs] by (metis (no_types, lifting) is_OK_returns_result_E known_ptrs get_parent_ok l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms node_ptr_casts_commute node_ptr_kinds_commutes) then show ?case proof (induct parent_opt) case None then have "ancestors = [child]" using step(2) step(3) s1 apply(simp add: get_ancestors_def) by(auto elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq) show ?case using step(2) step(3) apply(auto simp add: \ancestors = [child]\)[1] using assms(4) returns_result_eq by fastforce next case (Some parent) have "h \ Heap_Error_Monad.bind (check_in_heap child) (\_. Heap_Error_Monad.bind (case cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child of None \ return [] | Some node_ptr \ Heap_Error_Monad.bind (get_parent node_ptr) (\parent_ptr_opt. case parent_ptr_opt of None \ return [] | Some x \ get_ancestors x)) (\ancestors. return (child # ancestors))) \\<^sub>r ancestors" using step(2) by(simp add: get_ancestors_def) moreover obtain tl_ancestors where tl_ancestors: "ancestors = child # tl_ancestors" using calculation by(auto elim!: bind_returns_result_E2 split: option.splits) ultimately have "h \ get_ancestors parent \\<^sub>r tl_ancestors" using s1 Some by(auto elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq) show ?case using step(1)[OF s1[symmetric, simplified] Some \h \ get_ancestors parent \\<^sub>r tl_ancestors\] step(3) apply(auto simp add: tl_ancestors)[1] by (metis assms(4) insert_iff list.simps(15) local.step(2) returns_result_eq tl_ancestors) qed qed qed lemma get_ancestors_also_parent: assumes "heap_is_wellformed h" and "h \ get_ancestors some_ptr \\<^sub>r ancestors" and "cast child \ set ancestors" and "h \ get_parent child \\<^sub>r Some parent" and type_wf: "type_wf h" and known_ptrs: "known_ptrs h" shows "parent \ set ancestors" proof - obtain child_ancestors where child_ancestors: "h \ get_ancestors (cast child) \\<^sub>r child_ancestors" by (meson assms(1) assms(4) get_ancestors_ok is_OK_returns_result_I known_ptrs local.get_parent_ptr_in_heap node_ptr_kinds_commutes returns_result_select_result type_wf) then have "parent \ set child_ancestors" apply(simp add: get_ancestors_def) by(auto elim!: bind_returns_result_E2 split: option.splits dest!: returns_result_eq[OF assms(4)] get_ancestors_ptr) then show ?thesis using assms child_ancestors get_ancestors_subset by blast qed lemma get_ancestors_obtains_children: assumes "heap_is_wellformed h" and "ancestor \ ptr" and "ancestor \ set ancestors" and "h \ get_ancestors ptr \\<^sub>r ancestors" and type_wf: "type_wf h" and known_ptrs: "known_ptrs h" obtains children ancestor_child where "h \ get_child_nodes ancestor \\<^sub>r children" and "ancestor_child \ set children" and "cast ancestor_child \ set ancestors" proof - assume 0: "(\children ancestor_child. h \ get_child_nodes ancestor \\<^sub>r children \ ancestor_child \ set children \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ancestor_child \ set ancestors \ thesis)" have "\child. h \ get_parent child \\<^sub>r Some ancestor \ cast child \ set ancestors" proof (insert assms(1) assms(2) assms(3) assms(4), induct ptr arbitrary: ancestors rule: heap_wellformed_induct_rev) case (step child) have "child |\| object_ptr_kinds h" using get_ancestors_ptr_in_heap step(4) by auto show ?case proof (induct "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child") case None then have "ancestors = [child]" using step(3) step(4) by(auto simp add: get_ancestors_def elim!: bind_returns_result_E2) show ?case using step(2) step(3) step(4) by(auto simp add: \ancestors = [child]\) next case (Some child_node) note s1 = Some obtain parent_opt where parent_opt: "h \ get_parent child_node \\<^sub>r parent_opt" using \child |\| object_ptr_kinds h\ assms(1) Some[symmetric] using get_parent_ok known_ptrs type_wf by (metis (no_types, lifting) is_OK_returns_result_E node_ptr_casts_commute node_ptr_kinds_commutes) then show ?case proof (induct parent_opt) case None then have "ancestors = [child]" using step(2) step(3) step(4) s1 apply(simp add: get_ancestors_def) by(auto elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq) show ?case using step(2) step(3) step(4) by(auto simp add: \ancestors = [child]\) next case (Some parent) have "h \ Heap_Error_Monad.bind (check_in_heap child) (\_. Heap_Error_Monad.bind (case cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child of None \ return [] | Some node_ptr \ Heap_Error_Monad.bind (get_parent node_ptr) (\parent_ptr_opt. case parent_ptr_opt of None \ return [] | Some x \ get_ancestors x)) (\ancestors. return (child # ancestors))) \\<^sub>r ancestors" using step(4) by(simp add: get_ancestors_def) moreover obtain tl_ancestors where tl_ancestors: "ancestors = child # tl_ancestors" using calculation by(auto elim!: bind_returns_result_E2 split: option.splits) ultimately have "h \ get_ancestors parent \\<^sub>r tl_ancestors" using s1 Some by(auto elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq) (* have "ancestor \ parent" *) have "ancestor \ set tl_ancestors" using tl_ancestors step(2) step(3) by auto show ?case proof (cases "ancestor \ parent") case True show ?thesis using step(1)[OF s1[symmetric, simplified] Some True \ancestor \ set tl_ancestors\ \h \ get_ancestors parent \\<^sub>r tl_ancestors\] using tl_ancestors by auto next case False have "child \ set ancestors" using step(4) get_ancestors_ptr by simp then show ?thesis using Some False s1[symmetric] by(auto) qed qed qed qed then obtain child where child: "h \ get_parent child \\<^sub>r Some ancestor" and in_ancestors: "cast child \ set ancestors" by auto then obtain children where children: "h \ get_child_nodes ancestor \\<^sub>r children" and child_in_children: "child \ set children" using get_parent_child_dual by blast show thesis using 0[OF children child_in_children] child assms(3) in_ancestors by blast qed lemma get_ancestors_parent_child_rel: assumes "heap_is_wellformed h" and "h \ get_ancestors child \\<^sub>r ancestors" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "(ptr, child) \ (parent_child_rel h)\<^sup>* \ ptr \ set ancestors" proof (safe) assume 3: "(ptr, child) \ (parent_child_rel h)\<^sup>*" show "ptr \ set ancestors" proof (insert 3, induct ptr rule: heap_wellformed_induct[OF assms(1)]) case (1 ptr) then show ?case proof (cases "ptr = child") case True then show ?thesis by (metis (no_types, lifting) assms(2) bind_returns_result_E get_ancestors_def in_set_member member_rec(1) return_returns_result) next case False obtain ptr_child where ptr_child: "(ptr, ptr_child) \ (parent_child_rel h) \ (ptr_child, child) \ (parent_child_rel h)\<^sup>*" using converse_rtranclE[OF 1(2)] \ptr \ child\ by metis then obtain ptr_child_node where ptr_child_ptr_child_node: "ptr_child = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node" using ptr_child node_ptr_casts_commute3 parent_child_rel_node_ptr by (metis ) then obtain children where children: "h \ get_child_nodes ptr \\<^sub>r children" and ptr_child_node: "ptr_child_node \ set children" proof - assume a1: "\children. \h \ get_child_nodes ptr \\<^sub>r children; ptr_child_node \ set children\ \ thesis" have "ptr |\| object_ptr_kinds h" using local.parent_child_rel_parent_in_heap ptr_child by blast moreover have "ptr_child_node \ set |h \ get_child_nodes ptr|\<^sub>r" by (metis calculation known_ptrs local.get_child_nodes_ok local.known_ptrs_known_ptr local.parent_child_rel_child ptr_child ptr_child_ptr_child_node returns_result_select_result type_wf) ultimately show ?thesis using a1 get_child_nodes_ok type_wf known_ptrs by (meson local.known_ptrs_known_ptr returns_result_select_result) qed moreover have "(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node, child) \ (parent_child_rel h)\<^sup>*" using ptr_child ptr_child_ptr_child_node by auto ultimately have "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node \ set ancestors" using 1 by auto moreover have "h \ get_parent ptr_child_node \\<^sub>r Some ptr" using assms(1) children ptr_child_node child_parent_dual using known_ptrs type_wf by blast ultimately show ?thesis using get_ancestors_also_parent assms type_wf by blast qed qed next assume 3: "ptr \ set ancestors" show "(ptr, child) \ (parent_child_rel h)\<^sup>*" proof (insert 3, induct ptr rule: heap_wellformed_induct[OF assms(1)]) case (1 ptr) then show ?case proof (cases "ptr = child") case True then show ?thesis by simp next case False then obtain children ptr_child_node where children: "h \ get_child_nodes ptr \\<^sub>r children" and ptr_child_node: "ptr_child_node \ set children" and ptr_child_node_in_ancestors: "cast ptr_child_node \ set ancestors" using 1(2) assms(2) get_ancestors_obtains_children assms(1) using known_ptrs type_wf by blast then have "(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node, child) \ (parent_child_rel h)\<^sup>*" using 1(1) by blast moreover have "(ptr, cast ptr_child_node) \ parent_child_rel h" using children ptr_child_node assms(1) parent_child_rel_child_nodes2 using child_parent_dual known_ptrs parent_child_rel_parent type_wf by blast ultimately show ?thesis by auto qed qed qed lemma get_root_node_parent_child_rel: assumes "heap_is_wellformed h" and "h \ get_root_node child \\<^sub>r root" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "(root, child) \ (parent_child_rel h)\<^sup>*" using assms get_ancestors_parent_child_rel apply(auto simp add: get_root_node_def elim!: bind_returns_result_E2)[1] using get_ancestors_never_empty last_in_set by blast lemma get_ancestors_eq: assumes "heap_is_wellformed h" and "heap_is_wellformed h'" and "\object_ptr w. object_ptr \ ptr \ w \ get_child_nodes_locs object_ptr \ w h h'" and pointers_preserved: "\object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'" and known_ptrs: "known_ptrs h" and known_ptrs': "known_ptrs h'" and "h \ get_ancestors ptr \\<^sub>r ancestors" and type_wf: "type_wf h" and type_wf': "type_wf h'" shows "h' \ get_ancestors ptr \\<^sub>r ancestors" proof - have object_ptr_kinds_eq3: "object_ptr_kinds h = object_ptr_kinds h'" using pointers_preserved object_ptr_kinds_preserved_small by blast then have object_ptr_kinds_M_eq: "\ptrs. h \ object_ptr_kinds_M \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_eq: "|h \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" by(simp) have "h' \ ok (get_ancestors ptr)" using get_ancestors_ok get_ancestors_ptr_in_heap object_ptr_kinds_eq3 assms(1) known_ptrs known_ptrs' assms(2) assms(7) type_wf' by blast then obtain ancestors' where ancestors': "h' \ get_ancestors ptr \\<^sub>r ancestors'" by auto obtain root where root: "h \ get_root_node ptr \\<^sub>r root" proof - assume 0: "(\root. h \ get_root_node ptr \\<^sub>r root \ thesis)" show thesis apply(rule 0) using assms(7) by(auto simp add: get_root_node_def elim!: bind_returns_result_E2 split: option.splits) qed have children_eq: "\p children. p \ ptr \ h \ get_child_nodes p \\<^sub>r children = h' \ get_child_nodes p \\<^sub>r children" using get_child_nodes_reads assms(3) apply(simp add: reads_def reflp_def transp_def preserved_def) by blast have "acyclic (parent_child_rel h)" using assms(1) local.parent_child_rel_acyclic by auto have "acyclic (parent_child_rel h')" using assms(2) local.parent_child_rel_acyclic by blast have 2: "\c parent_opt. cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c \ set ancestors \ set ancestors' \ h \ get_parent c \\<^sub>r parent_opt = h' \ get_parent c \\<^sub>r parent_opt" proof - fix c parent_opt assume 1: " cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c \ set ancestors \ set ancestors'" obtain ptrs where ptrs: "h \ object_ptr_kinds_M \\<^sub>r ptrs" by simp let ?P = "(\ptr. Heap_Error_Monad.bind (get_child_nodes ptr) (\children. return (c \ set children)))" have children_eq_True: "\p. p \ set ptrs \ h \ ?P p \\<^sub>r True \ h' \ ?P p \\<^sub>r True" proof - fix p assume "p \ set ptrs" then show "h \ ?P p \\<^sub>r True \ h' \ ?P p \\<^sub>r True" proof (cases "p = ptr") case True have "(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c, ptr) \ (parent_child_rel h)\<^sup>*" using get_ancestors_parent_child_rel 1 assms by blast then have "(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \ (parent_child_rel h)" proof (cases "cast c = ptr") case True then show ?thesis using \acyclic (parent_child_rel h)\ by(auto simp add: acyclic_def) next case False then have "(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \ (parent_child_rel h)\<^sup>*" using \acyclic (parent_child_rel h)\ False rtrancl_eq_or_trancl rtrancl_trancl_trancl \(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c, ptr) \ (parent_child_rel h)\<^sup>*\ by (metis acyclic_def) then show ?thesis using r_into_rtrancl by auto qed obtain children where children: "h \ get_child_nodes ptr \\<^sub>r children" using type_wf by (metis \h' \ ok get_ancestors ptr\ assms(1) get_ancestors_ptr_in_heap get_child_nodes_ok heap_is_wellformed_def is_OK_returns_result_E known_ptrs local.known_ptrs_known_ptr object_ptr_kinds_eq3) then have "c \ set children" using \(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \ (parent_child_rel h)\ assms(1) using parent_child_rel_child_nodes2 using child_parent_dual known_ptrs parent_child_rel_parent type_wf by blast with children have "h \ ?P p \\<^sub>r False" by(auto simp add: True) moreover have "(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c, ptr) \ (parent_child_rel h')\<^sup>*" using get_ancestors_parent_child_rel assms(2) ancestors' 1 known_ptrs' type_wf type_wf' by blast then have "(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \ (parent_child_rel h')" proof (cases "cast c = ptr") case True then show ?thesis using \acyclic (parent_child_rel h')\ by(auto simp add: acyclic_def) next case False then have "(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \ (parent_child_rel h')\<^sup>*" using \acyclic (parent_child_rel h')\ False rtrancl_eq_or_trancl rtrancl_trancl_trancl \(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c, ptr) \ (parent_child_rel h')\<^sup>*\ by (metis acyclic_def) then show ?thesis using r_into_rtrancl by auto qed then have "(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \ (parent_child_rel h')" using r_into_rtrancl by auto obtain children' where children': "h' \ get_child_nodes ptr \\<^sub>r children'" using type_wf type_wf' by (meson \h' \ ok (get_ancestors ptr)\ assms(2) get_ancestors_ptr_in_heap get_child_nodes_ok is_OK_returns_result_E known_ptrs' local.known_ptrs_known_ptr) then have "c \ set children'" using \(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \ (parent_child_rel h')\ assms(2) type_wf type_wf' using parent_child_rel_child_nodes2 child_parent_dual known_ptrs' parent_child_rel_parent by auto with children' have "h' \ ?P p \\<^sub>r False" by(auto simp add: True) ultimately show ?thesis by (metis returns_result_eq) next case False then show ?thesis using children_eq ptrs by (metis (no_types, lifting) bind_pure_returns_result_I bind_returns_result_E get_child_nodes_pure return_returns_result) qed qed have "\pa. pa \ set ptrs \ h \ ok (get_child_nodes pa \ (\children. return (c \ set children))) = h' \ ok ( get_child_nodes pa \ (\children. return (c \ set children)))" using assms(1) assms(2) object_ptr_kinds_eq ptrs type_wf type_wf' by (metis (no_types, lifting) ObjectMonad.ptr_kinds_ptr_kinds_M bind_is_OK_pure_I get_child_nodes_ok get_child_nodes_pure known_ptrs' local.known_ptrs_known_ptr return_ok select_result_I2) have children_eq_False: "\pa. pa \ set ptrs \ h \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r False = h' \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r False" proof fix pa assume "pa \ set ptrs" and "h \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r False" have "h \ ok (get_child_nodes pa \ (\children. return (c \ set children))) \ h' \ ok ( get_child_nodes pa \ (\children. return (c \ set children)))" using \pa \ set ptrs\ \\pa. pa \ set ptrs \ h \ ok (get_child_nodes pa \ (\children. return (c \ set children))) = h' \ ok ( get_child_nodes pa \ (\children. return (c \ set children)))\ by auto moreover have "h \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r False \ h' \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r False" by (metis (mono_tags, lifting) \\pa. pa \ set ptrs \ h \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r True = h' \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r True\ \pa \ set ptrs\ calculation is_OK_returns_result_I returns_result_eq returns_result_select_result) ultimately show "h' \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r False" using \h \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r False\ by auto next fix pa assume "pa \ set ptrs" and "h' \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r False" have "h' \ ok (get_child_nodes pa \ (\children. return (c \ set children))) \ h \ ok ( get_child_nodes pa \ (\children. return (c \ set children)))" using \pa \ set ptrs\ \\pa. pa \ set ptrs \ h \ ok (get_child_nodes pa \ (\children. return (c \ set children))) = h' \ ok ( get_child_nodes pa \ (\children. return (c \ set children)))\ by auto moreover have "h' \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r False \ h \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r False" by (metis (mono_tags, lifting) \\pa. pa \ set ptrs \ h \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r True = h' \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r True\ \pa \ set ptrs\ calculation is_OK_returns_result_I returns_result_eq returns_result_select_result) ultimately show "h \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r False" using \h' \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r False\ by blast qed have filter_eq: "\xs. h \ filter_M ?P ptrs \\<^sub>r xs = h' \ filter_M ?P ptrs \\<^sub>r xs" proof (rule filter_M_eq) show "\xs x. pure (Heap_Error_Monad.bind (get_child_nodes x) (\children. return (c \ set children))) h" by(auto intro!: bind_pure_I) next show "\xs x. pure (Heap_Error_Monad.bind (get_child_nodes x) (\children. return (c \ set children))) h'" by(auto intro!: bind_pure_I) next fix xs b x assume 0: "x \ set ptrs" then show "h \ Heap_Error_Monad.bind (get_child_nodes x) (\children. return (c \ set children)) \\<^sub>r b = h' \ Heap_Error_Monad.bind (get_child_nodes x) (\children. return (c \ set children)) \\<^sub>r b" apply(induct b) using children_eq_True apply blast using children_eq_False apply blast done qed show "h \ get_parent c \\<^sub>r parent_opt = h' \ get_parent c \\<^sub>r parent_opt" apply(simp add: get_parent_def) apply(rule bind_cong_2) apply(simp) apply(simp) apply(simp add: check_in_heap_def node_ptr_kinds_def object_ptr_kinds_eq3) apply(rule bind_cong_2) apply(auto simp add: object_ptr_kinds_M_eq object_ptr_kinds_eq3)[1] apply(auto simp add: object_ptr_kinds_M_eq object_ptr_kinds_eq3)[1] apply(auto simp add: object_ptr_kinds_M_eq object_ptr_kinds_eq3)[1] apply(rule bind_cong_2) apply(auto intro!: filter_M_pure_I bind_pure_I)[1] apply(auto intro!: filter_M_pure_I bind_pure_I)[1] apply(auto simp add: filter_eq (* dest!: returns_result_eq[OF ptrs] *))[1] using filter_eq ptrs apply auto[1] using filter_eq ptrs by auto qed have "ancestors = ancestors'" proof(insert assms(1) assms(7) ancestors' 2, induct ptr arbitrary: ancestors ancestors' rule: heap_wellformed_induct_rev) case (step child) show ?case using step(2) step(3) step(4) apply(simp add: get_ancestors_def) apply(auto intro!: elim!: bind_returns_result_E2 split: option.splits)[1] using returns_result_eq apply fastforce apply (meson option.simps(3) returns_result_eq) by (metis IntD1 IntD2 option.inject returns_result_eq step.hyps) qed then show ?thesis using assms(5) ancestors' by simp qed lemma get_ancestors_remains_not_in_ancestors: assumes "heap_is_wellformed h" and "heap_is_wellformed h'" and "h \ get_ancestors ptr \\<^sub>r ancestors" and "h' \ get_ancestors ptr \\<^sub>r ancestors'" and "\p children children'. h \ get_child_nodes p \\<^sub>r children \ h' \ get_child_nodes p \\<^sub>r children' \ set children' \ set children" and "node \ set ancestors" and object_ptr_kinds_eq3: "object_ptr_kinds h = object_ptr_kinds h'" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" and type_wf': "type_wf h'" shows "node \ set ancestors'" proof - have object_ptr_kinds_M_eq: "\ptrs. h \ object_ptr_kinds_M \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" using object_ptr_kinds_eq3 by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_eq: "|h \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" by(simp) show ?thesis proof (insert assms(1) assms(3) assms(4) assms(6), induct ptr arbitrary: ancestors ancestors' rule: heap_wellformed_induct_rev) case (step child) have 1: "\p parent. h' \ get_parent p \\<^sub>r Some parent \ h \ get_parent p \\<^sub>r Some parent" proof - fix p parent assume "h' \ get_parent p \\<^sub>r Some parent" then obtain children' where children': "h' \ get_child_nodes parent \\<^sub>r children'" and p_in_children': "p \ set children'" using get_parent_child_dual by blast obtain children where children: "h \ get_child_nodes parent \\<^sub>r children" using get_child_nodes_ok assms(1) get_child_nodes_ptr_in_heap object_ptr_kinds_eq children' known_ptrs using type_wf type_wf' by (metis \h' \ get_parent p \\<^sub>r Some parent\ get_parent_parent_in_heap is_OK_returns_result_E local.known_ptrs_known_ptr object_ptr_kinds_eq3) have "p \ set children" using assms(5) children children' p_in_children' by blast then show "h \ get_parent p \\<^sub>r Some parent" using child_parent_dual assms(1) children known_ptrs type_wf by blast qed have "node \ child" using assms(1) get_ancestors_parent_child_rel step.prems(1) step.prems(3) known_ptrs using type_wf type_wf' by blast then show ?case using step(2) step(3) apply(simp add: get_ancestors_def) using step(4) apply(auto elim!: bind_returns_result_E2 split: option.splits)[1] using 1 apply (meson option.distinct(1) returns_result_eq) by (metis "1" option.inject returns_result_eq step.hyps) qed qed lemma get_ancestors_ptrs_in_heap: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_ancestors ptr \\<^sub>r ancestors" assumes "ptr' \ set ancestors" shows "ptr' |\| object_ptr_kinds h" proof (insert assms(4) assms(5), induct ancestors arbitrary: ptr) case Nil then show ?case by(auto) next case (Cons a ancestors) then obtain x where x: "h \ get_ancestors x \\<^sub>r a # ancestors" by(auto simp add: get_ancestors_def[of a] elim!: bind_returns_result_E2 split: option.splits) then have "x = a" by(auto simp add: get_ancestors_def[of x] elim!: bind_returns_result_E2 split: option.splits) then show ?case using Cons.hyps Cons.prems(2) get_ancestors_ptr_in_heap x by (metis assms(1) assms(2) assms(3) get_ancestors_obtains_children get_child_nodes_ptr_in_heap is_OK_returns_result_I) qed lemma get_ancestors_prefix: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_ancestors ptr \\<^sub>r ancestors" assumes "ptr' \ set ancestors" assumes "h \ get_ancestors ptr' \\<^sub>r ancestors'" shows "\pre. ancestors = pre @ ancestors'" proof (insert assms(1) assms(5) assms(6), induct ptr' arbitrary: ancestors' rule: heap_wellformed_induct) case (step parent) then show ?case proof (cases "parent \ ptr" ) case True then obtain children ancestor_child where "h \ get_child_nodes parent \\<^sub>r children" and "ancestor_child \ set children" and "cast ancestor_child \ set ancestors" using assms(1) assms(2) assms(3) assms(4) get_ancestors_obtains_children step.prems(1) by blast then have "h \ get_parent ancestor_child \\<^sub>r Some parent" using assms(1) assms(2) assms(3) child_parent_dual by blast then have "h \ get_ancestors (cast ancestor_child) \\<^sub>r cast ancestor_child # ancestors'" apply(simp add: get_ancestors_def) using \h \ get_ancestors parent \\<^sub>r ancestors'\ get_parent_ptr_in_heap by(auto simp add: check_in_heap_def is_OK_returns_result_I intro!: bind_pure_returns_result_I) then show ?thesis using step(1) \h \ get_child_nodes parent \\<^sub>r children\ \ancestor_child \ set children\ \cast ancestor_child \ set ancestors\ \h \ get_ancestors (cast ancestor_child) \\<^sub>r cast ancestor_child # ancestors'\ by fastforce next case False then show ?thesis by (metis append_Nil assms(4) returns_result_eq step.prems(2)) qed qed lemma get_ancestors_same_root_node: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_ancestors ptr \\<^sub>r ancestors" assumes "ptr' \ set ancestors" assumes "ptr'' \ set ancestors" shows "h \ get_root_node ptr' \\<^sub>r root_ptr \ h \ get_root_node ptr'' \\<^sub>r root_ptr" proof - have "ptr' |\| object_ptr_kinds h" by (metis assms(1) assms(2) assms(3) assms(4) assms(5) get_ancestors_obtains_children get_ancestors_ptr_in_heap get_child_nodes_ptr_in_heap is_OK_returns_result_I) then obtain ancestors' where ancestors': "h \ get_ancestors ptr' \\<^sub>r ancestors'" by (meson assms(1) assms(2) assms(3) get_ancestors_ok is_OK_returns_result_E) then have "\pre. ancestors = pre @ ancestors'" using get_ancestors_prefix assms by blast moreover have "ptr'' |\| object_ptr_kinds h" by (metis assms(1) assms(2) assms(3) assms(4) assms(6) get_ancestors_obtains_children get_ancestors_ptr_in_heap get_child_nodes_ptr_in_heap is_OK_returns_result_I) then obtain ancestors'' where ancestors'': "h \ get_ancestors ptr'' \\<^sub>r ancestors''" by (meson assms(1) assms(2) assms(3) get_ancestors_ok is_OK_returns_result_E) then have "\pre. ancestors = pre @ ancestors''" using get_ancestors_prefix assms by blast ultimately show ?thesis using ancestors' ancestors'' apply(auto simp add: get_root_node_def elim!: bind_returns_result_E2 intro!: bind_pure_returns_result_I)[1] apply (metis (no_types, lifting) assms(1) get_ancestors_never_empty last_appendR returns_result_eq) by (metis assms(1) get_ancestors_never_empty last_appendR returns_result_eq) qed lemma get_root_node_parent_same: assumes "h \ get_parent child \\<^sub>r Some ptr" shows "h \ get_root_node (cast child) \\<^sub>r root \ h \ get_root_node ptr \\<^sub>r root" proof assume 1: " h \ get_root_node (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r root" show "h \ get_root_node ptr \\<^sub>r root" using 1[unfolded get_root_node_def] assms apply(simp add: get_ancestors_def) apply(auto simp add: get_root_node_def dest: returns_result_eq elim!: bind_returns_result_E2 intro!: bind_pure_returns_result_I split: option.splits)[1] using returns_result_eq apply fastforce using get_ancestors_ptr by fastforce next assume 1: " h \ get_root_node ptr \\<^sub>r root" show "h \ get_root_node (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r root" apply(simp add: get_root_node_def) using assms 1 apply(simp add: get_ancestors_def) apply(auto simp add: get_root_node_def dest: returns_result_eq elim!: bind_returns_result_E2 intro!: bind_pure_returns_result_I split: option.splits)[1] apply (simp add: check_in_heap_def is_OK_returns_result_I) using get_ancestors_ptr get_parent_ptr_in_heap apply (simp add: is_OK_returns_result_I) by (meson list.distinct(1) list.set_cases local.get_ancestors_ptr) qed lemma get_root_node_same_no_parent: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_root_node ptr \\<^sub>r cast child" shows "h \ get_parent child \\<^sub>r None" proof (insert assms(1) assms(4), induct ptr rule: heap_wellformed_induct_rev) case (step c) then show ?case proof (cases "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r c") case None then have "c = cast child" using step(2) by(auto simp add: get_root_node_def get_ancestors_def[of c] elim!: bind_returns_result_E2) then show ?thesis using None by auto next case (Some child_node) note s = this then obtain parent_opt where parent_opt: "h \ get_parent child_node \\<^sub>r parent_opt" by (metis (no_types, lifting) assms(2) assms(3) get_root_node_ptr_in_heap is_OK_returns_result_I local.get_parent_ok node_ptr_casts_commute node_ptr_kinds_commutes returns_result_select_result step.prems) then show ?thesis proof(induct parent_opt) case None then show ?case using Some get_root_node_no_parent returns_result_eq step.prems by fastforce next case (Some parent) then show ?case using step s apply(auto simp add: get_root_node_def get_ancestors_def[of c] elim!: bind_returns_result_E2 split: option.splits list.splits)[1] using get_root_node_parent_same step.hyps step.prems by auto qed qed qed lemma get_root_node_not_node_same: assumes "ptr |\| object_ptr_kinds h" assumes "\is_node_ptr_kind ptr" shows "h \ get_root_node ptr \\<^sub>r ptr" using assms apply(simp add: get_root_node_def get_ancestors_def) by(auto simp add: get_root_node_def dest: returns_result_eq elim!: bind_returns_result_E2 intro!: bind_pure_returns_result_I split: option.splits) lemma get_root_node_root_in_heap: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_root_node ptr \\<^sub>r root" shows "root |\| object_ptr_kinds h" using assms apply(auto simp add: get_root_node_def elim!: bind_returns_result_E2)[1] by (simp add: get_ancestors_never_empty get_ancestors_ptrs_in_heap) lemma get_root_node_same_no_parent_parent_child_rel: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_root_node ptr' \\<^sub>r ptr'" shows "\(\p. (p, ptr') \ (parent_child_rel h))" by (metis (no_types, lifting) assms(1) assms(2) assms(3) assms(4) get_root_node_same_no_parent l_heap_is_wellformed.parent_child_rel_child local.child_parent_dual local.get_child_nodes_ok local.known_ptrs_known_ptr local.l_heap_is_wellformed_axioms local.parent_child_rel_node_ptr local.parent_child_rel_parent_in_heap node_ptr_casts_commute3 option.simps(3) returns_result_eq returns_result_select_result) end locale l_get_ancestors_wf = l_heap_is_wellformed_defs + l_known_ptrs + l_type_wf + l_get_ancestors_defs + l_get_child_nodes_defs + l_get_parent_defs + assumes get_ancestors_never_empty: "heap_is_wellformed h \ h \ get_ancestors child \\<^sub>r ancestors \ ancestors \ []" assumes get_ancestors_ok: "heap_is_wellformed h \ ptr |\| object_ptr_kinds h \ known_ptrs h \ type_wf h \ h \ ok (get_ancestors ptr)" assumes get_ancestors_reads: "heap_is_wellformed h \ reads get_ancestors_locs (get_ancestors node_ptr) h h'" assumes get_ancestors_ptrs_in_heap: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_ancestors ptr \\<^sub>r ancestors \ ptr' \ set ancestors \ ptr' |\| object_ptr_kinds h" assumes get_ancestors_remains_not_in_ancestors: "heap_is_wellformed h \ heap_is_wellformed h' \ h \ get_ancestors ptr \\<^sub>r ancestors \ h' \ get_ancestors ptr \\<^sub>r ancestors' \ (\p children children'. h \ get_child_nodes p \\<^sub>r children \ h' \ get_child_nodes p \\<^sub>r children' \ set children' \ set children) \ node \ set ancestors \ object_ptr_kinds h = object_ptr_kinds h' \ known_ptrs h \ type_wf h \ type_wf h' \ node \ set ancestors'" assumes get_ancestors_also_parent: "heap_is_wellformed h \ h \ get_ancestors some_ptr \\<^sub>r ancestors \ cast child_node \ set ancestors \ h \ get_parent child_node \\<^sub>r Some parent \ type_wf h \ known_ptrs h \ parent \ set ancestors" assumes get_ancestors_obtains_children: "heap_is_wellformed h \ ancestor \ ptr \ ancestor \ set ancestors \ h \ get_ancestors ptr \\<^sub>r ancestors \ type_wf h \ known_ptrs h \ (\children ancestor_child . h \ get_child_nodes ancestor \\<^sub>r children \ ancestor_child \ set children \ cast ancestor_child \ set ancestors \ thesis) \ thesis" assumes get_ancestors_parent_child_rel: "heap_is_wellformed h \ h \ get_ancestors child \\<^sub>r ancestors \ known_ptrs h \ type_wf h \ (ptr, child) \ (parent_child_rel h)\<^sup>* \ ptr \ set ancestors" locale l_get_root_node_wf = l_heap_is_wellformed_defs + l_get_root_node_defs + l_type_wf + l_known_ptrs + l_get_ancestors_defs + l_get_parent_defs + assumes get_root_node_ok: "heap_is_wellformed h \ known_ptrs h \ type_wf h \ ptr |\| object_ptr_kinds h \ h \ ok (get_root_node ptr)" assumes get_root_node_ptr_in_heap: "h \ ok (get_root_node ptr) \ ptr |\| object_ptr_kinds h" assumes get_root_node_root_in_heap: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_root_node ptr \\<^sub>r root \ root |\| object_ptr_kinds h" assumes get_ancestors_same_root_node: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_ancestors ptr \\<^sub>r ancestors \ ptr' \ set ancestors \ ptr'' \ set ancestors \ h \ get_root_node ptr' \\<^sub>r root_ptr \ h \ get_root_node ptr'' \\<^sub>r root_ptr" assumes get_root_node_same_no_parent: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_root_node ptr \\<^sub>r cast child \ h \ get_parent child \\<^sub>r None" assumes get_root_node_parent_same: "h \ get_parent child \\<^sub>r Some ptr \ h \ get_root_node (cast child) \\<^sub>r root \ h \ get_root_node ptr \\<^sub>r root" interpretation i_get_root_node_wf?: l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf known_ptrs heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_parent get_parent_locs get_ancestors get_ancestors_locs get_root_node get_root_node_locs using instances by(simp add: l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma get_ancestors_wf_is_l_get_ancestors_wf [instances]: "l_get_ancestors_wf heap_is_wellformed parent_child_rel known_ptr known_ptrs type_wf get_ancestors get_ancestors_locs get_child_nodes get_parent" using known_ptrs_is_l_known_ptrs apply(auto simp add: l_get_ancestors_wf_def l_get_ancestors_wf_axioms_def)[1] using get_ancestors_never_empty apply blast using get_ancestors_ok apply blast using get_ancestors_reads apply blast using get_ancestors_ptrs_in_heap apply blast using get_ancestors_remains_not_in_ancestors apply blast using get_ancestors_also_parent apply blast using get_ancestors_obtains_children apply blast using get_ancestors_parent_child_rel apply blast using get_ancestors_parent_child_rel apply blast done lemma get_root_node_wf_is_l_get_root_node_wf [instances]: "l_get_root_node_wf heap_is_wellformed get_root_node type_wf known_ptr known_ptrs get_ancestors get_parent" using known_ptrs_is_l_known_ptrs apply(auto simp add: l_get_root_node_wf_def l_get_root_node_wf_axioms_def)[1] using get_root_node_ok apply blast using get_root_node_ptr_in_heap apply blast using get_root_node_root_in_heap apply blast using get_ancestors_same_root_node apply(blast, blast) using get_root_node_same_no_parent apply blast using get_root_node_parent_same apply (blast, blast) done subsection \to\_tree\_order\ locale l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent + l_get_parent_wf + l_heap_is_wellformed (* l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M *) begin lemma to_tree_order_ptr_in_heap: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "h \ ok (to_tree_order ptr)" shows "ptr |\| object_ptr_kinds h" proof(insert assms(1) assms(4), induct rule: heap_wellformed_induct) case (step parent) then show ?case apply(auto simp add: to_tree_order_def[of parent] map_M_pure_I elim!: bind_is_OK_E3)[1] using get_child_nodes_ptr_in_heap by blast qed lemma to_tree_order_either_ptr_or_in_children: assumes "h \ to_tree_order ptr \\<^sub>r nodes" and "node \ set nodes" and "h \ get_child_nodes ptr \\<^sub>r children" and "node \ ptr" obtains child child_to where "child \ set children" and "h \ to_tree_order (cast child) \\<^sub>r child_to" and "node \ set child_to" proof - obtain treeorders where treeorders: "h \ map_M to_tree_order (map cast children) \\<^sub>r treeorders" using assms apply(auto simp add: to_tree_order_def elim!: bind_returns_result_E)[1] using pure_returns_heap_eq returns_result_eq by fastforce then have "node \ set (concat treeorders)" using assms[simplified to_tree_order_def] by(auto elim!: bind_returns_result_E4 dest: pure_returns_heap_eq) then obtain treeorder where "treeorder \ set treeorders" and node_in_treeorder: "node \ set treeorder" by auto then obtain child where "h \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r treeorder" and "child \ set children" using assms[simplified to_tree_order_def] treeorders by(auto elim!: map_M_pure_E2) then show ?thesis using node_in_treeorder returns_result_eq that by auto qed lemma to_tree_order_ptrs_in_heap: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ to_tree_order ptr \\<^sub>r to" assumes "ptr' \ set to" shows "ptr' |\| object_ptr_kinds h" proof(insert assms(1) assms(4) assms(5), induct ptr arbitrary: to rule: heap_wellformed_induct) case (step parent) have "parent |\| object_ptr_kinds h" using assms(1) assms(2) assms(3) step.prems(1) to_tree_order_ptr_in_heap by blast then obtain children where children: "h \ get_child_nodes parent \\<^sub>r children" by (meson assms(2) assms(3) get_child_nodes_ok is_OK_returns_result_E local.known_ptrs_known_ptr) then show ?case proof (cases "children = []") case True then have "to = [parent]" using step(2) children apply(auto simp add: to_tree_order_def[of parent] map_M_pure_I elim!: bind_returns_result_E2)[1] by (metis list.distinct(1) list.map_disc_iff list.set_cases map_M_pure_E2 returns_result_eq) then show ?thesis using \parent |\| object_ptr_kinds h\ step.prems(2) by auto next case False note f = this then show ?thesis using children step to_tree_order_either_ptr_or_in_children proof (cases "ptr' = parent") case True then show ?thesis using \parent |\| object_ptr_kinds h\ by blast next case False then show ?thesis using children step.hyps to_tree_order_either_ptr_or_in_children by (metis step.prems(1) step.prems(2)) qed qed qed lemma to_tree_order_ok: assumes wellformed: "heap_is_wellformed h" and "ptr |\| object_ptr_kinds h" and "known_ptrs h" and type_wf: "type_wf h" shows "h \ ok (to_tree_order ptr)" proof(insert assms(1) assms(2), induct rule: heap_wellformed_induct) case (step parent) then show ?case using assms(3) type_wf apply(simp add: to_tree_order_def) apply(auto simp add: heap_is_wellformed_def intro!: map_M_ok_I bind_is_OK_pure_I map_M_pure_I)[1] using get_child_nodes_ok known_ptrs_known_ptr apply blast by (simp add: local.heap_is_wellformed_children_in_heap local.to_tree_order_def wellformed) qed lemma to_tree_order_child_subset: assumes "heap_is_wellformed h" and "h \ to_tree_order ptr \\<^sub>r nodes" and "h \ get_child_nodes ptr \\<^sub>r children" and "node \ set children" and "h \ to_tree_order (cast node) \\<^sub>r nodes'" shows "set nodes' \ set nodes" proof fix x assume a1: "x \ set nodes'" moreover obtain treeorders where treeorders: "h \ map_M to_tree_order (map cast children) \\<^sub>r treeorders" using assms(2) assms(3) apply(auto simp add: to_tree_order_def elim!: bind_returns_result_E)[1] using pure_returns_heap_eq returns_result_eq by fastforce then have "nodes' \ set treeorders" using assms(4) assms(5) by(auto elim!: map_M_pure_E dest: returns_result_eq) moreover have "set (concat treeorders) \ set nodes" using treeorders assms(2) assms(3) by(auto simp add: to_tree_order_def elim!: bind_returns_result_E4 dest: pure_returns_heap_eq) ultimately show "x \ set nodes" by auto qed lemma to_tree_order_ptr_in_result: assumes "h \ to_tree_order ptr \\<^sub>r nodes" shows "ptr \ set nodes" using assms apply(simp add: to_tree_order_def) by(auto elim!: bind_returns_result_E2 intro!: map_M_pure_I bind_pure_I) lemma to_tree_order_subset: assumes "heap_is_wellformed h" and "h \ to_tree_order ptr \\<^sub>r nodes" and "node \ set nodes" and "h \ to_tree_order node \\<^sub>r nodes'" and "known_ptrs h" and type_wf: "type_wf h" shows "set nodes' \ set nodes" proof - have "\nodes. h \ to_tree_order ptr \\<^sub>r nodes \ (\node. node \ set nodes \ (\nodes'. h \ to_tree_order node \\<^sub>r nodes' \ set nodes' \ set nodes))" proof(insert assms(1), induct ptr rule: heap_wellformed_induct) case (step parent) then show ?case proof safe fix nodes node nodes' x assume 1: "(\children child. h \ get_child_nodes parent \\<^sub>r children \ child \ set children \ \nodes. h \ to_tree_order (cast child) \\<^sub>r nodes \ (\node. node \ set nodes \ (\nodes'. h \ to_tree_order node \\<^sub>r nodes' \ set nodes' \ set nodes)))" and 2: "h \ to_tree_order parent \\<^sub>r nodes" and 3: "node \ set nodes" and "h \ to_tree_order node \\<^sub>r nodes'" and "x \ set nodes'" have h1: "(\children child nodes node nodes'. h \ get_child_nodes parent \\<^sub>r children \ child \ set children \ h \ to_tree_order (cast child) \\<^sub>r nodes \ (node \ set nodes \ (h \ to_tree_order node \\<^sub>r nodes' \ set nodes' \ set nodes)))" using 1 by blast obtain children where children: "h \ get_child_nodes parent \\<^sub>r children" using 2 by(auto simp add: to_tree_order_def elim!: bind_returns_result_E) then have "set nodes' \ set nodes" proof (cases "children = []") case True then show ?thesis by (metis "2" "3" \h \ to_tree_order node \\<^sub>r nodes'\ children empty_iff list.set(1) subsetI to_tree_order_either_ptr_or_in_children) next case False then show ?thesis proof (cases "node = parent") case True then show ?thesis using "2" \h \ to_tree_order node \\<^sub>r nodes'\ returns_result_eq by fastforce next case False then obtain child nodes_of_child where "child \ set children" and "h \ to_tree_order (cast child) \\<^sub>r nodes_of_child" and "node \ set nodes_of_child" using 2[simplified to_tree_order_def] 3 to_tree_order_either_ptr_or_in_children[where node=node and ptr=parent] children apply(auto elim!: bind_returns_result_E2 intro: map_M_pure_I)[1] using is_OK_returns_result_E 2 a_all_ptrs_in_heap_def assms(1) heap_is_wellformed_def using "3" by blast then have "set nodes' \ set nodes_of_child" using h1 using \h \ to_tree_order node \\<^sub>r nodes'\ children by blast moreover have "set nodes_of_child \ set nodes" using "2" \child \ set children\ \h \ to_tree_order (cast child) \\<^sub>r nodes_of_child\ assms children to_tree_order_child_subset by auto ultimately show ?thesis by blast qed qed then show "x \ set nodes" using \x \ set nodes'\ by blast qed qed then show ?thesis using assms by blast qed lemma to_tree_order_parent: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ to_tree_order ptr \\<^sub>r nodes" assumes "h \ get_parent child \\<^sub>r Some parent" assumes "parent \ set nodes" shows "cast child \ set nodes" proof - obtain nodes' where nodes': "h \ to_tree_order parent \\<^sub>r nodes'" using assms to_tree_order_ok get_parent_parent_in_heap by (meson get_parent_parent_in_heap is_OK_returns_result_E) then have "set nodes' \ set nodes" using to_tree_order_subset assms by blast moreover obtain children where children: "h \ get_child_nodes parent \\<^sub>r children" and child: "child \ set children" using assms get_parent_child_dual by blast then obtain child_to where child_to: "h \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r child_to" by (meson assms(1) assms(2) assms(3) assms(5) is_OK_returns_result_E is_OK_returns_result_I get_parent_ptr_in_heap node_ptr_kinds_commutes to_tree_order_ok) then have "cast child \ set child_to" apply(simp add: to_tree_order_def) by(auto elim!: bind_returns_result_E2 map_M_pure_E dest!: bind_returns_result_E3[rotated, OF children, rotated] intro!: map_M_pure_I) have "cast child \ set nodes'" using nodes' child apply(simp add: to_tree_order_def) apply(auto elim!: bind_returns_result_E2 map_M_pure_E dest!: bind_returns_result_E3[rotated, OF children, rotated] intro!: map_M_pure_I)[1] using child_to \cast child \ set child_to\ returns_result_eq by fastforce ultimately show ?thesis by auto qed lemma to_tree_order_child: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ to_tree_order ptr \\<^sub>r nodes" assumes "h \ get_child_nodes parent \\<^sub>r children" assumes "cast child \ ptr" assumes "child \ set children" assumes "cast child \ set nodes" shows "parent \ set nodes" proof(insert assms(1) assms(4) assms(6) assms(8), induct ptr arbitrary: nodes rule: heap_wellformed_induct) case (step p) have "p |\| object_ptr_kinds h" using \h \ to_tree_order p \\<^sub>r nodes\ to_tree_order_ptr_in_heap using assms(1) assms(2) assms(3) by blast then obtain children where children: "h \ get_child_nodes p \\<^sub>r children" by (meson assms(2) assms(3) get_child_nodes_ok is_OK_returns_result_E local.known_ptrs_known_ptr) then show ?case proof (cases "children = []") case True then show ?thesis using step(2) step(3) step(4) children by(auto simp add: to_tree_order_def[of p] map_M_pure_I elim!: bind_returns_result_E2 dest!: bind_returns_result_E3[rotated, OF children, rotated]) next case False then obtain c child_to where child: "c \ set children" and child_to: "h \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \\<^sub>r child_to" and "cast child \ set child_to" using step(2) children apply(auto simp add: to_tree_order_def[of p] map_M_pure_I elim!: bind_returns_result_E2 dest!: bind_returns_result_E3[rotated, OF children, rotated])[1] by (metis (full_types) assms(1) assms(2) assms(3) get_parent_ptr_in_heap is_OK_returns_result_I l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.child_parent_dual l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms node_ptr_kinds_commutes returns_result_select_result step.prems(1) step.prems(2) step.prems(3) to_tree_order_either_ptr_or_in_children to_tree_order_ok) then have "set child_to \ set nodes" using assms(1) child children step.prems(1) to_tree_order_child_subset by auto show ?thesis proof (cases "c = child") case True then have "parent = p" using step(3) children child assms(5) assms(7) by (meson assms(1) assms(2) assms(3) child_parent_dual option.inject returns_result_eq) then show ?thesis using step.prems(1) to_tree_order_ptr_in_result by blast next case False then show ?thesis using step(1)[OF children child child_to] step(3) step(4) using \set child_to \ set nodes\ using \cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child \ set child_to\ by auto qed qed qed lemma to_tree_order_node_ptrs: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ to_tree_order ptr \\<^sub>r nodes" assumes "ptr' \ ptr" assumes "ptr' \ set nodes" shows "is_node_ptr_kind ptr'" proof(insert assms(1) assms(4) assms(5) assms(6), induct ptr arbitrary: nodes rule: heap_wellformed_induct) case (step p) have "p |\| object_ptr_kinds h" using \h \ to_tree_order p \\<^sub>r nodes\ to_tree_order_ptr_in_heap using assms(1) assms(2) assms(3) by blast then obtain children where children: "h \ get_child_nodes p \\<^sub>r children" by (meson assms(2) assms(3) get_child_nodes_ok is_OK_returns_result_E local.known_ptrs_known_ptr) then show ?case proof (cases "children = []") case True then show ?thesis using step(2) step(3) step(4) children by(auto simp add: to_tree_order_def[of p] map_M_pure_I elim!: bind_returns_result_E2 dest!: bind_returns_result_E3[rotated, OF children, rotated])[1] next case False then obtain c child_to where child: "c \ set children" and child_to: "h \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \\<^sub>r child_to" and "ptr' \ set child_to" using step(2) children apply(auto simp add: to_tree_order_def[of p] map_M_pure_I elim!: bind_returns_result_E2 dest!: bind_returns_result_E3[rotated, OF children, rotated])[1] using step.prems(1) step.prems(2) step.prems(3) to_tree_order_either_ptr_or_in_children by blast then have "set child_to \ set nodes" using assms(1) child children step.prems(1) to_tree_order_child_subset by auto show ?thesis proof (cases "cast c = ptr") case True then show ?thesis using step \ptr' \ set child_to\ assms(5) child child_to children by blast next case False then show ?thesis using \ptr' \ set child_to\ child child_to children is_node_ptr_kind_cast step.hyps by blast qed qed qed lemma to_tree_order_child2: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ to_tree_order ptr \\<^sub>r nodes" assumes "cast child \ ptr" assumes "cast child \ set nodes" obtains parent where "h \ get_parent child \\<^sub>r Some parent" and "parent \ set nodes" proof - assume 1: "(\parent. h \ get_parent child \\<^sub>r Some parent \ parent \ set nodes \ thesis)" show thesis proof(insert assms(1) assms(4) assms(5) assms(6) 1, induct ptr arbitrary: nodes rule: heap_wellformed_induct) case (step p) have "p |\| object_ptr_kinds h" using \h \ to_tree_order p \\<^sub>r nodes\ to_tree_order_ptr_in_heap using assms(1) assms(2) assms(3) by blast then obtain children where children: "h \ get_child_nodes p \\<^sub>r children" by (meson assms(2) assms(3) get_child_nodes_ok is_OK_returns_result_E local.known_ptrs_known_ptr) then show ?case proof (cases "children = []") case True then show ?thesis using step(2) step(3) step(4) children by(auto simp add: to_tree_order_def[of p] map_M_pure_I elim!: bind_returns_result_E2 dest!: bind_returns_result_E3[rotated, OF children, rotated]) next case False then obtain c child_to where child: "c \ set children" and child_to: "h \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \\<^sub>r child_to" and "cast child \ set child_to" using step(2) children apply(auto simp add: to_tree_order_def[of p] map_M_pure_I elim!: bind_returns_result_E2 dest!: bind_returns_result_E3[rotated, OF children, rotated])[1] using step.prems(1) step.prems(2) step.prems(3) to_tree_order_either_ptr_or_in_children by blast then have "set child_to \ set nodes" using assms(1) child children step.prems(1) to_tree_order_child_subset by auto have "cast child |\| object_ptr_kinds h" using assms(1) assms(2) assms(3) assms(4) assms(6) to_tree_order_ptrs_in_heap by blast then obtain parent_opt where parent_opt: "h \ get_parent child \\<^sub>r parent_opt" by (meson assms(2) assms(3) is_OK_returns_result_E get_parent_ok node_ptr_kinds_commutes) then show ?thesis proof (induct parent_opt) case None then show ?case by (metis \cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child \ set child_to\ assms(1) assms(2) assms(3) cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_inject child child_parent_dual child_to children option.distinct(1) returns_result_eq step.hyps) next case (Some option) then show ?case by (meson assms(1) assms(2) assms(3) get_parent_child_dual step.prems(1) step.prems(2) step.prems(3) step.prems(4) to_tree_order_child) qed qed qed qed lemma to_tree_order_parent_child_rel: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ to_tree_order ptr \\<^sub>r to" shows "(ptr, child) \ (parent_child_rel h)\<^sup>* \ child \ set to" proof assume 3: "(ptr, child) \ (parent_child_rel h)\<^sup>*" show "child \ set to" proof (insert 3, induct child rule: heap_wellformed_induct_rev[OF assms(1)]) case (1 child) then show ?case proof (cases "ptr = child") case True then show ?thesis using assms(4) apply(simp add: to_tree_order_def) by(auto simp add: map_M_pure_I elim!: bind_returns_result_E2) next case False obtain child_parent where "(ptr, child_parent) \ (parent_child_rel h)\<^sup>*" and "(child_parent, child) \ (parent_child_rel h)" using \ptr \ child\ by (metis "1.prems" rtranclE) obtain child_node where child_node: "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child_node = child" using \(child_parent, child) \ parent_child_rel h\ node_ptr_casts_commute3 parent_child_rel_node_ptr by blast then have "h \ get_parent child_node \\<^sub>r Some child_parent" using \(child_parent, child) \ (parent_child_rel h)\ by (meson assms(1) assms(2) assms(3) is_OK_returns_result_E l_get_parent_wf.child_parent_dual l_heap_is_wellformed.parent_child_rel_child local.get_child_nodes_ok local.known_ptrs_known_ptr local.l_get_parent_wf_axioms local.l_heap_is_wellformed_axioms local.parent_child_rel_parent_in_heap) then show ?thesis using 1(1) child_node \(ptr, child_parent) \ (parent_child_rel h)\<^sup>*\ using assms(1) assms(2) assms(3) assms(4) to_tree_order_parent by blast qed qed next assume "child \ set to" then show "(ptr, child) \ (parent_child_rel h)\<^sup>*" proof (induct child rule: heap_wellformed_induct_rev[OF assms(1)]) case (1 child) then show ?case proof (cases "ptr = child") case True then show ?thesis by simp next case False then have "\parent. (parent, child) \ (parent_child_rel h)" using 1(2) assms(4) to_tree_order_child2[OF assms(1) assms(2) assms(3) assms(4)] to_tree_order_node_ptrs by (metis assms(1) assms(2) assms(3) node_ptr_casts_commute3 parent_child_rel_parent) then obtain child_node where child_node: "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child_node = child" using node_ptr_casts_commute3 parent_child_rel_node_ptr by blast then obtain child_parent where child_parent: "h \ get_parent child_node \\<^sub>r Some child_parent" using \\parent. (parent, child) \ (parent_child_rel h)\ by (metis "1.prems" False assms(1) assms(2) assms(3) assms(4) to_tree_order_child2) then have "(child_parent, child) \ (parent_child_rel h)" using assms(1) child_node parent_child_rel_parent by blast moreover have "child_parent \ set to" by (metis "1.prems" False assms(1) assms(2) assms(3) assms(4) child_node child_parent get_parent_child_dual to_tree_order_child) then have "(ptr, child_parent) \ (parent_child_rel h)\<^sup>*" using 1 child_node child_parent by blast ultimately show ?thesis by auto qed qed qed end interpretation i_to_tree_order_wf?: l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs to_tree_order known_ptrs get_parent get_parent_locs heap_is_wellformed parent_child_rel get_disconnected_nodes get_disconnected_nodes_locs using instances apply(simp add: l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) done declare l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_to_tree_order_wf = l_heap_is_wellformed_defs + l_type_wf + l_known_ptrs + l_to_tree_order_defs + l_get_parent_defs + l_get_child_nodes_defs + assumes to_tree_order_ok: "heap_is_wellformed h \ ptr |\| object_ptr_kinds h \ known_ptrs h \ type_wf h \ h \ ok (to_tree_order ptr)" assumes to_tree_order_ptrs_in_heap: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ to_tree_order ptr \\<^sub>r to \ ptr' \ set to \ ptr' |\| object_ptr_kinds h" assumes to_tree_order_parent_child_rel: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ to_tree_order ptr \\<^sub>r to \ (ptr, child_ptr) \ (parent_child_rel h)\<^sup>* \ child_ptr \ set to" assumes to_tree_order_child2: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ to_tree_order ptr \\<^sub>r nodes \ cast child \ ptr \ cast child \ set nodes \ (\parent. h \ get_parent child \\<^sub>r Some parent \ parent \ set nodes \ thesis) \ thesis" assumes to_tree_order_node_ptrs: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ to_tree_order ptr \\<^sub>r nodes \ ptr' \ ptr \ ptr' \ set nodes \ is_node_ptr_kind ptr'" assumes to_tree_order_child: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ to_tree_order ptr \\<^sub>r nodes \ h \ get_child_nodes parent \\<^sub>r children \ cast child \ ptr \ child \ set children \ cast child \ set nodes \ parent \ set nodes" assumes to_tree_order_ptr_in_result: "h \ to_tree_order ptr \\<^sub>r nodes \ ptr \ set nodes" assumes to_tree_order_parent: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ to_tree_order ptr \\<^sub>r nodes \ h \ get_parent child \\<^sub>r Some parent \ parent \ set nodes \ cast child \ set nodes" assumes to_tree_order_subset: "heap_is_wellformed h \ h \ to_tree_order ptr \\<^sub>r nodes \ node \ set nodes \ h \ to_tree_order node \\<^sub>r nodes' \ known_ptrs h \ type_wf h \ set nodes' \ set nodes" lemma to_tree_order_wf_is_l_to_tree_order_wf [instances]: "l_to_tree_order_wf heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent get_child_nodes" using instances apply(auto simp add: l_to_tree_order_wf_def l_to_tree_order_wf_axioms_def)[1] using to_tree_order_ok apply blast using to_tree_order_ptrs_in_heap apply blast using to_tree_order_parent_child_rel apply(blast, blast) using to_tree_order_child2 apply blast using to_tree_order_node_ptrs apply blast using to_tree_order_child apply blast using to_tree_order_ptr_in_result apply blast using to_tree_order_parent apply blast using to_tree_order_subset apply blast done subsubsection \get\_root\_node\ locale l_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_to_tree_order_wf begin lemma to_tree_order_get_root_node: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ to_tree_order ptr \\<^sub>r to" assumes "ptr' \ set to" assumes "h \ get_root_node ptr' \\<^sub>r root_ptr" assumes "ptr'' \ set to" shows "h \ get_root_node ptr'' \\<^sub>r root_ptr" proof - obtain ancestors' where ancestors': "h \ get_ancestors ptr' \\<^sub>r ancestors'" by (meson assms(1) assms(2) assms(3) assms(4) assms(5) get_ancestors_ok is_OK_returns_result_E to_tree_order_ptrs_in_heap ) moreover have "ptr \ set ancestors'" using \h \ get_ancestors ptr' \\<^sub>r ancestors'\ using assms(1) assms(2) assms(3) assms(4) assms(5) get_ancestors_parent_child_rel to_tree_order_parent_child_rel by blast ultimately have "h \ get_root_node ptr \\<^sub>r root_ptr" using \h \ get_root_node ptr' \\<^sub>r root_ptr\ using assms(1) assms(2) assms(3) get_ancestors_ptr get_ancestors_same_root_node by blast obtain ancestors'' where ancestors'': "h \ get_ancestors ptr'' \\<^sub>r ancestors''" by (meson assms(1) assms(2) assms(3) assms(4) assms(7) get_ancestors_ok is_OK_returns_result_E to_tree_order_ptrs_in_heap) moreover have "ptr \ set ancestors''" using \h \ get_ancestors ptr'' \\<^sub>r ancestors''\ using assms(1) assms(2) assms(3) assms(4) assms(7) get_ancestors_parent_child_rel to_tree_order_parent_child_rel by blast ultimately show ?thesis using \h \ get_root_node ptr \\<^sub>r root_ptr\ assms(1) assms(2) assms(3) get_ancestors_ptr get_ancestors_same_root_node by blast qed lemma to_tree_order_same_root: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_root_node ptr \\<^sub>r root_ptr" assumes "h \ to_tree_order root_ptr \\<^sub>r to" assumes "ptr' \ set to" shows "h \ get_root_node ptr' \\<^sub>r root_ptr" proof (insert assms(1)(* assms(4) assms(5) *) assms(6), induct ptr' rule: heap_wellformed_induct_rev) case (step child) then show ?case proof (cases "h \ get_root_node child \\<^sub>r child") case True then have "child = root_ptr" using assms(1) assms(2) assms(3) assms(5) step.prems by (metis (no_types, lifting) get_root_node_same_no_parent node_ptr_casts_commute3 option.simps(3) returns_result_eq to_tree_order_child2 to_tree_order_node_ptrs) then show ?thesis using True by blast next case False then obtain child_node parent where "cast child_node = child" and "h \ get_parent child_node \\<^sub>r Some parent" by (metis assms(1) assms(2) assms(3) assms(4) assms(5) local.get_root_node_no_parent local.get_root_node_not_node_same local.get_root_node_same_no_parent local.to_tree_order_child2 local.to_tree_order_ptrs_in_heap node_ptr_casts_commute3 step.prems) then show ?thesis proof (cases "child = root_ptr") case True then have "h \ get_root_node root_ptr \\<^sub>r root_ptr" using assms(4) using \cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child_node = child\ assms(1) assms(2) assms(3) get_root_node_no_parent get_root_node_same_no_parent by blast then show ?thesis using step assms(4) using True by blast next case False then have "parent \ set to" using assms(5) step(2) to_tree_order_child \h \ get_parent child_node \\<^sub>r Some parent\ \cast child_node = child\ by (metis False assms(1) assms(2) assms(3) get_parent_child_dual) then show ?thesis using \cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child_node = child\ \h \ get_parent child_node \\<^sub>r Some parent\ get_root_node_parent_same using step.hyps by blast qed qed qed end interpretation i_to_tree_order_wf_get_root_node_wf?: l_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf known_ptrs heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_parent get_parent_locs get_ancestors get_ancestors_locs get_root_node get_root_node_locs to_tree_order using instances by(simp add: l_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) locale l_to_tree_order_wf_get_root_node_wf = l_type_wf + l_known_ptrs + l_to_tree_order_defs + l_get_root_node_defs + l_heap_is_wellformed_defs + assumes to_tree_order_get_root_node: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ to_tree_order ptr \\<^sub>r to \ ptr' \ set to \ h \ get_root_node ptr' \\<^sub>r root_ptr \ ptr'' \ set to \ h \ get_root_node ptr'' \\<^sub>r root_ptr" assumes to_tree_order_same_root: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_root_node ptr \\<^sub>r root_ptr \ h \ to_tree_order root_ptr \\<^sub>r to \ ptr' \ set to \ h \ get_root_node ptr' \\<^sub>r root_ptr" lemma to_tree_order_wf_get_root_node_wf_is_l_to_tree_order_wf_get_root_node_wf [instances]: "l_to_tree_order_wf_get_root_node_wf type_wf known_ptr known_ptrs to_tree_order get_root_node heap_is_wellformed" using instances apply(auto simp add: l_to_tree_order_wf_get_root_node_wf_def l_to_tree_order_wf_get_root_node_wf_axioms_def)[1] using to_tree_order_get_root_node apply blast using to_tree_order_same_root apply blast done subsection \get\_owner\_document\ locale l_get_owner_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_known_ptrs + l_heap_is_wellformed + l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_ancestors + l_get_ancestors_wf + l_get_parent + l_get_parent_wf + l_get_root_node_wf + l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_owner_document_disconnected_nodes: assumes "heap_is_wellformed h" assumes "h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes" assumes "node_ptr \ set disc_nodes" assumes known_ptrs: "known_ptrs h" assumes type_wf: "type_wf h" shows "h \ get_owner_document (cast node_ptr) \\<^sub>r document_ptr" proof - have 2: "node_ptr |\| node_ptr_kinds h" using assms heap_is_wellformed_disc_nodes_in_heap by blast have 3: "document_ptr |\| document_ptr_kinds h" using assms(2) get_disconnected_nodes_ptr_in_heap by blast have 0: "\!document_ptr\set |h \ document_ptr_kinds_M|\<^sub>r. node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r" by (metis (no_types, lifting) "3" DocumentMonad.ptr_kinds_ptr_kinds_M assms(1) assms(2) assms(3) disjoint_iff_not_equal l_heap_is_wellformed.heap_is_wellformed_one_disc_parent local.get_disconnected_nodes_ok local.l_heap_is_wellformed_axioms returns_result_select_result select_result_I2 type_wf) have "h \ get_parent node_ptr \\<^sub>r None" using heap_is_wellformed_children_disc_nodes_different child_parent_dual assms using "2" disjoint_iff_not_equal local.get_parent_child_dual local.get_parent_ok returns_result_select_result split_option_ex by (metis (no_types, lifting)) then have 4: "h \ get_root_node (cast node_ptr) \\<^sub>r cast node_ptr" using 2 get_root_node_no_parent by blast obtain document_ptrs where document_ptrs: "h \ document_ptr_kinds_M \\<^sub>r document_ptrs" by simp then have "h \ ok (filter_M (\document_ptr. do { disconnected_nodes \ get_disconnected_nodes document_ptr; return (((cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)) \ cast ` set disconnected_nodes) }) document_ptrs)" using assms(1) get_disconnected_nodes_ok type_wf unfolding heap_is_wellformed_def by(auto intro!: bind_is_OK_I2 filter_M_is_OK_I bind_pure_I) then obtain candidates where candidates: "h \ filter_M (\document_ptr. do { disconnected_nodes \ get_disconnected_nodes document_ptr; return (((cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)) \ cast ` set disconnected_nodes) }) document_ptrs \\<^sub>r candidates" by auto have eq: "\document_ptr. document_ptr |\| document_ptr_kinds h \ node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r \ |h \ do { disconnected_nodes \ get_disconnected_nodes document_ptr; return (((cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)) \ cast ` set disconnected_nodes) }|\<^sub>r" apply(auto dest!: get_disconnected_nodes_ok[OF type_wf] intro!: select_result_I[where P=id, simplified] elim!: bind_returns_result_E2)[1] apply(drule select_result_E[where P=id, simplified]) by(auto elim!: bind_returns_result_E2) have filter: "filter (\document_ptr. |h \ do { disconnected_nodes \ get_disconnected_nodes document_ptr; return (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr \ cast ` set disconnected_nodes) }|\<^sub>r) document_ptrs = [document_ptr]" apply(rule filter_ex1) using 0 document_ptrs apply(simp)[1] using eq using local.get_disconnected_nodes_ok apply auto[1] using assms(2) assms(3) apply(auto intro!: intro!: select_result_I[where P=id, simplified] elim!: bind_returns_result_E2)[1] using returns_result_eq apply fastforce using document_ptrs 3 apply(simp) using document_ptrs by simp have "h \ filter_M (\document_ptr. do { disconnected_nodes \ get_disconnected_nodes document_ptr; return (((cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)) \ cast ` set disconnected_nodes) }) document_ptrs \\<^sub>r [document_ptr]" apply(rule filter_M_filter2) using get_disconnected_nodes_ok document_ptrs 3 assms(1) type_wf filter unfolding heap_is_wellformed_def by(auto intro: bind_pure_I bind_is_OK_I2) with 4 document_ptrs have "h \ a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r document_ptr" by(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I split: option.splits)[1] moreover have "known_ptr (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)" using "4" assms(1) known_ptrs type_wf known_ptrs_known_ptr "2" node_ptr_kinds_commutes by blast ultimately show ?thesis using 2 apply(auto simp add: known_ptr_impl get_owner_document_def a_get_owner_document_tups_def)[1] apply(split invoke_splits, (rule conjI | rule impI)+)+ apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl]) apply(drule(1) known_ptr_not_character_data_ptr) apply(drule(1) known_ptr_not_element_ptr) apply(simp add: NodeClass.known_ptr_defs) by(auto split: option.splits intro!: bind_pure_returns_result_I) qed lemma in_disconnected_nodes_no_parent: assumes "heap_is_wellformed h" and "h \ get_parent node_ptr \\<^sub>r None" and "h \ get_owner_document (cast node_ptr) \\<^sub>r owner_document" and "h \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "node_ptr \ set disc_nodes" proof - have 2: "cast node_ptr |\| object_ptr_kinds h" using assms(3) get_owner_document_ptr_in_heap by fast then have 3: "h \ get_root_node (cast node_ptr) \\<^sub>r cast node_ptr" using assms(2) local.get_root_node_no_parent by blast have "\(\parent_ptr. parent_ptr |\| object_ptr_kinds h \ node_ptr \ set |h \ get_child_nodes parent_ptr|\<^sub>r)" apply(auto)[1] using assms(2) child_parent_dual[OF assms(1)] type_wf assms(1) assms(5) get_child_nodes_ok known_ptrs_known_ptr option.simps(3) returns_result_eq returns_result_select_result by (metis (no_types, hide_lams)) moreover have "node_ptr |\| node_ptr_kinds h" using assms(2) get_parent_ptr_in_heap by blast ultimately have 0: "\document_ptr\set |h \ document_ptr_kinds_M|\<^sub>r. node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r" by (metis DocumentMonad.ptr_kinds_ptr_kinds_M assms(1) finite_set_in heap_is_wellformed_children_disc_nodes) then obtain document_ptr where document_ptr: "document_ptr\set |h \ document_ptr_kinds_M|\<^sub>r" and node_ptr_in_disc_nodes: "node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r" by auto then show ?thesis using get_owner_document_disconnected_nodes known_ptrs type_wf assms using DocumentMonad.ptr_kinds_ptr_kinds_M assms(1) assms(3) assms(4) get_disconnected_nodes_ok returns_result_select_result select_result_I2 by (metis (no_types, hide_lams) ) qed lemma get_owner_document_owner_document_in_heap: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_owner_document ptr \\<^sub>r owner_document" shows "owner_document |\| document_ptr_kinds h" using assms apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1] apply(split invoke_split_asm)+ proof - assume "h \ invoke [] ptr () \\<^sub>r owner_document" then show "owner_document |\| document_ptr_kinds h" by (meson invoke_empty is_OK_returns_result_I) next assume "h \ Heap_Error_Monad.bind (check_in_heap ptr) (\_. (local.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r) ptr ()) \\<^sub>r owner_document" then show "owner_document |\| document_ptr_kinds h" by(auto simp add: a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 split: if_splits) next assume 0: "heap_is_wellformed h" and 1: "type_wf h" and 2: "known_ptrs h" and 3: "\ is_element_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 4: "is_character_data_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 5: "h \ Heap_Error_Monad.bind (check_in_heap ptr) (\_. (local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r) ptr ()) \\<^sub>r owner_document" then obtain root where root: "h \ get_root_node ptr \\<^sub>r root" by(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 split: option.splits) then show ?thesis proof (cases "is_document_ptr root") case True then show ?thesis using 4 5 root apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 intro!: filter_M_pure_I bind_pure_I split: option.splits)[1] apply(drule(1) returns_result_eq) apply(auto)[1] using "0" "1" "2" document_ptr_kinds_commutes local.get_root_node_root_in_heap by blast next case False have "known_ptr root" using "0" "1" "2" local.get_root_node_root_in_heap local.known_ptrs_known_ptr root by blast have "root |\| object_ptr_kinds h" using root using "0" "1" "2" local.get_root_node_root_in_heap by blast then have "is_node_ptr_kind root" using False \known_ptr root\ apply(simp add: known_ptr_impl known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs) using is_node_ptr_kind_none by force then have "(\document_ptr \ fset (document_ptr_kinds h). root \ cast ` set |h \ get_disconnected_nodes document_ptr|\<^sub>r)" by (metis (no_types, lifting) "0" "1" "2" \root |\| object_ptr_kinds h\ local.child_parent_dual local.get_child_nodes_ok local.get_root_node_same_no_parent local.heap_is_wellformed_children_disc_nodes local.known_ptrs_known_ptr node_ptr_casts_commute3 node_ptr_inclusion node_ptr_kinds_commutes notin_fset option.distinct(1) returns_result_eq returns_result_select_result root) then obtain some_owner_document where "some_owner_document |\| document_ptr_kinds h" and "root \ cast ` set |h \ get_disconnected_nodes some_owner_document|\<^sub>r" by auto then obtain candidates where candidates: "h \ filter_M (\document_ptr. Heap_Error_Monad.bind (get_disconnected_nodes document_ptr) (\disconnected_nodes. return (root \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ` set disconnected_nodes))) (sorted_list_of_set (fset (document_ptr_kinds h))) \\<^sub>r candidates" by (metis (no_types, lifting) "1" bind_is_OK_I2 bind_pure_I filter_M_is_OK_I finite_fset is_OK_returns_result_E local.get_disconnected_nodes_ok local.get_disconnected_nodes_pure notin_fset return_ok return_pure sorted_list_of_set(1)) then have "some_owner_document \ set candidates" apply(rule filter_M_in_result_if_ok) using \some_owner_document |\| document_ptr_kinds h\ \root \ cast ` set |h \ get_disconnected_nodes some_owner_document|\<^sub>r\ apply(auto intro!: bind_pure_I bind_pure_returns_result_I)[1] apply (simp add: \some_owner_document |\| document_ptr_kinds h\) using "1" \root \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ` set |h \ get_disconnected_nodes some_owner_document|\<^sub>r\ \some_owner_document |\| document_ptr_kinds h\ local.get_disconnected_nodes_ok by auto then have "candidates \ []" by auto then have "owner_document \ set candidates" using 5 root 4 apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 intro!: filter_M_pure_I bind_pure_I split: option.splits)[1] apply (metis candidates list.set_sel(1) returns_result_eq) by (metis \is_node_ptr_kind root\ node_ptr_no_document_ptr_cast returns_result_eq) then show ?thesis using candidates by (meson bind_pure_I bind_returns_result_E2 filter_M_holds_for_result is_OK_returns_result_I local.get_disconnected_nodes_ptr_in_heap local.get_disconnected_nodes_pure return_pure) qed next assume 0: "heap_is_wellformed h" and 1: "type_wf h" and 2: "known_ptrs h" and 3: "is_element_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 4: "h \ Heap_Error_Monad.bind (check_in_heap ptr) (\_. (local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r) ptr ()) \\<^sub>r owner_document" then obtain root where root: "h \ get_root_node ptr \\<^sub>r root" by(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 split: option.splits) then show ?thesis proof (cases "is_document_ptr root") case True then show ?thesis using 3 4 root apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 intro!: filter_M_pure_I bind_pure_I split: option.splits)[1] apply(drule(1) returns_result_eq) apply(auto)[1] using "0" "1" "2" document_ptr_kinds_commutes local.get_root_node_root_in_heap by blast next case False have "known_ptr root" using "0" "1" "2" local.get_root_node_root_in_heap local.known_ptrs_known_ptr root by blast have "root |\| object_ptr_kinds h" using root using "0" "1" "2" local.get_root_node_root_in_heap by blast then have "is_node_ptr_kind root" using False \known_ptr root\ apply(simp add: known_ptr_impl known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs) using is_node_ptr_kind_none by force then have "(\document_ptr \ fset (document_ptr_kinds h). root \ cast ` set |h \ get_disconnected_nodes document_ptr|\<^sub>r)" by (metis (no_types, lifting) "0" "1" "2" \root |\| object_ptr_kinds h\ local.child_parent_dual local.get_child_nodes_ok local.get_root_node_same_no_parent local.heap_is_wellformed_children_disc_nodes local.known_ptrs_known_ptr node_ptr_casts_commute3 node_ptr_inclusion node_ptr_kinds_commutes notin_fset option.distinct(1) returns_result_eq returns_result_select_result root) then obtain some_owner_document where "some_owner_document |\| document_ptr_kinds h" and "root \ cast ` set |h \ get_disconnected_nodes some_owner_document|\<^sub>r" by auto then obtain candidates where candidates: "h \ filter_M (\document_ptr. Heap_Error_Monad.bind (get_disconnected_nodes document_ptr) (\disconnected_nodes. return (root \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ` set disconnected_nodes))) (sorted_list_of_set (fset (document_ptr_kinds h))) \\<^sub>r candidates" by (metis (no_types, lifting) "1" bind_is_OK_I2 bind_pure_I filter_M_is_OK_I finite_fset is_OK_returns_result_E local.get_disconnected_nodes_ok local.get_disconnected_nodes_pure notin_fset return_ok return_pure sorted_list_of_set(1)) then have "some_owner_document \ set candidates" apply(rule filter_M_in_result_if_ok) using \some_owner_document |\| document_ptr_kinds h\ \root \ cast ` set |h \ get_disconnected_nodes some_owner_document|\<^sub>r\ apply(auto intro!: bind_pure_I bind_pure_returns_result_I)[1] using \some_owner_document |\| document_ptr_kinds h\ \root \ cast ` set |h \ get_disconnected_nodes some_owner_document|\<^sub>r\ apply(auto intro!: bind_pure_I bind_pure_returns_result_I)[1] using \some_owner_document |\| document_ptr_kinds h\ \root \ cast ` set |h \ get_disconnected_nodes some_owner_document|\<^sub>r\ apply(auto intro!: bind_pure_I bind_pure_returns_result_I)[1] by (simp add: "1" local.get_disconnected_nodes_ok) then have "candidates \ []" by auto then have "owner_document \ set candidates" using 4 root 3 apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 intro!: filter_M_pure_I bind_pure_I split: option.splits)[1] apply (metis candidates list.set_sel(1) returns_result_eq) by (metis \is_node_ptr_kind root\ node_ptr_no_document_ptr_cast returns_result_eq) then show ?thesis using candidates by (meson bind_pure_I bind_returns_result_E2 filter_M_holds_for_result is_OK_returns_result_I local.get_disconnected_nodes_ptr_in_heap local.get_disconnected_nodes_pure return_pure) qed qed lemma get_owner_document_ok: assumes "heap_is_wellformed h" "known_ptrs h" "type_wf h" assumes "ptr |\| object_ptr_kinds h" shows "h \ ok (get_owner_document ptr)" proof - have "known_ptr ptr" using assms(2) assms(4) local.known_ptrs_known_ptr by blast then show ?thesis apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1] apply(split invoke_splits, (rule conjI | rule impI)+)+ apply(auto simp add: known_ptr_impl)[1] using NodeClass.a_known_ptr_def known_ptr_not_character_data_ptr known_ptr_not_document_ptr known_ptr_not_element_ptr apply blast using assms(4) apply(auto simp add: a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_is_OK_pure_I)[1] apply (metis (no_types, lifting) document_ptr_casts_commute3 document_ptr_kinds_commutes is_document_ptr_kind_none option.case_eq_if) using assms(4) apply(auto simp add: a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_is_OK_pure_I)[1] apply (metis (no_types, lifting) assms(1) assms(2) assms(3) is_node_ptr_kind_none local.get_root_node_ok node_ptr_casts_commute3 option.case_eq_if) using assms(4) apply(auto simp add: a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_is_OK_pure_I)[1] apply(auto split: option.splits intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I filter_M_is_OK_I)[1] using assms(3) local.get_disconnected_nodes_ok apply blast apply (simp add: assms(1) assms(2) assms(3) local.get_root_node_ok) using assms(4) apply(auto simp add: a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_is_OK_pure_I)[1] apply(auto split: option.splits intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I filter_M_is_OK_I)[1] apply (simp add: assms(1) assms(2) assms(3) local.get_root_node_ok)[1] apply(auto split: option.splits intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I filter_M_is_OK_I)[1] using assms(3) local.get_disconnected_nodes_ok by blast qed lemma get_owner_document_child_same: assumes "heap_is_wellformed h" "known_ptrs h" "type_wf h" assumes "h \ get_child_nodes ptr \\<^sub>r children" assumes "child \ set children" shows "h \ get_owner_document ptr \\<^sub>r owner_document \ h \ get_owner_document (cast child) \\<^sub>r owner_document" proof - have "ptr |\| object_ptr_kinds h" by (meson assms(4) is_OK_returns_result_I local.get_child_nodes_ptr_in_heap) then have "known_ptr ptr" using assms(2) local.known_ptrs_known_ptr by blast have "cast child |\| object_ptr_kinds h" using assms(1) assms(4) assms(5) local.heap_is_wellformed_children_in_heap node_ptr_kinds_commutes by blast then have "known_ptr (cast child)" using assms(2) local.known_ptrs_known_ptr by blast obtain root where root: "h \ get_root_node ptr \\<^sub>r root" by (meson \ptr |\| object_ptr_kinds h\ assms(1) assms(2) assms(3) is_OK_returns_result_E local.get_root_node_ok) then have "h \ get_root_node (cast child) \\<^sub>r root" using assms(1) assms(2) assms(3) assms(4) assms(5) local.child_parent_dual local.get_root_node_parent_same by blast have "h \ get_owner_document ptr \\<^sub>r owner_document \ h \ a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child () \\<^sub>r owner_document" proof (cases "is_document_ptr ptr") case True then obtain document_ptr where document_ptr: "cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr = ptr" using case_optionE document_ptr_casts_commute by blast then have "root = cast document_ptr" using root by(auto simp add: get_root_node_def get_ancestors_def elim!: bind_returns_result_E2 split: option.splits) then have "h \ a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr () \\<^sub>r owner_document \ h \ a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child () \\<^sub>r owner_document" using document_ptr \h \ get_root_node (cast child) \\<^sub>r root\[simplified \root = cast document_ptr\ document_ptr] apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 dest!: bind_returns_result_E3[rotated, OF \h \ get_root_node (cast child) \\<^sub>r root\ [simplified \root = cast document_ptr\ document_ptr], rotated] intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I split: if_splits option.splits)[1] using \ptr |\| object_ptr_kinds h\ document_ptr_kinds_commutes by blast then show ?thesis using \known_ptr ptr\ apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def known_ptr_impl)[1] apply(split invoke_splits, ((rule conjI | rule impI)+)?)+ apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl]) apply(drule(1) known_ptr_not_character_data_ptr) apply(drule(1) known_ptr_not_element_ptr) apply(simp add: NodeClass.known_ptr_defs) using \ptr |\| object_ptr_kinds h\ True by(auto simp add: document_ptr[symmetric] intro!: bind_pure_returns_result_I split: option.splits) next case False then obtain node_ptr where node_ptr: "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr = ptr" using \known_ptr ptr\ by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits) then have "h \ a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r owner_document \ h \ a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child () \\<^sub>r owner_document" using root \h \ get_root_node (cast child) \\<^sub>r root\ unfolding a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def by (meson bind_pure_returns_result_I bind_returns_result_E3 local.get_root_node_pure) then show ?thesis using \known_ptr ptr\ apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def known_ptr_impl)[1] apply(split invoke_splits, ((rule conjI | rule impI)+)?)+ apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl]) apply(drule(1) known_ptr_not_character_data_ptr) apply(drule(1) known_ptr_not_element_ptr) apply(simp add: NodeClass.known_ptr_defs) using \cast child |\| object_ptr_kinds h\ \ptr |\| object_ptr_kinds h\ False apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1] using \cast child |\| object_ptr_kinds h\ \ptr |\| object_ptr_kinds h\ False apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1] using \cast child |\| object_ptr_kinds h\ \ptr |\| object_ptr_kinds h\ False apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1] using \cast child |\| object_ptr_kinds h\ \ptr |\| object_ptr_kinds h\ False apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1] apply(split invoke_splits, ((rule conjI | rule impI)+)?)+ apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1] apply (meson invoke_empty is_OK_returns_result_I) apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1] apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1] by(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1] qed then show ?thesis using \known_ptr (cast child)\ apply(auto simp add: get_owner_document_def[of "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child"] a_get_owner_document_tups_def known_ptr_impl)[1] apply(split invoke_splits, ((rule conjI | rule impI)+)?)+ apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl]) apply(drule(1) known_ptr_not_character_data_ptr) apply(drule(1) known_ptr_not_element_ptr) apply(simp add: NodeClass.known_ptr_defs) using \cast child |\| object_ptr_kinds h\ \ptr |\| object_ptr_kinds h\ apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1] using \cast child |\| object_ptr_kinds h\ \ptr |\| object_ptr_kinds h\ apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1] using \cast child |\| object_ptr_kinds h\ \ptr |\| object_ptr_kinds h\ apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1] using \cast child |\| object_ptr_kinds h\ \ptr |\| object_ptr_kinds h\ apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1] by (smt \cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child |\| object_ptr_kinds h\ cast_document_ptr_not_node_ptr(1) comp_apply invoke_empty invoke_not invoke_returns_result is_OK_returns_result_I node_ptr_casts_commute2 option.sel) qed end locale l_get_owner_document_wf = l_heap_is_wellformed_defs + l_type_wf + l_known_ptrs + l_get_disconnected_nodes_defs + l_get_owner_document_defs + l_get_parent_defs + assumes get_owner_document_disconnected_nodes: "heap_is_wellformed h \ known_ptrs h \ type_wf h \ h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes \ node_ptr \ set disc_nodes \ h \ get_owner_document (cast node_ptr) \\<^sub>r document_ptr" assumes in_disconnected_nodes_no_parent: "heap_is_wellformed h \ h \ get_parent node_ptr \\<^sub>r None\ h \ get_owner_document (cast node_ptr) \\<^sub>r owner_document \ h \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes \ known_ptrs h \ type_wf h\ node_ptr \ set disc_nodes" assumes get_owner_document_owner_document_in_heap: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_owner_document ptr \\<^sub>r owner_document \ owner_document |\| document_ptr_kinds h" assumes get_owner_document_ok: "heap_is_wellformed h \ known_ptrs h \ type_wf h \ ptr |\| object_ptr_kinds h \ h \ ok (get_owner_document ptr)" interpretation i_get_owner_document_wf?: l_get_owner_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr known_ptrs type_wf heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_parent get_parent_locs get_ancestors get_ancestors_locs get_root_node get_root_node_locs get_owner_document by(auto simp add: l_get_owner_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_owner_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_owner_document_wf_is_l_get_owner_document_wf [instances]: "l_get_owner_document_wf heap_is_wellformed type_wf known_ptr known_ptrs get_disconnected_nodes get_owner_document get_parent" using known_ptrs_is_l_known_ptrs apply(auto simp add: l_get_owner_document_wf_def l_get_owner_document_wf_axioms_def)[1] using get_owner_document_disconnected_nodes apply fast using in_disconnected_nodes_no_parent apply fast using get_owner_document_owner_document_in_heap apply fast using get_owner_document_ok apply fast done subsubsection \get\_root\_node\ locale l_get_owner_document_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_root_node_wf + l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_owner_document_wf begin lemma get_root_node_document: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_root_node ptr \\<^sub>r root" assumes "is_document_ptr_kind root" shows "h \ get_owner_document ptr \\<^sub>r the (cast root)" proof - have "ptr |\| object_ptr_kinds h" using assms(4) by (meson is_OK_returns_result_I local.get_root_node_ptr_in_heap) then have "known_ptr ptr" using assms(3) local.known_ptrs_known_ptr by blast { assume "is_document_ptr_kind ptr" then have "ptr = root" using assms(4) by(auto simp add: get_root_node_def get_ancestors_def elim!: bind_returns_result_E2 split: option.splits) then have ?thesis using \is_document_ptr_kind ptr\ \known_ptr ptr\ \ptr |\| object_ptr_kinds h\ apply(auto simp add: known_ptr_impl get_owner_document_def a_get_owner_document_tups_def)[1] apply(split invoke_splits, (rule conjI | rule impI)+)+ apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl]) apply(drule(1) known_ptr_not_character_data_ptr) apply(drule(1) known_ptr_not_element_ptr) apply(simp add: NodeClass.known_ptr_defs) by(auto simp add: a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_pure_returns_result_I split: option.splits) } moreover { assume "is_node_ptr_kind ptr" then have ?thesis using \known_ptr ptr\ \ptr |\| object_ptr_kinds h\ apply(auto simp add: known_ptr_impl get_owner_document_def a_get_owner_document_tups_def)[1] apply(split invoke_splits, (rule conjI | rule impI)+)+ apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl]) apply(drule(1) known_ptr_not_character_data_ptr) apply(drule(1) known_ptr_not_element_ptr) apply(simp add: NodeClass.known_ptr_defs) apply(auto split: option.splits)[1] using \h \ get_root_node ptr \\<^sub>r root\ assms(5) by(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def is_document_ptr_kind_def intro!: bind_pure_returns_result_I split: option.splits)[2] } ultimately show ?thesis using \known_ptr ptr\ by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits) qed lemma get_root_node_same_owner_document: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_root_node ptr \\<^sub>r root" shows "h \ get_owner_document ptr \\<^sub>r owner_document \ h \ get_owner_document root \\<^sub>r owner_document" proof - have "ptr |\| object_ptr_kinds h" by (meson assms(4) is_OK_returns_result_I local.get_root_node_ptr_in_heap) have "root |\| object_ptr_kinds h" using assms(1) assms(2) assms(3) assms(4) local.get_root_node_root_in_heap by blast have "known_ptr ptr" using \ptr |\| object_ptr_kinds h\ assms(3) local.known_ptrs_known_ptr by blast have "known_ptr root" using \root |\| object_ptr_kinds h\ assms(3) local.known_ptrs_known_ptr by blast show ?thesis proof (cases "is_document_ptr_kind ptr") case True then have "ptr = root" using assms(4) apply(auto simp add: get_root_node_def elim!: bind_returns_result_E2)[1] by (metis document_ptr_casts_commute3 last_ConsL local.get_ancestors_not_node node_ptr_no_document_ptr_cast) then show ?thesis by auto next case False then have "is_node_ptr_kind ptr" using \known_ptr ptr\ by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits) then obtain node_ptr where node_ptr: "ptr = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr" by (metis node_ptr_casts_commute3) show ?thesis proof assume "h \ get_owner_document ptr \\<^sub>r owner_document" then have "h \ local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r owner_document" using node_ptr apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1] apply(split invoke_splits)+ apply (meson invoke_empty is_OK_returns_result_I) by(auto elim!: bind_returns_result_E2 split: option.splits) show "h \ get_owner_document root \\<^sub>r owner_document" proof (cases "is_document_ptr_kind root") case True have "is_document_ptr root" using True \known_ptr root\ by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits) have "root = cast owner_document" using True by (smt \h \ get_owner_document ptr \\<^sub>r owner_document\ assms(1) assms(2) assms(3) assms(4) document_ptr_casts_commute3 get_root_node_document returns_result_eq) then show ?thesis apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1] apply(split invoke_splits, (rule conjI | rule impI)+)+ using \is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root\ apply blast using \root |\| object_ptr_kinds h\ by(auto simp add: a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def is_node_ptr_kind_none) next case False then have "is_node_ptr_kind root" using \known_ptr root\ by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits) then obtain root_node_ptr where root_node_ptr: "root = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr" by (metis node_ptr_casts_commute3) then have "h \ local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr () \\<^sub>r owner_document" using \h \ local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r owner_document\ assms(4) apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I split: option.splits)[1] apply (metis assms(1) assms(2) assms(3) local.get_root_node_no_parent local.get_root_node_same_no_parent node_ptr returns_result_eq) using \is_node_ptr_kind root\ node_ptr returns_result_eq by fastforce then show ?thesis apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1] apply(split invoke_splits, (rule conjI | rule impI)+)+ using \is_node_ptr_kind root\ \known_ptr root\ apply(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits)[1] apply(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits)[1] using \root |\| object_ptr_kinds h\ by(auto simp add: root_node_ptr) qed next assume "h \ get_owner_document root \\<^sub>r owner_document" show "h \ get_owner_document ptr \\<^sub>r owner_document" proof (cases "is_document_ptr_kind root") case True have "root = cast owner_document" using \h \ get_owner_document root \\<^sub>r owner_document\ apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1] apply(split invoke_splits)+ apply (meson invoke_empty is_OK_returns_result_I) apply(auto simp add: True a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 split: if_splits)[1] apply (metis True cast_document_ptr_not_node_ptr(2) is_document_ptr_kind_obtains is_node_ptr_kind_none node_ptr_casts_commute3 option.case_eq_if) by (metis True cast_document_ptr_not_node_ptr(1) document_ptr_casts_commute3 is_node_ptr_kind_none node_ptr_casts_commute3 option.case_eq_if) then show ?thesis using assms(1) assms(2) assms(3) assms(4) get_root_node_document by fastforce next case False then have "is_node_ptr_kind root" using \known_ptr root\ by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits) then obtain root_node_ptr where root_node_ptr: "root = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr" by (metis node_ptr_casts_commute3) then have "h \ local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr () \\<^sub>r owner_document" using \h \ get_owner_document root \\<^sub>r owner_document\ apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1] apply(split invoke_splits)+ apply (meson invoke_empty is_OK_returns_result_I) by(auto simp add: is_document_ptr_kind_none elim!: bind_returns_result_E2) then have "h \ local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r owner_document" apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I split: option.splits)[1] using assms(1) assms(2) assms(3) assms(4) local.get_root_node_no_parent local.get_root_node_same_no_parent node_ptr returns_result_eq root_node_ptr by fastforce+ then show ?thesis apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1] apply(split invoke_splits, (rule conjI | rule impI)+)+ using node_ptr \known_ptr ptr\ \ptr |\| object_ptr_kinds h\ by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs intro!: bind_pure_returns_result_I split: option.splits) qed qed qed qed end interpretation get_owner_document_wf_get_root_node_wf?: l_get_owner_document_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_ancestors get_ancestors_locs get_root_node get_root_node_locs heap_is_wellformed parent_child_rel get_disconnected_nodes get_disconnected_nodes_locs get_owner_document by(auto simp add: l_get_owner_document_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_owner_document_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_get_owner_document_wf_get_root_node_wf = l_heap_is_wellformed_defs + l_type_wf + l_known_ptrs + l_get_root_node_defs + l_get_owner_document_defs + assumes get_root_node_document: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_root_node ptr \\<^sub>r root \ is_document_ptr_kind root \ h \ get_owner_document ptr \\<^sub>r the (cast root)" assumes get_root_node_same_owner_document: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_root_node ptr \\<^sub>r root \ h \ get_owner_document ptr \\<^sub>r owner_document \ h \ get_owner_document root \\<^sub>r owner_document" lemma get_owner_document_wf_get_root_node_wf_is_l_get_owner_document_wf_get_root_node_wf [instances]: "l_get_owner_document_wf_get_root_node_wf heap_is_wellformed type_wf known_ptr known_ptrs get_root_node get_owner_document" apply(auto simp add: l_get_owner_document_wf_get_root_node_wf_def l_get_owner_document_wf_get_root_node_wf_axioms_def instances)[1] using get_root_node_document apply blast using get_root_node_same_owner_document apply (blast, blast) done subsection \Preserving heap-wellformedness\ subsection \set\_attribute\ locale l_set_attribute_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_attribute\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_attribute_get_disconnected_nodes + l_set_attribute_get_child_nodes begin lemma set_attribute_preserves_wellformedness: assumes "heap_is_wellformed h" and "h \ set_attribute element_ptr k v \\<^sub>h h'" shows "heap_is_wellformed h'" thm preserves_wellformedness_writes_needed apply(rule preserves_wellformedness_writes_needed[OF assms set_attribute_writes]) using set_attribute_get_child_nodes apply(fast) using set_attribute_get_disconnected_nodes apply(fast) by(auto simp add: all_args_def set_attribute_locs_def) end subsection \remove\_child\ locale l_remove_child_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_heap_is_wellformed + l_set_disconnected_nodes_get_child_nodes begin lemma remove_child_removes_parent: assumes wellformed: "heap_is_wellformed h" and remove_child: "h \ remove_child ptr child \\<^sub>h h2" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "h2 \ get_parent child \\<^sub>r None" proof - obtain children where children: "h \ get_child_nodes ptr \\<^sub>r children" using remove_child remove_child_def by auto then have "child \ set children" using remove_child remove_child_def by(auto elim!: bind_returns_heap_E dest: returns_result_eq split: if_splits) then have h1: "\other_ptr other_children. other_ptr \ ptr \ h \ get_child_nodes other_ptr \\<^sub>r other_children \ child \ set other_children" using assms(1) known_ptrs type_wf child_parent_dual by (meson child_parent_dual children option.inject returns_result_eq) have known_ptr: "known_ptr ptr" using known_ptrs by (meson is_OK_returns_heap_I l_known_ptrs.known_ptrs_known_ptr l_known_ptrs_axioms remove_child remove_child_ptr_in_heap) obtain owner_document disc_nodes h' where owner_document: "h \ get_owner_document (cast child) \\<^sub>r owner_document" and disc_nodes: "h \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes" and h': "h \ set_disconnected_nodes owner_document (child # disc_nodes) \\<^sub>h h'" and h2: "h' \ set_child_nodes ptr (remove1 child children) \\<^sub>h h2" using assms children unfolding remove_child_def apply(auto split: if_splits elim!: bind_returns_heap_E)[1] by (metis (full_types) get_child_nodes_pure get_disconnected_nodes_pure get_owner_document_pure pure_returns_heap_eq returns_result_eq) have "object_ptr_kinds h = object_ptr_kinds h2" using remove_child_writes remove_child unfolding remove_child_locs_def apply(rule writes_small_big) using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by(auto simp add: reflp_def transp_def) then have "|h \ object_ptr_kinds_M|\<^sub>r = |h2 \ object_ptr_kinds_M|\<^sub>r" unfolding object_ptr_kinds_M_defs by simp have "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h'] using set_disconnected_nodes_types_preserved type_wf by(auto simp add: reflp_def transp_def) have "type_wf h2" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF remove_child_writes remove_child] unfolding remove_child_locs_def using set_disconnected_nodes_types_preserved set_child_nodes_types_preserved type_wf apply(auto simp add: reflp_def transp_def)[1] by blast then obtain children' where children': "h2 \ get_child_nodes ptr \\<^sub>r children'" using h2 set_child_nodes_get_child_nodes known_ptr by (metis \object_ptr_kinds h = object_ptr_kinds h2\ children get_child_nodes_ok get_child_nodes_ptr_in_heap is_OK_returns_result_E is_OK_returns_result_I) have "child \ set children'" by (metis (mono_tags, lifting) \type_wf h'\ children children' distinct_remove1_removeAll h2 known_ptr local.heap_is_wellformed_children_distinct local.set_child_nodes_get_child_nodes member_remove remove_code(1) select_result_I2 wellformed) moreover have "\other_ptr other_children. other_ptr \ ptr \ h' \ get_child_nodes other_ptr \\<^sub>r other_children \ child \ set other_children" proof - fix other_ptr other_children assume a1: "other_ptr \ ptr" and a3: "h' \ get_child_nodes other_ptr \\<^sub>r other_children" have "h \ get_child_nodes other_ptr \\<^sub>r other_children" using get_child_nodes_reads set_disconnected_nodes_writes h' a3 apply(rule reads_writes_separate_backwards) using set_disconnected_nodes_get_child_nodes by fast show "child \ set other_children" using \h \ get_child_nodes other_ptr \\<^sub>r other_children\ a1 h1 by blast qed then have "\other_ptr other_children. other_ptr \ ptr \ h2 \ get_child_nodes other_ptr \\<^sub>r other_children \ child \ set other_children" proof - fix other_ptr other_children assume a1: "other_ptr \ ptr" and a3: "h2 \ get_child_nodes other_ptr \\<^sub>r other_children" have "h' \ get_child_nodes other_ptr \\<^sub>r other_children" using get_child_nodes_reads set_child_nodes_writes h2 a3 apply(rule reads_writes_separate_backwards) using set_disconnected_nodes_get_child_nodes a1 set_child_nodes_get_child_nodes_different_pointers by metis then show "child \ set other_children" using \\other_ptr other_children. \other_ptr \ ptr; h' \ get_child_nodes other_ptr \\<^sub>r other_children\ \ child \ set other_children\ a1 by blast qed ultimately have ha: "\other_ptr other_children. h2 \ get_child_nodes other_ptr \\<^sub>r other_children \ child \ set other_children" by (metis (full_types) children' returns_result_eq) moreover obtain ptrs where ptrs: "h2 \ object_ptr_kinds_M \\<^sub>r ptrs" by (simp add: object_ptr_kinds_M_defs) moreover have "\ptr. ptr \ set ptrs \ h2 \ ok (get_child_nodes ptr)" using \type_wf h2\ ptrs get_child_nodes_ok known_ptr using \object_ptr_kinds h = object_ptr_kinds h2\ known_ptrs local.known_ptrs_known_ptr by auto ultimately show "h2 \ get_parent child \\<^sub>r None" apply(auto simp add: get_parent_def intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I)[1] proof - have "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child |\| object_ptr_kinds h" using get_owner_document_ptr_in_heap owner_document by blast then show "h2 \ check_in_heap (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r ()" by (simp add: \object_ptr_kinds h = object_ptr_kinds h2\ check_in_heap_def) next show "(\other_ptr other_children. h2 \ get_child_nodes other_ptr \\<^sub>r other_children \ child \ set other_children) \ ptrs = sorted_list_of_set (fset (object_ptr_kinds h2)) \ (\ptr. ptr |\| object_ptr_kinds h2 \ h2 \ ok get_child_nodes ptr) \ h2 \ filter_M (\ptr. Heap_Error_Monad.bind (get_child_nodes ptr) (\children. return (child \ set children))) (sorted_list_of_set (fset (object_ptr_kinds h2))) \\<^sub>r []" by(auto intro!: filter_M_empty_I bind_pure_I) qed qed end locale l_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_remove_child_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma remove_child_parent_child_rel_subset: assumes "heap_is_wellformed h" and "h \ remove_child ptr child \\<^sub>h h'" and "known_ptrs h" and type_wf: "type_wf h" shows "parent_child_rel h' \ parent_child_rel h" proof (standard, safe) obtain owner_document children_h h2 disconnected_nodes_h where owner_document: "h \ get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r owner_document" and children_h: "h \ get_child_nodes ptr \\<^sub>r children_h" and child_in_children_h: "child \ set children_h" and disconnected_nodes_h: "h \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h" and h2: "h \ set_disconnected_nodes owner_document (child # disconnected_nodes_h) \\<^sub>h h2" and h': "h2 \ set_child_nodes ptr (remove1 child children_h) \\<^sub>h h'" using assms(2) apply(auto simp add: remove_child_def elim!: bind_returns_heap_E dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure] pure_returns_heap_eq[rotated, OF get_child_nodes_pure] split: if_splits)[1] using pure_returns_heap_eq by fastforce have object_ptr_kinds_eq3: "object_ptr_kinds h = object_ptr_kinds h'" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF remove_child_writes assms(2)]) unfolding remove_child_locs_def using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_eq: "\ptrs. h \ object_ptr_kinds_M \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" unfolding object_ptr_kinds_M_defs by simp then have object_ptr_kinds_eq2: "|h \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" using select_result_eq by force then have node_ptr_kinds_eq2: "|h \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by auto then have node_ptr_kinds_eq3: "node_ptr_kinds h = node_ptr_kinds h'" using node_ptr_kinds_M_eq by auto have document_ptr_kinds_eq2: "|h \ document_ptr_kinds_M|\<^sub>r = |h' \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_eq2 document_ptr_kinds_M_eq by auto then have document_ptr_kinds_eq3: "document_ptr_kinds h = document_ptr_kinds h'" using document_ptr_kinds_M_eq by auto have children_eq: "\ptr' children. ptr \ ptr' \ h \ get_child_nodes ptr' \\<^sub>r children =h' \ get_child_nodes ptr' \\<^sub>r children" apply(rule reads_writes_preserved[OF get_child_nodes_reads remove_child_writes assms(2)]) unfolding remove_child_locs_def using set_disconnected_nodes_get_child_nodes set_child_nodes_get_child_nodes_different_pointers by fast then have children_eq2: "\ptr' children. ptr \ ptr' \ |h \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq: "\document_ptr disconnected_nodes. document_ptr \ owner_document \ h \ get_disconnected_nodes document_ptr \\<^sub>r disconnected_nodes = h' \ get_disconnected_nodes document_ptr \\<^sub>r disconnected_nodes" apply(rule reads_writes_preserved[OF get_disconnected_nodes_reads remove_child_writes assms(2)]) unfolding remove_child_locs_def using set_child_nodes_get_disconnected_nodes set_disconnected_nodes_get_disconnected_nodes_different_pointers by (metis (no_types, lifting) Un_iff owner_document select_result_I2) then have disconnected_nodes_eq2: "\document_ptr. document_ptr \ owner_document \ |h \ get_disconnected_nodes document_ptr|\<^sub>r = |h' \ get_disconnected_nodes document_ptr|\<^sub>r" using select_result_eq by force have "h2 \ get_child_nodes ptr \\<^sub>r children_h" apply(rule reads_writes_separate_forwards[OF get_child_nodes_reads set_disconnected_nodes_writes h2 children_h] ) by (simp add: set_disconnected_nodes_get_child_nodes) have "known_ptr ptr" using assms(3) using children_h get_child_nodes_ptr_in_heap local.known_ptrs_known_ptr by blast have "type_wf h2" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h2] using set_disconnected_nodes_types_preserved type_wf by(auto simp add: reflp_def transp_def) then have "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_child_nodes_writes h'] using set_child_nodes_types_preserved by(auto simp add: reflp_def transp_def) have children_h': "h' \ get_child_nodes ptr \\<^sub>r remove1 child children_h" using assms(2) owner_document h2 disconnected_nodes_h children_h apply(auto simp add: remove_child_def split: if_splits)[1] apply(drule bind_returns_heap_E3) apply(auto split: if_splits)[1] apply(simp) apply(auto split: if_splits)[1] apply(drule bind_returns_heap_E3) apply(auto)[1] apply(simp) apply(drule bind_returns_heap_E3) apply(auto)[1] apply(simp) apply(drule bind_returns_heap_E4) apply(auto)[1] apply(simp) using \type_wf h2\ set_child_nodes_get_child_nodes \known_ptr ptr\ h' by blast fix parent child assume a1: "(parent, child) \ parent_child_rel h'" then show "(parent, child) \ parent_child_rel h" proof (cases "parent = ptr") case True then show ?thesis using a1 remove_child_removes_parent[OF assms(1) assms(2)] children_h children_h' get_child_nodes_ptr_in_heap apply(auto simp add: parent_child_rel_def object_ptr_kinds_eq )[1] by (metis notin_set_remove1) next case False then show ?thesis using a1 by(auto simp add: parent_child_rel_def object_ptr_kinds_eq3 children_eq2) qed qed lemma remove_child_heap_is_wellformed_preserved: assumes "heap_is_wellformed h" and "h \ remove_child ptr child \\<^sub>h h'" and "known_ptrs h" and type_wf: "type_wf h" shows "type_wf h'" and "known_ptrs h'" and "heap_is_wellformed h'" proof - obtain owner_document children_h h2 disconnected_nodes_h where owner_document: "h \ get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r owner_document" and children_h: "h \ get_child_nodes ptr \\<^sub>r children_h" and child_in_children_h: "child \ set children_h" and disconnected_nodes_h: "h \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h" and h2: "h \ set_disconnected_nodes owner_document (child # disconnected_nodes_h) \\<^sub>h h2" and h': "h2 \ set_child_nodes ptr (remove1 child children_h) \\<^sub>h h'" using assms(2) apply(auto simp add: remove_child_def elim!: bind_returns_heap_E dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure] pure_returns_heap_eq[rotated, OF get_child_nodes_pure] split: if_splits)[1] using pure_returns_heap_eq by fastforce have object_ptr_kinds_eq3: "object_ptr_kinds h = object_ptr_kinds h'" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF remove_child_writes assms(2)]) unfolding remove_child_locs_def using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_eq: "\ptrs. h \ object_ptr_kinds_M \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" unfolding object_ptr_kinds_M_defs by simp then have object_ptr_kinds_eq2: "|h \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" using select_result_eq by force then have node_ptr_kinds_eq2: "|h \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by auto then have node_ptr_kinds_eq3: "node_ptr_kinds h = node_ptr_kinds h'" using node_ptr_kinds_M_eq by auto have document_ptr_kinds_eq2: "|h \ document_ptr_kinds_M|\<^sub>r = |h' \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_eq2 document_ptr_kinds_M_eq by auto then have document_ptr_kinds_eq3: "document_ptr_kinds h = document_ptr_kinds h'" using document_ptr_kinds_M_eq by auto have children_eq: "\ptr' children. ptr \ ptr' \ h \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" apply(rule reads_writes_preserved[OF get_child_nodes_reads remove_child_writes assms(2)]) unfolding remove_child_locs_def using set_disconnected_nodes_get_child_nodes set_child_nodes_get_child_nodes_different_pointers by fast then have children_eq2: "\ptr' children. ptr \ ptr' \ |h \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq: "\document_ptr disconnected_nodes. document_ptr \ owner_document \ h \ get_disconnected_nodes document_ptr \\<^sub>r disconnected_nodes = h' \ get_disconnected_nodes document_ptr \\<^sub>r disconnected_nodes" apply(rule reads_writes_preserved[OF get_disconnected_nodes_reads remove_child_writes assms(2)]) unfolding remove_child_locs_def using set_child_nodes_get_disconnected_nodes set_disconnected_nodes_get_disconnected_nodes_different_pointers by (metis (no_types, lifting) Un_iff owner_document select_result_I2) then have disconnected_nodes_eq2: "\document_ptr. document_ptr \ owner_document \ |h \ get_disconnected_nodes document_ptr|\<^sub>r = |h' \ get_disconnected_nodes document_ptr|\<^sub>r" using select_result_eq by force have "h2 \ get_child_nodes ptr \\<^sub>r children_h" apply(rule reads_writes_separate_forwards[OF get_child_nodes_reads set_disconnected_nodes_writes h2 children_h] ) by (simp add: set_disconnected_nodes_get_child_nodes) show "known_ptrs h'" using object_ptr_kinds_eq3 known_ptrs_preserved \known_ptrs h\ by blast have "known_ptr ptr" using assms(3) using children_h get_child_nodes_ptr_in_heap local.known_ptrs_known_ptr by blast have "type_wf h2" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h2] using set_disconnected_nodes_types_preserved type_wf by(auto simp add: reflp_def transp_def) then show "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_child_nodes_writes h'] using set_child_nodes_types_preserved by(auto simp add: reflp_def transp_def) have children_h': "h' \ get_child_nodes ptr \\<^sub>r remove1 child children_h" using assms(2) owner_document h2 disconnected_nodes_h children_h apply(auto simp add: remove_child_def split: if_splits)[1] apply(drule bind_returns_heap_E3) apply(auto split: if_splits)[1] apply(simp) apply(auto split: if_splits)[1] apply(drule bind_returns_heap_E3) apply(auto)[1] apply(simp) apply(drule bind_returns_heap_E3) apply(auto)[1] apply(simp) apply(drule bind_returns_heap_E4) apply(auto)[1] apply simp using \type_wf h2\ set_child_nodes_get_child_nodes \known_ptr ptr\ h' by blast have disconnected_nodes_h2: "h2 \ get_disconnected_nodes owner_document \\<^sub>r child # disconnected_nodes_h" using owner_document assms(2) h2 disconnected_nodes_h apply (auto simp add: remove_child_def split: if_splits)[1] apply(drule bind_returns_heap_E2) apply(auto split: if_splits)[1] apply(simp) by(auto simp add: local.set_disconnected_nodes_get_disconnected_nodes split: if_splits) then have disconnected_nodes_h': "h' \ get_disconnected_nodes owner_document \\<^sub>r child # disconnected_nodes_h" apply(rule reads_writes_separate_forwards[OF get_disconnected_nodes_reads set_child_nodes_writes h']) by (simp add: set_child_nodes_get_disconnected_nodes) moreover have "a_acyclic_heap h" using assms(1) by (simp add: heap_is_wellformed_def) have "parent_child_rel h' \ parent_child_rel h" proof (standard, safe) fix parent child assume a1: "(parent, child) \ parent_child_rel h'" then show "(parent, child) \ parent_child_rel h" proof (cases "parent = ptr") case True then show ?thesis using a1 remove_child_removes_parent[OF assms(1) assms(2)] children_h children_h' get_child_nodes_ptr_in_heap apply(auto simp add: parent_child_rel_def object_ptr_kinds_eq )[1] by (metis imageI notin_set_remove1) next case False then show ?thesis using a1 by(auto simp add: parent_child_rel_def object_ptr_kinds_eq3 children_eq2) qed qed then have "a_acyclic_heap h'" using \a_acyclic_heap h\ acyclic_heap_def acyclic_subset by blast moreover have "a_all_ptrs_in_heap h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_all_ptrs_in_heap h'" apply(auto simp add: a_all_ptrs_in_heap_def node_ptr_kinds_eq3 disconnected_nodes_eq)[1] apply (metis (no_types, lifting) \type_wf h'\ assms(2) assms(3) local.get_child_nodes_ok local.known_ptrs_known_ptr local.remove_child_children_subset notin_fset object_ptr_kinds_eq3 returns_result_select_result subset_code(1) type_wf) apply (metis (no_types, lifting) assms(2) disconnected_nodes_eq2 disconnected_nodes_h disconnected_nodes_h' document_ptr_kinds_eq3 finite_set_in local.remove_child_child_in_heap node_ptr_kinds_eq3 select_result_I2 set_ConsD subset_code(1)) done moreover have "a_owner_document_valid h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_owner_document_valid h'" apply(auto simp add: a_owner_document_valid_def object_ptr_kinds_eq3 document_ptr_kinds_eq3 node_ptr_kinds_eq3)[1] proof - fix node_ptr assume 0: "\node_ptr\fset (node_ptr_kinds h'). (\document_ptr. document_ptr |\| document_ptr_kinds h' \ node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r) \ (\parent_ptr. parent_ptr |\| object_ptr_kinds h' \ node_ptr \ set |h \ get_child_nodes parent_ptr|\<^sub>r)" and 1: "node_ptr |\| node_ptr_kinds h'" and 2: "\parent_ptr. parent_ptr |\| object_ptr_kinds h' \ node_ptr \ set |h' \ get_child_nodes parent_ptr|\<^sub>r" then show "\document_ptr. document_ptr |\| document_ptr_kinds h' \ node_ptr \ set |h' \ get_disconnected_nodes document_ptr|\<^sub>r" proof (cases "node_ptr = child") case True show ?thesis apply(rule exI[where x=owner_document]) using children_eq2 disconnected_nodes_eq2 children_h children_h' disconnected_nodes_h' True by (metis (no_types, lifting) get_disconnected_nodes_ptr_in_heap is_OK_returns_result_I list.set_intros(1) select_result_I2) next case False then show ?thesis using 0 1 2 children_eq2 children_h children_h' disconnected_nodes_eq2 disconnected_nodes_h disconnected_nodes_h' apply(auto simp add: children_eq2 disconnected_nodes_eq2 dest!: select_result_I2)[1] by (metis children_eq2 disconnected_nodes_eq2 finite_set_in in_set_remove1 list.set_intros(2)) qed qed moreover { have h0: "a_distinct_lists h" using assms(1) by (simp add: heap_is_wellformed_def) moreover have ha1: "(\x\set |h \ object_ptr_kinds_M|\<^sub>r. set |h \ get_child_nodes x|\<^sub>r) \ (\x\set |h \ document_ptr_kinds_M|\<^sub>r. set |h \ get_disconnected_nodes x|\<^sub>r) = {}" using \a_distinct_lists h\ unfolding a_distinct_lists_def by(auto) have ha2: "ptr |\| object_ptr_kinds h" using children_h get_child_nodes_ptr_in_heap by blast have ha3: "child \ set |h \ get_child_nodes ptr|\<^sub>r" using child_in_children_h children_h by(simp) have child_not_in: "\document_ptr. document_ptr |\| document_ptr_kinds h \ child \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r" using ha1 ha2 ha3 apply(simp) using IntI by fastforce moreover have "distinct |h \ object_ptr_kinds_M|\<^sub>r" apply(rule select_result_I) by(auto simp add: object_ptr_kinds_M_defs) moreover have "distinct |h \ document_ptr_kinds_M|\<^sub>r" apply(rule select_result_I) by(auto simp add: document_ptr_kinds_M_defs) ultimately have "a_distinct_lists h'" proof(simp (no_asm) add: a_distinct_lists_def, safe) assume 1: "a_distinct_lists h" and 3: "distinct |h \ object_ptr_kinds_M|\<^sub>r" assume 1: "a_distinct_lists h" and 3: "distinct |h \ object_ptr_kinds_M|\<^sub>r" have 4: "distinct (concat ((map (\ptr. |h \ get_child_nodes ptr|\<^sub>r) |h \ object_ptr_kinds_M|\<^sub>r)))" using 1 by(auto simp add: a_distinct_lists_def) show "distinct (concat (map (\ptr. |h' \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h')))))" proof(rule distinct_concat_map_I[OF 3[unfolded object_ptr_kinds_eq2], simplified]) fix x assume 5: "x |\| object_ptr_kinds h'" then have 6: "distinct |h \ get_child_nodes x|\<^sub>r" using 4 distinct_concat_map_E object_ptr_kinds_eq2 by fastforce obtain children where children: "h \ get_child_nodes x \\<^sub>r children" and distinct_children: "distinct children" by (metis "5" "6" type_wf assms(3) get_child_nodes_ok local.known_ptrs_known_ptr object_ptr_kinds_eq3 select_result_I) obtain children' where children': "h' \ get_child_nodes x \\<^sub>r children'" using children children_eq children_h' by fastforce then have "distinct children'" proof (cases "ptr = x") case True then show ?thesis using children distinct_children children_h children_h' by (metis children' distinct_remove1 returns_result_eq) next case False then show ?thesis using children distinct_children children_eq[OF False] using children' distinct_lists_children h0 using select_result_I2 by fastforce qed then show "distinct |h' \ get_child_nodes x|\<^sub>r" using children' by(auto simp add: ) next fix x y assume 5: "x |\| object_ptr_kinds h'" and 6: "y |\| object_ptr_kinds h'" and 7: "x \ y" obtain children_x where children_x: "h \ get_child_nodes x \\<^sub>r children_x" by (metis "5" type_wf assms(3) get_child_nodes_ok is_OK_returns_result_E local.known_ptrs_known_ptr object_ptr_kinds_eq3) obtain children_y where children_y: "h \ get_child_nodes y \\<^sub>r children_y" by (metis "6" type_wf assms(3) get_child_nodes_ok is_OK_returns_result_E local.known_ptrs_known_ptr object_ptr_kinds_eq3) obtain children_x' where children_x': "h' \ get_child_nodes x \\<^sub>r children_x'" using children_eq children_h' children_x by fastforce obtain children_y' where children_y': "h' \ get_child_nodes y \\<^sub>r children_y'" using children_eq children_h' children_y by fastforce have "distinct (concat (map (\ptr. |h \ get_child_nodes ptr|\<^sub>r) |h \ object_ptr_kinds_M|\<^sub>r))" using h0 by(auto simp add: a_distinct_lists_def) then have 8: "set children_x \ set children_y = {}" using "7" assms(1) children_x children_y local.heap_is_wellformed_one_parent by blast have "set children_x' \ set children_y' = {}" proof (cases "ptr = x") case True then have "ptr \ y" by(simp add: 7) have "children_x' = remove1 child children_x" using children_h children_h' children_x children_x' True returns_result_eq by fastforce moreover have "children_y' = children_y" using children_y children_y' children_eq[OF \ptr \ y\] by auto ultimately show ?thesis using 8 set_remove1_subset by fastforce next case False then show ?thesis proof (cases "ptr = y") case True have "children_y' = remove1 child children_y" using children_h children_h' children_y children_y' True returns_result_eq by fastforce moreover have "children_x' = children_x" using children_x children_x' children_eq[OF \ptr \ x\] by auto ultimately show ?thesis using 8 set_remove1_subset by fastforce next case False have "children_x' = children_x" using children_x children_x' children_eq[OF \ptr \ x\] by auto moreover have "children_y' = children_y" using children_y children_y' children_eq[OF \ptr \ y\] by auto ultimately show ?thesis using 8 by simp qed qed then show "set |h' \ get_child_nodes x|\<^sub>r \ set |h' \ get_child_nodes y|\<^sub>r = {}" using children_x' children_y' by (metis (no_types, lifting) select_result_I2) qed next assume 2: "distinct |h \ document_ptr_kinds_M|\<^sub>r" then have 4: "distinct (sorted_list_of_set (fset (document_ptr_kinds h')))" by simp have 3: "distinct (concat (map (\document_ptr. |h \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h')))))" using h0 by(simp add: a_distinct_lists_def document_ptr_kinds_eq3) show "distinct (concat (map (\document_ptr. |h' \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h')))))" proof(rule distinct_concat_map_I[OF 4[unfolded document_ptr_kinds_eq3]]) fix x assume 4: "x \ set (sorted_list_of_set (fset (document_ptr_kinds h')))" have 5: "distinct |h \ get_disconnected_nodes x|\<^sub>r" using distinct_lists_disconnected_nodes[OF h0] 4 get_disconnected_nodes_ok by (simp add: type_wf document_ptr_kinds_eq3 select_result_I) show "distinct |h' \ get_disconnected_nodes x|\<^sub>r" proof (cases "x = owner_document") case True have "child \ set |h \ get_disconnected_nodes x|\<^sub>r" using child_not_in document_ptr_kinds_eq2 "4" by fastforce moreover have "|h' \ get_disconnected_nodes x|\<^sub>r = child # |h \ get_disconnected_nodes x|\<^sub>r" using disconnected_nodes_h' disconnected_nodes_h unfolding True by(simp) ultimately show ?thesis using 5 unfolding True by simp next case False show ?thesis using "5" False disconnected_nodes_eq2 by auto qed next fix x y assume 4: "x \ set (sorted_list_of_set (fset (document_ptr_kinds h')))" and 5: "y \ set (sorted_list_of_set (fset (document_ptr_kinds h')))" and "x \ y" obtain disc_nodes_x where disc_nodes_x: "h \ get_disconnected_nodes x \\<^sub>r disc_nodes_x" using 4 get_disconnected_nodes_ok[OF \type_wf h\, of x] document_ptr_kinds_eq2 by auto obtain disc_nodes_y where disc_nodes_y: "h \ get_disconnected_nodes y \\<^sub>r disc_nodes_y" using 5 get_disconnected_nodes_ok[OF \type_wf h\, of y] document_ptr_kinds_eq2 by auto obtain disc_nodes_x' where disc_nodes_x': "h' \ get_disconnected_nodes x \\<^sub>r disc_nodes_x'" using 4 get_disconnected_nodes_ok[OF \type_wf h'\, of x] document_ptr_kinds_eq2 by auto obtain disc_nodes_y' where disc_nodes_y': "h' \ get_disconnected_nodes y \\<^sub>r disc_nodes_y'" using 5 get_disconnected_nodes_ok[OF \type_wf h'\, of y] document_ptr_kinds_eq2 by auto have "distinct (concat (map (\document_ptr. |h \ get_disconnected_nodes document_ptr|\<^sub>r) |h \ document_ptr_kinds_M|\<^sub>r))" using h0 by (simp add: a_distinct_lists_def) then have 6: "set disc_nodes_x \ set disc_nodes_y = {}" using \x \ y\ assms(1) disc_nodes_x disc_nodes_y local.heap_is_wellformed_one_disc_parent by blast have "set disc_nodes_x' \ set disc_nodes_y' = {}" proof (cases "x = owner_document") case True then have "y \ owner_document" using \x \ y\ by simp then have "disc_nodes_y' = disc_nodes_y" using disconnected_nodes_eq[OF \y \ owner_document\] disc_nodes_y disc_nodes_y' by auto have "disc_nodes_x' = child # disc_nodes_x" using disconnected_nodes_h' disc_nodes_x disc_nodes_x' True disconnected_nodes_h returns_result_eq by fastforce have "child \ set disc_nodes_y" using child_not_in disc_nodes_y 5 using document_ptr_kinds_eq2 by fastforce then show ?thesis apply(unfold \disc_nodes_x' = child # disc_nodes_x\ \disc_nodes_y' = disc_nodes_y\) using 6 by auto next case False then show ?thesis proof (cases "y = owner_document") case True then have "disc_nodes_x' = disc_nodes_x" using disconnected_nodes_eq[OF \x \ owner_document\] disc_nodes_x disc_nodes_x' by auto have "disc_nodes_y' = child # disc_nodes_y" using disconnected_nodes_h' disc_nodes_y disc_nodes_y' True disconnected_nodes_h returns_result_eq by fastforce have "child \ set disc_nodes_x" using child_not_in disc_nodes_x 4 using document_ptr_kinds_eq2 by fastforce then show ?thesis apply(unfold \disc_nodes_y' = child # disc_nodes_y\ \disc_nodes_x' = disc_nodes_x\) using 6 by auto next case False have "disc_nodes_x' = disc_nodes_x" using disconnected_nodes_eq[OF \x \ owner_document\] disc_nodes_x disc_nodes_x' by auto have "disc_nodes_y' = disc_nodes_y" using disconnected_nodes_eq[OF \y \ owner_document\] disc_nodes_y disc_nodes_y' by auto then show ?thesis apply(unfold \disc_nodes_y' = disc_nodes_y\ \disc_nodes_x' = disc_nodes_x\) using 6 by auto qed qed then show "set |h' \ get_disconnected_nodes x|\<^sub>r \ set |h' \ get_disconnected_nodes y|\<^sub>r = {}" using disc_nodes_x' disc_nodes_y' by auto qed next fix x xa xb assume 1: "xa \ fset (object_ptr_kinds h')" and 2: "x \ set |h' \ get_child_nodes xa|\<^sub>r" and 3: "xb \ fset (document_ptr_kinds h')" and 4: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" obtain disc_nodes where disc_nodes: "h \ get_disconnected_nodes xb \\<^sub>r disc_nodes" using 3 get_disconnected_nodes_ok[OF \type_wf h\, of xb] document_ptr_kinds_eq2 by auto obtain disc_nodes' where disc_nodes': "h' \ get_disconnected_nodes xb \\<^sub>r disc_nodes'" using 3 get_disconnected_nodes_ok[OF \type_wf h'\, of xb] document_ptr_kinds_eq2 by auto obtain children where children: "h \ get_child_nodes xa \\<^sub>r children" by (metis "1" type_wf assms(3) finite_set_in get_child_nodes_ok is_OK_returns_result_E local.known_ptrs_known_ptr object_ptr_kinds_eq3) obtain children' where children': "h' \ get_child_nodes xa \\<^sub>r children'" using children children_eq children_h' by fastforce have "\x. x \ set |h \ get_child_nodes xa|\<^sub>r \ x \ set |h \ get_disconnected_nodes xb|\<^sub>r \ False" using 1 3 apply(fold \ object_ptr_kinds h = object_ptr_kinds h'\) apply(fold \ document_ptr_kinds h = document_ptr_kinds h'\) using children disc_nodes h0 apply(auto simp add: a_distinct_lists_def)[1] by (metis (no_types, lifting) h0 local.distinct_lists_no_parent select_result_I2) then have 5: "\x. x \ set children \ x \ set disc_nodes \ False" using children disc_nodes by fastforce have 6: "|h' \ get_child_nodes xa|\<^sub>r = children'" using children' by (simp add: ) have 7: "|h' \ get_disconnected_nodes xb|\<^sub>r = disc_nodes'" using disc_nodes' by (simp add: ) have "False" proof (cases "xa = ptr") case True have "distinct children_h" using children_h distinct_lists_children h0 \known_ptr ptr\ by blast have "|h' \ get_child_nodes ptr|\<^sub>r = remove1 child children_h" using children_h' by(simp add: ) have "children = children_h" using True children children_h by auto show ?thesis using disc_nodes' children' 5 2 4 children_h \distinct children_h\ disconnected_nodes_h' apply(auto simp add: 6 7 \xa = ptr\ \|h' \ get_child_nodes ptr|\<^sub>r = remove1 child children_h\ \children = children_h\)[1] by (metis (no_types, lifting) disc_nodes disconnected_nodes_eq2 disconnected_nodes_h select_result_I2 set_ConsD) next case False have "children' = children" using children' children children_eq[OF False[symmetric]] by auto then show ?thesis proof (cases "xb = owner_document") case True then show ?thesis using disc_nodes disconnected_nodes_h disconnected_nodes_h' using "2" "4" "5" "6" "7" False \children' = children\ assms(1) child_in_children_h child_parent_dual children children_h disc_nodes' get_child_nodes_ptr_in_heap list.set_cases list.simps(3) option.simps(1) returns_result_eq set_ConsD by (metis (no_types, hide_lams) assms(3) type_wf) next case False then show ?thesis using "2" "4" "5" "6" "7" \children' = children\ disc_nodes disc_nodes' disconnected_nodes_eq returns_result_eq by metis qed qed then show "x \ {}" by simp qed } ultimately show "heap_is_wellformed h'" using heap_is_wellformed_def by blast qed lemma remove_heap_is_wellformed_preserved: assumes "heap_is_wellformed h" and "h \ remove child \\<^sub>h h'" and "known_ptrs h" and type_wf: "type_wf h" shows "type_wf h'" and "known_ptrs h'" and "heap_is_wellformed h'" using assms by(auto simp add: remove_def intro: remove_child_heap_is_wellformed_preserved elim!: bind_returns_heap_E2 split: option.splits) lemma remove_child_removes_child: assumes wellformed: "heap_is_wellformed h" and remove_child: "h \ remove_child ptr' child \\<^sub>h h'" and children: "h' \ get_child_nodes ptr \\<^sub>r children" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "child \ set children" proof - obtain owner_document children_h h2 disconnected_nodes_h where owner_document: "h \ get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r owner_document" and children_h: "h \ get_child_nodes ptr' \\<^sub>r children_h" and child_in_children_h: "child \ set children_h" and disconnected_nodes_h: "h \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h" and h2: "h \ set_disconnected_nodes owner_document (child # disconnected_nodes_h) \\<^sub>h h2" and h': "h2 \ set_child_nodes ptr' (remove1 child children_h) \\<^sub>h h'" using assms(2) apply(auto simp add: remove_child_def elim!: bind_returns_heap_E dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure] pure_returns_heap_eq[rotated, OF get_child_nodes_pure] split: if_splits)[1] using pure_returns_heap_eq by fastforce have "object_ptr_kinds h = object_ptr_kinds h'" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF remove_child_writes remove_child]) unfolding remove_child_locs_def using set_child_nodes_pointers_preserved set_disconnected_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) moreover have "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF remove_child_writes assms(2)] using set_child_nodes_types_preserved set_disconnected_nodes_types_preserved type_wf unfolding remove_child_locs_def apply(auto simp add: reflp_def transp_def)[1] by blast ultimately show ?thesis using remove_child_removes_parent remove_child_heap_is_wellformed_preserved child_parent_dual by (meson children known_ptrs local.known_ptrs_preserved option.distinct(1) remove_child returns_result_eq type_wf wellformed) qed lemma remove_child_removes_first_child: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_child_nodes ptr \\<^sub>r node_ptr # children" assumes "h \ remove_child ptr node_ptr \\<^sub>h h'" shows "h' \ get_child_nodes ptr \\<^sub>r children" proof - obtain h2 disc_nodes owner_document where "h \ get_owner_document (cast node_ptr) \\<^sub>r owner_document" and "h \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes" and h2: "h \ set_disconnected_nodes owner_document (node_ptr # disc_nodes) \\<^sub>h h2" and "h2 \ set_child_nodes ptr children \\<^sub>h h'" using assms(5) apply(auto simp add: remove_child_def dest!: bind_returns_heap_E3[rotated, OF assms(4) get_child_nodes_pure, rotated])[1] by(auto elim!: bind_returns_heap_E bind_returns_heap_E2[rotated,OF get_owner_document_pure, rotated] bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated]) have "known_ptr ptr" by (meson assms(3) assms(4) is_OK_returns_result_I get_child_nodes_ptr_in_heap known_ptrs_known_ptr) moreover have "h2 \ get_child_nodes ptr \\<^sub>r node_ptr # children" apply(rule reads_writes_separate_forwards[OF get_child_nodes_reads set_disconnected_nodes_writes h2 assms(4)]) using set_disconnected_nodes_get_child_nodes set_child_nodes_get_child_nodes_different_pointers by fast moreover have "type_wf h2" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h2] using \type_wf h\ set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) ultimately show ?thesis using set_child_nodes_get_child_nodes\h2 \ set_child_nodes ptr children \\<^sub>h h'\ by fast qed lemma remove_removes_child: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_child_nodes ptr \\<^sub>r node_ptr # children" assumes "h \ remove node_ptr \\<^sub>h h'" shows "h' \ get_child_nodes ptr \\<^sub>r children" proof - have "h \ get_parent node_ptr \\<^sub>r Some ptr" using child_parent_dual assms by fastforce show ?thesis using assms remove_child_removes_first_child by(auto simp add: remove_def dest!: bind_returns_heap_E3[rotated, OF \h \ get_parent node_ptr \\<^sub>r Some ptr\, rotated] bind_returns_heap_E3[rotated, OF assms(4) get_child_nodes_pure, rotated]) qed lemma remove_for_all_empty_children: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_child_nodes ptr \\<^sub>r children" assumes "h \ forall_M remove children \\<^sub>h h'" shows "h' \ get_child_nodes ptr \\<^sub>r []" using assms proof(induct children arbitrary: h h') case Nil then show ?case by simp next case (Cons a children) have "h \ get_parent a \\<^sub>r Some ptr" using child_parent_dual Cons by fastforce with Cons show ?case proof(auto elim!: bind_returns_heap_E)[1] fix h2 assume 0: "(\h h'. heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_child_nodes ptr \\<^sub>r children \ h \ forall_M remove children \\<^sub>h h' \ h' \ get_child_nodes ptr \\<^sub>r [])" and 1: "heap_is_wellformed h" and 2: "type_wf h" and 3: "known_ptrs h" and 4: "h \ get_child_nodes ptr \\<^sub>r a # children" and 5: "h \ get_parent a \\<^sub>r Some ptr" and 7: "h \ remove a \\<^sub>h h2" and 8: "h2 \ forall_M remove children \\<^sub>h h'" then have "h2 \ get_child_nodes ptr \\<^sub>r children" using remove_removes_child by blast moreover have "heap_is_wellformed h2" using 7 1 2 3 remove_child_heap_is_wellformed_preserved(3) by(auto simp add: remove_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_parent_pure, rotated] split: option.splits) moreover have "type_wf h2" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF remove_writes 7] using \type_wf h\ remove_child_types_preserved by(auto simp add: a_remove_child_locs_def reflp_def transp_def) moreover have "object_ptr_kinds h = object_ptr_kinds h2" using 7 apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF remove_writes]) using remove_child_pointers_preserved by (auto simp add: reflp_def transp_def) then have "known_ptrs h2" using 3 known_ptrs_preserved by blast ultimately show "h' \ get_child_nodes ptr \\<^sub>r []" using 0 8 by fast qed qed end locale l_remove_child_wf2 = l_type_wf + l_known_ptrs + l_remove_child_defs + l_heap_is_wellformed_defs + l_get_child_nodes_defs + l_remove_defs + assumes remove_child_preserves_type_wf: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ remove_child ptr child \\<^sub>h h' \ type_wf h'" assumes remove_child_preserves_known_ptrs: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ remove_child ptr child \\<^sub>h h' \ known_ptrs h'" assumes remove_child_heap_is_wellformed_preserved: "type_wf h \ known_ptrs h \ heap_is_wellformed h \ h \ remove_child ptr child \\<^sub>h h' \ heap_is_wellformed h'" assumes remove_preserves_type_wf: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ remove child \\<^sub>h h' \ type_wf h'" assumes remove_preserves_known_ptrs: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ remove child \\<^sub>h h' \ known_ptrs h'" assumes remove_heap_is_wellformed_preserved: "type_wf h \ known_ptrs h \ heap_is_wellformed h \ h \ remove child \\<^sub>h h' \ heap_is_wellformed h'" assumes remove_child_removes_child: "heap_is_wellformed h \ h \ remove_child ptr' child \\<^sub>h h' \ h' \ get_child_nodes ptr \\<^sub>r children \ known_ptrs h \ type_wf h \ child \ set children" assumes remove_child_removes_first_child: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_child_nodes ptr \\<^sub>r node_ptr # children \ h \ remove_child ptr node_ptr \\<^sub>h h' \ h' \ get_child_nodes ptr \\<^sub>r children" assumes remove_removes_child: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_child_nodes ptr \\<^sub>r node_ptr # children \ h \ remove node_ptr \\<^sub>h h' \ h' \ get_child_nodes ptr \\<^sub>r children" assumes remove_for_all_empty_children: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_child_nodes ptr \\<^sub>r children \ h \ forall_M remove children \\<^sub>h h' \ h' \ get_child_nodes ptr \\<^sub>r []" interpretation i_remove_child_wf2?: l_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs set_child_nodes set_child_nodes_locs get_parent get_parent_locs get_owner_document get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs remove_child remove_child_locs remove type_wf known_ptr known_ptrs heap_is_wellformed parent_child_rel by unfold_locales lemma remove_child_wf2_is_l_remove_child_wf2 [instances]: "l_remove_child_wf2 type_wf known_ptr known_ptrs remove_child heap_is_wellformed get_child_nodes remove" apply(auto simp add: l_remove_child_wf2_def l_remove_child_wf2_axioms_def instances)[1] using remove_child_heap_is_wellformed_preserved apply(fast, fast, fast) using remove_heap_is_wellformed_preserved apply(fast, fast, fast) using remove_child_removes_child apply fast using remove_child_removes_first_child apply fast using remove_removes_child apply fast using remove_for_all_empty_children apply fast done subsection \adopt\_node\ locale l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent_wf + l_get_owner_document_wf + l_remove_child_wf2 + l_heap_is_wellformed begin lemma adopt_node_removes_first_child: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ adopt_node owner_document node \\<^sub>h h'" assumes "h \ get_child_nodes ptr' \\<^sub>r node # children" shows "h' \ get_child_nodes ptr' \\<^sub>r children" proof - obtain old_document parent_opt h2 where old_document: "h \ get_owner_document (cast node) \\<^sub>r old_document" and parent_opt: "h \ get_parent node \\<^sub>r parent_opt" and h2: "h \ (case parent_opt of Some parent \ do { remove_child parent node } | None \ do { return ()}) \\<^sub>h h2" and h': "h2 \ (if owner_document \ old_document then do { old_disc_nodes \ get_disconnected_nodes old_document; set_disconnected_nodes old_document (remove1 node old_disc_nodes); disc_nodes \ get_disconnected_nodes owner_document; set_disconnected_nodes owner_document (node # disc_nodes) } else do { return () }) \\<^sub>h h'" using assms(4) by(auto simp add: adopt_node_def elim!: bind_returns_heap_E dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure] pure_returns_heap_eq[rotated, OF get_parent_pure]) have "h2 \ get_child_nodes ptr' \\<^sub>r children" using h2 remove_child_removes_first_child assms(1) assms(2) assms(3) assms(5) by (metis list.set_intros(1) local.child_parent_dual option.simps(5) parent_opt returns_result_eq) then show ?thesis using h' by(auto elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] dest!: reads_writes_separate_forwards[OF get_child_nodes_reads set_disconnected_nodes_writes] split: if_splits) qed lemma adopt_node_document_in_heap: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "h \ ok (adopt_node owner_document node)" shows "owner_document |\| document_ptr_kinds h" proof - obtain old_document parent_opt h2 h' where old_document: "h \ get_owner_document (cast node) \\<^sub>r old_document" and parent_opt: "h \ get_parent node \\<^sub>r parent_opt" and h2: "h \ (case parent_opt of Some parent \ do { remove_child parent node } | None \ do { return ()}) \\<^sub>h h2" and h': "h2 \ (if owner_document \ old_document then do { old_disc_nodes \ get_disconnected_nodes old_document; set_disconnected_nodes old_document (remove1 node old_disc_nodes); disc_nodes \ get_disconnected_nodes owner_document; set_disconnected_nodes owner_document (node # disc_nodes) } else do { return () }) \\<^sub>h h'" using assms(4) by(auto simp add: adopt_node_def elim!: bind_returns_heap_E dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure] pure_returns_heap_eq[rotated, OF get_parent_pure]) show ?thesis proof (cases "owner_document = old_document") case True then show ?thesis using old_document get_owner_document_owner_document_in_heap assms(1) assms(2) assms(3) by auto next case False then obtain h3 old_disc_nodes disc_nodes where old_disc_nodes: "h2 \ get_disconnected_nodes old_document \\<^sub>r old_disc_nodes" and h3: "h2 \ set_disconnected_nodes old_document (remove1 node old_disc_nodes) \\<^sub>h h3" and old_disc_nodes: "h3 \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes" and h': "h3 \ set_disconnected_nodes owner_document (node # disc_nodes) \\<^sub>h h'" using h' by(auto elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] ) then have "owner_document |\| document_ptr_kinds h3" by (meson is_OK_returns_result_I local.get_disconnected_nodes_ptr_in_heap) moreover have "object_ptr_kinds h = object_ptr_kinds h2" using h2 apply(simp split: option.splits) apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF remove_child_writes]) using remove_child_pointers_preserved by (auto simp add: reflp_def transp_def) moreover have "object_ptr_kinds h2 = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes h3]) using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) ultimately show ?thesis by(auto simp add: document_ptr_kinds_def) qed qed end locale l_adopt_node_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_root_node + l_get_owner_document_wf + l_remove_child_wf2 + l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma adopt_node_removes_child_step: assumes wellformed: "heap_is_wellformed h" and adopt_node: "h \ adopt_node owner_document node_ptr \\<^sub>h h2" and children: "h2 \ get_child_nodes ptr \\<^sub>r children" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "node_ptr \ set children" proof - obtain old_document parent_opt h' where old_document: "h \ get_owner_document (cast node_ptr) \\<^sub>r old_document" and parent_opt: "h \ get_parent node_ptr \\<^sub>r parent_opt" and h': "h \ (case parent_opt of Some parent \ remove_child parent node_ptr | None \ return () ) \\<^sub>h h'" using adopt_node get_parent_pure by(auto simp add: adopt_node_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated] bind_returns_heap_E2[rotated, OF get_parent_pure, rotated] bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] split: if_splits) then have "h' \ get_child_nodes ptr \\<^sub>r children" using adopt_node apply(auto simp add: adopt_node_def dest!: bind_returns_heap_E3[rotated, OF old_document, rotated] bind_returns_heap_E3[rotated, OF parent_opt, rotated] elim!: bind_returns_heap_E4[rotated, OF h', rotated])[1] apply(auto split: if_splits elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated])[1] apply (simp add: set_disconnected_nodes_get_child_nodes children reads_writes_preserved[OF get_child_nodes_reads set_disconnected_nodes_writes]) using children by blast show ?thesis proof(insert parent_opt h', induct parent_opt) case None then show ?case using child_parent_dual wellformed known_ptrs type_wf \h' \ get_child_nodes ptr \\<^sub>r children\ returns_result_eq by fastforce next case (Some option) then show ?case using remove_child_removes_child \h' \ get_child_nodes ptr \\<^sub>r children\ known_ptrs type_wf wellformed by auto qed qed lemma adopt_node_removes_child: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "h \ adopt_node owner_document node_ptr \\<^sub>h h'" shows "\ptr' children'. h' \ get_child_nodes ptr' \\<^sub>r children' \ node_ptr \ set children'" using adopt_node_removes_child_step assms by blast lemma adopt_node_preserves_wellformedness: assumes "heap_is_wellformed h" and "h \ adopt_node document_ptr child \\<^sub>h h'" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "heap_is_wellformed h'" and "known_ptrs h'" and "type_wf h'" proof - obtain old_document parent_opt h2 where old_document: "h \ get_owner_document (cast child) \\<^sub>r old_document" and parent_opt: "h \ get_parent child \\<^sub>r parent_opt" and h2: "h \ (case parent_opt of Some parent \ remove_child parent child | None \ return ()) \\<^sub>h h2" and h': "h2 \ (if document_ptr \ old_document then do { old_disc_nodes \ get_disconnected_nodes old_document; set_disconnected_nodes old_document (remove1 child old_disc_nodes); disc_nodes \ get_disconnected_nodes document_ptr; set_disconnected_nodes document_ptr (child # disc_nodes) } else do { return () }) \\<^sub>h h'" using assms(2) by(auto simp add: adopt_node_def elim!: bind_returns_heap_E dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure] pure_returns_heap_eq[rotated, OF get_parent_pure]) have object_ptr_kinds_h_eq3: "object_ptr_kinds h = object_ptr_kinds h2" using h2 apply(simp split: option.splits) apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF remove_child_writes]) using remove_child_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h: "\ptrs. h \ object_ptr_kinds_M \\<^sub>r ptrs = h2 \ object_ptr_kinds_M \\<^sub>r ptrs" unfolding object_ptr_kinds_M_defs by simp then have object_ptr_kinds_eq_h: "|h \ object_ptr_kinds_M|\<^sub>r = |h2 \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq_h: "|h \ node_ptr_kinds_M|\<^sub>r = |h2 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have wellformed_h2: "heap_is_wellformed h2" using h2 remove_child_heap_is_wellformed_preserved known_ptrs type_wf by (metis (no_types, lifting) assms(1) option.case_eq_if pure_returns_heap_eq return_pure) have "type_wf h2" using h2 remove_child_preserves_type_wf known_ptrs type_wf by (metis (no_types, lifting) assms(1) option.case_eq_if pure_returns_heap_eq return_pure) have "known_ptrs h2" using h2 remove_child_preserves_known_ptrs known_ptrs type_wf by (metis (no_types, lifting) assms(1) option.case_eq_if pure_returns_heap_eq return_pure) have "heap_is_wellformed h' \ known_ptrs h' \ type_wf h'" proof(cases "document_ptr = old_document") case True then show ?thesis using h' wellformed_h2 \type_wf h2\ \known_ptrs h2\ by auto next case False then obtain h3 old_disc_nodes disc_nodes_document_ptr_h3 where docs_neq: "document_ptr \ old_document" and old_disc_nodes: "h2 \ get_disconnected_nodes old_document \\<^sub>r old_disc_nodes" and h3: "h2 \ set_disconnected_nodes old_document (remove1 child old_disc_nodes) \\<^sub>h h3" and disc_nodes_document_ptr_h3: "h3 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_document_ptr_h3" and h': "h3 \ set_disconnected_nodes document_ptr (child # disc_nodes_document_ptr_h3) \\<^sub>h h'" using h' by(auto elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] ) have object_ptr_kinds_h2_eq3: "object_ptr_kinds h2 = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes h3]) using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h2: "\ptrs. h2 \ object_ptr_kinds_M \\<^sub>r ptrs = h3 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_eq_h2: "|h2 \ object_ptr_kinds_M|\<^sub>r = |h3 \ object_ptr_kinds_M|\<^sub>r" by(simp) then have node_ptr_kinds_eq_h2: "|h2 \ node_ptr_kinds_M|\<^sub>r = |h3 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast then have node_ptr_kinds_eq3_h2: "node_ptr_kinds h2 = node_ptr_kinds h3" by auto have document_ptr_kinds_eq2_h2: "|h2 \ document_ptr_kinds_M|\<^sub>r = |h3 \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_eq_h2 document_ptr_kinds_M_eq by auto then have document_ptr_kinds_eq3_h2: "document_ptr_kinds h2 = document_ptr_kinds h3" using object_ptr_kinds_eq_h2 document_ptr_kinds_M_eq by auto have children_eq_h2: "\ptr children. h2 \ get_child_nodes ptr \\<^sub>r children = h3 \ get_child_nodes ptr \\<^sub>r children" using get_child_nodes_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by (simp add: set_disconnected_nodes_get_child_nodes) then have children_eq2_h2: "\ptr. |h2 \ get_child_nodes ptr|\<^sub>r = |h3 \ get_child_nodes ptr|\<^sub>r" using select_result_eq by force have object_ptr_kinds_h3_eq3: "object_ptr_kinds h3 = object_ptr_kinds h'" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes h']) using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h3: "\ptrs. h3 \ object_ptr_kinds_M \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_eq_h3: "|h3 \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" by(simp) then have node_ptr_kinds_eq_h3: "|h3 \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast then have node_ptr_kinds_eq3_h3: "node_ptr_kinds h3 = node_ptr_kinds h'" by auto have document_ptr_kinds_eq2_h3: "|h3 \ document_ptr_kinds_M|\<^sub>r = |h' \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_eq_h3 document_ptr_kinds_M_eq by auto then have document_ptr_kinds_eq3_h3: "document_ptr_kinds h3 = document_ptr_kinds h'" using object_ptr_kinds_eq_h3 document_ptr_kinds_M_eq by auto have children_eq_h3: "\ptr children. h3 \ get_child_nodes ptr \\<^sub>r children = h' \ get_child_nodes ptr \\<^sub>r children" using get_child_nodes_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by (simp add: set_disconnected_nodes_get_child_nodes) then have children_eq2_h3: "\ptr. |h3 \ get_child_nodes ptr|\<^sub>r = |h' \ get_child_nodes ptr|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. old_document \ doc_ptr \ h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by (simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers) then have disconnected_nodes_eq2_h2: "\doc_ptr. old_document \ doc_ptr \ |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force obtain disc_nodes_old_document_h2 where disc_nodes_old_document_h2: "h2 \ get_disconnected_nodes old_document \\<^sub>r disc_nodes_old_document_h2" using old_disc_nodes by blast then have disc_nodes_old_document_h3: "h3 \ get_disconnected_nodes old_document \\<^sub>r remove1 child disc_nodes_old_document_h2" using h3 old_disc_nodes returns_result_eq set_disconnected_nodes_get_disconnected_nodes by fastforce have "distinct disc_nodes_old_document_h2" using disc_nodes_old_document_h2 local.heap_is_wellformed_disconnected_nodes_distinct wellformed_h2 by blast have "type_wf h2" proof (insert h2, induct parent_opt) case None then show ?case using type_wf by simp next case (Some option) then show ?case using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF remove_child_writes] type_wf remove_child_types_preserved by (simp add: reflp_def transp_def) qed then have "type_wf h3" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h3] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) then have "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h'] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) have "known_ptrs h3" using known_ptrs local.known_ptrs_preserved object_ptr_kinds_h2_eq3 object_ptr_kinds_h_eq3 by blast then have "known_ptrs h'" using local.known_ptrs_preserved object_ptr_kinds_h3_eq3 by blast have disconnected_nodes_eq_h3: "\doc_ptr disc_nodes. document_ptr \ doc_ptr \ h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by (simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers) then have disconnected_nodes_eq2_h3: "\doc_ptr. document_ptr \ doc_ptr \ |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have disc_nodes_document_ptr_h2: "h2 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_document_ptr_h3" using disconnected_nodes_eq_h2 docs_neq disc_nodes_document_ptr_h3 by auto have disc_nodes_document_ptr_h': " h' \ get_disconnected_nodes document_ptr \\<^sub>r child # disc_nodes_document_ptr_h3" using h' disc_nodes_document_ptr_h3 using set_disconnected_nodes_get_disconnected_nodes by blast have document_ptr_in_heap: "document_ptr |\| document_ptr_kinds h2" using disc_nodes_document_ptr_h3 document_ptr_kinds_eq2_h2 get_disconnected_nodes_ok assms(1) unfolding heap_is_wellformed_def using disc_nodes_document_ptr_h2 get_disconnected_nodes_ptr_in_heap by blast have old_document_in_heap: "old_document |\| document_ptr_kinds h2" using disc_nodes_old_document_h3 document_ptr_kinds_eq2_h2 get_disconnected_nodes_ok assms(1) unfolding heap_is_wellformed_def using get_disconnected_nodes_ptr_in_heap old_disc_nodes by blast have "child \ set disc_nodes_old_document_h2" proof (insert parent_opt h2, induct parent_opt) case None then have "h = h2" by(auto) moreover have "a_owner_document_valid h" using assms(1) heap_is_wellformed_def by(simp add: heap_is_wellformed_def) ultimately show ?case using old_document disc_nodes_old_document_h2 None(1) child_parent_dual[OF assms(1)] in_disconnected_nodes_no_parent assms(1) known_ptrs type_wf by blast next case (Some option) then show ?case apply(simp split: option.splits) using assms(1) disc_nodes_old_document_h2 old_document remove_child_in_disconnected_nodes known_ptrs by blast qed have "child \ set (remove1 child disc_nodes_old_document_h2)" using disc_nodes_old_document_h3 h3 known_ptrs wellformed_h2 \distinct disc_nodes_old_document_h2\ by auto have "child \ set disc_nodes_document_ptr_h3" proof - have "a_distinct_lists h2" using heap_is_wellformed_def wellformed_h2 by blast then have 0: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>r) |h2 \ document_ptr_kinds_M|\<^sub>r))" by(simp add: a_distinct_lists_def) show ?thesis using distinct_concat_map_E(1)[OF 0] \child \ set disc_nodes_old_document_h2\ disc_nodes_old_document_h2 disc_nodes_document_ptr_h2 by (meson \type_wf h2\ docs_neq known_ptrs local.get_owner_document_disconnected_nodes local.known_ptrs_preserved object_ptr_kinds_h_eq3 returns_result_eq wellformed_h2) qed have child_in_heap: "child |\| node_ptr_kinds h" using get_owner_document_ptr_in_heap[OF is_OK_returns_result_I[OF old_document]] node_ptr_kinds_commutes by blast have "a_acyclic_heap h2" using wellformed_h2 by (simp add: heap_is_wellformed_def) have "parent_child_rel h' \ parent_child_rel h2" proof fix x assume "x \ parent_child_rel h'" then show "x \ parent_child_rel h2" using object_ptr_kinds_h2_eq3 object_ptr_kinds_h3_eq3 children_eq2_h2 children_eq2_h3 mem_Collect_eq object_ptr_kinds_M_eq_h3 select_result_eq split_cong unfolding parent_child_rel_def by(simp) qed then have "a_acyclic_heap h'" using \a_acyclic_heap h2\ acyclic_heap_def acyclic_subset by blast moreover have "a_all_ptrs_in_heap h2" using wellformed_h2 by (simp add: heap_is_wellformed_def) then have "a_all_ptrs_in_heap h3" apply(auto simp add: a_all_ptrs_in_heap_def node_ptr_kinds_eq3_h2 children_eq_h2)[1] apply (simp add: children_eq2_h2 object_ptr_kinds_h2_eq3 subset_code(1)) by (metis (no_types, lifting) \child \ set disc_nodes_old_document_h2\ \type_wf h2\ disc_nodes_old_document_h2 disc_nodes_old_document_h3 disconnected_nodes_eq2_h2 document_ptr_kinds_eq3_h2 in_set_remove1 local.get_disconnected_nodes_ok local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq3_h2 returns_result_select_result select_result_I2 wellformed_h2) then have "a_all_ptrs_in_heap h'" apply(auto simp add: a_all_ptrs_in_heap_def node_ptr_kinds_eq3_h3 children_eq_h3)[1] apply (simp add: children_eq2_h3 object_ptr_kinds_h3_eq3 subset_code(1)) by (metis (no_types, lifting) \child \ set disc_nodes_old_document_h2\ disc_nodes_document_ptr_h' disc_nodes_document_ptr_h2 disc_nodes_old_document_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq3_h3 finite_set_in local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq3_h2 node_ptr_kinds_eq3_h3 select_result_I2 set_ConsD subset_code(1) wellformed_h2) moreover have "a_owner_document_valid h2" using wellformed_h2 by (simp add: heap_is_wellformed_def) then have "a_owner_document_valid h'" apply(simp add: a_owner_document_valid_def node_ptr_kinds_eq_h2 node_ptr_kinds_eq3_h3 object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 children_eq2_h2 children_eq2_h3 ) by (smt disc_nodes_document_ptr_h' disc_nodes_document_ptr_h2 disc_nodes_old_document_h2 disc_nodes_old_document_h3 disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_in_heap document_ptr_kinds_eq3_h2 document_ptr_kinds_eq3_h3 in_set_remove1 list.set_intros(1) list.set_intros(2) node_ptr_kinds_eq3_h2 node_ptr_kinds_eq3_h3 object_ptr_kinds_h2_eq3 object_ptr_kinds_h3_eq3 select_result_I2) have a_distinct_lists_h2: "a_distinct_lists h2" using wellformed_h2 by (simp add: heap_is_wellformed_def) then have "a_distinct_lists h'" apply(auto simp add: a_distinct_lists_def object_ptr_kinds_eq_h3 object_ptr_kinds_eq_h2 children_eq2_h2 children_eq2_h3)[1] proof - assume 1: "distinct (concat (map (\ptr. |h' \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h')))))" and 2: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h2)))))" and 3: "(\x\fset (object_ptr_kinds h'). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h2). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" show "distinct (concat (map (\document_ptr. |h' \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h')))))" proof(rule distinct_concat_map_I) show "distinct (sorted_list_of_set (fset (document_ptr_kinds h')))" by(auto simp add: document_ptr_kinds_M_def ) next fix x assume a1: "x \ set (sorted_list_of_set (fset (document_ptr_kinds h')))" have 4: "distinct |h2 \ get_disconnected_nodes x|\<^sub>r" using a_distinct_lists_h2 "2" a1 concat_map_all_distinct document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 by fastforce then show "distinct |h' \ get_disconnected_nodes x|\<^sub>r" proof (cases "old_document \ x") case True then show ?thesis proof (cases "document_ptr \ x") case True then show ?thesis using disconnected_nodes_eq2_h2[OF \old_document \ x\] disconnected_nodes_eq2_h3[OF \document_ptr \ x\] 4 by(auto) next case False then show ?thesis using disc_nodes_document_ptr_h3 disc_nodes_document_ptr_h' 4 \child \ set disc_nodes_document_ptr_h3\ by(auto simp add: disconnected_nodes_eq2_h2[OF \old_document \ x\] ) qed next case False then show ?thesis by (metis (no_types, hide_lams) \distinct disc_nodes_old_document_h2\ disc_nodes_old_document_h3 disconnected_nodes_eq2_h3 distinct_remove1 docs_neq select_result_I2) qed next fix x y assume a0: "x \ set (sorted_list_of_set (fset (document_ptr_kinds h')))" and a1: "y \ set (sorted_list_of_set (fset (document_ptr_kinds h')))" and a2: "x \ y" moreover have 5: "set |h2 \ get_disconnected_nodes x|\<^sub>r \ set |h2 \ get_disconnected_nodes y|\<^sub>r = {}" using 2 calculation by (auto simp add: document_ptr_kinds_eq3_h2 document_ptr_kinds_eq3_h3 dest: distinct_concat_map_E(1)) ultimately show "set |h' \ get_disconnected_nodes x|\<^sub>r \ set |h' \ get_disconnected_nodes y|\<^sub>r = {}" proof(cases "old_document = x") case True have "old_document \ y" using \x \ y\ \old_document = x\ by simp have "document_ptr \ x" using docs_neq \old_document = x\ by auto show ?thesis proof(cases "document_ptr = y") case True then show ?thesis using 5 True select_result_I2[OF disc_nodes_document_ptr_h'] select_result_I2[OF disc_nodes_document_ptr_h2] select_result_I2[OF disc_nodes_old_document_h2] select_result_I2[OF disc_nodes_old_document_h3] \old_document = x\ by (metis (no_types, lifting) \child \ set (remove1 child disc_nodes_old_document_h2)\ \document_ptr \ x\ disconnected_nodes_eq2_h3 disjoint_iff_not_equal notin_set_remove1 set_ConsD) next case False then show ?thesis using 5 select_result_I2[OF disc_nodes_document_ptr_h'] select_result_I2[OF disc_nodes_document_ptr_h2] select_result_I2[OF disc_nodes_old_document_h2] select_result_I2[OF disc_nodes_old_document_h3] disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 \old_document = x\ docs_neq \old_document \ y\ by (metis (no_types, lifting) disjoint_iff_not_equal notin_set_remove1) qed next case False then show ?thesis proof(cases "old_document = y") case True then show ?thesis proof(cases "document_ptr = x") case True show ?thesis using 5 select_result_I2[OF disc_nodes_document_ptr_h'] select_result_I2[OF disc_nodes_document_ptr_h2] select_result_I2[OF disc_nodes_old_document_h2] select_result_I2[OF disc_nodes_old_document_h3] \old_document \ x\ \old_document = y\ \document_ptr = x\ apply(simp) by (metis (no_types, lifting) \child \ set (remove1 child disc_nodes_old_document_h2)\ disconnected_nodes_eq2_h3 disjoint_iff_not_equal notin_set_remove1) next case False then show ?thesis using 5 select_result_I2[OF disc_nodes_document_ptr_h'] select_result_I2[OF disc_nodes_document_ptr_h2] select_result_I2[OF disc_nodes_old_document_h2] select_result_I2[OF disc_nodes_old_document_h3] \old_document \ x\ \old_document = y\ \document_ptr \ x\ by (metis (no_types, lifting) disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disjoint_iff_not_equal docs_neq notin_set_remove1) qed next case False have "set |h2 \ get_disconnected_nodes y|\<^sub>r \ set disc_nodes_old_document_h2 = {}" by (metis DocumentMonad.ptr_kinds_M_ok DocumentMonad.ptr_kinds_M_ptr_kinds False \type_wf h2\ a1 disc_nodes_old_document_h2 document_ptr_kinds_M_def document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 l_ptr_kinds_M.ptr_kinds_ptr_kinds_M local.get_disconnected_nodes_ok local.heap_is_wellformed_one_disc_parent returns_result_select_result wellformed_h2) then show ?thesis proof(cases "document_ptr = x") case True then have "document_ptr \ y" using \x \ y\ by auto have "set |h2 \ get_disconnected_nodes y|\<^sub>r \ set disc_nodes_old_document_h2 = {}" using \set |h2 \ get_disconnected_nodes y|\<^sub>r \ set disc_nodes_old_document_h2 = {}\ by blast then show ?thesis using 5 select_result_I2[OF disc_nodes_document_ptr_h'] select_result_I2[OF disc_nodes_document_ptr_h2] select_result_I2[OF disc_nodes_old_document_h2] select_result_I2[OF disc_nodes_old_document_h3] \old_document \ x\ \old_document \ y\ \document_ptr = x\ \document_ptr \ y\ \child \ set disc_nodes_old_document_h2\ disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 \set |h2 \ get_disconnected_nodes y|\<^sub>r \ set disc_nodes_old_document_h2 = {}\ by(auto) next case False then show ?thesis proof(cases "document_ptr = y") case True have f1: "set |h2 \ get_disconnected_nodes x|\<^sub>r \ set disc_nodes_document_ptr_h3 = {}" using 2 a1 document_ptr_in_heap document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 \document_ptr \ x\ select_result_I2[OF disc_nodes_document_ptr_h3, symmetric] disconnected_nodes_eq2_h2[OF docs_neq[symmetric], symmetric] by (simp add: "5" True) moreover have f1: "set |h2 \ get_disconnected_nodes x|\<^sub>r \ set |h2 \ get_disconnected_nodes old_document|\<^sub>r = {}" using 2 a1 old_document_in_heap document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 \old_document \ x\ by (metis (no_types, lifting) a0 distinct_concat_map_E(1) document_ptr_kinds_eq3_h2 document_ptr_kinds_eq3_h3 finite_fset fmember.rep_eq set_sorted_list_of_set) ultimately show ?thesis using 5 select_result_I2[OF disc_nodes_document_ptr_h'] select_result_I2[OF disc_nodes_old_document_h2] \old_document \ x\ \document_ptr \ x\ \document_ptr = y\ \child \ set disc_nodes_old_document_h2\ disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 by auto next case False then show ?thesis using 5 select_result_I2[OF disc_nodes_old_document_h2] \old_document \ x\ \document_ptr \ x\ \document_ptr \ y\ \child \ set disc_nodes_old_document_h2\ disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 by (metis \set |h2 \ get_disconnected_nodes y|\<^sub>r \ set disc_nodes_old_document_h2 = {}\ empty_iff inf.idem) qed qed qed qed qed next fix x xa xb assume 0: "distinct (concat (map (\ptr. |h' \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h')))))" and 1: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h2)))))" and 2: "(\x\fset (object_ptr_kinds h'). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h2). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" and 3: "xa |\| object_ptr_kinds h'" and 4: "x \ set |h' \ get_child_nodes xa|\<^sub>r" and 5: "xb |\| document_ptr_kinds h'" and 6: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" then show False using \child \ set disc_nodes_old_document_h2\ disc_nodes_document_ptr_h' disc_nodes_document_ptr_h2 disc_nodes_old_document_h2 disc_nodes_old_document_h3 disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 old_document_in_heap apply(auto)[1] apply(cases "xb = old_document") proof - assume a1: "xb = old_document" assume a2: "h2 \ get_disconnected_nodes old_document \\<^sub>r disc_nodes_old_document_h2" assume a3: "h3 \ get_disconnected_nodes old_document \\<^sub>r remove1 child disc_nodes_old_document_h2" assume a4: "x \ set |h' \ get_child_nodes xa|\<^sub>r" assume "document_ptr_kinds h2 = document_ptr_kinds h'" assume a5: "(\x\fset (object_ptr_kinds h'). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h'). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" have f6: "old_document |\| document_ptr_kinds h'" using a1 \xb |\| document_ptr_kinds h'\ by blast have f7: "|h2 \ get_disconnected_nodes old_document|\<^sub>r = disc_nodes_old_document_h2" using a2 by simp have "x \ set disc_nodes_old_document_h2" using f6 a3 a1 by (metis (no_types) \type_wf h'\ \x \ set |h' \ get_disconnected_nodes xb|\<^sub>r\ disconnected_nodes_eq_h3 docs_neq get_disconnected_nodes_ok returns_result_eq returns_result_select_result set_remove1_subset subsetCE) then have "set |h' \ get_child_nodes xa|\<^sub>r \ set |h2 \ get_disconnected_nodes xb|\<^sub>r = {}" using f7 f6 a5 a4 \xa |\| object_ptr_kinds h'\ by fastforce then show ?thesis using \x \ set disc_nodes_old_document_h2\ a1 a4 f7 by blast next assume a1: "xb \ old_document" assume a2: "h2 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_document_ptr_h3" assume a3: "h2 \ get_disconnected_nodes old_document \\<^sub>r disc_nodes_old_document_h2" assume a4: "xa |\| object_ptr_kinds h'" assume a5: "h' \ get_disconnected_nodes document_ptr \\<^sub>r child # disc_nodes_document_ptr_h3" assume a6: "old_document |\| document_ptr_kinds h'" assume a7: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" assume a8: "x \ set |h' \ get_child_nodes xa|\<^sub>r" assume a9: "document_ptr_kinds h2 = document_ptr_kinds h'" assume a10: "\doc_ptr. old_document \ doc_ptr \ |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" assume a11: "\doc_ptr. document_ptr \ doc_ptr \ |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" assume a12: "(\x\fset (object_ptr_kinds h'). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h'). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" have f13: "\d. d \ set |h' \ document_ptr_kinds_M|\<^sub>r \ h2 \ ok get_disconnected_nodes d" using a9 \type_wf h2\ get_disconnected_nodes_ok by simp then have f14: "|h2 \ get_disconnected_nodes old_document|\<^sub>r = disc_nodes_old_document_h2" using a6 a3 by simp have "x \ set |h2 \ get_disconnected_nodes xb|\<^sub>r" using a12 a8 a4 \xb |\| document_ptr_kinds h'\ by (meson UN_I disjoint_iff_not_equal fmember.rep_eq) then have "x = child" using f13 a11 a10 a7 a5 a2 a1 by (metis (no_types, lifting) select_result_I2 set_ConsD) then have "child \ set disc_nodes_old_document_h2" using f14 a12 a8 a6 a4 by (metis \type_wf h'\ adopt_node_removes_child assms(1) assms(2) type_wf get_child_nodes_ok known_ptrs local.known_ptrs_known_ptr object_ptr_kinds_h2_eq3 object_ptr_kinds_h3_eq3 object_ptr_kinds_h_eq3 returns_result_select_result) then show ?thesis using \child \ set disc_nodes_old_document_h2\ by fastforce qed qed ultimately show ?thesis using \type_wf h'\ \known_ptrs h'\ \a_owner_document_valid h'\ heap_is_wellformed_def by blast qed then show "heap_is_wellformed h'" and "known_ptrs h'" and "type_wf h'" by auto qed lemma adopt_node_node_in_disconnected_nodes: assumes wellformed: "heap_is_wellformed h" and adopt_node: "h \ adopt_node owner_document node_ptr \\<^sub>h h'" and "h' \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "node_ptr \ set disc_nodes" proof - obtain old_document parent_opt h2 where old_document: "h \ get_owner_document (cast node_ptr) \\<^sub>r old_document" and parent_opt: "h \ get_parent node_ptr \\<^sub>r parent_opt" and h2: "h \ (case parent_opt of Some parent \ remove_child parent node_ptr | None \ return ()) \\<^sub>h h2" and h': "h2 \ (if owner_document \ old_document then do { old_disc_nodes \ get_disconnected_nodes old_document; set_disconnected_nodes old_document (remove1 node_ptr old_disc_nodes); disc_nodes \ get_disconnected_nodes owner_document; set_disconnected_nodes owner_document (node_ptr # disc_nodes) } else do { return () }) \\<^sub>h h'" using assms(2) by(auto simp add: adopt_node_def elim!: bind_returns_heap_E dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure] pure_returns_heap_eq[rotated, OF get_parent_pure]) show ?thesis proof (cases "owner_document = old_document") case True then show ?thesis proof (insert parent_opt h2, induct parent_opt) case None then have "h = h'" using h2 h' by(auto) then show ?case using in_disconnected_nodes_no_parent assms None old_document by blast next case (Some parent) then show ?case using remove_child_in_disconnected_nodes known_ptrs True h' assms(3) old_document by auto qed next case False then show ?thesis using assms(3) h' list.set_intros(1) select_result_I2 set_disconnected_nodes_get_disconnected_nodes apply(auto elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated])[1] proof - fix x and h'a and xb assume a1: "h' \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes" assume a2: "\h document_ptr disc_nodes h'. h \ set_disconnected_nodes document_ptr disc_nodes \\<^sub>h h' \ h' \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes" assume "h'a \ set_disconnected_nodes owner_document (node_ptr # xb) \\<^sub>h h'" then have "node_ptr # xb = disc_nodes" using a2 a1 by (meson returns_result_eq) then show ?thesis by (meson list.set_intros(1)) qed qed qed end interpretation i_adopt_node_wf?: l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_owner_document get_parent get_parent_locs remove_child remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs adopt_node adopt_node_locs known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs set_child_nodes set_child_nodes_locs remove heap_is_wellformed parent_child_rel by(simp add: l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] interpretation i_adopt_node_wf2?: l_adopt_node_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_owner_document get_parent get_parent_locs remove_child remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs adopt_node adopt_node_locs known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs set_child_nodes set_child_nodes_locs remove heap_is_wellformed parent_child_rel get_root_node get_root_node_locs by(simp add: l_adopt_node_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_adopt_node_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_adopt_node_wf = l_heap_is_wellformed + l_known_ptrs + l_type_wf + l_adopt_node_defs + l_get_child_nodes_defs + l_get_disconnected_nodes_defs + assumes adopt_node_preserves_wellformedness: "heap_is_wellformed h \ h \ adopt_node document_ptr child \\<^sub>h h' \ known_ptrs h \ type_wf h \ heap_is_wellformed h'" assumes adopt_node_removes_child: "heap_is_wellformed h \ h \ adopt_node owner_document node_ptr \\<^sub>h h2 \ h2 \ get_child_nodes ptr \\<^sub>r children \ known_ptrs h \ type_wf h \ node_ptr \ set children" assumes adopt_node_node_in_disconnected_nodes: "heap_is_wellformed h \ h \ adopt_node owner_document node_ptr \\<^sub>h h' \ h' \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes \ known_ptrs h \ type_wf h \ node_ptr \ set disc_nodes" assumes adopt_node_removes_first_child: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ adopt_node owner_document node \\<^sub>h h' \ h \ get_child_nodes ptr' \\<^sub>r node # children \ h' \ get_child_nodes ptr' \\<^sub>r children" assumes adopt_node_document_in_heap: "heap_is_wellformed h \ known_ptrs h \ type_wf h \ h \ ok (adopt_node owner_document node) \ owner_document |\| document_ptr_kinds h" assumes adopt_node_preserves_type_wf: "heap_is_wellformed h \ h \ adopt_node document_ptr child \\<^sub>h h' \ known_ptrs h \ type_wf h \ type_wf h'" assumes adopt_node_preserves_known_ptrs: "heap_is_wellformed h \ h \ adopt_node document_ptr child \\<^sub>h h' \ known_ptrs h \ type_wf h \ known_ptrs h'" lemma adopt_node_wf_is_l_adopt_node_wf [instances]: "l_adopt_node_wf type_wf known_ptr heap_is_wellformed parent_child_rel get_child_nodes get_disconnected_nodes known_ptrs adopt_node" using heap_is_wellformed_is_l_heap_is_wellformed known_ptrs_is_l_known_ptrs apply(auto simp add: l_adopt_node_wf_def l_adopt_node_wf_axioms_def)[1] using adopt_node_preserves_wellformedness apply blast using adopt_node_removes_child apply blast using adopt_node_node_in_disconnected_nodes apply blast using adopt_node_removes_first_child apply blast using adopt_node_document_in_heap apply blast using adopt_node_preserves_wellformedness apply blast using adopt_node_preserves_wellformedness apply blast done subsection \insert\_before\ locale l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_adopt_node_wf + l_set_disconnected_nodes_get_child_nodes + l_heap_is_wellformed begin lemma insert_before_removes_child: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "ptr \ ptr'" assumes "h \ insert_before ptr node child \\<^sub>h h'" assumes "h \ get_child_nodes ptr' \\<^sub>r node # children" shows "h' \ get_child_nodes ptr' \\<^sub>r children" proof - obtain owner_document h2 h3 disc_nodes reference_child where "h \ (if Some node = child then a_next_sibling node else return child) \\<^sub>r reference_child" and "h \ get_owner_document ptr \\<^sub>r owner_document" and h2: "h \ adopt_node owner_document node \\<^sub>h h2" and "h2 \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes" and h3: "h2 \ set_disconnected_nodes owner_document (remove1 node disc_nodes) \\<^sub>h h3" and h': "h3 \ a_insert_node ptr node reference_child \\<^sub>h h'" using assms(5) by(auto simp add: insert_before_def a_ensure_pre_insertion_validity_def elim!: bind_returns_heap_E bind_returns_result_E bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated] bind_returns_heap_E2[rotated, OF get_parent_pure, rotated] bind_returns_heap_E2[rotated, OF get_ancestors_pure, rotated] bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated] bind_returns_heap_E2[rotated, OF next_sibling_pure, rotated] bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] split: if_splits option.splits) have "h2 \ get_child_nodes ptr' \\<^sub>r children" using h2 adopt_node_removes_first_child assms(1) assms(2) assms(3) assms(6) by simp then have "h3 \ get_child_nodes ptr' \\<^sub>r children" using h3 by(auto simp add: set_disconnected_nodes_get_child_nodes dest!: reads_writes_separate_forwards[OF get_child_nodes_reads set_disconnected_nodes_writes]) then show ?thesis using h' assms(4) apply(auto simp add: a_insert_node_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated])[1] by(auto simp add: set_child_nodes_get_child_nodes_different_pointers elim!: reads_writes_separate_forwards[OF get_child_nodes_reads set_child_nodes_writes]) qed end locale l_insert_before_wf = l_heap_is_wellformed_defs + l_type_wf + l_known_ptrs + l_insert_before_defs + l_get_child_nodes_defs + assumes insert_before_removes_child: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ ptr \ ptr' \ h \ insert_before ptr node child \\<^sub>h h' \ h \ get_child_nodes ptr' \\<^sub>r node # children \ h' \ get_child_nodes ptr' \\<^sub>r children" interpretation i_insert_before_wf?: l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_parent get_parent_locs get_child_nodes get_child_nodes_locs set_child_nodes set_child_nodes_locs get_ancestors get_ancestors_locs adopt_node adopt_node_locs set_disconnected_nodes set_disconnected_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_owner_document insert_before insert_before_locs append_child type_wf known_ptr known_ptrs heap_is_wellformed parent_child_rel by(simp add: l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma insert_before_wf_is_l_insert_before_wf [instances]: "l_insert_before_wf heap_is_wellformed type_wf known_ptr known_ptrs insert_before get_child_nodes" apply(auto simp add: l_insert_before_wf_def l_insert_before_wf_axioms_def instances)[1] using insert_before_removes_child apply fast done locale l_insert_before_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_child_nodes_get_disconnected_nodes + l_remove_child + l_get_root_node_wf + l_set_disconnected_nodes_get_disconnected_nodes_wf + l_set_disconnected_nodes_get_ancestors + l_get_ancestors_wf + l_get_owner_document + l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_owner_document_wf begin lemma insert_before_preserves_acyclitity: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "h \ insert_before ptr node child \\<^sub>h h'" shows "acyclic (parent_child_rel h')" proof - obtain ancestors reference_child owner_document h2 h3 disconnected_nodes_h2 where ancestors: "h \ get_ancestors ptr \\<^sub>r ancestors" and node_not_in_ancestors: "cast node \ set ancestors" and reference_child: "h \ (if Some node = child then a_next_sibling node else return child) \\<^sub>r reference_child" and owner_document: "h \ get_owner_document ptr \\<^sub>r owner_document" and h2: "h \ adopt_node owner_document node \\<^sub>h h2" and disconnected_nodes_h2: "h2 \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h2" and h3: "h2 \ set_disconnected_nodes owner_document (remove1 node disconnected_nodes_h2) \\<^sub>h h3" and h': "h3 \ a_insert_node ptr node reference_child \\<^sub>h h'" using assms(4) by(auto simp add: insert_before_def a_ensure_pre_insertion_validity_def elim!: bind_returns_heap_E bind_returns_result_E bind_returns_heap_E2[rotated, OF get_parent_pure, rotated] bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated] bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] bind_returns_heap_E2[rotated, OF get_ancestors_pure, rotated] bind_returns_heap_E2[rotated, OF next_sibling_pure, rotated] bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated] split: if_splits option.splits) have "known_ptr ptr" by (meson get_owner_document_ptr_in_heap is_OK_returns_result_I assms l_known_ptrs.known_ptrs_known_ptr l_known_ptrs_axioms owner_document) have "type_wf h2" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF adopt_node_writes h2] using assms adopt_node_types_preserved by(auto simp add: a_remove_child_locs_def reflp_def transp_def) then have "type_wf h3" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h3] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) then have "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF insert_node_writes h'] using set_child_nodes_types_preserved by(auto simp add: reflp_def transp_def) have object_ptr_kinds_M_eq3_h: "object_ptr_kinds h = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF adopt_node_writes h2]) using adopt_node_pointers_preserved apply blast by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h: "\ptrs. h \ object_ptr_kinds_M \\<^sub>r ptrs = h2 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs ) then have object_ptr_kinds_M_eq2_h: "|h \ object_ptr_kinds_M|\<^sub>r = |h2 \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h: "|h \ node_ptr_kinds_M|\<^sub>r = |h2 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have "known_ptrs h2" using assms object_ptr_kinds_M_eq3_h known_ptrs_preserved by blast have wellformed_h2: "heap_is_wellformed h2" using adopt_node_preserves_wellformedness[OF assms(1) h2] assms by simp have object_ptr_kinds_M_eq3_h2: "object_ptr_kinds h2 = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes h3]) unfolding a_remove_child_locs_def using set_disconnected_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h2: "\ptrs. h2 \ object_ptr_kinds_M \\<^sub>r ptrs = h3 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_M_eq2_h2: "|h2 \ object_ptr_kinds_M|\<^sub>r = |h3 \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h2: "|h2 \ node_ptr_kinds_M|\<^sub>r = |h3 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have document_ptr_kinds_eq2_h2: "|h2 \ document_ptr_kinds_M|\<^sub>r = |h3 \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_M_eq2_h2 document_ptr_kinds_M_eq by auto have "known_ptrs h3" using object_ptr_kinds_M_eq3_h2 known_ptrs_preserved \known_ptrs h2\ by blast have object_ptr_kinds_M_eq3_h': "object_ptr_kinds h3 = object_ptr_kinds h'" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF insert_node_writes h']) unfolding a_remove_child_locs_def using set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h3: "\ptrs. h3 \ object_ptr_kinds_M \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_M_eq2_h3: "|h3 \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h3: "|h3 \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have document_ptr_kinds_eq2_h3: "|h3 \ document_ptr_kinds_M|\<^sub>r = |h' \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_M_eq2_h3 document_ptr_kinds_M_eq by auto have "known_ptrs h'" using object_ptr_kinds_M_eq3_h' known_ptrs_preserved \known_ptrs h3\ by blast have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. owner_document \ doc_ptr \ h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by (auto simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers) then have disconnected_nodes_eq2_h2: "\doc_ptr. doc_ptr \ owner_document \ |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have disconnected_nodes_h3: "h3 \ get_disconnected_nodes owner_document \\<^sub>r remove1 node disconnected_nodes_h2" using h3 set_disconnected_nodes_get_disconnected_nodes by blast have disconnected_nodes_eq_h3: "\doc_ptr disc_nodes. h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads insert_node_writes h' apply(rule reads_writes_preserved) using set_child_nodes_get_disconnected_nodes by fast then have disconnected_nodes_eq2_h3: "\doc_ptr. |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. h2 \ get_child_nodes ptr' \\<^sub>r children = h3 \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by (auto simp add: set_disconnected_nodes_get_child_nodes) then have children_eq2_h2: "\ptr'. |h2 \ get_child_nodes ptr'|\<^sub>r = |h3 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have children_eq_h3: "\ptr' children. ptr \ ptr' \ h3 \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads insert_node_writes h' apply(rule reads_writes_preserved) by (auto simp add: set_child_nodes_get_child_nodes_different_pointers) then have children_eq2_h3: "\ptr'. ptr \ ptr' \ |h3 \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force obtain children_h3 where children_h3: "h3 \ get_child_nodes ptr \\<^sub>r children_h3" using h' a_insert_node_def by auto have children_h': "h' \ get_child_nodes ptr \\<^sub>r insert_before_list node reference_child children_h3" using h' \type_wf h3\ \known_ptr ptr\ by(auto simp add: a_insert_node_def elim!: bind_returns_heap_E2 dest!: set_child_nodes_get_child_nodes returns_result_eq[OF children_h3]) have ptr_in_heap: "ptr |\| object_ptr_kinds h3" using children_h3 get_child_nodes_ptr_in_heap by blast have node_in_heap: "node |\| node_ptr_kinds h" using h2 adopt_node_child_in_heap by fast have child_not_in_any_children: "\p children. h2 \ get_child_nodes p \\<^sub>r children \ node \ set children" using assms h2 adopt_node_removes_child by auto have "node \ set disconnected_nodes_h2" using disconnected_nodes_h2 h2 adopt_node_node_in_disconnected_nodes assms(1) \type_wf h\ \known_ptrs h\ by blast have node_not_in_disconnected_nodes: "\d. d |\| document_ptr_kinds h3 \ node \ set |h3 \ get_disconnected_nodes d|\<^sub>r" proof - fix d assume "d |\| document_ptr_kinds h3" show "node \ set |h3 \ get_disconnected_nodes d|\<^sub>r" proof (cases "d = owner_document") case True then show ?thesis using disconnected_nodes_h2 wellformed_h2 h3 remove_from_disconnected_nodes_removes wellformed_h2 \d |\| document_ptr_kinds h3\ disconnected_nodes_h3 by fastforce next case False then have "set |h2 \ get_disconnected_nodes d|\<^sub>r \ set |h2 \ get_disconnected_nodes owner_document|\<^sub>r = {}" using distinct_concat_map_E(1) wellformed_h2 by (metis (no_types, lifting) \d |\| document_ptr_kinds h3\ \type_wf h2\ disconnected_nodes_h2 document_ptr_kinds_M_def document_ptr_kinds_eq2_h2 l_ptr_kinds_M.ptr_kinds_ptr_kinds_M local.get_disconnected_nodes_ok local.heap_is_wellformed_one_disc_parent returns_result_select_result select_result_I2) then show ?thesis using disconnected_nodes_eq2_h2[OF False] \node \ set disconnected_nodes_h2\ disconnected_nodes_h2 by fastforce qed qed have "cast node \ ptr" using ancestors node_not_in_ancestors get_ancestors_ptr by fast obtain ancestors_h2 where ancestors_h2: "h2 \ get_ancestors ptr \\<^sub>r ancestors_h2" using get_ancestors_ok object_ptr_kinds_M_eq2_h2 \known_ptrs h2\ \type_wf h2\ by (metis is_OK_returns_result_E object_ptr_kinds_M_eq3_h2 ptr_in_heap wellformed_h2) have ancestors_h3: "h3 \ get_ancestors ptr \\<^sub>r ancestors_h2" using get_ancestors_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_separate_forwards) using \heap_is_wellformed h2\ ancestors_h2 by (auto simp add: set_disconnected_nodes_get_ancestors) have node_not_in_ancestors_h2: "cast node \ set ancestors_h2" apply(rule get_ancestors_remains_not_in_ancestors[OF assms(1) wellformed_h2 ancestors ancestors_h2]) using adopt_node_children_subset using h2 \known_ptrs h\ \ type_wf h\ apply(blast) using node_not_in_ancestors apply(blast) using object_ptr_kinds_M_eq3_h apply(blast) using \known_ptrs h\ apply(blast) using \type_wf h\ apply(blast) using \type_wf h2\ by blast have "acyclic (parent_child_rel h2)" using wellformed_h2 by (simp add: heap_is_wellformed_def acyclic_heap_def) then have "acyclic (parent_child_rel h3)" by(auto simp add: parent_child_rel_def object_ptr_kinds_M_eq3_h2 children_eq2_h2) moreover have "cast node \ {x. (x, ptr) \ (parent_child_rel h2)\<^sup>*}" using adopt_node_removes_child using ancestors node_not_in_ancestors using \known_ptrs h2\ \type_wf h2\ ancestors_h2 local.get_ancestors_parent_child_rel node_not_in_ancestors_h2 wellformed_h2 by blast then have "cast node \ {x. (x, ptr) \ (parent_child_rel h3)\<^sup>*}" by(auto simp add: parent_child_rel_def object_ptr_kinds_M_eq3_h2 children_eq2_h2) moreover have "parent_child_rel h' = insert (ptr, cast node) ((parent_child_rel h3))" using children_h3 children_h' ptr_in_heap apply(auto simp add: parent_child_rel_def object_ptr_kinds_M_eq3_h' children_eq2_h3 insert_before_list_node_in_set)[1] apply (metis (no_types, lifting) children_eq2_h3 insert_before_list_in_set select_result_I2) by (metis (no_types, lifting) children_eq2_h3 imageI insert_before_list_in_set select_result_I2) ultimately show "acyclic (parent_child_rel h')" by (auto simp add: heap_is_wellformed_def) qed lemma insert_before_heap_is_wellformed_preserved: assumes wellformed: "heap_is_wellformed h" and insert_before: "h \ insert_before ptr node child \\<^sub>h h'" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "heap_is_wellformed h'" and "type_wf h'" and "known_ptrs h'" proof - obtain ancestors reference_child owner_document h2 h3 disconnected_nodes_h2 where ancestors: "h \ get_ancestors ptr \\<^sub>r ancestors" and node_not_in_ancestors: "cast node \ set ancestors" and reference_child: "h \ (if Some node = child then a_next_sibling node else return child) \\<^sub>r reference_child" and owner_document: "h \ get_owner_document ptr \\<^sub>r owner_document" and h2: "h \ adopt_node owner_document node \\<^sub>h h2" and disconnected_nodes_h2: "h2 \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h2" and h3: "h2 \ set_disconnected_nodes owner_document (remove1 node disconnected_nodes_h2) \\<^sub>h h3" and h': "h3 \ a_insert_node ptr node reference_child \\<^sub>h h'" using assms(2) by(auto simp add: insert_before_def a_ensure_pre_insertion_validity_def elim!: bind_returns_heap_E bind_returns_result_E bind_returns_heap_E2[rotated, OF get_parent_pure, rotated] bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated] bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] bind_returns_heap_E2[rotated, OF get_ancestors_pure, rotated] bind_returns_heap_E2[rotated, OF next_sibling_pure, rotated] bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated] split: if_splits option.splits) have "known_ptr ptr" by (meson get_owner_document_ptr_in_heap is_OK_returns_result_I known_ptrs l_known_ptrs.known_ptrs_known_ptr l_known_ptrs_axioms owner_document) have "type_wf h2" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF adopt_node_writes h2] using type_wf adopt_node_types_preserved by(auto simp add: a_remove_child_locs_def reflp_def transp_def) then have "type_wf h3" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h3] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) then show "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF insert_node_writes h'] using set_child_nodes_types_preserved by(auto simp add: reflp_def transp_def) have object_ptr_kinds_M_eq3_h: "object_ptr_kinds h = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF adopt_node_writes h2]) using adopt_node_pointers_preserved apply blast by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h: "\ptrs. h \ object_ptr_kinds_M \\<^sub>r ptrs = h2 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs ) then have object_ptr_kinds_M_eq2_h: "|h \ object_ptr_kinds_M|\<^sub>r = |h2 \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h: "|h \ node_ptr_kinds_M|\<^sub>r = |h2 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have "known_ptrs h2" using known_ptrs object_ptr_kinds_M_eq3_h known_ptrs_preserved by blast have wellformed_h2: "heap_is_wellformed h2" using adopt_node_preserves_wellformedness[OF wellformed h2] known_ptrs type_wf . have object_ptr_kinds_M_eq3_h2: "object_ptr_kinds h2 = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes h3]) unfolding a_remove_child_locs_def using set_disconnected_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h2: "\ptrs. h2 \ object_ptr_kinds_M \\<^sub>r ptrs = h3 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_M_eq2_h2: "|h2 \ object_ptr_kinds_M|\<^sub>r = |h3 \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h2: "|h2 \ node_ptr_kinds_M|\<^sub>r = |h3 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have document_ptr_kinds_eq2_h2: "|h2 \ document_ptr_kinds_M|\<^sub>r = |h3 \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_M_eq2_h2 document_ptr_kinds_M_eq by auto have "known_ptrs h3" using object_ptr_kinds_M_eq3_h2 known_ptrs_preserved \known_ptrs h2\ by blast have object_ptr_kinds_M_eq3_h': "object_ptr_kinds h3 = object_ptr_kinds h'" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF insert_node_writes h']) unfolding a_remove_child_locs_def using set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h3: "\ptrs. h3 \ object_ptr_kinds_M \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_M_eq2_h3: "|h3 \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h3: "|h3 \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have document_ptr_kinds_eq2_h3: "|h3 \ document_ptr_kinds_M|\<^sub>r = |h' \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_M_eq2_h3 document_ptr_kinds_M_eq by auto show "known_ptrs h'" using object_ptr_kinds_M_eq3_h' known_ptrs_preserved \known_ptrs h3\ by blast have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. owner_document \ doc_ptr \ h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by (auto simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers) then have disconnected_nodes_eq2_h2: "\doc_ptr. doc_ptr \ owner_document \ |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have disconnected_nodes_h3: "h3 \ get_disconnected_nodes owner_document \\<^sub>r remove1 node disconnected_nodes_h2" using h3 set_disconnected_nodes_get_disconnected_nodes by blast have disconnected_nodes_eq_h3: "\doc_ptr disc_nodes. h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads insert_node_writes h' apply(rule reads_writes_preserved) using set_child_nodes_get_disconnected_nodes by fast then have disconnected_nodes_eq2_h3: "\doc_ptr. |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. h2 \ get_child_nodes ptr' \\<^sub>r children = h3 \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by (auto simp add: set_disconnected_nodes_get_child_nodes) then have children_eq2_h2: "\ptr'. |h2 \ get_child_nodes ptr'|\<^sub>r = |h3 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have children_eq_h3: "\ptr' children. ptr \ ptr' \ h3 \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads insert_node_writes h' apply(rule reads_writes_preserved) by (auto simp add: set_child_nodes_get_child_nodes_different_pointers) then have children_eq2_h3: "\ptr'. ptr \ ptr' \ |h3 \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force obtain children_h3 where children_h3: "h3 \ get_child_nodes ptr \\<^sub>r children_h3" using h' a_insert_node_def by auto have children_h': "h' \ get_child_nodes ptr \\<^sub>r insert_before_list node reference_child children_h3" using h' \type_wf h3\ \known_ptr ptr\ by(auto simp add: a_insert_node_def elim!: bind_returns_heap_E2 dest!: set_child_nodes_get_child_nodes returns_result_eq[OF children_h3]) have ptr_in_heap: "ptr |\| object_ptr_kinds h3" using children_h3 get_child_nodes_ptr_in_heap by blast have node_in_heap: "node |\| node_ptr_kinds h" using h2 adopt_node_child_in_heap by fast have child_not_in_any_children: "\p children. h2 \ get_child_nodes p \\<^sub>r children \ node \ set children" using wellformed h2 adopt_node_removes_child \type_wf h\ \known_ptrs h\ by auto have "node \ set disconnected_nodes_h2" using disconnected_nodes_h2 h2 adopt_node_node_in_disconnected_nodes assms(1) \type_wf h\ \known_ptrs h\ by blast have node_not_in_disconnected_nodes: "\d. d |\| document_ptr_kinds h3 \ node \ set |h3 \ get_disconnected_nodes d|\<^sub>r" proof - fix d assume "d |\| document_ptr_kinds h3" show "node \ set |h3 \ get_disconnected_nodes d|\<^sub>r" proof (cases "d = owner_document") case True then show ?thesis using disconnected_nodes_h2 wellformed_h2 h3 remove_from_disconnected_nodes_removes wellformed_h2 \d |\| document_ptr_kinds h3\ disconnected_nodes_h3 by fastforce next case False then have "set |h2 \ get_disconnected_nodes d|\<^sub>r \ set |h2 \ get_disconnected_nodes owner_document|\<^sub>r = {}" using distinct_concat_map_E(1) wellformed_h2 by (metis (no_types, lifting) \d |\| document_ptr_kinds h3\ \type_wf h2\ disconnected_nodes_h2 document_ptr_kinds_M_def document_ptr_kinds_eq2_h2 l_ptr_kinds_M.ptr_kinds_ptr_kinds_M local.get_disconnected_nodes_ok local.heap_is_wellformed_one_disc_parent returns_result_select_result select_result_I2) then show ?thesis using disconnected_nodes_eq2_h2[OF False] \node \ set disconnected_nodes_h2\ disconnected_nodes_h2 by fastforce qed qed have "cast node \ ptr" using ancestors node_not_in_ancestors get_ancestors_ptr by fast obtain ancestors_h2 where ancestors_h2: "h2 \ get_ancestors ptr \\<^sub>r ancestors_h2" using get_ancestors_ok object_ptr_kinds_M_eq2_h2 \known_ptrs h2\ \type_wf h2\ by (metis is_OK_returns_result_E object_ptr_kinds_M_eq3_h2 ptr_in_heap wellformed_h2) have ancestors_h3: "h3 \ get_ancestors ptr \\<^sub>r ancestors_h2" using get_ancestors_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_separate_forwards) using \heap_is_wellformed h2\ ancestors_h2 by (auto simp add: set_disconnected_nodes_get_ancestors) have node_not_in_ancestors_h2: "cast node \ set ancestors_h2" apply(rule get_ancestors_remains_not_in_ancestors[OF assms(1) wellformed_h2 ancestors ancestors_h2]) using adopt_node_children_subset using h2 \known_ptrs h\ \ type_wf h\ apply(blast) using node_not_in_ancestors apply(blast) using object_ptr_kinds_M_eq3_h apply(blast) using \known_ptrs h\ apply(blast) using \type_wf h\ apply(blast) using \type_wf h2\ by blast moreover have "a_acyclic_heap h'" proof - have "acyclic (parent_child_rel h2)" using wellformed_h2 by (simp add: heap_is_wellformed_def acyclic_heap_def) then have "acyclic (parent_child_rel h3)" by(auto simp add: parent_child_rel_def object_ptr_kinds_M_eq3_h2 children_eq2_h2) moreover have "cast node \ {x. (x, ptr) \ (parent_child_rel h2)\<^sup>*}" using get_ancestors_parent_child_rel node_not_in_ancestors_h2 \known_ptrs h2\ \type_wf h2\ using ancestors_h2 wellformed_h2 by blast then have "cast node \ {x. (x, ptr) \ (parent_child_rel h3)\<^sup>*}" by(auto simp add: parent_child_rel_def object_ptr_kinds_M_eq3_h2 children_eq2_h2) moreover have "parent_child_rel h' = insert (ptr, cast node) ((parent_child_rel h3))" using children_h3 children_h' ptr_in_heap apply(auto simp add: parent_child_rel_def object_ptr_kinds_M_eq3_h' children_eq2_h3 insert_before_list_node_in_set)[1] apply (metis (no_types, lifting) children_eq2_h3 insert_before_list_in_set select_result_I2) by (metis (no_types, lifting) children_eq2_h3 imageI insert_before_list_in_set select_result_I2) ultimately show ?thesis by(auto simp add: acyclic_heap_def) qed moreover have "a_all_ptrs_in_heap h2" using wellformed_h2 by (simp add: heap_is_wellformed_def) have "a_all_ptrs_in_heap h'" proof - have "a_all_ptrs_in_heap h3" using \a_all_ptrs_in_heap h2\ apply(auto simp add: a_all_ptrs_in_heap_def object_ptr_kinds_M_eq2_h2 node_ptr_kinds_eq2_h2 children_eq_h2)[1] using disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3 using node_ptr_kinds_eq2_h2 apply auto[1] apply (metis \known_ptrs h2\ \type_wf h3\ children_eq_h2 local.get_child_nodes_ok local.heap_is_wellformed_children_in_heap local.known_ptrs_known_ptr object_ptr_kinds_M_eq3_h2 returns_result_select_result wellformed_h2) by (metis (no_types, lifting) disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3 document_ptr_kinds_commutes finite_set_in node_ptr_kinds_commutes object_ptr_kinds_M_eq3_h2 select_result_I2 set_remove1_subset subsetD) have "set children_h3 \ set |h' \ node_ptr_kinds_M|\<^sub>r" using children_h3 \a_all_ptrs_in_heap h3\ apply(auto simp add: a_all_ptrs_in_heap_def node_ptr_kinds_eq2_h3)[1] by (metis children_eq_h2 l_heap_is_wellformed.heap_is_wellformed_children_in_heap local.l_heap_is_wellformed_axioms node_ptr_kinds_commutes object_ptr_kinds_M_eq3_h' object_ptr_kinds_M_eq3_h2 wellformed_h2) then have "set (insert_before_list node reference_child children_h3) \ set |h' \ node_ptr_kinds_M|\<^sub>r" using node_in_heap apply(auto simp add: node_ptr_kinds_eq2_h node_ptr_kinds_eq2_h2 node_ptr_kinds_eq2_h3)[1] by (metis (no_types, hide_lams) contra_subsetD finite_set_in insert_before_list_in_set node_ptr_kinds_commutes object_ptr_kinds_M_eq3_h object_ptr_kinds_M_eq3_h' object_ptr_kinds_M_eq3_h2) then show ?thesis using \a_all_ptrs_in_heap h3\ apply(auto simp add: object_ptr_kinds_M_eq3_h' a_all_ptrs_in_heap_def node_ptr_kinds_def node_ptr_kinds_eq2_h3 disconnected_nodes_eq_h3)[1] using children_eq_h3 children_h' apply (metis (no_types, lifting) children_eq2_h3 finite_set_in select_result_I2 subsetD) by (metis (no_types) \type_wf h'\ disconnected_nodes_eq2_h3 disconnected_nodes_eq_h3 finite_set_in is_OK_returns_result_I local.get_disconnected_nodes_ok local.get_disconnected_nodes_ptr_in_heap returns_result_select_result subsetD) qed moreover have "a_distinct_lists h2" using wellformed_h2 by (simp add: heap_is_wellformed_def) then have "a_distinct_lists h3" proof(auto simp add: a_distinct_lists_def object_ptr_kinds_M_eq2_h2 document_ptr_kinds_eq2_h2 children_eq2_h2 intro!: distinct_concat_map_I)[1] fix x assume 1: "x |\| document_ptr_kinds h3" and 2: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" show "distinct |h3 \ get_disconnected_nodes x|\<^sub>r" using distinct_concat_map_E(2)[OF 2] select_result_I2[OF disconnected_nodes_h3] disconnected_nodes_eq2_h2 select_result_I2[OF disconnected_nodes_h2] 1 by (metis (full_types) distinct_remove1 finite_fset fmember.rep_eq set_sorted_list_of_set) next fix x y xa assume 1: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" and 2: "x |\| document_ptr_kinds h3" and 3: "y |\| document_ptr_kinds h3" and 4: "x \ y" and 5: "xa \ set |h3 \ get_disconnected_nodes x|\<^sub>r" and 6: "xa \ set |h3 \ get_disconnected_nodes y|\<^sub>r" show False proof (cases "x = owner_document") case True then have "y \ owner_document" using 4 by simp show ?thesis using distinct_concat_map_E(1)[OF 1] using 2 3 4 5 6 select_result_I2[OF disconnected_nodes_h3] select_result_I2[OF disconnected_nodes_h2] apply(auto simp add: True disconnected_nodes_eq2_h2[OF \y \ owner_document\])[1] by (metis (no_types, hide_lams) disconnected_nodes_eq2_h2 disjoint_iff_not_equal notin_set_remove1) next case False then show ?thesis proof (cases "y = owner_document") case True then show ?thesis using distinct_concat_map_E(1)[OF 1] using 2 3 4 5 6 select_result_I2[OF disconnected_nodes_h3] select_result_I2[OF disconnected_nodes_h2] apply(auto simp add: True disconnected_nodes_eq2_h2[OF \x \ owner_document\])[1] by (metis (no_types, hide_lams) disconnected_nodes_eq2_h2 disjoint_iff_not_equal notin_set_remove1) next case False then show ?thesis using distinct_concat_map_E(1)[OF 1, simplified, OF 2 3 4] 5 6 using disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3 disjoint_iff_not_equal finite_fset fmember.rep_eq notin_set_remove1 select_result_I2 set_sorted_list_of_set by (metis (no_types, lifting)) qed qed next fix x xa xb assume 1: "(\x\fset (object_ptr_kinds h3). set |h3 \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h3). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" and 2: "xa |\| object_ptr_kinds h3" and 3: "x \ set |h3 \ get_child_nodes xa|\<^sub>r" and 4: "xb |\| document_ptr_kinds h3" and 5: "x \ set |h3 \ get_disconnected_nodes xb|\<^sub>r" have 6: "set |h3 \ get_child_nodes xa|\<^sub>r \ set |h2 \ get_disconnected_nodes xb|\<^sub>r = {}" using 1 2 4 by (metis \type_wf h2\ children_eq2_h2 document_ptr_kinds_commutes known_ptrs local.get_child_nodes_ok local.get_disconnected_nodes_ok local.heap_is_wellformed_children_disc_nodes_different local.known_ptrs_known_ptr object_ptr_kinds_M_eq3_h object_ptr_kinds_M_eq3_h2 returns_result_select_result wellformed_h2) show False proof (cases "xb = owner_document") case True then show ?thesis using select_result_I2[OF disconnected_nodes_h3,folded select_result_I2[OF disconnected_nodes_h2]] by (metis (no_types, lifting) "3" "5" "6" disjoint_iff_not_equal notin_set_remove1) next case False show ?thesis using 2 3 4 5 6 unfolding disconnected_nodes_eq2_h2[OF False] by auto qed qed then have "a_distinct_lists h'" proof(auto simp add: a_distinct_lists_def document_ptr_kinds_eq2_h3 object_ptr_kinds_M_eq2_h3 disconnected_nodes_eq2_h3 intro!: distinct_concat_map_I)[1] fix x assume 1: "distinct (concat (map (\ptr. |h3 \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h')))))" and 2: "x |\| object_ptr_kinds h'" have 3: "\p. p |\| object_ptr_kinds h' \ distinct |h3 \ get_child_nodes p|\<^sub>r" using 1 by (auto elim: distinct_concat_map_E) show "distinct |h' \ get_child_nodes x|\<^sub>r" proof(cases "ptr = x") case True show ?thesis using 3[OF 2] children_h3 children_h' by(auto simp add: True insert_before_list_distinct dest: child_not_in_any_children[unfolded children_eq_h2]) next case False show ?thesis using children_eq2_h3[OF False] 3[OF 2] by auto qed next fix x y xa assume 1: "distinct (concat (map (\ptr. |h3 \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h')))))" and 2: "x |\| object_ptr_kinds h'" and 3: "y |\| object_ptr_kinds h'" and 4: "x \ y" and 5: "xa \ set |h' \ get_child_nodes x|\<^sub>r" and 6: "xa \ set |h' \ get_child_nodes y|\<^sub>r" have 7:"set |h3 \ get_child_nodes x|\<^sub>r \ set |h3 \ get_child_nodes y|\<^sub>r = {}" using distinct_concat_map_E(1)[OF 1] 2 3 4 by auto show False proof (cases "ptr = x") case True then have "ptr \ y" using 4 by simp then show ?thesis using children_h3 children_h' child_not_in_any_children[unfolded children_eq_h2] 5 6 apply(auto simp add: True children_eq2_h3[OF \ptr \ y\])[1] by (metis (no_types, hide_lams) "3" "7" \type_wf h3\ children_eq2_h3 disjoint_iff_not_equal get_child_nodes_ok insert_before_list_in_set known_ptrs local.known_ptrs_known_ptr object_ptr_kinds_M_eq3_h object_ptr_kinds_M_eq3_h' object_ptr_kinds_M_eq3_h2 returns_result_select_result select_result_I2) next case False then show ?thesis proof (cases "ptr = y") case True then show ?thesis using children_h3 children_h' child_not_in_any_children[unfolded children_eq_h2] 5 6 apply(auto simp add: True children_eq2_h3[OF \ptr \ x\])[1] by (metis (no_types, hide_lams) "2" "4" "7" IntI \known_ptrs h3\ \type_wf h'\ children_eq_h3 empty_iff insert_before_list_in_set local.get_child_nodes_ok local.known_ptrs_known_ptr object_ptr_kinds_M_eq3_h' returns_result_select_result select_result_I2) next case False then show ?thesis using children_eq2_h3[OF \ptr \ x\] children_eq2_h3[OF \ptr \ y\] 5 6 7 by auto qed qed next fix x xa xb assume 1: " (\x\fset (object_ptr_kinds h'). set |h3 \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h'). set |h' \ get_disconnected_nodes x|\<^sub>r) = {} " and 2: "xa |\| object_ptr_kinds h'" and 3: "x \ set |h' \ get_child_nodes xa|\<^sub>r" and 4: "xb |\| document_ptr_kinds h'" and 5: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" have 6: "set |h3 \ get_child_nodes xa|\<^sub>r \ set |h' \ get_disconnected_nodes xb|\<^sub>r = {}" using 1 2 3 4 5 proof - have "\h d. \ type_wf h \ d |\| document_ptr_kinds h \ h \ ok get_disconnected_nodes d" using local.get_disconnected_nodes_ok by satx then have "h' \ ok get_disconnected_nodes xb" using "4" \type_wf h'\ by fastforce then have f1: "h3 \ get_disconnected_nodes xb \\<^sub>r |h' \ get_disconnected_nodes xb|\<^sub>r" by (simp add: disconnected_nodes_eq_h3) have "xa |\| object_ptr_kinds h3" using "2" object_ptr_kinds_M_eq3_h' by blast then show ?thesis using f1 \local.a_distinct_lists h3\ local.distinct_lists_no_parent by fastforce qed show False proof (cases "ptr = xa") case True show ?thesis using 6 node_not_in_disconnected_nodes 3 4 5 select_result_I2[OF children_h'] select_result_I2[OF children_h3] True disconnected_nodes_eq2_h3 by (metis (no_types, lifting) "2" DocumentMonad.ptr_kinds_ptr_kinds_M \a_distinct_lists h3\ \type_wf h'\ disconnected_nodes_eq_h3 distinct_lists_no_parent document_ptr_kinds_eq2_h3 get_disconnected_nodes_ok insert_before_list_in_set object_ptr_kinds_M_eq3_h' returns_result_select_result) next case False then show ?thesis using 1 2 3 4 5 children_eq2_h3[OF False] by fastforce qed qed moreover have "a_owner_document_valid h2" using wellformed_h2 by (simp add: heap_is_wellformed_def) then have "a_owner_document_valid h'" apply(auto simp add: a_owner_document_valid_def object_ptr_kinds_M_eq2_h2 object_ptr_kinds_M_eq2_h3 node_ptr_kinds_eq2_h2 node_ptr_kinds_eq2_h3 document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 children_eq2_h2)[1] apply(auto simp add: document_ptr_kinds_eq2_h2[simplified] document_ptr_kinds_eq2_h3[simplified] object_ptr_kinds_M_eq2_h2[simplified] object_ptr_kinds_M_eq2_h3[simplified] node_ptr_kinds_eq2_h2[simplified] node_ptr_kinds_eq2_h3[simplified])[1] apply(auto simp add: disconnected_nodes_eq2_h3[symmetric])[1] by (smt children_eq2_h3 children_h' children_h3 disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3 finite_set_in in_set_remove1 insert_before_list_in_set object_ptr_kinds_M_eq3_h' ptr_in_heap select_result_I2) ultimately show "heap_is_wellformed h'" by (simp add: heap_is_wellformed_def) qed lemma adopt_node_children_remain_distinct: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "h \ adopt_node owner_document node_ptr \\<^sub>h h'" shows "\ptr' children'. h' \ get_child_nodes ptr' \\<^sub>r children' \ distinct children'" using assms(1) assms(2) assms(3) assms(4) local.adopt_node_preserves_wellformedness local.heap_is_wellformed_children_distinct by blast lemma insert_node_children_remain_distinct: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "h \ a_insert_node ptr new_child reference_child_opt \\<^sub>h h'" assumes "h \ get_child_nodes ptr \\<^sub>r children" assumes "new_child \ set children" shows "\children'. h' \ get_child_nodes ptr \\<^sub>r children' \ distinct children'" proof - fix children' assume a1: "h' \ get_child_nodes ptr \\<^sub>r children'" have "h' \ get_child_nodes ptr \\<^sub>r (insert_before_list new_child reference_child_opt children)" using assms(4) assms(5) apply(auto simp add: a_insert_node_def elim!: bind_returns_heap_E)[1] using returns_result_eq set_child_nodes_get_child_nodes assms(2) assms(3) by (metis is_OK_returns_result_I local.get_child_nodes_ptr_in_heap local.get_child_nodes_pure local.known_ptrs_known_ptr pure_returns_heap_eq) moreover have "a_distinct_lists h" using assms local.heap_is_wellformed_def by blast then have "\children. h \ get_child_nodes ptr \\<^sub>r children \ distinct children" using assms local.heap_is_wellformed_children_distinct by blast ultimately show "h' \ get_child_nodes ptr \\<^sub>r children' \ distinct children'" using assms(5) assms(6) insert_before_list_distinct returns_result_eq by fastforce qed lemma insert_before_children_remain_distinct: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "h \ insert_before ptr new_child child_opt \\<^sub>h h'" shows "\ptr' children'. h' \ get_child_nodes ptr' \\<^sub>r children' \ distinct children'" proof - obtain reference_child owner_document h2 h3 disconnected_nodes_h2 where reference_child: "h \ (if Some new_child = child_opt then a_next_sibling new_child else return child_opt) \\<^sub>r reference_child" and owner_document: "h \ get_owner_document ptr \\<^sub>r owner_document" and h2: "h \ adopt_node owner_document new_child \\<^sub>h h2" and disconnected_nodes_h2: "h2 \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h2" and h3: "h2 \ set_disconnected_nodes owner_document (remove1 new_child disconnected_nodes_h2) \\<^sub>h h3" and h': "h3 \ a_insert_node ptr new_child reference_child \\<^sub>h h'" using assms(4) by(auto simp add: insert_before_def a_ensure_pre_insertion_validity_def elim!: bind_returns_heap_E bind_returns_result_E bind_returns_heap_E2[rotated, OF get_parent_pure, rotated] bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated] bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] bind_returns_heap_E2[rotated, OF get_ancestors_pure, rotated] bind_returns_heap_E2[rotated, OF next_sibling_pure, rotated] bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated] split: if_splits option.splits) have "\ptr children. h2 \ get_child_nodes ptr \\<^sub>r children \ distinct children" using adopt_node_children_remain_distinct using assms(1) assms(2) assms(3) h2 by blast moreover have "\ptr children. h2 \ get_child_nodes ptr \\<^sub>r children \ new_child \ set children" using adopt_node_removes_child using assms(1) assms(2) assms(3) h2 by blast moreover have "\ptr children. h2 \ get_child_nodes ptr \\<^sub>r children = h3 \ get_child_nodes ptr \\<^sub>r children" using get_child_nodes_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by (auto simp add: set_disconnected_nodes_get_child_nodes) ultimately show "\ptr children. h' \ get_child_nodes ptr \\<^sub>r children \ distinct children" using insert_node_children_remain_distinct by (meson assms(1) assms(2) assms(3) assms(4) insert_before_heap_is_wellformed_preserved(1) local.heap_is_wellformed_children_distinct) qed lemma insert_before_removes_child: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "h \ insert_before ptr node child \\<^sub>h h'" assumes "ptr \ ptr'" shows "\children'. h' \ get_child_nodes ptr' \\<^sub>r children' \ node \ set children'" proof - fix children' assume a1: "h' \ get_child_nodes ptr' \\<^sub>r children'" obtain ancestors reference_child owner_document h2 h3 disconnected_nodes_h2 where ancestors: "h \ get_ancestors ptr \\<^sub>r ancestors" and node_not_in_ancestors: "cast node \ set ancestors" and reference_child: "h \ (if Some node = child then a_next_sibling node else return child) \\<^sub>r reference_child" and owner_document: "h \ get_owner_document ptr \\<^sub>r owner_document" and h2: "h \ adopt_node owner_document node \\<^sub>h h2" and disconnected_nodes_h2: "h2 \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h2" and h3: "h2 \ set_disconnected_nodes owner_document (remove1 node disconnected_nodes_h2) \\<^sub>h h3" and h': "h3 \ a_insert_node ptr node reference_child \\<^sub>h h'" using assms(4) by(auto simp add: insert_before_def a_ensure_pre_insertion_validity_def elim!: bind_returns_heap_E bind_returns_result_E bind_returns_heap_E2[rotated, OF get_parent_pure, rotated] bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated] bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] bind_returns_heap_E2[rotated, OF get_ancestors_pure, rotated] bind_returns_heap_E2[rotated, OF next_sibling_pure, rotated] bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated] split: if_splits option.splits) have "known_ptr ptr" by (meson get_owner_document_ptr_in_heap is_OK_returns_result_I assms(2) l_known_ptrs.known_ptrs_known_ptr l_known_ptrs_axioms owner_document) have "type_wf h2" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF adopt_node_writes h2] using assms(3) adopt_node_types_preserved by(auto simp add: a_remove_child_locs_def reflp_def transp_def) then have "type_wf h3" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h3] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) then have "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF insert_node_writes h'] using set_child_nodes_types_preserved by(auto simp add: reflp_def transp_def) have object_ptr_kinds_M_eq3_h: "object_ptr_kinds h = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF adopt_node_writes h2]) using adopt_node_pointers_preserved apply blast by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h: "\ptrs. h \ object_ptr_kinds_M \\<^sub>r ptrs = h2 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs ) then have object_ptr_kinds_M_eq2_h: "|h \ object_ptr_kinds_M|\<^sub>r = |h2 \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h: "|h \ node_ptr_kinds_M|\<^sub>r = |h2 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have "known_ptrs h2" using assms object_ptr_kinds_M_eq3_h known_ptrs_preserved by blast have wellformed_h2: "heap_is_wellformed h2" using adopt_node_preserves_wellformedness[OF assms(1) h2] assms by simp have object_ptr_kinds_M_eq3_h2: "object_ptr_kinds h2 = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes h3]) unfolding a_remove_child_locs_def using set_disconnected_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h2: "\ptrs. h2 \ object_ptr_kinds_M \\<^sub>r ptrs = h3 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_M_eq2_h2: "|h2 \ object_ptr_kinds_M|\<^sub>r = |h3 \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h2: "|h2 \ node_ptr_kinds_M|\<^sub>r = |h3 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have document_ptr_kinds_eq2_h2: "|h2 \ document_ptr_kinds_M|\<^sub>r = |h3 \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_M_eq2_h2 document_ptr_kinds_M_eq by auto have "known_ptrs h3" using object_ptr_kinds_M_eq3_h2 known_ptrs_preserved \known_ptrs h2\ by blast have object_ptr_kinds_M_eq3_h': "object_ptr_kinds h3 = object_ptr_kinds h'" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF insert_node_writes h']) unfolding a_remove_child_locs_def using set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h3: "\ptrs. h3 \ object_ptr_kinds_M \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_M_eq2_h3: "|h3 \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h3: "|h3 \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have document_ptr_kinds_eq2_h3: "|h3 \ document_ptr_kinds_M|\<^sub>r = |h' \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_M_eq2_h3 document_ptr_kinds_M_eq by auto have "known_ptrs h'" using object_ptr_kinds_M_eq3_h' known_ptrs_preserved \known_ptrs h3\ by blast have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. owner_document \ doc_ptr \ h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by (auto simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers) then have disconnected_nodes_eq2_h2: "\doc_ptr. doc_ptr \ owner_document \ |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have disconnected_nodes_h3: "h3 \ get_disconnected_nodes owner_document \\<^sub>r remove1 node disconnected_nodes_h2" using h3 set_disconnected_nodes_get_disconnected_nodes by blast have disconnected_nodes_eq_h3: "\doc_ptr disc_nodes. h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads insert_node_writes h' apply(rule reads_writes_preserved) using set_child_nodes_get_disconnected_nodes by fast then have disconnected_nodes_eq2_h3: "\doc_ptr. |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. h2 \ get_child_nodes ptr' \\<^sub>r children = h3 \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by (auto simp add: set_disconnected_nodes_get_child_nodes) then have children_eq2_h2: "\ptr'. |h2 \ get_child_nodes ptr'|\<^sub>r = |h3 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have children_eq_h3: "\ptr' children. ptr \ ptr' \ h3 \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads insert_node_writes h' apply(rule reads_writes_preserved) by (auto simp add: set_child_nodes_get_child_nodes_different_pointers) then have children_eq2_h3: "\ptr'. ptr \ ptr' \ |h3 \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force obtain children_h3 where children_h3: "h3 \ get_child_nodes ptr \\<^sub>r children_h3" using h' a_insert_node_def by auto have children_h': "h' \ get_child_nodes ptr \\<^sub>r insert_before_list node reference_child children_h3" using h' \type_wf h3\ \known_ptr ptr\ by(auto simp add: a_insert_node_def elim!: bind_returns_heap_E2 dest!: set_child_nodes_get_child_nodes returns_result_eq[OF children_h3]) have ptr_in_heap: "ptr |\| object_ptr_kinds h3" using children_h3 get_child_nodes_ptr_in_heap by blast have node_in_heap: "node |\| node_ptr_kinds h" using h2 adopt_node_child_in_heap by fast have child_not_in_any_children: "\p children. h2 \ get_child_nodes p \\<^sub>r children \ node \ set children" using assms(1) assms(2) assms(3) h2 local.adopt_node_removes_child by blast show "node \ set children'" using a1 assms(5) child_not_in_any_children children_eq_h2 children_eq_h3 by blast qed lemma ensure_pre_insertion_validity_ok: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "ptr |\| object_ptr_kinds h" assumes "\is_character_data_ptr_kind parent" assumes "cast node \ set |h \ get_ancestors parent|\<^sub>r" assumes "h \ get_parent ref \\<^sub>r Some parent" assumes "is_document_ptr parent \ h \ get_child_nodes parent \\<^sub>r []" assumes "is_document_ptr parent \ \is_character_data_ptr_kind node" shows "h \ ok (a_ensure_pre_insertion_validity node parent (Some ref))" proof - have "h \ (if is_character_data_ptr_kind parent then error HierarchyRequestError else return ()) \\<^sub>r ()" using assms by (simp add: assms(4)) moreover have "h \ do { ancestors \ get_ancestors parent; (if cast node \ set ancestors then error HierarchyRequestError else return ()) } \\<^sub>r ()" using assms(6) apply(auto intro!: bind_pure_returns_result_I)[1] using assms(1) assms(2) assms(3) assms(7) local.get_ancestors_ok local.get_parent_parent_in_heap by auto moreover have "h \ do { (case Some ref of Some child \ do { child_parent \ get_parent child; (if child_parent \ Some parent then error NotFoundError else return ())} | None \ return ()) } \\<^sub>r ()" using assms(7) by(auto split: option.splits) moreover have "h \ do { children \ get_child_nodes parent; (if children \ [] \ is_document_ptr parent then error HierarchyRequestError else return ()) } \\<^sub>r ()" using assms(8) by (smt assms(5) assms(7) bind_pure_returns_result_I2 calculation(1) is_OK_returns_result_I local.get_child_nodes_pure local.get_parent_child_dual returns_result_eq) moreover have "h \ do { (if is_character_data_ptr node \ is_document_ptr parent then error HierarchyRequestError else return ()) } \\<^sub>r ()" using assms using is_character_data_ptr_kind_none by force ultimately show ?thesis unfolding a_ensure_pre_insertion_validity_def apply(intro bind_is_OK_pure_I) apply auto[1] apply auto[1] apply auto[1] using assms(1) assms(2) assms(3) assms(7) local.get_ancestors_ok local.get_parent_parent_in_heap apply blast apply auto[1] apply auto[1] using assms(6) apply auto[1] using assms(1) assms(2) assms(3) assms(7) local.get_ancestors_ok local.get_parent_parent_in_heap apply auto[1] apply (smt bind_returns_heap_E is_OK_returns_heap_E local.get_parent_pure pure_def pure_returns_heap_eq return_returns_heap returns_result_eq) apply(blast) using local.get_child_nodes_pure apply blast apply (meson assms(7) is_OK_returns_result_I local.get_parent_child_dual) apply (simp) apply (smt assms(5) assms(8) is_OK_returns_result_I returns_result_eq) by(auto) qed end locale l_insert_before_wf2 = l_type_wf + l_known_ptrs + l_insert_before_defs + l_heap_is_wellformed_defs + l_get_child_nodes_defs + l_remove_defs + assumes insert_before_preserves_type_wf: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ insert_before ptr child ref \\<^sub>h h' \ type_wf h'" assumes insert_before_preserves_known_ptrs: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ insert_before ptr child ref \\<^sub>h h' \ known_ptrs h'" assumes insert_before_heap_is_wellformed_preserved: "type_wf h \ known_ptrs h \ heap_is_wellformed h \ h \ insert_before ptr child ref \\<^sub>h h' \ heap_is_wellformed h'" interpretation i_insert_before_wf2?: l_insert_before_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_parent get_parent_locs get_child_nodes get_child_nodes_locs set_child_nodes set_child_nodes_locs get_ancestors get_ancestors_locs adopt_node adopt_node_locs set_disconnected_nodes set_disconnected_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_owner_document insert_before insert_before_locs append_child type_wf known_ptr known_ptrs heap_is_wellformed parent_child_rel remove_child remove_child_locs get_root_node get_root_node_locs by(simp add: l_insert_before_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_insert_before_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma insert_before_wf2_is_l_insert_before_wf2 [instances]: "l_insert_before_wf2 type_wf known_ptr known_ptrs insert_before heap_is_wellformed" apply(auto simp add: l_insert_before_wf2_def l_insert_before_wf2_axioms_def instances)[1] using insert_before_heap_is_wellformed_preserved apply(fast, fast, fast) done locale l_insert_before_wf3\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_insert_before_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_child_nodes_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_remove_child_wf2 begin lemma next_sibling_ok: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "node_ptr |\| node_ptr_kinds h" shows "h \ ok (a_next_sibling node_ptr)" proof - have "known_ptr (cast node_ptr)" using assms(2) assms(4) local.known_ptrs_known_ptr node_ptr_kinds_commutes by blast then show ?thesis using assms apply(auto simp add: a_next_sibling_def intro!: bind_is_OK_pure_I split: option.splits list.splits)[1] using get_child_nodes_ok local.get_parent_parent_in_heap local.known_ptrs_known_ptr by blast qed lemma remove_child_ok: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "h \ get_child_nodes ptr \\<^sub>r children" assumes "child \ set children" shows "h \ ok (remove_child ptr child)" proof - have "ptr |\| object_ptr_kinds h" using assms(4) local.get_child_nodes_ptr_in_heap by blast have "child |\| node_ptr_kinds h" using assms(1) assms(4) assms(5) local.heap_is_wellformed_children_in_heap by blast have "\is_character_data_ptr ptr" proof (rule ccontr, simp) assume "is_character_data_ptr ptr" then have "h \ get_child_nodes ptr \\<^sub>r []" using \ptr |\| object_ptr_kinds h\ apply(simp add: get_child_nodes_def a_get_child_nodes_tups_def) apply(split invoke_splits)+ by(auto simp add: get_child_nodes\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_pure_returns_result_I split: option.splits) then show False using assms returns_result_eq by fastforce qed have "is_character_data_ptr child \ \is_document_ptr_kind ptr" proof (rule ccontr, simp) assume "is_character_data_ptr\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child" and "is_document_ptr_kind ptr" then show False using assms using \ptr |\| object_ptr_kinds h\ apply(simp add: get_child_nodes_def a_get_child_nodes_tups_def) apply(split invoke_splits)+ apply(auto split: option.splits)[1] apply (meson invoke_empty is_OK_returns_result_I) apply (meson invoke_empty is_OK_returns_result_I) by(auto simp add: get_child_nodes\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 split: option.splits) qed obtain owner_document where owner_document: "h \ get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r owner_document" by (meson \child |\| node_ptr_kinds h\ assms(1) assms(2) assms(3) is_OK_returns_result_E local.get_owner_document_ok node_ptr_kinds_commutes) obtain disconnected_nodes_h where disconnected_nodes_h: "h \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h" by (meson assms(1) assms(2) assms(3) is_OK_returns_result_E local.get_disconnected_nodes_ok local.get_owner_document_owner_document_in_heap owner_document) obtain h2 where h2: "h \ set_disconnected_nodes owner_document (child # disconnected_nodes_h) \\<^sub>h h2" by (meson assms(1) assms(2) assms(3) is_OK_returns_heap_E l_set_disconnected_nodes.set_disconnected_nodes_ok local.get_owner_document_owner_document_in_heap local.l_set_disconnected_nodes_axioms owner_document) have "known_ptr ptr" using assms(2) assms(4) local.known_ptrs_known_ptr using \ptr |\| object_ptr_kinds h\ by blast have "type_wf h2" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h2] using set_disconnected_nodes_types_preserved assms(3) by(auto simp add: reflp_def transp_def) have "object_ptr_kinds h = object_ptr_kinds h2" using h2 apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes]) using set_disconnected_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) have "h2 \ ok (set_child_nodes ptr (remove1 child children))" proof (cases "is_element_ptr_kind ptr") case True then show ?thesis using set_child_nodes_element_ok \known_ptr ptr\ \object_ptr_kinds h = object_ptr_kinds h2\ \type_wf h2\ assms(4) using \ptr |\| object_ptr_kinds h\ by blast next case False then have "is_document_ptr_kind ptr" using \known_ptr ptr\ \ptr |\| object_ptr_kinds h\ \\is_character_data_ptr ptr\ by(auto simp add:known_ptr_impl known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits) moreover have "is_document_ptr ptr" using \known_ptr ptr\ \ptr |\| object_ptr_kinds h\ False \\is_character_data_ptr ptr\ by(auto simp add: known_ptr_impl known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits) ultimately show ?thesis using assms(4) apply(auto simp add: get_child_nodes_def a_get_child_nodes_tups_def)[1] apply(split invoke_splits)+ apply(auto elim!: bind_returns_result_E2 split: option.splits)[1] apply(auto simp add: get_child_nodes\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 split: option.splits)[1] using assms(5) apply auto[1] using \is_document_ptr_kind ptr\ \known_ptr ptr\ \object_ptr_kinds h = object_ptr_kinds h2\ \ptr |\| object_ptr_kinds h\ \type_wf h2\ local.set_child_nodes_document1_ok apply blast using \is_document_ptr_kind ptr\ \known_ptr ptr\ \object_ptr_kinds h = object_ptr_kinds h2\ \ptr |\| object_ptr_kinds h\ \type_wf h2\ is_element_ptr_kind_cast local.set_child_nodes_document2_ok apply blast using \\ is_character_data_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr\ apply blast by (metis False is_element_ptr_implies_kind option.case_eq_if) qed then obtain h' where h': "h2 \ set_child_nodes ptr (remove1 child children) \\<^sub>h h'" by auto show ?thesis using assms apply(auto simp add: remove_child_def simp add: is_OK_returns_heap_I[OF h2] is_OK_returns_heap_I[OF h'] is_OK_returns_result_I[OF assms(4)] is_OK_returns_result_I[OF owner_document] is_OK_returns_result_I[OF disconnected_nodes_h] intro!: bind_is_OK_pure_I[OF get_owner_document_pure] bind_is_OK_pure_I[OF get_child_nodes_pure] bind_is_OK_pure_I[OF get_disconnected_nodes_pure] bind_is_OK_I[rotated, OF h2] dest!: returns_result_eq[OF assms(4)] returns_result_eq[OF owner_document] returns_result_eq[OF disconnected_nodes_h] )[1] using h2 returns_result_select_result by force qed lemma adopt_node_ok: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "document_ptr |\| document_ptr_kinds h" assumes "child |\| node_ptr_kinds h" shows "h \ ok (adopt_node document_ptr child)" proof - obtain old_document where old_document: "h \ get_owner_document (cast child) \\<^sub>r old_document" by (meson assms(1) assms(2) assms(3) assms(5) is_OK_returns_result_E local.get_owner_document_ok node_ptr_kinds_commutes) then have "h \ ok (get_owner_document (cast child))" by auto obtain parent_opt where parent_opt: "h \ get_parent child \\<^sub>r parent_opt" by (meson assms(2) assms(3) is_OK_returns_result_I l_get_owner_document.get_owner_document_ptr_in_heap local.get_parent_ok local.l_get_owner_document_axioms node_ptr_kinds_commutes old_document returns_result_select_result) then have "h \ ok (get_parent child)" by auto have "h \ ok (case parent_opt of Some parent \ remove_child parent child | None \ return ())" apply(auto split: option.splits)[1] using remove_child_ok by (metis assms(1) assms(2) assms(3) local.get_parent_child_dual parent_opt) then obtain h2 where h2: "h \ (case parent_opt of Some parent \ remove_child parent child | None \ return ()) \\<^sub>h h2" by auto have "object_ptr_kinds h = object_ptr_kinds h2" using h2 apply(simp split: option.splits) apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF remove_child_writes]) using remove_child_pointers_preserved by (auto simp add: reflp_def transp_def) then have "old_document |\| document_ptr_kinds h2" using assms(1) assms(2) assms(3) document_ptr_kinds_commutes local.get_owner_document_owner_document_in_heap old_document by blast have wellformed_h2: "heap_is_wellformed h2" using h2 remove_child_heap_is_wellformed_preserved assms by(auto split: option.splits) have "type_wf h2" using h2 remove_child_preserves_type_wf assms by(auto split: option.splits) have "known_ptrs h2" using h2 remove_child_preserves_known_ptrs assms by(auto split: option.splits) have "object_ptr_kinds h = object_ptr_kinds h2" using h2 apply(simp split: option.splits) apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF remove_child_writes]) using remove_child_pointers_preserved by (auto simp add: reflp_def transp_def) then have "document_ptr_kinds h = document_ptr_kinds h2" by(auto simp add: document_ptr_kinds_def) have "h2 \ ok (if document_ptr \ old_document then do { old_disc_nodes \ get_disconnected_nodes old_document; set_disconnected_nodes old_document (remove1 child old_disc_nodes); disc_nodes \ get_disconnected_nodes document_ptr; set_disconnected_nodes document_ptr (child # disc_nodes) } else do { return () })" proof(cases "document_ptr = old_document") case True then show ?thesis by simp next case False then have "h2 \ ok (get_disconnected_nodes old_document)" by (simp add: \old_document |\| document_ptr_kinds h2\ \type_wf h2\ local.get_disconnected_nodes_ok) then obtain old_disc_nodes where old_disc_nodes: "h2 \ get_disconnected_nodes old_document \\<^sub>r old_disc_nodes" by auto have "h2 \ ok (set_disconnected_nodes old_document (remove1 child old_disc_nodes))" by (simp add: \old_document |\| document_ptr_kinds h2\ \type_wf h2\ local.set_disconnected_nodes_ok) then obtain h3 where h3: "h2 \ set_disconnected_nodes old_document (remove1 child old_disc_nodes) \\<^sub>h h3" by auto have object_ptr_kinds_h2_eq3: "object_ptr_kinds h2 = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes h3]) using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h2: "\ptrs. h2 \ object_ptr_kinds_M \\<^sub>r ptrs = h3 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_eq_h2: "|h2 \ object_ptr_kinds_M|\<^sub>r = |h3 \ object_ptr_kinds_M|\<^sub>r" by(simp) then have node_ptr_kinds_eq_h2: "|h2 \ node_ptr_kinds_M|\<^sub>r = |h3 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast then have node_ptr_kinds_eq3_h2: "node_ptr_kinds h2 = node_ptr_kinds h3" by auto have document_ptr_kinds_eq2_h2: "|h2 \ document_ptr_kinds_M|\<^sub>r = |h3 \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_eq_h2 document_ptr_kinds_M_eq by auto then have document_ptr_kinds_eq3_h2: "document_ptr_kinds h2 = document_ptr_kinds h3" using object_ptr_kinds_eq_h2 document_ptr_kinds_M_eq by auto have children_eq_h2: "\ptr children. h2 \ get_child_nodes ptr \\<^sub>r children = h3 \ get_child_nodes ptr \\<^sub>r children" using get_child_nodes_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by (simp add: set_disconnected_nodes_get_child_nodes) then have children_eq2_h2: "\ptr. |h2 \ get_child_nodes ptr|\<^sub>r = |h3 \ get_child_nodes ptr|\<^sub>r" using select_result_eq by force have "type_wf h3" using \type_wf h2\ using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h3] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) moreover have "document_ptr |\| document_ptr_kinds h3" using \document_ptr_kinds h = document_ptr_kinds h2\ assms(4) document_ptr_kinds_eq3_h2 by auto ultimately have "h3 \ ok (get_disconnected_nodes document_ptr)" by (simp add: local.get_disconnected_nodes_ok) then obtain disc_nodes where disc_nodes: "h3 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes" by auto have "h3 \ ok (set_disconnected_nodes document_ptr (child # disc_nodes))" using \document_ptr |\| document_ptr_kinds h3\ \type_wf h3\ local.set_disconnected_nodes_ok by auto then obtain h' where h': "h3 \ set_disconnected_nodes document_ptr (child # disc_nodes) \\<^sub>h h'" by auto then show ?thesis using False using \h2 \ ok get_disconnected_nodes old_document\ using \h3 \ ok get_disconnected_nodes document_ptr\ apply(auto dest!: returns_result_eq[OF old_disc_nodes] returns_result_eq[OF disc_nodes] intro!: bind_is_OK_I[rotated, OF h3] bind_is_OK_pure_I[OF get_disconnected_nodes_pure] )[1] using \h2 \ ok set_disconnected_nodes old_document (remove1 child old_disc_nodes)\ by auto qed then obtain h' where h': "h2 \ (if document_ptr \ old_document then do { old_disc_nodes \ get_disconnected_nodes old_document; set_disconnected_nodes old_document (remove1 child old_disc_nodes); disc_nodes \ get_disconnected_nodes document_ptr; set_disconnected_nodes document_ptr (child # disc_nodes) } else do { return () }) \\<^sub>h h'" by auto show ?thesis using \h \ ok (get_owner_document (cast child))\ using \h \ ok (get_parent child)\ using h2 h' apply(auto simp add: adopt_node_def simp add: is_OK_returns_heap_I[OF h2] intro!: bind_is_OK_pure_I[OF get_owner_document_pure] bind_is_OK_pure_I[OF get_parent_pure] bind_is_OK_I[rotated, OF h2] dest!: returns_result_eq[OF parent_opt] returns_result_eq[OF old_document])[1] using \h \ ok (case parent_opt of None \ return () | Some parent \ remove_child parent child)\ by auto qed lemma insert_node_ok: assumes "known_ptr parent" and "type_wf h" assumes "parent |\| object_ptr_kinds h" assumes "\is_character_data_ptr_kind parent" assumes "is_document_ptr parent \ h \ get_child_nodes parent \\<^sub>r []" assumes "is_document_ptr parent \ \is_character_data_ptr_kind node" assumes "known_ptr (cast node)" shows "h \ ok (a_insert_node parent node ref)" proof(auto simp add: a_insert_node_def get_child_nodes_ok[OF assms(1) assms(2) assms(3)] intro!: bind_is_OK_pure_I) fix children' assume "h \ get_child_nodes parent \\<^sub>r children'" show "h \ ok set_child_nodes parent (insert_before_list node ref children')" proof (cases "is_element_ptr_kind parent") case True then show ?thesis using set_child_nodes_element_ok using assms(1) assms(2) assms(3) by blast next case False then have "is_document_ptr_kind parent" using assms(4) assms(1) by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits) then have "is_document_ptr parent" using assms(1) by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits) then obtain children where children: "h \ get_child_nodes parent \\<^sub>r children" and "children = []" using assms(5) by blast have "insert_before_list node ref children' = [node]" by (metis \children = []\ \h \ get_child_nodes parent \\<^sub>r children'\ append.left_neutral children insert_Nil l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.insert_before_list.elims l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.insert_before_list.simps(3) neq_Nil_conv returns_result_eq) moreover have "\is_character_data_ptr_kind node" using \is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r parent\ assms(6) by blast then have "is_element_ptr_kind node" by (metis (no_types, lifting) CharacterDataClass.a_known_ptr_def DocumentClass.a_known_ptr_def ElementClass.a_known_ptr_def NodeClass.a_known_ptr_def assms(7) cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_inject document_ptr_no_node_ptr_cast is_character_data_ptr_kind_none is_document_ptr_kind_none is_element_ptr_implies_kind is_node_ptr_kind_cast local.known_ptr_impl node_ptr_casts_commute3 option.case_eq_if) ultimately show ?thesis using set_child_nodes_document2_ok by (metis \is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r parent\ assms(1) assms(2) assms(3) assms(5) is_document_ptr_kind_none option.case_eq_if) qed qed lemma insert_before_ok: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "parent |\| object_ptr_kinds h" assumes "node |\| node_ptr_kinds h" assumes "\is_character_data_ptr_kind parent" assumes "cast node \ set |h \ get_ancestors parent|\<^sub>r" assumes "h \ get_parent ref \\<^sub>r Some parent" assumes "is_document_ptr parent \ h \ get_child_nodes parent \\<^sub>r []" assumes "is_document_ptr parent \ \is_character_data_ptr_kind node" shows "h \ ok (insert_before parent node (Some ref))" proof - have "h \ ok (a_ensure_pre_insertion_validity node parent (Some ref))" using assms ensure_pre_insertion_validity_ok by blast have "h \ ok (if Some node = Some ref then a_next_sibling node else return (Some ref))" (is "h \ ok ?P") apply(auto split: if_splits)[1] using assms(1) assms(2) assms(3) assms(5) next_sibling_ok by blast then obtain reference_child where reference_child: "h \ ?P \\<^sub>r reference_child" by auto obtain owner_document where owner_document: "h \ get_owner_document parent \\<^sub>r owner_document" using assms get_owner_document_ok by (meson returns_result_select_result) then have "h \ ok (get_owner_document parent)" by auto have "owner_document |\| document_ptr_kinds h" using assms(1) assms(2) assms(3) local.get_owner_document_owner_document_in_heap owner_document by blast obtain h2 where h2: "h \ adopt_node owner_document node \\<^sub>h h2" by (meson assms(1) assms(2) assms(3) assms(5) is_OK_returns_heap_E adopt_node_ok l_insert_before_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms local.get_owner_document_owner_document_in_heap owner_document) then have "h \ ok (adopt_node owner_document node)" by auto have "object_ptr_kinds h = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF adopt_node_writes h2]) using adopt_node_pointers_preserved apply blast by (auto simp add: reflp_def transp_def) then have "document_ptr_kinds h = document_ptr_kinds h2" by(auto simp add: document_ptr_kinds_def) have "heap_is_wellformed h2" using h2 adopt_node_preserves_wellformedness assms by blast have "known_ptrs h2" using h2 adopt_node_preserves_known_ptrs assms by blast have "type_wf h2" using h2 adopt_node_preserves_type_wf assms by blast obtain disconnected_nodes_h2 where disconnected_nodes_h2: "h2 \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h2" by (metis \document_ptr_kinds h = document_ptr_kinds h2\ \type_wf h2\ assms(1) assms(2) assms(3) is_OK_returns_result_E local.get_disconnected_nodes_ok local.get_owner_document_owner_document_in_heap owner_document) obtain h3 where h3: "h2 \ set_disconnected_nodes owner_document (remove1 node disconnected_nodes_h2) \\<^sub>h h3" by (metis \document_ptr_kinds h = document_ptr_kinds h2\ \owner_document |\| document_ptr_kinds h\ \type_wf h2\ document_ptr_kinds_def is_OK_returns_heap_E l_set_disconnected_nodes.set_disconnected_nodes_ok local.l_set_disconnected_nodes_axioms) have "type_wf h3" using \type_wf h2\ using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h3] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) have object_ptr_kinds_M_eq3_h2: "object_ptr_kinds h2 = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes h3]) unfolding a_remove_child_locs_def using set_disconnected_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) have "parent |\| object_ptr_kinds h3" using \object_ptr_kinds h = object_ptr_kinds h2\ assms(4) object_ptr_kinds_M_eq3_h2 by blast moreover have "known_ptr parent" using assms(2) assms(4) local.known_ptrs_known_ptr by blast moreover have "known_ptr (cast node)" using assms(2) assms(5) local.known_ptrs_known_ptr node_ptr_kinds_commutes by blast moreover have "is_document_ptr parent \ h3 \ get_child_nodes parent \\<^sub>r []" by (metis assms(8) assms(9) distinct.simps(2) distinct_singleton local.get_parent_child_dual returns_result_eq) ultimately obtain h' where h': "h3 \ a_insert_node parent node reference_child \\<^sub>h h'" using insert_node_ok \type_wf h3\ assms by blast show ?thesis using \h \ ok (a_ensure_pre_insertion_validity node parent (Some ref))\ using reference_child \h \ ok (get_owner_document parent)\ \h \ ok (adopt_node owner_document node)\ h3 h' apply(auto simp add: insert_before_def simp add: is_OK_returns_result_I[OF disconnected_nodes_h2] simp add: is_OK_returns_heap_I[OF h3] is_OK_returns_heap_I[OF h'] intro!: bind_is_OK_I2 bind_is_OK_pure_I[OF ensure_pre_insertion_validity_pure] bind_is_OK_pure_I[OF next_sibling_pure] bind_is_OK_pure_I[OF get_owner_document_pure] bind_is_OK_pure_I[OF get_disconnected_nodes_pure] dest!: returns_result_eq[OF owner_document] returns_result_eq[OF disconnected_nodes_h2] returns_heap_eq[OF h2] returns_heap_eq[OF h3] dest!: sym[of node ref] )[1] using returns_result_eq by fastforce qed end interpretation i_insert_before_wf3?: l_insert_before_wf3\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_parent get_parent_locs get_child_nodes get_child_nodes_locs set_child_nodes set_child_nodes_locs get_ancestors get_ancestors_locs adopt_node adopt_node_locs set_disconnected_nodes set_disconnected_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_owner_document insert_before insert_before_locs append_child type_wf known_ptr known_ptrs heap_is_wellformed parent_child_rel remove_child remove_child_locs get_root_node get_root_node_locs remove by(auto simp add: l_insert_before_wf3\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_insert_before_wf3\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_append_child_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_insert_before_wf + l_insert_before_wf2 + l_get_child_nodes begin lemma append_child_heap_is_wellformed_preserved: assumes wellformed: "heap_is_wellformed h" and append_child: "h \ append_child ptr node \\<^sub>h h'" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "heap_is_wellformed h'" and "type_wf h'" and "known_ptrs h'" using assms by(auto simp add: append_child_def intro: insert_before_preserves_type_wf insert_before_preserves_known_ptrs insert_before_heap_is_wellformed_preserved) lemma append_child_children: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_child_nodes ptr \\<^sub>r xs" assumes "h \ append_child ptr node \\<^sub>h h'" assumes "node \ set xs" shows "h' \ get_child_nodes ptr \\<^sub>r xs @ [node]" proof - obtain ancestors owner_document h2 h3 disconnected_nodes_h2 where ancestors: "h \ get_ancestors ptr \\<^sub>r ancestors" and node_not_in_ancestors: "cast node \ set ancestors" and owner_document: "h \ get_owner_document ptr \\<^sub>r owner_document" and h2: "h \ adopt_node owner_document node \\<^sub>h h2" and disconnected_nodes_h2: "h2 \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h2" and h3: "h2 \ set_disconnected_nodes owner_document (remove1 node disconnected_nodes_h2) \\<^sub>h h3" and h': "h3 \ a_insert_node ptr node None \\<^sub>h h'" using assms(5) by(auto simp add: append_child_def insert_before_def a_ensure_pre_insertion_validity_def elim!: bind_returns_heap_E bind_returns_result_E bind_returns_heap_E2[rotated, OF get_parent_pure, rotated] bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated] bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] bind_returns_heap_E2[rotated, OF get_ancestors_pure, rotated] bind_returns_heap_E2[rotated, OF next_sibling_pure, rotated] bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated] split: if_splits option.splits) have "\parent. |h \ get_parent node|\<^sub>r = Some parent \ parent \ ptr" using assms(1) assms(4) assms(6) by (metis (no_types, lifting) assms(2) assms(3) h2 is_OK_returns_heap_I is_OK_returns_result_E local.adopt_node_child_in_heap local.get_parent_child_dual local.get_parent_ok select_result_I2) have "h2 \ get_child_nodes ptr \\<^sub>r xs" using get_child_nodes_reads adopt_node_writes h2 assms(4) apply(rule reads_writes_separate_forwards) using \\parent. |h \ get_parent node|\<^sub>r = Some parent \ parent \ ptr\ apply(auto simp add: adopt_node_locs_def remove_child_locs_def)[1] by (meson local.set_child_nodes_get_child_nodes_different_pointers) have "h3 \ get_child_nodes ptr \\<^sub>r xs" using get_child_nodes_reads set_disconnected_nodes_writes h3 \h2 \ get_child_nodes ptr \\<^sub>r xs\ apply(rule reads_writes_separate_forwards) by(auto) have "ptr |\| object_ptr_kinds h" by (meson ancestors is_OK_returns_result_I local.get_ancestors_ptr_in_heap) then have "known_ptr ptr" using assms(3) using local.known_ptrs_known_ptr by blast have "type_wf h2" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF adopt_node_writes h2] using adopt_node_types_preserved \type_wf h\ by(auto simp add: adopt_node_locs_def remove_child_locs_def reflp_def transp_def split: if_splits) then have "type_wf h3" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h3] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) show "h' \ get_child_nodes ptr \\<^sub>r xs@[node]" using h' apply(auto simp add: a_insert_node_def dest!: bind_returns_heap_E3[rotated, OF \h3 \ get_child_nodes ptr \\<^sub>r xs\ get_child_nodes_pure, rotated])[1] using \type_wf h3\ set_child_nodes_get_child_nodes \known_ptr ptr\ by metis qed lemma append_child_for_all_on_children: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_child_nodes ptr \\<^sub>r xs" assumes "h \ forall_M (append_child ptr) nodes \\<^sub>h h'" assumes "set nodes \ set xs = {}" assumes "distinct nodes" shows "h' \ get_child_nodes ptr \\<^sub>r xs@nodes" using assms apply(induct nodes arbitrary: h xs) apply(simp) proof(auto elim!: bind_returns_heap_E)[1]fix a nodes h xs h'a assume 0: "(\h xs. heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_child_nodes ptr \\<^sub>r xs \ h \ forall_M (append_child ptr) nodes \\<^sub>h h' \ set nodes \ set xs = {} \ h' \ get_child_nodes ptr \\<^sub>r xs @ nodes)" and 1: "heap_is_wellformed h" and 2: "type_wf h" and 3: "known_ptrs h" and 4: "h \ get_child_nodes ptr \\<^sub>r xs" and 5: "h \ append_child ptr a \\<^sub>r ()" and 6: "h \ append_child ptr a \\<^sub>h h'a" and 7: "h'a \ forall_M (append_child ptr) nodes \\<^sub>h h'" and 8: "a \ set xs" and 9: "set nodes \ set xs = {}" and 10: "a \ set nodes" and 11: "distinct nodes" then have "h'a \ get_child_nodes ptr \\<^sub>r xs @ [a]" using append_child_children 6 using "1" "2" "3" "4" "8" by blast moreover have "heap_is_wellformed h'a" and "type_wf h'a" and "known_ptrs h'a" using insert_before_heap_is_wellformed_preserved insert_before_preserves_known_ptrs insert_before_preserves_type_wf 1 2 3 6 append_child_def by metis+ moreover have "set nodes \ set (xs @ [a]) = {}" using 9 10 by auto ultimately show "h' \ get_child_nodes ptr \\<^sub>r xs @ a # nodes" using 0 7 by fastforce qed lemma append_child_for_all_on_no_children: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_child_nodes ptr \\<^sub>r []" assumes "h \ forall_M (append_child ptr) nodes \\<^sub>h h'" assumes "distinct nodes" shows "h' \ get_child_nodes ptr \\<^sub>r nodes" using assms append_child_for_all_on_children by force end locale l_append_child_wf = l_type_wf + l_known_ptrs + l_append_child_defs + l_heap_is_wellformed_defs + assumes append_child_preserves_type_wf: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ append_child ptr child \\<^sub>h h' \ type_wf h'" assumes append_child_preserves_known_ptrs: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ append_child ptr child \\<^sub>h h' \ known_ptrs h'" assumes append_child_heap_is_wellformed_preserved: "type_wf h \ known_ptrs h \ heap_is_wellformed h \ h \ append_child ptr child \\<^sub>h h' \ heap_is_wellformed h'" interpretation i_append_child_wf?: l_append_child_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_owner_document get_parent get_parent_locs remove_child remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs adopt_node adopt_node_locs known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs set_child_nodes set_child_nodes_locs remove get_ancestors get_ancestors_locs insert_before insert_before_locs append_child heap_is_wellformed parent_child_rel by(auto simp add: l_append_child_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) lemma append_child_wf_is_l_append_child_wf [instances]: "l_append_child_wf type_wf known_ptr known_ptrs append_child heap_is_wellformed" apply(auto simp add: l_append_child_wf_def l_append_child_wf_axioms_def instances)[1] using append_child_heap_is_wellformed_preserved by fast+ subsection \create\_element\ locale l_create_element_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed parent_child_rel + l_new_element_get_disconnected_nodes get_disconnected_nodes get_disconnected_nodes_locs + l_set_tag_name_get_disconnected_nodes type_wf set_tag_name set_tag_name_locs get_disconnected_nodes get_disconnected_nodes_locs + l_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs set_tag_name set_tag_name_locs type_wf create_element known_ptr + l_new_element_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs + l_set_tag_name_get_child_nodes type_wf set_tag_name set_tag_name_locs known_ptr get_child_nodes get_child_nodes_locs + l_set_disconnected_nodes_get_child_nodes set_disconnected_nodes set_disconnected_nodes_locs get_child_nodes get_child_nodes_locs + l_set_disconnected_nodes type_wf set_disconnected_nodes set_disconnected_nodes_locs + l_set_disconnected_nodes_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs + l_new_element type_wf + l_known_ptrs known_ptr known_ptrs for known_ptr :: "(_::linorder) object_ptr \ bool" and known_ptrs :: "(_) heap \ bool" and type_wf :: "(_) heap \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and set_tag_name :: "(_) element_ptr \ char list \ ((_) heap, exception, unit) prog" and set_tag_name_locs :: "(_) element_ptr \ ((_) heap, exception, unit) prog set" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" and create_element :: "(_) document_ptr \ char list \ ((_) heap, exception, (_) element_ptr) prog" begin lemma create_element_preserves_wellformedness: assumes "heap_is_wellformed h" and "h \ create_element document_ptr tag \\<^sub>h h'" and "type_wf h" and "known_ptrs h" shows "heap_is_wellformed h'" and "type_wf h'" and "known_ptrs h'" proof - obtain new_element_ptr h2 h3 disc_nodes_h3 where new_element_ptr: "h \ new_element \\<^sub>r new_element_ptr" and h2: "h \ new_element \\<^sub>h h2" and h3: "h2 \ set_tag_name new_element_ptr tag \\<^sub>h h3" and disc_nodes_h3: "h3 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" and h': "h3 \ set_disconnected_nodes document_ptr (cast new_element_ptr # disc_nodes_h3) \\<^sub>h h'" using assms(2) by(auto simp add: create_element_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] ) then have "h \ create_element document_ptr tag \\<^sub>r new_element_ptr" apply(auto simp add: create_element_def intro!: bind_returns_result_I)[1] apply (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust) apply (metis is_OK_returns_heap_E is_OK_returns_result_I local.get_disconnected_nodes_pure pure_returns_heap_eq) by (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust) have "new_element_ptr \ set |h \ element_ptr_kinds_M|\<^sub>r" using new_element_ptr ElementMonad.ptr_kinds_ptr_kinds_M h2 using new_element_ptr_not_in_heap by blast then have "cast new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r" by simp then have "cast new_element_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r" by simp have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\| {|cast new_element_ptr|}" using new_element_new_ptr h2 new_element_ptr by blast then have node_ptr_kinds_eq_h: "node_ptr_kinds h2 = node_ptr_kinds h |\| {|cast new_element_ptr|}" apply(simp add: node_ptr_kinds_def) by force then have element_ptr_kinds_eq_h: "element_ptr_kinds h2 = element_ptr_kinds h |\| {|new_element_ptr|}" apply(simp add: element_ptr_kinds_def) by force have character_data_ptr_kinds_eq_h: "character_data_ptr_kinds h2 = character_data_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: node_ptr_kinds_def character_data_ptr_kinds_def) have document_ptr_kinds_eq_h: "document_ptr_kinds h2 = document_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: document_ptr_kinds_def) have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_tag_name_writes h3]) using set_tag_name_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2" using object_ptr_kinds_eq_h2 by(auto simp add: node_ptr_kinds_def) have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_disconnected_nodes_writes h']) using set_disconnected_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3" using object_ptr_kinds_eq_h3 by(auto simp add: node_ptr_kinds_def) have "known_ptr (cast new_element_ptr)" using \h \ create_element document_ptr tag \\<^sub>r new_element_ptr\ local.create_element_known_ptr by blast then have "known_ptrs h2" using known_ptrs_new_ptr object_ptr_kinds_eq_h \known_ptrs h\ h2 by blast then have "known_ptrs h3" using known_ptrs_preserved object_ptr_kinds_eq_h2 by blast then show "known_ptrs h'" using known_ptrs_preserved object_ptr_kinds_eq_h3 by blast have "document_ptr |\| document_ptr_kinds h" using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2 get_disconnected_nodes_ptr_in_heap \type_wf h\ document_ptr_kinds_def by (metis is_OK_returns_result_I) have children_eq_h: "\(ptr'::(_) object_ptr) children. ptr' \ cast new_element_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads h2 get_child_nodes_new_element[rotated, OF new_element_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have children_eq2_h: "\ptr'. ptr' \ cast new_element_ptr \ |h \ get_child_nodes ptr'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have "h2 \ get_child_nodes (cast new_element_ptr) \\<^sub>r []" using new_element_ptr h2 new_element_ptr_in_heap[OF h2 new_element_ptr] new_element_is_element_ptr[OF new_element_ptr] new_element_no_child_nodes by blast have disconnected_nodes_eq_h: "\doc_ptr disc_nodes. h \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads h2 get_disconnected_nodes_new_element[OF new_element_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have disconnected_nodes_eq2_h: "\doc_ptr. |h \ get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. h2 \ get_child_nodes ptr' \\<^sub>r children = h3 \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads set_tag_name_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_tag_name_get_child_nodes) then have children_eq2_h2: "\ptr'. |h2 \ get_child_nodes ptr'|\<^sub>r = |h3 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_tag_name_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_tag_name_get_disconnected_nodes) then have disconnected_nodes_eq2_h2: "\doc_ptr. |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have "type_wf h2" using \type_wf h\ new_element_types_preserved h2 by blast then have "type_wf h3" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_tag_name_writes h3] using set_tag_name_types_preserved by(auto simp add: reflp_def transp_def) then show "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h'] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) have children_eq_h3: "\ptr' children. h3 \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_disconnected_nodes_get_child_nodes) then have children_eq2_h3: "\ptr'. |h3 \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h3: "\doc_ptr disc_nodes. document_ptr \ doc_ptr \ h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers) then have disconnected_nodes_eq2_h3: "\doc_ptr. document_ptr \ doc_ptr \ |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have disc_nodes_document_ptr_h2: "h2 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" using disconnected_nodes_eq_h2 disc_nodes_h3 by auto then have disc_nodes_document_ptr_h: "h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" using disconnected_nodes_eq_h by auto then have "cast new_element_ptr \ set disc_nodes_h3" using \heap_is_wellformed h\ using \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ a_all_ptrs_in_heap_def heap_is_wellformed_def using NodeMonad.ptr_kinds_ptr_kinds_M local.heap_is_wellformed_disc_nodes_in_heap by blast have "acyclic (parent_child_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def acyclic_heap_def) also have "parent_child_rel h = parent_child_rel h2" proof(auto simp add: parent_child_rel_def)[1] fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h2" by (simp add: object_ptr_kinds_eq_h) next fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "x \ set |h2 \ get_child_nodes a|\<^sub>r" by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2_h) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h" using object_ptr_kinds_eq_h \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ by(auto) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "x \ set |h \ get_child_nodes a|\<^sub>r" by (metis (no_types, lifting) \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ children_eq2_h empty_iff empty_set image_eqI select_result_I2) qed also have "\ = parent_child_rel h3" by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2) also have "\ = parent_child_rel h'" by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3) finally have "a_acyclic_heap h'" by (simp add: acyclic_heap_def) have "a_all_ptrs_in_heap h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) then have "a_all_ptrs_in_heap h2" apply(auto simp add: a_all_ptrs_in_heap_def)[1] apply (metis \known_ptrs h2\ \parent_child_rel h = parent_child_rel h2\ \type_wf h2\ assms(1) assms(3) funion_iff local.get_child_nodes_ok local.known_ptrs_known_ptr local.parent_child_rel_child_in_heap local.parent_child_rel_child_nodes2 node_ptr_kinds_commutes node_ptr_kinds_eq_h returns_result_select_result) by (metis assms(1) assms(3) disconnected_nodes_eq2_h document_ptr_kinds_eq_h funion_iff local.get_disconnected_nodes_ok local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq_h returns_result_select_result) then have "a_all_ptrs_in_heap h3" by (simp add: children_eq2_h2 disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 local.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 object_ptr_kinds_eq_h2) then have "a_all_ptrs_in_heap h'" by (smt \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ children_eq2_h3 disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 finite_set_in h' is_OK_returns_result_I set_disconnected_nodes_get_disconnected_nodes local.a_all_ptrs_in_heap_def local.get_child_nodes_ptr_in_heap node_ptr_kinds_commutes object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 select_result_I2 set_ConsD subset_code(1)) have "\p. p |\| object_ptr_kinds h \ cast new_element_ptr \ set |h \ get_child_nodes p|\<^sub>r" using \heap_is_wellformed h\ \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ heap_is_wellformed_children_in_heap by (meson NodeMonad.ptr_kinds_ptr_kinds_M a_all_ptrs_in_heap_def assms(3) assms(4) fset_mp fset_of_list_elem get_child_nodes_ok known_ptrs_known_ptr returns_result_select_result) then have "\p. p |\| object_ptr_kinds h2 \ cast new_element_ptr \ set |h2 \ get_child_nodes p|\<^sub>r" using children_eq2_h apply(auto simp add: object_ptr_kinds_eq_h)[1] using \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ apply auto[1] by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\) then have "\p. p |\| object_ptr_kinds h3 \ cast new_element_ptr \ set |h3 \ get_child_nodes p|\<^sub>r" using object_ptr_kinds_eq_h2 children_eq2_h2 by auto then have new_element_ptr_not_in_any_children: "\p. p |\| object_ptr_kinds h' \ cast new_element_ptr \ set |h' \ get_child_nodes p|\<^sub>r" using object_ptr_kinds_eq_h3 children_eq2_h3 by auto have "a_distinct_lists h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) then have "a_distinct_lists h2" using \h2 \ get_child_nodes (cast new_element_ptr) \\<^sub>r []\ apply(auto simp add: a_distinct_lists_def object_ptr_kinds_eq_h document_ptr_kinds_eq_h disconnected_nodes_eq2_h intro!: distinct_concat_map_I)[1] - apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert) + apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert_remove) apply(case_tac "x=cast new_element_ptr") apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply (metis IntI assms(1) assms(3) assms(4) empty_iff local.get_child_nodes_ok local.heap_is_wellformed_one_parent local.known_ptrs_known_ptr returns_result_select_result) apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] by (metis \local.a_distinct_lists h\ \type_wf h2\ disconnected_nodes_eq_h document_ptr_kinds_eq_h local.distinct_lists_no_parent local.get_disconnected_nodes_ok returns_result_select_result) then have "a_distinct_lists h3" by(auto simp add: a_distinct_lists_def disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 children_eq2_h2 object_ptr_kinds_eq_h2) then have "a_distinct_lists h'" proof(auto simp add: a_distinct_lists_def disconnected_nodes_eq2_h3 children_eq2_h3 object_ptr_kinds_eq_h3 document_ptr_kinds_eq_h3 intro!: distinct_concat_map_I)[1] fix x assume "distinct (concat (map (\document_ptr. |h3 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" and "x |\| document_ptr_kinds h3" then show "distinct |h' \ get_disconnected_nodes x|\<^sub>r" using document_ptr_kinds_eq_h3 disconnected_nodes_eq_h3 h' set_disconnected_nodes_get_disconnected_nodes by (metis (no_types, lifting) \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set disc_nodes_h3\ \a_distinct_lists h3\ \type_wf h'\ disc_nodes_h3 distinct.simps(2) distinct_lists_disconnected_nodes get_disconnected_nodes_ok returns_result_eq returns_result_select_result) next fix x y xa assume "distinct (concat (map (\document_ptr. |h3 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" and "x |\| document_ptr_kinds h3" and "y |\| document_ptr_kinds h3" and "x \ y" and "xa \ set |h' \ get_disconnected_nodes x|\<^sub>r" and "xa \ set |h' \ get_disconnected_nodes y|\<^sub>r" moreover have "set |h3 \ get_disconnected_nodes x|\<^sub>r \ set |h3 \ get_disconnected_nodes y|\<^sub>r = {}" using calculation by(auto dest: distinct_concat_map_E(1)) ultimately show "False" apply(-) apply(cases "x = document_ptr") apply (smt NodeMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ \local.a_all_ptrs_in_heap h\ disc_nodes_h3 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 finite_set_in h' l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms select_result_I2 set_ConsD subsetD) by (smt NodeMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ \local.a_all_ptrs_in_heap h\ disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 finite_set_in h' l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms select_result_I2 set_ConsD subsetD) next fix x xa xb assume 2: "(\x\fset (object_ptr_kinds h3). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h3). set |h3 \ get_disconnected_nodes x|\<^sub>r) = {}" and 3: "xa |\| object_ptr_kinds h3" and 4: "x \ set |h' \ get_child_nodes xa|\<^sub>r" and 5: "xb |\| document_ptr_kinds h3" and 6: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" show "False" using disc_nodes_document_ptr_h disconnected_nodes_eq2_h3 apply - apply(cases "xb = document_ptr") apply (metis (no_types, hide_lams) "3" "4" "6" \\p. p |\| object_ptr_kinds h3 \ cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h3 \ get_child_nodes p|\<^sub>r\ \a_distinct_lists h3\ children_eq2_h3 disc_nodes_h3 distinct_lists_no_parent h' select_result_I2 set_ConsD set_disconnected_nodes_get_disconnected_nodes) by (metis "3" "4" "5" "6" \a_distinct_lists h3\ \type_wf h3\ children_eq2_h3 distinct_lists_no_parent get_disconnected_nodes_ok returns_result_select_result) qed have "a_owner_document_valid h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) then have "a_owner_document_valid h'" using disc_nodes_h3 \document_ptr |\| document_ptr_kinds h\ apply(auto simp add: a_owner_document_valid_def)[1] apply(auto simp add: object_ptr_kinds_eq_h object_ptr_kinds_eq_h3 )[1] apply(auto simp add: object_ptr_kinds_eq_h2)[1] apply(auto simp add: document_ptr_kinds_eq_h document_ptr_kinds_eq_h3 )[1] apply(auto simp add: document_ptr_kinds_eq_h2)[1] apply(auto simp add: node_ptr_kinds_eq_h node_ptr_kinds_eq_h3 )[1] apply(auto simp add: node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h )[1] apply(auto simp add: children_eq2_h2[symmetric] children_eq2_h3[symmetric] disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3)[1] apply (metis (no_types, lifting) document_ptr_kinds_eq_h h' list.set_intros(1) local.set_disconnected_nodes_get_disconnected_nodes select_result_I2) apply(simp add: object_ptr_kinds_eq_h) by (smt ObjectMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2_h children_eq2_h2 children_eq2_h3 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h finite_set_in h' list.set_intros(2) local.set_disconnected_nodes_get_disconnected_nodes select_result_I2) show "heap_is_wellformed h'" using \a_acyclic_heap h'\ \a_all_ptrs_in_heap h'\ \a_distinct_lists h'\ \a_owner_document_valid h'\ by(simp add: heap_is_wellformed_def) qed end interpretation i_create_element_wf?: l_create_element_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr known_ptrs type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed parent_child_rel set_tag_name set_tag_name_locs set_disconnected_nodes set_disconnected_nodes_locs create_element using instances by(auto simp add: l_create_element_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_create_element_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsection \create\_character\_data\ locale l_create_character_data_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed parent_child_rel + l_new_character_data_get_disconnected_nodes get_disconnected_nodes get_disconnected_nodes_locs + l_set_val_get_disconnected_nodes type_wf set_val set_val_locs get_disconnected_nodes get_disconnected_nodes_locs + l_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs set_val set_val_locs type_wf create_character_data known_ptr + l_new_character_data_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs + l_set_val_get_child_nodes type_wf set_val set_val_locs known_ptr get_child_nodes get_child_nodes_locs + l_set_disconnected_nodes_get_child_nodes set_disconnected_nodes set_disconnected_nodes_locs get_child_nodes get_child_nodes_locs + l_set_disconnected_nodes type_wf set_disconnected_nodes set_disconnected_nodes_locs + l_set_disconnected_nodes_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs + l_new_character_data type_wf + l_known_ptrs known_ptr known_ptrs for known_ptr :: "(_::linorder) object_ptr \ bool" and type_wf :: "(_) heap \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and set_val :: "(_) character_data_ptr \ char list \ ((_) heap, exception, unit) prog" and set_val_locs :: "(_) character_data_ptr \ ((_) heap, exception, unit) prog set" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" and create_character_data :: "(_) document_ptr \ char list \ ((_) heap, exception, (_) character_data_ptr) prog" and known_ptrs :: "(_) heap \ bool" begin lemma create_character_data_preserves_wellformedness: assumes "heap_is_wellformed h" and "h \ create_character_data document_ptr text \\<^sub>h h'" and "type_wf h" and "known_ptrs h" shows "heap_is_wellformed h'" and "type_wf h'" and "known_ptrs h'" proof - obtain new_character_data_ptr h2 h3 disc_nodes_h3 where new_character_data_ptr: "h \ new_character_data \\<^sub>r new_character_data_ptr" and h2: "h \ new_character_data \\<^sub>h h2" and h3: "h2 \ set_val new_character_data_ptr text \\<^sub>h h3" and disc_nodes_h3: "h3 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" and h': "h3 \ set_disconnected_nodes document_ptr (cast new_character_data_ptr # disc_nodes_h3) \\<^sub>h h'" using assms(2) by(auto simp add: create_character_data_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] ) then have "h \ create_character_data document_ptr text \\<^sub>r new_character_data_ptr" apply(auto simp add: create_character_data_def intro!: bind_returns_result_I)[1] apply (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust) apply (metis is_OK_returns_heap_E is_OK_returns_result_I local.get_disconnected_nodes_pure pure_returns_heap_eq) by (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust) have "new_character_data_ptr \ set |h \ character_data_ptr_kinds_M|\<^sub>r" using new_character_data_ptr CharacterDataMonad.ptr_kinds_ptr_kinds_M h2 using new_character_data_ptr_not_in_heap by blast then have "cast new_character_data_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r" by simp then have "cast new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r" by simp have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\| {|cast new_character_data_ptr|}" using new_character_data_new_ptr h2 new_character_data_ptr by blast then have node_ptr_kinds_eq_h: "node_ptr_kinds h2 = node_ptr_kinds h |\| {|cast new_character_data_ptr|}" apply(simp add: node_ptr_kinds_def) by force then have character_data_ptr_kinds_eq_h: "character_data_ptr_kinds h2 = character_data_ptr_kinds h |\| {|new_character_data_ptr|}" apply(simp add: character_data_ptr_kinds_def) by force have element_ptr_kinds_eq_h: "element_ptr_kinds h2 = element_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: node_ptr_kinds_def element_ptr_kinds_def) have document_ptr_kinds_eq_h: "document_ptr_kinds h2 = document_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: document_ptr_kinds_def) have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_val_writes h3]) using set_val_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2" using object_ptr_kinds_eq_h2 by(auto simp add: node_ptr_kinds_def) have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_disconnected_nodes_writes h']) using set_disconnected_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3" using object_ptr_kinds_eq_h3 by(auto simp add: node_ptr_kinds_def) have "document_ptr |\| document_ptr_kinds h" using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2 get_disconnected_nodes_ptr_in_heap \type_wf h\ document_ptr_kinds_def by (metis is_OK_returns_result_I) have children_eq_h: "\(ptr'::(_) object_ptr) children. ptr' \ cast new_character_data_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads h2 get_child_nodes_new_character_data[rotated, OF new_character_data_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have children_eq2_h: "\ptr'. ptr' \ cast new_character_data_ptr \ |h \ get_child_nodes ptr'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\| {|cast new_character_data_ptr|}" using new_character_data_new_ptr h2 new_character_data_ptr by blast then have node_ptr_kinds_eq_h: "node_ptr_kinds h2 = node_ptr_kinds h |\| {|cast new_character_data_ptr|}" apply(simp add: node_ptr_kinds_def) by force then have character_data_ptr_kinds_eq_h: "character_data_ptr_kinds h2 = character_data_ptr_kinds h |\| {|new_character_data_ptr|}" apply(simp add: character_data_ptr_kinds_def) by force have element_ptr_kinds_eq_h: "element_ptr_kinds h2 = element_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: node_ptr_kinds_def element_ptr_kinds_def) have document_ptr_kinds_eq_h: "document_ptr_kinds h2 = document_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: document_ptr_kinds_def) have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_val_writes h3]) using set_val_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2" using object_ptr_kinds_eq_h2 by(auto simp add: node_ptr_kinds_def) have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_disconnected_nodes_writes h']) using set_disconnected_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3" using object_ptr_kinds_eq_h3 by(auto simp add: node_ptr_kinds_def) have "document_ptr |\| document_ptr_kinds h" using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2 get_disconnected_nodes_ptr_in_heap \type_wf h\ document_ptr_kinds_def by (metis is_OK_returns_result_I) have children_eq_h: "\(ptr'::(_) object_ptr) children. ptr' \ cast new_character_data_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads h2 get_child_nodes_new_character_data[rotated, OF new_character_data_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have children_eq2_h: "\ptr'. ptr' \ cast new_character_data_ptr \ |h \ get_child_nodes ptr'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have "h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []" using new_character_data_ptr h2 new_character_data_ptr_in_heap[OF h2 new_character_data_ptr] new_character_data_is_character_data_ptr[OF new_character_data_ptr] new_character_data_no_child_nodes by blast have disconnected_nodes_eq_h: "\doc_ptr disc_nodes. h \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads h2 get_disconnected_nodes_new_character_data[OF new_character_data_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have disconnected_nodes_eq2_h: "\doc_ptr. |h \ get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. h2 \ get_child_nodes ptr' \\<^sub>r children = h3 \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads set_val_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_val_get_child_nodes) then have children_eq2_h2: "\ptr'. |h2 \ get_child_nodes ptr'|\<^sub>r = |h3 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_val_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_val_get_disconnected_nodes) then have disconnected_nodes_eq2_h2: "\doc_ptr. |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have "type_wf h2" using \type_wf h\ new_character_data_types_preserved h2 by blast then have "type_wf h3" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_val_writes h3] using set_val_types_preserved by(auto simp add: reflp_def transp_def) then show "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h'] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) have children_eq_h3: "\ptr' children. h3 \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_disconnected_nodes_get_child_nodes) then have children_eq2_h3: " \ptr'. |h3 \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h3: "\doc_ptr disc_nodes. document_ptr \ doc_ptr \ h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers) then have disconnected_nodes_eq2_h3: "\doc_ptr. document_ptr \ doc_ptr \ |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have disc_nodes_document_ptr_h2: "h2 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" using disconnected_nodes_eq_h2 disc_nodes_h3 by auto then have disc_nodes_document_ptr_h: "h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" using disconnected_nodes_eq_h by auto then have "cast new_character_data_ptr \ set disc_nodes_h3" using \heap_is_wellformed h\ using \cast new_character_data_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ a_all_ptrs_in_heap_def heap_is_wellformed_def using NodeMonad.ptr_kinds_ptr_kinds_M local.heap_is_wellformed_disc_nodes_in_heap by blast have "acyclic (parent_child_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def acyclic_heap_def) also have "parent_child_rel h = parent_child_rel h2" proof(auto simp add: parent_child_rel_def)[1] fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h2" by (simp add: object_ptr_kinds_eq_h) next fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "x \ set |h2 \ get_child_nodes a|\<^sub>r" by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2_h) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h" using object_ptr_kinds_eq_h \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ by(auto) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "x \ set |h \ get_child_nodes a|\<^sub>r" by (metis (no_types, lifting) \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ children_eq2_h empty_iff empty_set image_eqI select_result_I2) qed also have "\ = parent_child_rel h3" by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2) also have "\ = parent_child_rel h'" by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3) finally have "a_acyclic_heap h'" by (simp add: acyclic_heap_def) have "a_all_ptrs_in_heap h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) then have "a_all_ptrs_in_heap h2" apply(auto simp add: a_all_ptrs_in_heap_def)[1] using node_ptr_kinds_eq_h \cast new_character_data_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ apply (metis (no_types, lifting) NodeMonad.ptr_kinds_ptr_kinds_M \parent_child_rel h = parent_child_rel h2\ children_eq2_h finite_set_in finsert_iff funion_finsert_right local.parent_child_rel_child local.parent_child_rel_parent_in_heap node_ptr_kinds_commutes object_ptr_kinds_eq_h select_result_I2 subsetD sup_bot.right_neutral) by (metis assms(1) assms(3) disconnected_nodes_eq2_h document_ptr_kinds_eq_h funionI1 local.get_disconnected_nodes_ok local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq_h returns_result_select_result) then have "a_all_ptrs_in_heap h3" by (simp add: children_eq2_h2 disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 local.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 object_ptr_kinds_eq_h2) then have "a_all_ptrs_in_heap h'" by (smt character_data_ptr_kinds_commutes children_eq2_h3 disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 finite_set_in h' h2 local.a_all_ptrs_in_heap_def local.set_disconnected_nodes_get_disconnected_nodes new_character_data_ptr new_character_data_ptr_in_heap node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h3 object_ptr_kinds_eq_h3 select_result_I2 set_ConsD subset_code(1)) have "\p. p |\| object_ptr_kinds h \ cast new_character_data_ptr \ set |h \ get_child_nodes p|\<^sub>r" using \heap_is_wellformed h\ \cast new_character_data_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ heap_is_wellformed_children_in_heap by (meson NodeMonad.ptr_kinds_ptr_kinds_M a_all_ptrs_in_heap_def assms(3) assms(4) fset_mp fset_of_list_elem get_child_nodes_ok known_ptrs_known_ptr returns_result_select_result) then have "\p. p |\| object_ptr_kinds h2 \ cast new_character_data_ptr \ set |h2 \ get_child_nodes p|\<^sub>r" using children_eq2_h apply(auto simp add: object_ptr_kinds_eq_h)[1] using \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ apply auto[1] by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\) then have "\p. p |\| object_ptr_kinds h3 \ cast new_character_data_ptr \ set |h3 \ get_child_nodes p|\<^sub>r" using object_ptr_kinds_eq_h2 children_eq2_h2 by auto then have new_character_data_ptr_not_in_any_children: "\p. p |\| object_ptr_kinds h' \ cast new_character_data_ptr \ set |h' \ get_child_nodes p|\<^sub>r" using object_ptr_kinds_eq_h3 children_eq2_h3 by auto have "a_distinct_lists h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) then have "a_distinct_lists h2" using \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ apply(auto simp add: a_distinct_lists_def object_ptr_kinds_eq_h document_ptr_kinds_eq_h disconnected_nodes_eq2_h intro!: distinct_concat_map_I)[1] - apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert) + apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert_remove) apply(case_tac "x=cast new_character_data_ptr") apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply (metis IntI assms(1) assms(3) assms(4) empty_iff local.get_child_nodes_ok local.heap_is_wellformed_one_parent local.known_ptrs_known_ptr returns_result_select_result) apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] by (metis \local.a_distinct_lists h\ \type_wf h2\ disconnected_nodes_eq_h document_ptr_kinds_eq_h local.distinct_lists_no_parent local.get_disconnected_nodes_ok returns_result_select_result) then have "a_distinct_lists h3" by(auto simp add: a_distinct_lists_def disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 children_eq2_h2 object_ptr_kinds_eq_h2)[1] then have "a_distinct_lists h'" proof(auto simp add: a_distinct_lists_def disconnected_nodes_eq2_h3 children_eq2_h3 object_ptr_kinds_eq_h3 document_ptr_kinds_eq_h3 intro!: distinct_concat_map_I)[1] fix x assume "distinct (concat (map (\document_ptr. |h3 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" and "x |\| document_ptr_kinds h3" then show "distinct |h' \ get_disconnected_nodes x|\<^sub>r" using document_ptr_kinds_eq_h3 disconnected_nodes_eq_h3 h' set_disconnected_nodes_get_disconnected_nodes by (metis (no_types, lifting) \cast new_character_data_ptr \ set disc_nodes_h3\ \a_distinct_lists h3\ \type_wf h'\ disc_nodes_h3 distinct.simps(2) distinct_lists_disconnected_nodes get_disconnected_nodes_ok returns_result_eq returns_result_select_result) next fix x y xa assume "distinct (concat (map (\document_ptr. |h3 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" and "x |\| document_ptr_kinds h3" and "y |\| document_ptr_kinds h3" and "x \ y" and "xa \ set |h' \ get_disconnected_nodes x|\<^sub>r" and "xa \ set |h' \ get_disconnected_nodes y|\<^sub>r" moreover have "set |h3 \ get_disconnected_nodes x|\<^sub>r \ set |h3 \ get_disconnected_nodes y|\<^sub>r = {}" using calculation by(auto dest: distinct_concat_map_E(1)) ultimately show "False" by (smt NodeMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_character_data_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ \local.a_all_ptrs_in_heap h\ disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 finite_set_in h' l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms select_result_I2 set_ConsD subsetD) next fix x xa xb assume 2: "(\x\fset (object_ptr_kinds h3). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h3). set |h3 \ get_disconnected_nodes x|\<^sub>r) = {}" and 3: "xa |\| object_ptr_kinds h3" and 4: "x \ set |h' \ get_child_nodes xa|\<^sub>r" and 5: "xb |\| document_ptr_kinds h3" and 6: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" show "False" using disc_nodes_document_ptr_h disconnected_nodes_eq2_h3 apply(cases "xb = document_ptr") apply (metis (no_types, hide_lams) "3" "4" "6" \\p. p |\| object_ptr_kinds h3 \ cast new_character_data_ptr \ set |h3 \ get_child_nodes p|\<^sub>r\ \a_distinct_lists h3\ children_eq2_h3 disc_nodes_h3 distinct_lists_no_parent h' select_result_I2 set_ConsD set_disconnected_nodes_get_disconnected_nodes) by (metis "3" "4" "5" "6" \a_distinct_lists h3\ \type_wf h3\ children_eq2_h3 distinct_lists_no_parent get_disconnected_nodes_ok returns_result_select_result) qed have "a_owner_document_valid h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) then have "a_owner_document_valid h'" using disc_nodes_h3 \document_ptr |\| document_ptr_kinds h\ apply(simp add: a_owner_document_valid_def) apply(simp add: object_ptr_kinds_eq_h object_ptr_kinds_eq_h3 ) apply(simp add: object_ptr_kinds_eq_h2) apply(simp add: document_ptr_kinds_eq_h document_ptr_kinds_eq_h3 ) apply(simp add: document_ptr_kinds_eq_h2) apply(simp add: node_ptr_kinds_eq_h node_ptr_kinds_eq_h3 ) apply(simp add: node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h ) apply(auto simp add: children_eq2_h2[symmetric] children_eq2_h3[symmetric] disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3)[1] apply (metis (no_types, lifting) document_ptr_kinds_eq_h h' list.set_intros(1) local.set_disconnected_nodes_get_disconnected_nodes select_result_I2) apply(simp add: object_ptr_kinds_eq_h) by (smt ObjectMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2_h disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h finite_set_in h' list.set_intros(2) local.set_disconnected_nodes_get_disconnected_nodes select_result_I2) have "known_ptr (cast new_character_data_ptr)" using \h \ create_character_data document_ptr text \\<^sub>r new_character_data_ptr\ local.create_character_data_known_ptr by blast then have "known_ptrs h2" using known_ptrs_new_ptr object_ptr_kinds_eq_h \known_ptrs h\ h2 by blast then have "known_ptrs h3" using known_ptrs_preserved object_ptr_kinds_eq_h2 by blast then show "known_ptrs h'" using known_ptrs_preserved object_ptr_kinds_eq_h3 by blast show "heap_is_wellformed h'" using \a_acyclic_heap h'\ \a_all_ptrs_in_heap h'\ \a_distinct_lists h'\ \a_owner_document_valid h'\ by(simp add: heap_is_wellformed_def) qed end interpretation i_create_character_data_wf?: l_create_character_data_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed parent_child_rel set_val set_val_locs set_disconnected_nodes set_disconnected_nodes_locs create_character_data known_ptrs using instances by (auto simp add: l_create_character_data_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_create_character_data_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsection \create\_document\ locale l_create_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed parent_child_rel + l_new_document_get_disconnected_nodes get_disconnected_nodes get_disconnected_nodes_locs + l_create_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M create_document + l_new_document_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs + l_new_document type_wf + l_known_ptrs known_ptr known_ptrs for known_ptr :: "(_::linorder) object_ptr \ bool" and type_wf :: "(_) heap \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and set_val :: "(_) character_data_ptr \ char list \ ((_) heap, exception, unit) prog" and set_val_locs :: "(_) character_data_ptr \ ((_) heap, exception, unit) prog set" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" and create_document :: "((_) heap, exception, (_) document_ptr) prog" and known_ptrs :: "(_) heap \ bool" begin lemma create_document_preserves_wellformedness: assumes "heap_is_wellformed h" and "h \ create_document \\<^sub>h h'" and "type_wf h" and "known_ptrs h" shows "heap_is_wellformed h'" proof - obtain new_document_ptr where new_document_ptr: "h \ new_document \\<^sub>r new_document_ptr" and h': "h \ new_document \\<^sub>h h'" using assms(2) apply(simp add: create_document_def) using new_document_ok by blast have "new_document_ptr \ set |h \ document_ptr_kinds_M|\<^sub>r" using new_document_ptr DocumentMonad.ptr_kinds_ptr_kinds_M using new_document_ptr_not_in_heap h' by blast then have "cast new_document_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r" by simp have "new_document_ptr |\| document_ptr_kinds h" using new_document_ptr DocumentMonad.ptr_kinds_ptr_kinds_M using new_document_ptr_not_in_heap h' by blast then have "cast new_document_ptr |\| object_ptr_kinds h" by simp have object_ptr_kinds_eq: "object_ptr_kinds h' = object_ptr_kinds h |\| {|cast new_document_ptr|}" using new_document_new_ptr h' new_document_ptr by blast then have node_ptr_kinds_eq: "node_ptr_kinds h' = node_ptr_kinds h" apply(simp add: node_ptr_kinds_def) by force then have character_data_ptr_kinds_eq_h: "character_data_ptr_kinds h' = character_data_ptr_kinds h" by(simp add: character_data_ptr_kinds_def) have element_ptr_kinds_eq_h: "element_ptr_kinds h' = element_ptr_kinds h" using object_ptr_kinds_eq by(auto simp add: node_ptr_kinds_def element_ptr_kinds_def) have document_ptr_kinds_eq_h: "document_ptr_kinds h' = document_ptr_kinds h |\| {|new_document_ptr|}" using object_ptr_kinds_eq apply(auto simp add: document_ptr_kinds_def)[1] by (metis (no_types, lifting) document_ptr_kinds_commutes document_ptr_kinds_def finsertI1 fset.map_comp) have children_eq: "\(ptr'::(_) object_ptr) children. ptr' \ cast new_document_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads h' get_child_nodes_new_document[rotated, OF new_document_ptr h'] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have children_eq2: "\ptr'. ptr' \ cast new_document_ptr \ |h \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have "h' \ get_child_nodes (cast new_document_ptr) \\<^sub>r []" using new_document_ptr h' new_document_ptr_in_heap[OF h' new_document_ptr] new_document_is_document_ptr[OF new_document_ptr] new_document_no_child_nodes by blast have disconnected_nodes_eq_h: "\doc_ptr disc_nodes. doc_ptr \ new_document_ptr \ h \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads h' get_disconnected_nodes_new_document_different_pointers new_document_ptr apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by (metis(full_types) \\thesis. (\new_document_ptr. \h \ new_document \\<^sub>r new_document_ptr; h \ new_document \\<^sub>h h'\ \ thesis) \ thesis\ local.get_disconnected_nodes_new_document_different_pointers new_document_ptr)+ then have disconnected_nodes_eq2_h: "\doc_ptr. doc_ptr \ new_document_ptr \ |h \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have "h' \ get_disconnected_nodes new_document_ptr \\<^sub>r []" using h' local.new_document_no_disconnected_nodes new_document_ptr by blast have "type_wf h'" using \type_wf h\ new_document_types_preserved h' by blast have "acyclic (parent_child_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def acyclic_heap_def) also have "parent_child_rel h = parent_child_rel h'" proof(auto simp add: parent_child_rel_def)[1] fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h'" by (simp add: object_ptr_kinds_eq) next fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "x \ set |h' \ get_child_nodes a|\<^sub>r" by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast new_document_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2) next fix a x assume 0: "a |\| object_ptr_kinds h'" and 1: "x \ set |h' \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h" using object_ptr_kinds_eq \h' \ get_child_nodes (cast new_document_ptr) \\<^sub>r []\ by(auto) next fix a x assume 0: "a |\| object_ptr_kinds h'" and 1: "x \ set |h' \ get_child_nodes a|\<^sub>r" then show "x \ set |h \ get_child_nodes a|\<^sub>r" by (metis (no_types, lifting) \h' \ get_child_nodes (cast new_document_ptr) \\<^sub>r []\ children_eq2 empty_iff empty_set image_eqI select_result_I2) qed finally have "a_acyclic_heap h'" by (simp add: acyclic_heap_def) have "a_all_ptrs_in_heap h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) then have "a_all_ptrs_in_heap h'" apply(auto simp add: a_all_ptrs_in_heap_def)[1] using ObjectMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ \parent_child_rel h = parent_child_rel h'\ assms(1) children_eq fset_of_list_elem local.heap_is_wellformed_children_in_heap local.parent_child_rel_child local.parent_child_rel_parent_in_heap node_ptr_kinds_eq apply (metis (no_types, lifting) \h' \ get_child_nodes (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr) \\<^sub>r []\ children_eq2 finite_set_in finsert_iff funion_finsert_right object_ptr_kinds_eq select_result_I2 subsetD sup_bot.right_neutral) by (metis (no_types, lifting) \cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr |\| object_ptr_kinds h\ \h' \ get_child_nodes (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr) \\<^sub>r []\ \h' \ get_disconnected_nodes new_document_ptr \\<^sub>r []\ \parent_child_rel h = parent_child_rel h'\ \type_wf h'\ assms(1) disconnected_nodes_eq_h local.get_disconnected_nodes_ok local.heap_is_wellformed_disc_nodes_in_heap local.parent_child_rel_child local.parent_child_rel_parent_in_heap node_ptr_kinds_eq returns_result_select_result select_result_I2) have "a_distinct_lists h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) then have "a_distinct_lists h'" using \h' \ get_disconnected_nodes new_document_ptr \\<^sub>r []\ \h' \ get_child_nodes (cast new_document_ptr) \\<^sub>r []\ apply(auto simp add: children_eq2[symmetric] a_distinct_lists_def insort_split object_ptr_kinds_eq document_ptr_kinds_eq_h disconnected_nodes_eq2_h intro!: distinct_concat_map_I)[1] - apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert) + apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert_remove) apply(auto simp add: dest: distinct_concat_map_E)[1] apply(auto simp add: dest: distinct_concat_map_E)[1] using \new_document_ptr |\| document_ptr_kinds h\ apply(auto simp add: distinct_insort dest: distinct_concat_map_E)[1] using disconnected_nodes_eq_h apply (metis assms(1) assms(3) disconnected_nodes_eq2_h local.get_disconnected_nodes_ok local.heap_is_wellformed_disconnected_nodes_distinct returns_result_select_result) proof - fix x :: "(_) document_ptr" and y :: "(_) document_ptr" and xa :: "(_) node_ptr" assume a1: "x \ y" assume a2: "x |\| document_ptr_kinds h" assume a3: "x \ new_document_ptr" assume a4: "y |\| document_ptr_kinds h" assume a5: "y \ new_document_ptr" assume a6: "distinct (concat (map (\document_ptr. |h \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h)))))" assume a7: "xa \ set |h' \ get_disconnected_nodes x|\<^sub>r" assume a8: "xa \ set |h' \ get_disconnected_nodes y|\<^sub>r" have f9: "xa \ set |h \ get_disconnected_nodes x|\<^sub>r" using a7 a3 disconnected_nodes_eq2_h by presburger have f10: "xa \ set |h \ get_disconnected_nodes y|\<^sub>r" using a8 a5 disconnected_nodes_eq2_h by presburger have f11: "y \ set (sorted_list_of_set (fset (document_ptr_kinds h)))" using a4 by simp have "x \ set (sorted_list_of_set (fset (document_ptr_kinds h)))" using a2 by simp then show False using f11 f10 f9 a6 a1 by (meson disjoint_iff_not_equal distinct_concat_map_E(1)) next fix x xa xb assume 0: "h' \ get_disconnected_nodes new_document_ptr \\<^sub>r []" and 1: "h' \ get_child_nodes (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr) \\<^sub>r []" and 2: "distinct (concat (map (\ptr. |h \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h)))))" and 3: "distinct (concat (map (\document_ptr. |h \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h)))))" and 4: "(\x\fset (object_ptr_kinds h). set |h \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h). set |h \ get_disconnected_nodes x|\<^sub>r) = {}" and 5: "x \ set |h \ get_child_nodes xa|\<^sub>r" and 6: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" and 7: "xa |\| object_ptr_kinds h" and 8: "xa \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr" and 9: "xb |\| document_ptr_kinds h" and 10: "xb \ new_document_ptr" then show "False" by (metis \local.a_distinct_lists h\ assms(3) disconnected_nodes_eq2_h local.distinct_lists_no_parent local.get_disconnected_nodes_ok returns_result_select_result) qed have "a_owner_document_valid h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) then have "a_owner_document_valid h'" apply(auto simp add: a_owner_document_valid_def)[1] by (metis \cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr |\| object_ptr_kinds h\ children_eq2 disconnected_nodes_eq2_h document_ptr_kinds_commutes finite_set_in funion_iff node_ptr_kinds_eq object_ptr_kinds_eq) show "heap_is_wellformed h'" using \a_acyclic_heap h'\ \a_all_ptrs_in_heap h'\ \a_distinct_lists h'\ \a_owner_document_valid h'\ by(simp add: heap_is_wellformed_def) qed end interpretation i_create_document_wf?: l_create_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed parent_child_rel set_val set_val_locs set_disconnected_nodes set_disconnected_nodes_locs create_document known_ptrs using instances by (auto simp add: l_create_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_create_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] end diff --git a/thys/Core_SC_DOM/safely_composable/Core_DOM_Heap_WF.thy b/thys/Core_SC_DOM/safely_composable/Core_DOM_Heap_WF.thy --- a/thys/Core_SC_DOM/safely_composable/Core_DOM_Heap_WF.thy +++ b/thys/Core_SC_DOM/safely_composable/Core_DOM_Heap_WF.thy @@ -1,6965 +1,6965 @@ (*********************************************************************************** * Copyright (c) 2016-2018 The University of Sheffield, UK * * 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. * * 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 HOLDER 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. * * SPDX-License-Identifier: BSD-2-Clause ***********************************************************************************) section\Wellformedness\ text\In this theory, we discuss the wellformedness of the DOM. First, we define wellformedness and, second, we show for all functions for querying and modifying the DOM to what extend they preserve wellformendess.\ theory Core_DOM_Heap_WF imports "Core_DOM_Functions" begin locale l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_child_nodes_defs get_child_nodes get_child_nodes_locs + l_get_disconnected_nodes_defs get_disconnected_nodes get_disconnected_nodes_locs for get_child_nodes :: "(_::linorder) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" begin definition a_owner_document_valid :: "(_) heap \ bool" where "a_owner_document_valid h \ (\node_ptr \ fset (node_ptr_kinds h). ((\document_ptr. document_ptr |\| document_ptr_kinds h \ node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r) \ (\parent_ptr. parent_ptr |\| object_ptr_kinds h \ node_ptr \ set |h \ get_child_nodes parent_ptr|\<^sub>r)))" lemma a_owner_document_valid_code [code]: "a_owner_document_valid h \ node_ptr_kinds h |\| fset_of_list (concat (map (\parent. |h \ get_child_nodes parent|\<^sub>r) (sorted_list_of_fset (object_ptr_kinds h)) @ map (\parent. |h \ get_disconnected_nodes parent|\<^sub>r) (sorted_list_of_fset (document_ptr_kinds h)))) " apply(auto simp add: a_owner_document_valid_def l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_owner_document_valid_def)[1] proof - fix x assume 1: " \node_ptr\fset (node_ptr_kinds h). (\document_ptr. document_ptr |\| document_ptr_kinds h \ node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r) \ (\parent_ptr. parent_ptr |\| object_ptr_kinds h \ node_ptr \ set |h \ get_child_nodes parent_ptr|\<^sub>r)" assume 2: "x |\| node_ptr_kinds h" assume 3: "x |\| fset_of_list (concat (map (\parent. |h \ get_disconnected_nodes parent|\<^sub>r) (sorted_list_of_fset (document_ptr_kinds h))))" have "\(\document_ptr. document_ptr |\| document_ptr_kinds h \ x \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r)" using 1 2 3 by (smt UN_I fset_of_list_elem image_eqI notin_fset set_concat set_map sorted_list_of_fset_simps(1)) then have "(\parent_ptr. parent_ptr |\| object_ptr_kinds h \ x \ set |h \ get_child_nodes parent_ptr|\<^sub>r)" using 1 2 by auto then obtain parent_ptr where parent_ptr: "parent_ptr |\| object_ptr_kinds h \ x \ set |h \ get_child_nodes parent_ptr|\<^sub>r" by auto moreover have "parent_ptr \ set (sorted_list_of_fset (object_ptr_kinds h))" using parent_ptr by auto moreover have "|h \ get_child_nodes parent_ptr|\<^sub>r \ set (map (\parent. |h \ get_child_nodes parent|\<^sub>r) (sorted_list_of_fset (object_ptr_kinds h)))" using calculation(2) by auto ultimately show "x |\| fset_of_list (concat (map (\parent. |h \ get_child_nodes parent|\<^sub>r) (sorted_list_of_fset (object_ptr_kinds h))))" using fset_of_list_elem by fastforce next fix node_ptr assume 1: "node_ptr_kinds h |\| fset_of_list (concat (map (\parent. |h \ get_child_nodes parent|\<^sub>r) (sorted_list_of_fset (object_ptr_kinds h)))) |\| fset_of_list (concat (map (\parent. |h \ get_disconnected_nodes parent|\<^sub>r) (sorted_list_of_fset (document_ptr_kinds h))))" assume 2: "node_ptr |\| node_ptr_kinds h" assume 3: "\parent_ptr. parent_ptr |\| object_ptr_kinds h \ node_ptr \ set |h \ get_child_nodes parent_ptr|\<^sub>r" have "node_ptr \ set (concat (map (\parent. |h \ get_child_nodes parent|\<^sub>r) (sorted_list_of_fset (object_ptr_kinds h)))) \ node_ptr \ set (concat (map (\parent. |h \ get_disconnected_nodes parent|\<^sub>r) (sorted_list_of_fset (document_ptr_kinds h))))" using 1 2 by (meson fin_mono fset_of_list_elem funion_iff) then show "\document_ptr. document_ptr |\| document_ptr_kinds h \ node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r" using 3 by auto qed definition a_parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" where "a_parent_child_rel h = {(parent, child). parent |\| object_ptr_kinds h \ child \ cast ` set |h \ get_child_nodes parent|\<^sub>r}" lemma a_parent_child_rel_code [code]: "a_parent_child_rel h = set (concat (map (\parent. map (\child. (parent, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child)) |h \ get_child_nodes parent|\<^sub>r) (sorted_list_of_fset (object_ptr_kinds h))) )" by(auto simp add: a_parent_child_rel_def) definition a_acyclic_heap :: "(_) heap \ bool" where "a_acyclic_heap h = acyclic (a_parent_child_rel h)" definition a_all_ptrs_in_heap :: "(_) heap \ bool" where "a_all_ptrs_in_heap h \ (\ptr \ fset (object_ptr_kinds h). set |h \ get_child_nodes ptr|\<^sub>r \ fset (node_ptr_kinds h)) \ (\document_ptr \ fset (document_ptr_kinds h). set |h \ get_disconnected_nodes document_ptr|\<^sub>r \ fset (node_ptr_kinds h))" definition a_distinct_lists :: "(_) heap \ bool" where "a_distinct_lists h = distinct (concat ( (map (\ptr. |h \ get_child_nodes ptr|\<^sub>r) |h \ object_ptr_kinds_M|\<^sub>r) @ (map (\document_ptr. |h \ get_disconnected_nodes document_ptr|\<^sub>r) |h \ document_ptr_kinds_M|\<^sub>r) ))" definition a_heap_is_wellformed :: "(_) heap \ bool" where "a_heap_is_wellformed h \ a_acyclic_heap h \ a_all_ptrs_in_heap h \ a_distinct_lists h \ a_owner_document_valid h" end locale l_heap_is_wellformed_defs = fixes heap_is_wellformed :: "(_) heap \ bool" fixes parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" global_interpretation l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs defines heap_is_wellformed = "l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_heap_is_wellformed get_child_nodes get_disconnected_nodes" and parent_child_rel = "l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_parent_child_rel get_child_nodes" and acyclic_heap = a_acyclic_heap and all_ptrs_in_heap = a_all_ptrs_in_heap and distinct_lists = a_distinct_lists and owner_document_valid = a_owner_document_valid . locale l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs + l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs + l_heap_is_wellformed_defs heap_is_wellformed parent_child_rel + l_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs for known_ptr :: "(_::linorder) object_ptr \ bool" and type_wf :: "(_) heap \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" + assumes heap_is_wellformed_impl: "heap_is_wellformed = a_heap_is_wellformed" assumes parent_child_rel_impl: "parent_child_rel = a_parent_child_rel" begin lemmas heap_is_wellformed_def = heap_is_wellformed_impl[unfolded a_heap_is_wellformed_def] lemmas parent_child_rel_def = parent_child_rel_impl[unfolded a_parent_child_rel_def] lemmas acyclic_heap_def = a_acyclic_heap_def[folded parent_child_rel_impl] lemma parent_child_rel_node_ptr: "(parent, child) \ parent_child_rel h \ is_node_ptr_kind child" by(auto simp add: parent_child_rel_def) lemma parent_child_rel_child_nodes: assumes "known_ptr parent" and "h \ get_child_nodes parent \\<^sub>r children" and "child \ set children" shows "(parent, cast child) \ parent_child_rel h" using assms apply(auto simp add: parent_child_rel_def is_OK_returns_result_I )[1] using get_child_nodes_ptr_in_heap by blast lemma parent_child_rel_child_nodes2: assumes "known_ptr parent" and "h \ get_child_nodes parent \\<^sub>r children" and "child \ set children" and "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child = child_obj" shows "(parent, child_obj) \ parent_child_rel h" using assms parent_child_rel_child_nodes by blast lemma parent_child_rel_finite: "finite (parent_child_rel h)" proof - have "parent_child_rel h = (\ptr \ set |h \ object_ptr_kinds_M|\<^sub>r. (\child \ set |h \ get_child_nodes ptr|\<^sub>r. {(ptr, cast child)}))" by(auto simp add: parent_child_rel_def) moreover have "finite (\ptr \ set |h \ object_ptr_kinds_M|\<^sub>r. (\child \ set |h \ get_child_nodes ptr|\<^sub>r. {(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child)}))" by simp ultimately show ?thesis by simp qed lemma distinct_lists_no_parent: assumes "a_distinct_lists h" assumes "h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes" assumes "node_ptr \ set disc_nodes" shows "\(\parent_ptr. parent_ptr |\| object_ptr_kinds h \ node_ptr \ set |h \ get_child_nodes parent_ptr|\<^sub>r)" using assms apply(auto simp add: a_distinct_lists_def)[1] proof - fix parent_ptr :: "(_) object_ptr" assume a1: "parent_ptr |\| object_ptr_kinds h" assume a2: "(\x\fset (object_ptr_kinds h). set |h \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h). set |h \ get_disconnected_nodes x|\<^sub>r) = {}" assume a3: "h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes" assume a4: "node_ptr \ set disc_nodes" assume a5: "node_ptr \ set |h \ get_child_nodes parent_ptr|\<^sub>r" have f6: "parent_ptr \ fset (object_ptr_kinds h)" using a1 by auto have f7: "document_ptr \ fset (document_ptr_kinds h)" using a3 by (meson fmember.rep_eq get_disconnected_nodes_ptr_in_heap is_OK_returns_result_I) have "|h \ get_disconnected_nodes document_ptr|\<^sub>r = disc_nodes" using a3 by simp then show False using f7 f6 a5 a4 a2 by blast qed lemma distinct_lists_disconnected_nodes: assumes "a_distinct_lists h" and "h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes" shows "distinct disc_nodes" proof - have h1: "distinct (concat (map (\document_ptr. |h \ get_disconnected_nodes document_ptr|\<^sub>r) |h \ document_ptr_kinds_M|\<^sub>r))" using assms(1) by(simp add: a_distinct_lists_def) then show ?thesis using concat_map_all_distinct[OF h1] assms(2) is_OK_returns_result_I get_disconnected_nodes_ok by (metis (no_types, lifting) DocumentMonad.ptr_kinds_ptr_kinds_M l_get_disconnected_nodes.get_disconnected_nodes_ptr_in_heap l_get_disconnected_nodes_axioms select_result_I2) qed lemma distinct_lists_children: assumes "a_distinct_lists h" and "known_ptr ptr" and "h \ get_child_nodes ptr \\<^sub>r children" shows "distinct children" proof (cases "children = []", simp) assume "children \ []" have h1: "distinct (concat ((map (\ptr. |h \ get_child_nodes ptr|\<^sub>r) |h \ object_ptr_kinds_M|\<^sub>r)))" using assms(1) by(simp add: a_distinct_lists_def) show ?thesis using concat_map_all_distinct[OF h1] assms(2) assms(3) by (metis (no_types, lifting) ObjectMonad.ptr_kinds_ptr_kinds_M get_child_nodes_ptr_in_heap is_OK_returns_result_I select_result_I2) qed lemma heap_is_wellformed_children_in_heap: assumes "heap_is_wellformed h" assumes "h \ get_child_nodes ptr \\<^sub>r children" assumes "child \ set children" shows "child |\| node_ptr_kinds h" using assms apply(auto simp add: heap_is_wellformed_def a_all_ptrs_in_heap_def)[1] by (metis (no_types, lifting) finite_set_in is_OK_returns_result_I local.get_child_nodes_ptr_in_heap select_result_I2 subsetD) lemma heap_is_wellformed_one_parent: assumes "heap_is_wellformed h" assumes "h \ get_child_nodes ptr \\<^sub>r children" assumes "h \ get_child_nodes ptr' \\<^sub>r children'" assumes "set children \ set children' \ {}" shows "ptr = ptr'" using assms proof (auto simp add: heap_is_wellformed_def a_distinct_lists_def)[1] fix x :: "(_) node_ptr" assume a1: "ptr \ ptr'" assume a2: "h \ get_child_nodes ptr \\<^sub>r children" assume a3: "h \ get_child_nodes ptr' \\<^sub>r children'" assume a4: "distinct (concat (map (\ptr. |h \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h)))))" have f5: "|h \ get_child_nodes ptr|\<^sub>r = children" using a2 by simp have "|h \ get_child_nodes ptr'|\<^sub>r = children'" using a3 by (meson select_result_I2) then have "ptr \ set (sorted_list_of_set (fset (object_ptr_kinds h))) \ ptr' \ set (sorted_list_of_set (fset (object_ptr_kinds h))) \ set children \ set children' = {}" using f5 a4 a1 by (meson distinct_concat_map_E(1)) then show False using a3 a2 by (metis (no_types) assms(4) finite_fset fmember.rep_eq is_OK_returns_result_I local.get_child_nodes_ptr_in_heap set_sorted_list_of_set) qed lemma parent_child_rel_child: "h \ get_child_nodes ptr \\<^sub>r children \ child \ set children \ (ptr, cast child) \ parent_child_rel h" by (simp add: is_OK_returns_result_I get_child_nodes_ptr_in_heap parent_child_rel_def) lemma parent_child_rel_acyclic: "heap_is_wellformed h \ acyclic (parent_child_rel h)" by (simp add: acyclic_heap_def local.heap_is_wellformed_def) lemma heap_is_wellformed_disconnected_nodes_distinct: "heap_is_wellformed h \ h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes \ distinct disc_nodes" using distinct_lists_disconnected_nodes local.heap_is_wellformed_def by blast lemma parent_child_rel_parent_in_heap: "(parent, child_ptr) \ parent_child_rel h \ parent |\| object_ptr_kinds h" using local.parent_child_rel_def by blast lemma parent_child_rel_child_in_heap: "heap_is_wellformed h \ type_wf h \ known_ptr parent \ (parent, child_ptr) \ parent_child_rel h \ child_ptr |\| object_ptr_kinds h" apply(auto simp add: heap_is_wellformed_def parent_child_rel_def a_all_ptrs_in_heap_def)[1] using get_child_nodes_ok by (meson finite_set_in subsetD) lemma heap_is_wellformed_disc_nodes_in_heap: "heap_is_wellformed h \ h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes \ node \ set disc_nodes \ node |\| node_ptr_kinds h" by (metis (no_types, lifting) finite_set_in is_OK_returns_result_I local.a_all_ptrs_in_heap_def local.get_disconnected_nodes_ptr_in_heap local.heap_is_wellformed_def select_result_I2 subsetD) lemma heap_is_wellformed_one_disc_parent: "heap_is_wellformed h \ h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes \ h \ get_disconnected_nodes document_ptr' \\<^sub>r disc_nodes' \ set disc_nodes \ set disc_nodes' \ {} \ document_ptr = document_ptr'" using DocumentMonad.ptr_kinds_ptr_kinds_M concat_append distinct_append distinct_concat_map_E(1) is_OK_returns_result_I local.a_distinct_lists_def local.get_disconnected_nodes_ptr_in_heap local.heap_is_wellformed_def select_result_I2 proof - assume a1: "heap_is_wellformed h" assume a2: "h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes" assume a3: "h \ get_disconnected_nodes document_ptr' \\<^sub>r disc_nodes'" assume a4: "set disc_nodes \ set disc_nodes' \ {}" have f5: "|h \ get_disconnected_nodes document_ptr|\<^sub>r = disc_nodes" using a2 by (meson select_result_I2) have f6: "|h \ get_disconnected_nodes document_ptr'|\<^sub>r = disc_nodes'" using a3 by (meson select_result_I2) have "\nss nssa. \ distinct (concat (nss @ nssa)) \ distinct (concat nssa::(_) node_ptr list)" by (metis (no_types) concat_append distinct_append) then have "distinct (concat (map (\d. |h \ get_disconnected_nodes d|\<^sub>r) |h \ document_ptr_kinds_M|\<^sub>r))" using a1 local.a_distinct_lists_def local.heap_is_wellformed_def by blast then show ?thesis using f6 f5 a4 a3 a2 by (meson DocumentMonad.ptr_kinds_ptr_kinds_M distinct_concat_map_E(1) is_OK_returns_result_I local.get_disconnected_nodes_ptr_in_heap) qed lemma heap_is_wellformed_children_disc_nodes_different: "heap_is_wellformed h \ h \ get_child_nodes ptr \\<^sub>r children \ h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes \ set children \ set disc_nodes = {}" by (metis (no_types, hide_lams) disjoint_iff_not_equal distinct_lists_no_parent is_OK_returns_result_I local.get_child_nodes_ptr_in_heap local.heap_is_wellformed_def select_result_I2) lemma heap_is_wellformed_children_disc_nodes: "heap_is_wellformed h \ node_ptr |\| node_ptr_kinds h \ \(\parent \ fset (object_ptr_kinds h). node_ptr \ set |h \ get_child_nodes parent|\<^sub>r) \ (\document_ptr \ fset (document_ptr_kinds h). node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r)" apply(auto simp add: heap_is_wellformed_def a_distinct_lists_def a_owner_document_valid_def)[1] by (meson fmember.rep_eq) lemma heap_is_wellformed_children_distinct: "heap_is_wellformed h \ h \ get_child_nodes ptr \\<^sub>r children \ distinct children" by (metis (no_types, lifting) ObjectMonad.ptr_kinds_ptr_kinds_M concat_append distinct_append distinct_concat_map_E(2) is_OK_returns_result_I local.a_distinct_lists_def local.get_child_nodes_ptr_in_heap local.heap_is_wellformed_def select_result_I2) end locale l_heap_is_wellformed = l_type_wf + l_known_ptr + l_heap_is_wellformed_defs + l_get_child_nodes_defs + l_get_disconnected_nodes_defs + assumes heap_is_wellformed_children_in_heap: "heap_is_wellformed h \ h \ get_child_nodes ptr \\<^sub>r children \ child \ set children \ child |\| node_ptr_kinds h" assumes heap_is_wellformed_disc_nodes_in_heap: "heap_is_wellformed h \ h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes \ node \ set disc_nodes \ node |\| node_ptr_kinds h" assumes heap_is_wellformed_one_parent: "heap_is_wellformed h \ h \ get_child_nodes ptr \\<^sub>r children \ h \ get_child_nodes ptr' \\<^sub>r children' \ set children \ set children' \ {} \ ptr = ptr'" assumes heap_is_wellformed_one_disc_parent: "heap_is_wellformed h \ h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes \ h \ get_disconnected_nodes document_ptr' \\<^sub>r disc_nodes' \ set disc_nodes \ set disc_nodes' \ {} \ document_ptr = document_ptr'" assumes heap_is_wellformed_children_disc_nodes_different: "heap_is_wellformed h \ h \ get_child_nodes ptr \\<^sub>r children \ h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes \ set children \ set disc_nodes = {}" assumes heap_is_wellformed_disconnected_nodes_distinct: "heap_is_wellformed h \ h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes \ distinct disc_nodes" assumes heap_is_wellformed_children_distinct: "heap_is_wellformed h \ h \ get_child_nodes ptr \\<^sub>r children \ distinct children" assumes heap_is_wellformed_children_disc_nodes: "heap_is_wellformed h \ node_ptr |\| node_ptr_kinds h \ \(\parent \ fset (object_ptr_kinds h). node_ptr \ set |h \ get_child_nodes parent|\<^sub>r) \ (\document_ptr \ fset (document_ptr_kinds h). node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r)" assumes parent_child_rel_child: "h \ get_child_nodes ptr \\<^sub>r children \ child \ set children \ (ptr, cast child) \ parent_child_rel h" assumes parent_child_rel_finite: "heap_is_wellformed h \ finite (parent_child_rel h)" assumes parent_child_rel_acyclic: "heap_is_wellformed h \ acyclic (parent_child_rel h)" assumes parent_child_rel_node_ptr: "(parent, child_ptr) \ parent_child_rel h \ is_node_ptr_kind child_ptr" assumes parent_child_rel_parent_in_heap: "(parent, child_ptr) \ parent_child_rel h \ parent |\| object_ptr_kinds h" assumes parent_child_rel_child_in_heap: "heap_is_wellformed h \ type_wf h \ known_ptr parent \ (parent, child_ptr) \ parent_child_rel h \ child_ptr |\| object_ptr_kinds h" interpretation i_heap_is_wellformed?: l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed parent_child_rel apply(unfold_locales) by(auto simp add: heap_is_wellformed_def parent_child_rel_def) declare l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma heap_is_wellformed_is_l_heap_is_wellformed [instances]: "l_heap_is_wellformed type_wf known_ptr heap_is_wellformed parent_child_rel get_child_nodes get_disconnected_nodes" apply(auto simp add: l_heap_is_wellformed_def)[1] using heap_is_wellformed_children_in_heap apply blast using heap_is_wellformed_disc_nodes_in_heap apply blast using heap_is_wellformed_one_parent apply blast using heap_is_wellformed_one_disc_parent apply blast using heap_is_wellformed_children_disc_nodes_different apply blast using heap_is_wellformed_disconnected_nodes_distinct apply blast using heap_is_wellformed_children_distinct apply blast using heap_is_wellformed_children_disc_nodes apply blast using parent_child_rel_child apply (blast) using parent_child_rel_child apply(blast) using parent_child_rel_finite apply blast using parent_child_rel_acyclic apply blast using parent_child_rel_node_ptr apply blast using parent_child_rel_parent_in_heap apply blast using parent_child_rel_child_in_heap apply blast done subsection \get\_parent\ locale l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs get_parent get_parent_locs + l_heap_is_wellformed type_wf known_ptr heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs for known_ptr :: "(_::linorder) object_ptr \ bool" and type_wf :: "(_) heap \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and known_ptrs :: "(_) heap \ bool" and get_parent :: "(_) node_ptr \ ((_) heap, exception, (_) object_ptr option) prog" and get_parent_locs :: "((_) heap \ (_) heap \ bool) set" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma child_parent_dual: assumes heap_is_wellformed: "heap_is_wellformed h" assumes "h \ get_child_nodes ptr \\<^sub>r children" assumes "child \ set children" assumes "known_ptrs h" assumes type_wf: "type_wf h" shows "h \ get_parent child \\<^sub>r Some ptr" proof - obtain ptrs where ptrs: "h \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have h1: "ptr \ set ptrs" using get_child_nodes_ok assms(2) is_OK_returns_result_I by (metis (no_types, hide_lams) ObjectMonad.ptr_kinds_ptr_kinds_M \\thesis. (\ptrs. h \ object_ptr_kinds_M \\<^sub>r ptrs \ thesis) \ thesis\ get_child_nodes_ptr_in_heap returns_result_eq select_result_I2) let ?P = "(\ptr. get_child_nodes ptr \ (\children. return (child \ set children)))" let ?filter = "filter_M ?P ptrs" have "h \ ok ?filter" using ptrs type_wf using get_child_nodes_ok apply(auto intro!: filter_M_is_OK_I bind_is_OK_pure_I get_child_nodes_ok simp add: bind_pure_I)[1] using assms(4) local.known_ptrs_known_ptr by blast then obtain parent_ptrs where parent_ptrs: "h \ ?filter \\<^sub>r parent_ptrs" by auto have h5: "\!x. x \ set ptrs \ h \ Heap_Error_Monad.bind (get_child_nodes x) (\children. return (child \ set children)) \\<^sub>r True" apply(auto intro!: bind_pure_returns_result_I)[1] using heap_is_wellformed_one_parent proof - have "h \ (return (child \ set children)::((_) heap, exception, bool) prog) \\<^sub>r True" by (simp add: assms(3)) then show "\z. z \ set ptrs \ h \ Heap_Error_Monad.bind (get_child_nodes z) (\ns. return (child \ set ns)) \\<^sub>r True" by (metis (no_types) assms(2) bind_pure_returns_result_I2 h1 is_OK_returns_result_I local.get_child_nodes_pure select_result_I2) next fix x y assume 0: "x \ set ptrs" and 1: "h \ Heap_Error_Monad.bind (get_child_nodes x) (\children. return (child \ set children)) \\<^sub>r True" and 2: "y \ set ptrs" and 3: "h \ Heap_Error_Monad.bind (get_child_nodes y) (\children. return (child \ set children)) \\<^sub>r True" and 4: "(\h ptr children ptr' children'. heap_is_wellformed h \ h \ get_child_nodes ptr \\<^sub>r children \ h \ get_child_nodes ptr' \\<^sub>r children' \ set children \ set children' \ {} \ ptr = ptr')" then show "x = y" by (metis (no_types, lifting) bind_returns_result_E disjoint_iff_not_equal heap_is_wellformed return_returns_result) qed have "child |\| node_ptr_kinds h" using heap_is_wellformed_children_in_heap heap_is_wellformed assms(2) assms(3) by fast moreover have "parent_ptrs = [ptr]" apply(rule filter_M_ex1[OF parent_ptrs h1 h5]) using ptrs assms(2) assms(3) by(auto simp add: object_ptr_kinds_M_defs bind_pure_I intro!: bind_pure_returns_result_I) ultimately show ?thesis using ptrs parent_ptrs by(auto simp add: bind_pure_I get_parent_def elim!: bind_returns_result_E2 intro!: bind_pure_returns_result_I filter_M_pure_I) (*slow, ca 1min *) qed lemma parent_child_rel_parent: assumes "heap_is_wellformed h" and "h \ get_parent child_node \\<^sub>r Some parent" shows "(parent, cast child_node) \ parent_child_rel h" using assms parent_child_rel_child get_parent_child_dual by auto lemma heap_wellformed_induct [consumes 1, case_names step]: assumes "heap_is_wellformed h" and step: "\parent. (\children child. h \ get_child_nodes parent \\<^sub>r children \ child \ set children \ P (cast child)) \ P parent" shows "P ptr" proof - fix ptr have "wf ((parent_child_rel h)\)" by (simp add: assms(1) finite_acyclic_wf_converse parent_child_rel_acyclic parent_child_rel_finite) then show "?thesis" proof (induct rule: wf_induct_rule) case (less parent) then show ?case using assms parent_child_rel_child by (meson converse_iff) qed qed lemma heap_wellformed_induct2 [consumes 3, case_names not_in_heap empty_children step]: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and not_in_heap: "\parent. parent |\| object_ptr_kinds h \ P parent" and empty_children: "\parent. h \ get_child_nodes parent \\<^sub>r [] \ P parent" and step: "\parent children child. h \ get_child_nodes parent \\<^sub>r children \ child \ set children \ P (cast child) \ P parent" shows "P ptr" proof(insert assms(1), induct rule: heap_wellformed_induct) case (step parent) then show ?case proof(cases "parent |\| object_ptr_kinds h") case True then obtain children where children: "h \ get_child_nodes parent \\<^sub>r children" using get_child_nodes_ok assms(2) assms(3) by (meson is_OK_returns_result_E local.known_ptrs_known_ptr) then show ?thesis proof (cases "children = []") case True then show ?thesis using children empty_children by simp next case False then show ?thesis using assms(6) children last_in_set step.hyps by blast qed next case False then show ?thesis by (simp add: not_in_heap) qed qed lemma heap_wellformed_induct_rev [consumes 1, case_names step]: assumes "heap_is_wellformed h" and step: "\child. (\parent child_node. cast child_node = child \ h \ get_parent child_node \\<^sub>r Some parent \ P parent) \ P child" shows "P ptr" proof - fix ptr have "wf ((parent_child_rel h))" by (simp add: assms(1) local.parent_child_rel_acyclic local.parent_child_rel_finite wf_iff_acyclic_if_finite) then show "?thesis" proof (induct rule: wf_induct_rule) case (less child) show ?case using assms get_parent_child_dual by (metis less.hyps parent_child_rel_parent) qed qed end interpretation i_get_parent_wf?: l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs get_parent get_parent_locs heap_is_wellformed parent_child_rel get_disconnected_nodes using instances by(simp add: l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs get_parent get_parent_locs heap_is_wellformed parent_child_rel get_disconnected_nodes get_disconnected_nodes_locs + l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed parent_child_rel for known_ptr :: "(_::linorder) object_ptr \ bool" and type_wf :: "(_) heap \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and known_ptrs :: "(_) heap \ bool" and get_parent :: "(_) node_ptr \ ((_) heap, exception, (_) object_ptr option) prog" and get_parent_locs :: "((_) heap \ (_) heap \ bool) set" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma preserves_wellformedness_writes_needed: assumes heap_is_wellformed: "heap_is_wellformed h" and "h \ f \\<^sub>h h'" and "writes SW f h h'" and preserved_get_child_nodes: "\h h' w. w \ SW \ h \ w \\<^sub>h h' \ \object_ptr. \r \ get_child_nodes_locs object_ptr. r h h'" and preserved_get_disconnected_nodes: "\h h' w. w \ SW \ h \ w \\<^sub>h h' \ \document_ptr. \r \ get_disconnected_nodes_locs document_ptr. r h h'" and preserved_object_pointers: "\h h' w. w \ SW \ h \ w \\<^sub>h h' \ \object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'" shows "heap_is_wellformed h'" proof - have object_ptr_kinds_eq3: "object_ptr_kinds h = object_ptr_kinds h'" using assms(2) assms(3) object_ptr_kinds_preserved preserved_object_pointers by blast then have object_ptr_kinds_eq: "\ptrs. h \ object_ptr_kinds_M \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" unfolding object_ptr_kinds_M_defs by simp then have object_ptr_kinds_eq2: "|h \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" using select_result_eq by force then have node_ptr_kinds_eq2: "|h \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by auto then have node_ptr_kinds_eq3: "node_ptr_kinds h = node_ptr_kinds h'" by auto have document_ptr_kinds_eq2: "|h \ document_ptr_kinds_M|\<^sub>r = |h' \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_eq2 document_ptr_kinds_M_eq by auto then have document_ptr_kinds_eq3: "document_ptr_kinds h = document_ptr_kinds h'" by auto have children_eq: "\ptr children. h \ get_child_nodes ptr \\<^sub>r children = h' \ get_child_nodes ptr \\<^sub>r children" apply(rule reads_writes_preserved[OF get_child_nodes_reads assms(3) assms(2)]) using preserved_get_child_nodes by fast then have children_eq2: "\ptr. |h \ get_child_nodes ptr|\<^sub>r = |h' \ get_child_nodes ptr|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq: "\document_ptr disconnected_nodes. h \ get_disconnected_nodes document_ptr \\<^sub>r disconnected_nodes = h' \ get_disconnected_nodes document_ptr \\<^sub>r disconnected_nodes" apply(rule reads_writes_preserved[OF get_disconnected_nodes_reads assms(3) assms(2)]) using preserved_get_disconnected_nodes by fast then have disconnected_nodes_eq2: "\document_ptr. |h \ get_disconnected_nodes document_ptr|\<^sub>r = |h' \ get_disconnected_nodes document_ptr|\<^sub>r" using select_result_eq by force have get_parent_eq: "\ptr parent. h \ get_parent ptr \\<^sub>r parent = h' \ get_parent ptr \\<^sub>r parent" apply(rule reads_writes_preserved[OF get_parent_reads assms(3) assms(2)]) using preserved_get_child_nodes preserved_object_pointers unfolding get_parent_locs_def by fast have "a_acyclic_heap h" using heap_is_wellformed by (simp add: heap_is_wellformed_def) have "parent_child_rel h' \ parent_child_rel h" proof fix x assume "x \ parent_child_rel h'" then show "x \ parent_child_rel h" by(simp add: parent_child_rel_def children_eq2 object_ptr_kinds_eq3) qed then have "a_acyclic_heap h'" using \a_acyclic_heap h\ acyclic_heap_def acyclic_subset by blast moreover have "a_all_ptrs_in_heap h" using heap_is_wellformed by (simp add: heap_is_wellformed_def) then have "a_all_ptrs_in_heap h'" by (simp add: children_eq2 disconnected_nodes_eq2 document_ptr_kinds_eq3 l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_all_ptrs_in_heap_def node_ptr_kinds_eq3 object_ptr_kinds_eq3) moreover have h0: "a_distinct_lists h" using heap_is_wellformed by (simp add: heap_is_wellformed_def) have h1: "map (\ptr. |h \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h))) = map (\ptr. |h' \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h')))" by (simp add: children_eq2 object_ptr_kinds_eq3) have h2: "map (\document_ptr. |h \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h))) = map (\document_ptr. |h' \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h')))" using disconnected_nodes_eq document_ptr_kinds_eq2 select_result_eq by force have "a_distinct_lists h'" using h0 by(simp add: a_distinct_lists_def h1 h2) moreover have "a_owner_document_valid h" using heap_is_wellformed by (simp add: heap_is_wellformed_def) then have "a_owner_document_valid h'" by(auto simp add: a_owner_document_valid_def children_eq2 disconnected_nodes_eq2 object_ptr_kinds_eq3 node_ptr_kinds_eq3 document_ptr_kinds_eq3) ultimately show ?thesis by (simp add: heap_is_wellformed_def) qed end interpretation i_get_parent_wf2?: l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs get_parent get_parent_locs heap_is_wellformed parent_child_rel get_disconnected_nodes get_disconnected_nodes_locs using l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms by (simp add: l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_get_parent_wf = l_type_wf + l_known_ptrs + l_heap_is_wellformed_defs + l_get_child_nodes_defs + l_get_parent_defs + assumes child_parent_dual: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_child_nodes ptr \\<^sub>r children \ child \ set children \ h \ get_parent child \\<^sub>r Some ptr" assumes heap_wellformed_induct [consumes 1, case_names step]: "heap_is_wellformed h \ (\parent. (\children child. h \ get_child_nodes parent \\<^sub>r children \ child \ set children \ P (cast child)) \ P parent) \ P ptr" assumes heap_wellformed_induct_rev [consumes 1, case_names step]: "heap_is_wellformed h \ (\child. (\parent child_node. cast child_node = child \ h \ get_parent child_node \\<^sub>r Some parent \ P parent) \ P child) \ P ptr" assumes parent_child_rel_parent: "heap_is_wellformed h \ h \ get_parent child_node \\<^sub>r Some parent \ (parent, cast child_node) \ parent_child_rel h" lemma get_parent_wf_is_l_get_parent_wf [instances]: "l_get_parent_wf type_wf known_ptr known_ptrs heap_is_wellformed parent_child_rel get_child_nodes get_parent" using known_ptrs_is_l_known_ptrs apply(auto simp add: l_get_parent_wf_def l_get_parent_wf_axioms_def)[1] using child_parent_dual heap_wellformed_induct heap_wellformed_induct_rev parent_child_rel_parent by metis+ subsection \get\_disconnected\_nodes\ subsection \set\_disconnected\_nodes\ subsubsection \get\_disconnected\_nodes\ locale l_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_disconnected_nodes_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs + l_heap_is_wellformed type_wf known_ptr heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs for known_ptr :: "(_) object_ptr \ bool" and type_wf :: "(_) heap \ bool" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma remove_from_disconnected_nodes_removes: assumes "heap_is_wellformed h" assumes "h \ get_disconnected_nodes ptr \\<^sub>r disc_nodes" assumes "h \ set_disconnected_nodes ptr (remove1 node_ptr disc_nodes) \\<^sub>h h'" assumes "h' \ get_disconnected_nodes ptr \\<^sub>r disc_nodes'" shows "node_ptr \ set disc_nodes'" using assms by (metis distinct_remove1_removeAll heap_is_wellformed_disconnected_nodes_distinct set_disconnected_nodes_get_disconnected_nodes member_remove remove_code(1) returns_result_eq) end locale l_set_disconnected_nodes_get_disconnected_nodes_wf = l_heap_is_wellformed + l_set_disconnected_nodes_get_disconnected_nodes + assumes remove_from_disconnected_nodes_removes: "heap_is_wellformed h \ h \ get_disconnected_nodes ptr \\<^sub>r disc_nodes \ h \ set_disconnected_nodes ptr (remove1 node_ptr disc_nodes) \\<^sub>h h' \ h' \ get_disconnected_nodes ptr \\<^sub>r disc_nodes' \ node_ptr \ set disc_nodes'" interpretation i_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M?: l_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs heap_is_wellformed parent_child_rel get_child_nodes using instances by (simp add: l_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_disconnected_nodes_get_disconnected_nodes_wf_is_l_set_disconnected_nodes_get_disconnected_nodes_wf [instances]: "l_set_disconnected_nodes_get_disconnected_nodes_wf type_wf known_ptr heap_is_wellformed parent_child_rel get_child_nodes get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs" apply(auto simp add: l_set_disconnected_nodes_get_disconnected_nodes_wf_def l_set_disconnected_nodes_get_disconnected_nodes_wf_axioms_def instances)[1] using remove_from_disconnected_nodes_removes apply fast done subsection \get\_root\_node\ locale l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed type_wf known_ptr heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs + l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs get_parent get_parent_locs + l_get_parent_wf type_wf known_ptr known_ptrs heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs get_parent get_parent_locs + l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_ancestors get_ancestors_locs get_root_node get_root_node_locs for known_ptr :: "(_::linorder) object_ptr \ bool" and type_wf :: "(_) heap \ bool" and known_ptrs :: "(_) heap \ bool" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and get_parent :: "(_) node_ptr \ ((_) heap, exception, (_) object_ptr option) prog" and get_parent_locs :: "((_) heap \ (_) heap \ bool) set" and get_ancestors :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr list) prog" and get_ancestors_locs :: "((_) heap \ (_) heap \ bool) set" and get_root_node :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr) prog" and get_root_node_locs :: "((_) heap \ (_) heap \ bool) set" begin lemma get_ancestors_reads: assumes "heap_is_wellformed h" shows "reads get_ancestors_locs (get_ancestors node_ptr) h h'" proof (insert assms(1), induct rule: heap_wellformed_induct_rev) case (step child) then show ?case using [[simproc del: Product_Type.unit_eq]] get_parent_reads[unfolded reads_def] apply(simp (no_asm) add: get_ancestors_def) by(auto simp add: get_ancestors_locs_def reads_subset[OF return_reads] get_parent_reads_pointers intro!: reads_bind_pure reads_subset[OF check_in_heap_reads] reads_subset[OF get_parent_reads] reads_subset[OF get_child_nodes_reads] split: option.splits) qed lemma get_ancestors_ok: assumes "heap_is_wellformed h" and "ptr |\| object_ptr_kinds h" and "known_ptrs h" and type_wf: "type_wf h" shows "h \ ok (get_ancestors ptr)" proof (insert assms(1) assms(2), induct rule: heap_wellformed_induct_rev) case (step child) then show ?case using assms(3) assms(4) apply(simp (no_asm) add: get_ancestors_def) apply(simp add: assms(1) get_parent_parent_in_heap) by(auto intro!: bind_is_OK_pure_I bind_pure_I get_parent_ok split: option.splits) qed lemma get_root_node_ptr_in_heap: assumes "h \ ok (get_root_node ptr)" shows "ptr |\| object_ptr_kinds h" using assms unfolding get_root_node_def using get_ancestors_ptr_in_heap by auto lemma get_root_node_ok: assumes "heap_is_wellformed h" "known_ptrs h" "type_wf h" and "ptr |\| object_ptr_kinds h" shows "h \ ok (get_root_node ptr)" unfolding get_root_node_def using assms get_ancestors_ok by auto lemma get_ancestors_parent: assumes "heap_is_wellformed h" and "h \ get_parent child \\<^sub>r Some parent" shows "h \ get_ancestors (cast child) \\<^sub>r (cast child) # parent # ancestors \ h \ get_ancestors parent \\<^sub>r parent # ancestors" proof assume a1: "h \ get_ancestors (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child # parent # ancestors" then have "h \ Heap_Error_Monad.bind (check_in_heap (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child)) (\_. Heap_Error_Monad.bind (get_parent child) (\x. Heap_Error_Monad.bind (case x of None \ return [] | Some x \ get_ancestors x) (\ancestors. return (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child # ancestors)))) \\<^sub>r cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child # parent # ancestors" by(simp add: get_ancestors_def) then show "h \ get_ancestors parent \\<^sub>r parent # ancestors" using assms(2) apply(auto elim!: bind_returns_result_E2 split: option.splits)[1] using returns_result_eq by fastforce next assume "h \ get_ancestors parent \\<^sub>r parent # ancestors" then show "h \ get_ancestors (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child # parent # ancestors" using assms(2) apply(simp (no_asm) add: get_ancestors_def) apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1] by (metis (full_types) assms(2) check_in_heap_ptr_in_heap is_OK_returns_result_I local.get_parent_ptr_in_heap node_ptr_kinds_commutes old.unit.exhaust select_result_I) qed lemma get_ancestors_never_empty: assumes "heap_is_wellformed h" and "h \ get_ancestors child \\<^sub>r ancestors" shows "ancestors \ []" proof(insert assms(2), induct arbitrary: ancestors rule: heap_wellformed_induct_rev[OF assms(1)]) case (1 child) then show ?case proof (induct "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child") case None then show ?case apply(simp add: get_ancestors_def) by(auto elim!: bind_returns_result_E2 split: option.splits) next case (Some child_node) then obtain parent_opt where parent_opt: "h \ get_parent child_node \\<^sub>r parent_opt" apply(simp add: get_ancestors_def) by(auto elim!: bind_returns_result_E2 split: option.splits) with Some show ?case proof(induct parent_opt) case None then show ?case apply(simp add: get_ancestors_def) by(auto elim!: bind_returns_result_E2 split: option.splits) next case (Some option) then show ?case apply(simp add: get_ancestors_def) by(auto elim!: bind_returns_result_E2 split: option.splits) qed qed qed lemma get_ancestors_subset: assumes "heap_is_wellformed h" and "h \ get_ancestors ptr \\<^sub>r ancestors" and "ancestor \ set ancestors" and "h \ get_ancestors ancestor \\<^sub>r ancestor_ancestors" and type_wf: "type_wf h" and known_ptrs: "known_ptrs h" shows "set ancestor_ancestors \ set ancestors" proof (insert assms(1) assms(2) assms(3), induct ptr arbitrary: ancestors rule: heap_wellformed_induct_rev) case (step child) have "child |\| object_ptr_kinds h" using get_ancestors_ptr_in_heap step(2) by auto (* then have "h \ check_in_heap child \\<^sub>r ()" using returns_result_select_result by force *) show ?case proof (induct "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child") case None then have "ancestors = [child]" using step(2) step(3) by(auto simp add: get_ancestors_def elim!: bind_returns_result_E2) show ?case using step(2) step(3) apply(auto simp add: \ancestors = [child]\)[1] using assms(4) returns_result_eq by fastforce next case (Some child_node) note s1 = Some obtain parent_opt where parent_opt: "h \ get_parent child_node \\<^sub>r parent_opt" using \child |\| object_ptr_kinds h\ assms(1) Some[symmetric] get_parent_ok[OF type_wf known_ptrs] by (metis (no_types, lifting) is_OK_returns_result_E known_ptrs get_parent_ok l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms node_ptr_casts_commute node_ptr_kinds_commutes) then show ?case proof (induct parent_opt) case None then have "ancestors = [child]" using step(2) step(3) s1 apply(simp add: get_ancestors_def) by(auto elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq) show ?case using step(2) step(3) apply(auto simp add: \ancestors = [child]\)[1] using assms(4) returns_result_eq by fastforce next case (Some parent) have "h \ Heap_Error_Monad.bind (check_in_heap child) (\_. Heap_Error_Monad.bind (case cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child of None \ return [] | Some node_ptr \ Heap_Error_Monad.bind (get_parent node_ptr) (\parent_ptr_opt. case parent_ptr_opt of None \ return [] | Some x \ get_ancestors x)) (\ancestors. return (child # ancestors))) \\<^sub>r ancestors" using step(2) by(simp add: get_ancestors_def) moreover obtain tl_ancestors where tl_ancestors: "ancestors = child # tl_ancestors" using calculation by(auto elim!: bind_returns_result_E2 split: option.splits) ultimately have "h \ get_ancestors parent \\<^sub>r tl_ancestors" using s1 Some by(auto elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq) show ?case using step(1)[OF s1[symmetric, simplified] Some \h \ get_ancestors parent \\<^sub>r tl_ancestors\] step(3) apply(auto simp add: tl_ancestors)[1] by (metis assms(4) insert_iff list.simps(15) local.step(2) returns_result_eq tl_ancestors) qed qed qed lemma get_ancestors_also_parent: assumes "heap_is_wellformed h" and "h \ get_ancestors some_ptr \\<^sub>r ancestors" and "cast child \ set ancestors" and "h \ get_parent child \\<^sub>r Some parent" and type_wf: "type_wf h" and known_ptrs: "known_ptrs h" shows "parent \ set ancestors" proof - obtain child_ancestors where child_ancestors: "h \ get_ancestors (cast child) \\<^sub>r child_ancestors" by (meson assms(1) assms(4) get_ancestors_ok is_OK_returns_result_I known_ptrs local.get_parent_ptr_in_heap node_ptr_kinds_commutes returns_result_select_result type_wf) then have "parent \ set child_ancestors" apply(simp add: get_ancestors_def) by(auto elim!: bind_returns_result_E2 split: option.splits dest!: returns_result_eq[OF assms(4)] get_ancestors_ptr) then show ?thesis using assms child_ancestors get_ancestors_subset by blast qed lemma get_ancestors_obtains_children: assumes "heap_is_wellformed h" and "ancestor \ ptr" and "ancestor \ set ancestors" and "h \ get_ancestors ptr \\<^sub>r ancestors" and type_wf: "type_wf h" and known_ptrs: "known_ptrs h" obtains children ancestor_child where "h \ get_child_nodes ancestor \\<^sub>r children" and "ancestor_child \ set children" and "cast ancestor_child \ set ancestors" proof - assume 0: "(\children ancestor_child. h \ get_child_nodes ancestor \\<^sub>r children \ ancestor_child \ set children \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ancestor_child \ set ancestors \ thesis)" have "\child. h \ get_parent child \\<^sub>r Some ancestor \ cast child \ set ancestors" proof (insert assms(1) assms(2) assms(3) assms(4), induct ptr arbitrary: ancestors rule: heap_wellformed_induct_rev) case (step child) have "child |\| object_ptr_kinds h" using get_ancestors_ptr_in_heap step(4) by auto show ?case proof (induct "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child") case None then have "ancestors = [child]" using step(3) step(4) by(auto simp add: get_ancestors_def elim!: bind_returns_result_E2) show ?case using step(2) step(3) step(4) by(auto simp add: \ancestors = [child]\) next case (Some child_node) note s1 = Some obtain parent_opt where parent_opt: "h \ get_parent child_node \\<^sub>r parent_opt" using \child |\| object_ptr_kinds h\ assms(1) Some[symmetric] using get_parent_ok known_ptrs type_wf by (metis (no_types, lifting) is_OK_returns_result_E node_ptr_casts_commute node_ptr_kinds_commutes) then show ?case proof (induct parent_opt) case None then have "ancestors = [child]" using step(2) step(3) step(4) s1 apply(simp add: get_ancestors_def) by(auto elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq) show ?case using step(2) step(3) step(4) by(auto simp add: \ancestors = [child]\) next case (Some parent) have "h \ Heap_Error_Monad.bind (check_in_heap child) (\_. Heap_Error_Monad.bind (case cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child of None \ return [] | Some node_ptr \ Heap_Error_Monad.bind (get_parent node_ptr) (\parent_ptr_opt. case parent_ptr_opt of None \ return [] | Some x \ get_ancestors x)) (\ancestors. return (child # ancestors))) \\<^sub>r ancestors" using step(4) by(simp add: get_ancestors_def) moreover obtain tl_ancestors where tl_ancestors: "ancestors = child # tl_ancestors" using calculation by(auto elim!: bind_returns_result_E2 split: option.splits) ultimately have "h \ get_ancestors parent \\<^sub>r tl_ancestors" using s1 Some by(auto elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq) (* have "ancestor \ parent" *) have "ancestor \ set tl_ancestors" using tl_ancestors step(2) step(3) by auto show ?case proof (cases "ancestor \ parent") case True show ?thesis using step(1)[OF s1[symmetric, simplified] Some True \ancestor \ set tl_ancestors\ \h \ get_ancestors parent \\<^sub>r tl_ancestors\] using tl_ancestors by auto next case False have "child \ set ancestors" using step(4) get_ancestors_ptr by simp then show ?thesis using Some False s1[symmetric] by(auto) qed qed qed qed then obtain child where child: "h \ get_parent child \\<^sub>r Some ancestor" and in_ancestors: "cast child \ set ancestors" by auto then obtain children where children: "h \ get_child_nodes ancestor \\<^sub>r children" and child_in_children: "child \ set children" using get_parent_child_dual by blast show thesis using 0[OF children child_in_children] child assms(3) in_ancestors by blast qed lemma get_ancestors_parent_child_rel: assumes "heap_is_wellformed h" and "h \ get_ancestors child \\<^sub>r ancestors" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "(ptr, child) \ (parent_child_rel h)\<^sup>* \ ptr \ set ancestors" proof (safe) assume 3: "(ptr, child) \ (parent_child_rel h)\<^sup>*" show "ptr \ set ancestors" proof (insert 3, induct ptr rule: heap_wellformed_induct[OF assms(1)]) case (1 ptr) then show ?case proof (cases "ptr = child") case True then show ?thesis by (metis (no_types, lifting) assms(2) bind_returns_result_E get_ancestors_def in_set_member member_rec(1) return_returns_result) next case False obtain ptr_child where ptr_child: "(ptr, ptr_child) \ (parent_child_rel h) \ (ptr_child, child) \ (parent_child_rel h)\<^sup>*" using converse_rtranclE[OF 1(2)] \ptr \ child\ by metis then obtain ptr_child_node where ptr_child_ptr_child_node: "ptr_child = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node" using ptr_child node_ptr_casts_commute3 parent_child_rel_node_ptr by (metis ) then obtain children where children: "h \ get_child_nodes ptr \\<^sub>r children" and ptr_child_node: "ptr_child_node \ set children" proof - assume a1: "\children. \h \ get_child_nodes ptr \\<^sub>r children; ptr_child_node \ set children\ \ thesis" have "ptr |\| object_ptr_kinds h" using local.parent_child_rel_parent_in_heap ptr_child by blast moreover have "ptr_child_node \ set |h \ get_child_nodes ptr|\<^sub>r" by (metis calculation known_ptrs local.get_child_nodes_ok local.known_ptrs_known_ptr local.parent_child_rel_child ptr_child ptr_child_ptr_child_node returns_result_select_result type_wf) ultimately show ?thesis using a1 get_child_nodes_ok type_wf known_ptrs by (meson local.known_ptrs_known_ptr returns_result_select_result) qed moreover have "(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node, child) \ (parent_child_rel h)\<^sup>*" using ptr_child ptr_child_ptr_child_node by auto ultimately have "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node \ set ancestors" using 1 by auto moreover have "h \ get_parent ptr_child_node \\<^sub>r Some ptr" using assms(1) children ptr_child_node child_parent_dual using known_ptrs type_wf by blast ultimately show ?thesis using get_ancestors_also_parent assms type_wf by blast qed qed next assume 3: "ptr \ set ancestors" show "(ptr, child) \ (parent_child_rel h)\<^sup>*" proof (insert 3, induct ptr rule: heap_wellformed_induct[OF assms(1)]) case (1 ptr) then show ?case proof (cases "ptr = child") case True then show ?thesis by simp next case False then obtain children ptr_child_node where children: "h \ get_child_nodes ptr \\<^sub>r children" and ptr_child_node: "ptr_child_node \ set children" and ptr_child_node_in_ancestors: "cast ptr_child_node \ set ancestors" using 1(2) assms(2) get_ancestors_obtains_children assms(1) using known_ptrs type_wf by blast then have "(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node, child) \ (parent_child_rel h)\<^sup>*" using 1(1) by blast moreover have "(ptr, cast ptr_child_node) \ parent_child_rel h" using children ptr_child_node assms(1) parent_child_rel_child_nodes2 using child_parent_dual known_ptrs parent_child_rel_parent type_wf by blast ultimately show ?thesis by auto qed qed qed lemma get_root_node_parent_child_rel: assumes "heap_is_wellformed h" and "h \ get_root_node child \\<^sub>r root" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "(root, child) \ (parent_child_rel h)\<^sup>*" using assms get_ancestors_parent_child_rel apply(auto simp add: get_root_node_def elim!: bind_returns_result_E2)[1] using get_ancestors_never_empty last_in_set by blast lemma get_ancestors_eq: assumes "heap_is_wellformed h" and "heap_is_wellformed h'" and "\object_ptr w. object_ptr \ ptr \ w \ get_child_nodes_locs object_ptr \ w h h'" and pointers_preserved: "\object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'" and known_ptrs: "known_ptrs h" and known_ptrs': "known_ptrs h'" and "h \ get_ancestors ptr \\<^sub>r ancestors" and type_wf: "type_wf h" and type_wf': "type_wf h'" shows "h' \ get_ancestors ptr \\<^sub>r ancestors" proof - have object_ptr_kinds_eq3: "object_ptr_kinds h = object_ptr_kinds h'" using pointers_preserved object_ptr_kinds_preserved_small by blast then have object_ptr_kinds_M_eq: "\ptrs. h \ object_ptr_kinds_M \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_eq: "|h \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" by(simp) have "h' \ ok (get_ancestors ptr)" using get_ancestors_ok get_ancestors_ptr_in_heap object_ptr_kinds_eq3 assms(1) known_ptrs known_ptrs' assms(2) assms(7) type_wf' by blast then obtain ancestors' where ancestors': "h' \ get_ancestors ptr \\<^sub>r ancestors'" by auto obtain root where root: "h \ get_root_node ptr \\<^sub>r root" proof - assume 0: "(\root. h \ get_root_node ptr \\<^sub>r root \ thesis)" show thesis apply(rule 0) using assms(7) by(auto simp add: get_root_node_def elim!: bind_returns_result_E2 split: option.splits) qed have children_eq: "\p children. p \ ptr \ h \ get_child_nodes p \\<^sub>r children = h' \ get_child_nodes p \\<^sub>r children" using get_child_nodes_reads assms(3) apply(simp add: reads_def reflp_def transp_def preserved_def) by blast have "acyclic (parent_child_rel h)" using assms(1) local.parent_child_rel_acyclic by auto have "acyclic (parent_child_rel h')" using assms(2) local.parent_child_rel_acyclic by blast have 2: "\c parent_opt. cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c \ set ancestors \ set ancestors' \ h \ get_parent c \\<^sub>r parent_opt = h' \ get_parent c \\<^sub>r parent_opt" proof - fix c parent_opt assume 1: " cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c \ set ancestors \ set ancestors'" obtain ptrs where ptrs: "h \ object_ptr_kinds_M \\<^sub>r ptrs" by simp let ?P = "(\ptr. Heap_Error_Monad.bind (get_child_nodes ptr) (\children. return (c \ set children)))" have children_eq_True: "\p. p \ set ptrs \ h \ ?P p \\<^sub>r True \ h' \ ?P p \\<^sub>r True" proof - fix p assume "p \ set ptrs" then show "h \ ?P p \\<^sub>r True \ h' \ ?P p \\<^sub>r True" proof (cases "p = ptr") case True have "(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c, ptr) \ (parent_child_rel h)\<^sup>*" using get_ancestors_parent_child_rel 1 assms by blast then have "(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \ (parent_child_rel h)" proof (cases "cast c = ptr") case True then show ?thesis using \acyclic (parent_child_rel h)\ by(auto simp add: acyclic_def) next case False then have "(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \ (parent_child_rel h)\<^sup>*" using \acyclic (parent_child_rel h)\ False rtrancl_eq_or_trancl rtrancl_trancl_trancl \(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c, ptr) \ (parent_child_rel h)\<^sup>*\ by (metis acyclic_def) then show ?thesis using r_into_rtrancl by auto qed obtain children where children: "h \ get_child_nodes ptr \\<^sub>r children" using type_wf by (metis \h' \ ok get_ancestors ptr\ assms(1) get_ancestors_ptr_in_heap get_child_nodes_ok heap_is_wellformed_def is_OK_returns_result_E known_ptrs local.known_ptrs_known_ptr object_ptr_kinds_eq3) then have "c \ set children" using \(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \ (parent_child_rel h)\ assms(1) using parent_child_rel_child_nodes2 using child_parent_dual known_ptrs parent_child_rel_parent type_wf by blast with children have "h \ ?P p \\<^sub>r False" by(auto simp add: True) moreover have "(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c, ptr) \ (parent_child_rel h')\<^sup>*" using get_ancestors_parent_child_rel assms(2) ancestors' 1 known_ptrs' type_wf type_wf' by blast then have "(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \ (parent_child_rel h')" proof (cases "cast c = ptr") case True then show ?thesis using \acyclic (parent_child_rel h')\ by(auto simp add: acyclic_def) next case False then have "(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \ (parent_child_rel h')\<^sup>*" using \acyclic (parent_child_rel h')\ False rtrancl_eq_or_trancl rtrancl_trancl_trancl \(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c, ptr) \ (parent_child_rel h')\<^sup>*\ by (metis acyclic_def) then show ?thesis using r_into_rtrancl by auto qed then have "(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \ (parent_child_rel h')" using r_into_rtrancl by auto obtain children' where children': "h' \ get_child_nodes ptr \\<^sub>r children'" using type_wf type_wf' by (meson \h' \ ok (get_ancestors ptr)\ assms(2) get_ancestors_ptr_in_heap get_child_nodes_ok is_OK_returns_result_E known_ptrs' local.known_ptrs_known_ptr) then have "c \ set children'" using \(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \ (parent_child_rel h')\ assms(2) type_wf type_wf' using parent_child_rel_child_nodes2 child_parent_dual known_ptrs' parent_child_rel_parent by auto with children' have "h' \ ?P p \\<^sub>r False" by(auto simp add: True) ultimately show ?thesis by (metis returns_result_eq) next case False then show ?thesis using children_eq ptrs by (metis (no_types, lifting) bind_pure_returns_result_I bind_returns_result_E get_child_nodes_pure return_returns_result) qed qed have "\pa. pa \ set ptrs \ h \ ok (get_child_nodes pa \ (\children. return (c \ set children))) = h' \ ok ( get_child_nodes pa \ (\children. return (c \ set children)))" using assms(1) assms(2) object_ptr_kinds_eq ptrs type_wf type_wf' by (metis (no_types, lifting) ObjectMonad.ptr_kinds_ptr_kinds_M bind_is_OK_pure_I get_child_nodes_ok get_child_nodes_pure known_ptrs' local.known_ptrs_known_ptr return_ok select_result_I2) have children_eq_False: "\pa. pa \ set ptrs \ h \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r False = h' \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r False" proof fix pa assume "pa \ set ptrs" and "h \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r False" have "h \ ok (get_child_nodes pa \ (\children. return (c \ set children))) \ h' \ ok ( get_child_nodes pa \ (\children. return (c \ set children)))" using \pa \ set ptrs\ \\pa. pa \ set ptrs \ h \ ok (get_child_nodes pa \ (\children. return (c \ set children))) = h' \ ok ( get_child_nodes pa \ (\children. return (c \ set children)))\ by auto moreover have "h \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r False \ h' \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r False" by (metis (mono_tags, lifting) \\pa. pa \ set ptrs \ h \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r True = h' \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r True\ \pa \ set ptrs\ calculation is_OK_returns_result_I returns_result_eq returns_result_select_result) ultimately show "h' \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r False" using \h \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r False\ by auto next fix pa assume "pa \ set ptrs" and "h' \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r False" have "h' \ ok (get_child_nodes pa \ (\children. return (c \ set children))) \ h \ ok ( get_child_nodes pa \ (\children. return (c \ set children)))" using \pa \ set ptrs\ \\pa. pa \ set ptrs \ h \ ok (get_child_nodes pa \ (\children. return (c \ set children))) = h' \ ok ( get_child_nodes pa \ (\children. return (c \ set children)))\ by auto moreover have "h' \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r False \ h \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r False" by (metis (mono_tags, lifting) \\pa. pa \ set ptrs \ h \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r True = h' \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r True\ \pa \ set ptrs\ calculation is_OK_returns_result_I returns_result_eq returns_result_select_result) ultimately show "h \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r False" using \h' \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r False\ by blast qed have filter_eq: "\xs. h \ filter_M ?P ptrs \\<^sub>r xs = h' \ filter_M ?P ptrs \\<^sub>r xs" proof (rule filter_M_eq) show "\xs x. pure (Heap_Error_Monad.bind (get_child_nodes x) (\children. return (c \ set children))) h" by(auto intro!: bind_pure_I) next show "\xs x. pure (Heap_Error_Monad.bind (get_child_nodes x) (\children. return (c \ set children))) h'" by(auto intro!: bind_pure_I) next fix xs b x assume 0: "x \ set ptrs" then show "h \ Heap_Error_Monad.bind (get_child_nodes x) (\children. return (c \ set children)) \\<^sub>r b = h' \ Heap_Error_Monad.bind (get_child_nodes x) (\children. return (c \ set children)) \\<^sub>r b" apply(induct b) using children_eq_True apply blast using children_eq_False apply blast done qed show "h \ get_parent c \\<^sub>r parent_opt = h' \ get_parent c \\<^sub>r parent_opt" apply(simp add: get_parent_def) apply(rule bind_cong_2) apply(simp) apply(simp) apply(simp add: check_in_heap_def node_ptr_kinds_def object_ptr_kinds_eq3) apply(rule bind_cong_2) apply(auto simp add: object_ptr_kinds_M_eq object_ptr_kinds_eq3)[1] apply(auto simp add: object_ptr_kinds_M_eq object_ptr_kinds_eq3)[1] apply(auto simp add: object_ptr_kinds_M_eq object_ptr_kinds_eq3)[1] apply(rule bind_cong_2) apply(auto intro!: filter_M_pure_I bind_pure_I)[1] apply(auto intro!: filter_M_pure_I bind_pure_I)[1] apply(auto simp add: filter_eq (* dest!: returns_result_eq[OF ptrs] *))[1] using filter_eq ptrs apply auto[1] using filter_eq ptrs by auto qed have "ancestors = ancestors'" proof(insert assms(1) assms(7) ancestors' 2, induct ptr arbitrary: ancestors ancestors' rule: heap_wellformed_induct_rev) case (step child) show ?case using step(2) step(3) step(4) apply(simp add: get_ancestors_def) apply(auto intro!: elim!: bind_returns_result_E2 split: option.splits)[1] using returns_result_eq apply fastforce apply (meson option.simps(3) returns_result_eq) by (metis IntD1 IntD2 option.inject returns_result_eq step.hyps) qed then show ?thesis using assms(5) ancestors' by simp qed lemma get_ancestors_remains_not_in_ancestors: assumes "heap_is_wellformed h" and "heap_is_wellformed h'" and "h \ get_ancestors ptr \\<^sub>r ancestors" and "h' \ get_ancestors ptr \\<^sub>r ancestors'" and "\p children children'. h \ get_child_nodes p \\<^sub>r children \ h' \ get_child_nodes p \\<^sub>r children' \ set children' \ set children" and "node \ set ancestors" and object_ptr_kinds_eq3: "object_ptr_kinds h = object_ptr_kinds h'" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" and type_wf': "type_wf h'" shows "node \ set ancestors'" proof - have object_ptr_kinds_M_eq: "\ptrs. h \ object_ptr_kinds_M \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" using object_ptr_kinds_eq3 by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_eq: "|h \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" by(simp) show ?thesis proof (insert assms(1) assms(3) assms(4) assms(6), induct ptr arbitrary: ancestors ancestors' rule: heap_wellformed_induct_rev) case (step child) have 1: "\p parent. h' \ get_parent p \\<^sub>r Some parent \ h \ get_parent p \\<^sub>r Some parent" proof - fix p parent assume "h' \ get_parent p \\<^sub>r Some parent" then obtain children' where children': "h' \ get_child_nodes parent \\<^sub>r children'" and p_in_children': "p \ set children'" using get_parent_child_dual by blast obtain children where children: "h \ get_child_nodes parent \\<^sub>r children" using get_child_nodes_ok assms(1) get_child_nodes_ptr_in_heap object_ptr_kinds_eq children' known_ptrs using type_wf type_wf' by (metis \h' \ get_parent p \\<^sub>r Some parent\ get_parent_parent_in_heap is_OK_returns_result_E local.known_ptrs_known_ptr object_ptr_kinds_eq3) have "p \ set children" using assms(5) children children' p_in_children' by blast then show "h \ get_parent p \\<^sub>r Some parent" using child_parent_dual assms(1) children known_ptrs type_wf by blast qed have "node \ child" using assms(1) get_ancestors_parent_child_rel step.prems(1) step.prems(3) known_ptrs using type_wf type_wf' by blast then show ?case using step(2) step(3) apply(simp add: get_ancestors_def) using step(4) apply(auto elim!: bind_returns_result_E2 split: option.splits)[1] using 1 apply (meson option.distinct(1) returns_result_eq) by (metis "1" option.inject returns_result_eq step.hyps) qed qed lemma get_ancestors_ptrs_in_heap: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_ancestors ptr \\<^sub>r ancestors" assumes "ptr' \ set ancestors" shows "ptr' |\| object_ptr_kinds h" proof (insert assms(4) assms(5), induct ancestors arbitrary: ptr) case Nil then show ?case by(auto) next case (Cons a ancestors) then obtain x where x: "h \ get_ancestors x \\<^sub>r a # ancestors" by(auto simp add: get_ancestors_def[of a] elim!: bind_returns_result_E2 split: option.splits) then have "x = a" by(auto simp add: get_ancestors_def[of x] elim!: bind_returns_result_E2 split: option.splits) then show ?case using Cons.hyps Cons.prems(2) get_ancestors_ptr_in_heap x by (metis assms(1) assms(2) assms(3) get_ancestors_obtains_children get_child_nodes_ptr_in_heap is_OK_returns_result_I) qed lemma get_ancestors_prefix: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_ancestors ptr \\<^sub>r ancestors" assumes "ptr' \ set ancestors" assumes "h \ get_ancestors ptr' \\<^sub>r ancestors'" shows "\pre. ancestors = pre @ ancestors'" proof (insert assms(1) assms(5) assms(6), induct ptr' arbitrary: ancestors' rule: heap_wellformed_induct) case (step parent) then show ?case proof (cases "parent \ ptr" ) case True then obtain children ancestor_child where "h \ get_child_nodes parent \\<^sub>r children" and "ancestor_child \ set children" and "cast ancestor_child \ set ancestors" using assms(1) assms(2) assms(3) assms(4) get_ancestors_obtains_children step.prems(1) by blast then have "h \ get_parent ancestor_child \\<^sub>r Some parent" using assms(1) assms(2) assms(3) child_parent_dual by blast then have "h \ get_ancestors (cast ancestor_child) \\<^sub>r cast ancestor_child # ancestors'" apply(simp add: get_ancestors_def) using \h \ get_ancestors parent \\<^sub>r ancestors'\ get_parent_ptr_in_heap by(auto simp add: check_in_heap_def is_OK_returns_result_I intro!: bind_pure_returns_result_I) then show ?thesis using step(1) \h \ get_child_nodes parent \\<^sub>r children\ \ancestor_child \ set children\ \cast ancestor_child \ set ancestors\ \h \ get_ancestors (cast ancestor_child) \\<^sub>r cast ancestor_child # ancestors'\ by fastforce next case False then show ?thesis by (metis append_Nil assms(4) returns_result_eq step.prems(2)) qed qed lemma get_ancestors_same_root_node: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_ancestors ptr \\<^sub>r ancestors" assumes "ptr' \ set ancestors" assumes "ptr'' \ set ancestors" shows "h \ get_root_node ptr' \\<^sub>r root_ptr \ h \ get_root_node ptr'' \\<^sub>r root_ptr" proof - have "ptr' |\| object_ptr_kinds h" by (metis assms(1) assms(2) assms(3) assms(4) assms(5) get_ancestors_obtains_children get_ancestors_ptr_in_heap get_child_nodes_ptr_in_heap is_OK_returns_result_I) then obtain ancestors' where ancestors': "h \ get_ancestors ptr' \\<^sub>r ancestors'" by (meson assms(1) assms(2) assms(3) get_ancestors_ok is_OK_returns_result_E) then have "\pre. ancestors = pre @ ancestors'" using get_ancestors_prefix assms by blast moreover have "ptr'' |\| object_ptr_kinds h" by (metis assms(1) assms(2) assms(3) assms(4) assms(6) get_ancestors_obtains_children get_ancestors_ptr_in_heap get_child_nodes_ptr_in_heap is_OK_returns_result_I) then obtain ancestors'' where ancestors'': "h \ get_ancestors ptr'' \\<^sub>r ancestors''" by (meson assms(1) assms(2) assms(3) get_ancestors_ok is_OK_returns_result_E) then have "\pre. ancestors = pre @ ancestors''" using get_ancestors_prefix assms by blast ultimately show ?thesis using ancestors' ancestors'' apply(auto simp add: get_root_node_def elim!: bind_returns_result_E2 intro!: bind_pure_returns_result_I)[1] apply (metis (no_types, lifting) assms(1) get_ancestors_never_empty last_appendR returns_result_eq) by (metis assms(1) get_ancestors_never_empty last_appendR returns_result_eq) qed lemma get_root_node_parent_same: assumes "h \ get_parent child \\<^sub>r Some ptr" shows "h \ get_root_node (cast child) \\<^sub>r root \ h \ get_root_node ptr \\<^sub>r root" proof assume 1: " h \ get_root_node (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r root" show "h \ get_root_node ptr \\<^sub>r root" using 1[unfolded get_root_node_def] assms apply(simp add: get_ancestors_def) apply(auto simp add: get_root_node_def dest: returns_result_eq elim!: bind_returns_result_E2 intro!: bind_pure_returns_result_I split: option.splits)[1] using returns_result_eq apply fastforce using get_ancestors_ptr by fastforce next assume 1: " h \ get_root_node ptr \\<^sub>r root" show "h \ get_root_node (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r root" apply(simp add: get_root_node_def) using assms 1 apply(simp add: get_ancestors_def) apply(auto simp add: get_root_node_def dest: returns_result_eq elim!: bind_returns_result_E2 intro!: bind_pure_returns_result_I split: option.splits)[1] apply (simp add: check_in_heap_def is_OK_returns_result_I) using get_ancestors_ptr get_parent_ptr_in_heap apply (simp add: is_OK_returns_result_I) by (meson list.distinct(1) list.set_cases local.get_ancestors_ptr) qed lemma get_root_node_same_no_parent: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_root_node ptr \\<^sub>r cast child" shows "h \ get_parent child \\<^sub>r None" proof (insert assms(1) assms(4), induct ptr rule: heap_wellformed_induct_rev) case (step c) then show ?case proof (cases "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r c") case None then have "c = cast child" using step(2) by(auto simp add: get_root_node_def get_ancestors_def[of c] elim!: bind_returns_result_E2) then show ?thesis using None by auto next case (Some child_node) note s = this then obtain parent_opt where parent_opt: "h \ get_parent child_node \\<^sub>r parent_opt" by (metis (no_types, lifting) assms(2) assms(3) get_root_node_ptr_in_heap is_OK_returns_result_I local.get_parent_ok node_ptr_casts_commute node_ptr_kinds_commutes returns_result_select_result step.prems) then show ?thesis proof(induct parent_opt) case None then show ?case using Some get_root_node_no_parent returns_result_eq step.prems by fastforce next case (Some parent) then show ?case using step s apply(auto simp add: get_root_node_def get_ancestors_def[of c] elim!: bind_returns_result_E2 split: option.splits list.splits)[1] using get_root_node_parent_same step.hyps step.prems by auto qed qed qed lemma get_root_node_not_node_same: assumes "ptr |\| object_ptr_kinds h" assumes "\is_node_ptr_kind ptr" shows "h \ get_root_node ptr \\<^sub>r ptr" using assms apply(simp add: get_root_node_def get_ancestors_def) by(auto simp add: get_root_node_def dest: returns_result_eq elim!: bind_returns_result_E2 intro!: bind_pure_returns_result_I split: option.splits) lemma get_root_node_root_in_heap: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_root_node ptr \\<^sub>r root" shows "root |\| object_ptr_kinds h" using assms apply(auto simp add: get_root_node_def elim!: bind_returns_result_E2)[1] by (simp add: get_ancestors_never_empty get_ancestors_ptrs_in_heap) lemma get_root_node_same_no_parent_parent_child_rel: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_root_node ptr' \\<^sub>r ptr'" shows "\(\p. (p, ptr') \ (parent_child_rel h))" by (metis (no_types, lifting) assms(1) assms(2) assms(3) assms(4) get_root_node_same_no_parent l_heap_is_wellformed.parent_child_rel_child local.child_parent_dual local.get_child_nodes_ok local.known_ptrs_known_ptr local.l_heap_is_wellformed_axioms local.parent_child_rel_node_ptr local.parent_child_rel_parent_in_heap node_ptr_casts_commute3 option.simps(3) returns_result_eq returns_result_select_result) end locale l_get_ancestors_wf = l_heap_is_wellformed_defs + l_known_ptrs + l_type_wf + l_get_ancestors_defs + l_get_child_nodes_defs + l_get_parent_defs + assumes get_ancestors_never_empty: "heap_is_wellformed h \ h \ get_ancestors child \\<^sub>r ancestors \ ancestors \ []" assumes get_ancestors_ok: "heap_is_wellformed h \ ptr |\| object_ptr_kinds h \ known_ptrs h \ type_wf h \ h \ ok (get_ancestors ptr)" assumes get_ancestors_reads: "heap_is_wellformed h \ reads get_ancestors_locs (get_ancestors node_ptr) h h'" assumes get_ancestors_ptrs_in_heap: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_ancestors ptr \\<^sub>r ancestors \ ptr' \ set ancestors \ ptr' |\| object_ptr_kinds h" assumes get_ancestors_remains_not_in_ancestors: "heap_is_wellformed h \ heap_is_wellformed h' \ h \ get_ancestors ptr \\<^sub>r ancestors \ h' \ get_ancestors ptr \\<^sub>r ancestors' \ (\p children children'. h \ get_child_nodes p \\<^sub>r children \ h' \ get_child_nodes p \\<^sub>r children' \ set children' \ set children) \ node \ set ancestors \ object_ptr_kinds h = object_ptr_kinds h' \ known_ptrs h \ type_wf h \ type_wf h' \ node \ set ancestors'" assumes get_ancestors_also_parent: "heap_is_wellformed h \ h \ get_ancestors some_ptr \\<^sub>r ancestors \ cast child_node \ set ancestors \ h \ get_parent child_node \\<^sub>r Some parent \ type_wf h \ known_ptrs h \ parent \ set ancestors" assumes get_ancestors_obtains_children: "heap_is_wellformed h \ ancestor \ ptr \ ancestor \ set ancestors \ h \ get_ancestors ptr \\<^sub>r ancestors \ type_wf h \ known_ptrs h \ (\children ancestor_child . h \ get_child_nodes ancestor \\<^sub>r children \ ancestor_child \ set children \ cast ancestor_child \ set ancestors \ thesis) \ thesis" assumes get_ancestors_parent_child_rel: "heap_is_wellformed h \ h \ get_ancestors child \\<^sub>r ancestors \ known_ptrs h \ type_wf h \ (ptr, child) \ (parent_child_rel h)\<^sup>* \ ptr \ set ancestors" locale l_get_root_node_wf = l_heap_is_wellformed_defs + l_get_root_node_defs + l_type_wf + l_known_ptrs + l_get_ancestors_defs + l_get_parent_defs + assumes get_root_node_ok: "heap_is_wellformed h \ known_ptrs h \ type_wf h \ ptr |\| object_ptr_kinds h \ h \ ok (get_root_node ptr)" assumes get_root_node_ptr_in_heap: "h \ ok (get_root_node ptr) \ ptr |\| object_ptr_kinds h" assumes get_root_node_root_in_heap: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_root_node ptr \\<^sub>r root \ root |\| object_ptr_kinds h" assumes get_ancestors_same_root_node: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_ancestors ptr \\<^sub>r ancestors \ ptr' \ set ancestors \ ptr'' \ set ancestors \ h \ get_root_node ptr' \\<^sub>r root_ptr \ h \ get_root_node ptr'' \\<^sub>r root_ptr" assumes get_root_node_same_no_parent: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_root_node ptr \\<^sub>r cast child \ h \ get_parent child \\<^sub>r None" assumes get_root_node_parent_same: "h \ get_parent child \\<^sub>r Some ptr \ h \ get_root_node (cast child) \\<^sub>r root \ h \ get_root_node ptr \\<^sub>r root" interpretation i_get_root_node_wf?: l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf known_ptrs heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_parent get_parent_locs get_ancestors get_ancestors_locs get_root_node get_root_node_locs using instances by(simp add: l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma get_ancestors_wf_is_l_get_ancestors_wf [instances]: "l_get_ancestors_wf heap_is_wellformed parent_child_rel known_ptr known_ptrs type_wf get_ancestors get_ancestors_locs get_child_nodes get_parent" using known_ptrs_is_l_known_ptrs apply(auto simp add: l_get_ancestors_wf_def l_get_ancestors_wf_axioms_def)[1] using get_ancestors_never_empty apply blast using get_ancestors_ok apply blast using get_ancestors_reads apply blast using get_ancestors_ptrs_in_heap apply blast using get_ancestors_remains_not_in_ancestors apply blast using get_ancestors_also_parent apply blast using get_ancestors_obtains_children apply blast using get_ancestors_parent_child_rel apply blast using get_ancestors_parent_child_rel apply blast done lemma get_root_node_wf_is_l_get_root_node_wf [instances]: "l_get_root_node_wf heap_is_wellformed get_root_node type_wf known_ptr known_ptrs get_ancestors get_parent" using known_ptrs_is_l_known_ptrs apply(auto simp add: l_get_root_node_wf_def l_get_root_node_wf_axioms_def)[1] using get_root_node_ok apply blast using get_root_node_ptr_in_heap apply blast using get_root_node_root_in_heap apply blast using get_ancestors_same_root_node apply(blast, blast) using get_root_node_same_no_parent apply blast using get_root_node_parent_same apply (blast, blast) done subsection \to\_tree\_order\ locale l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent + l_get_parent_wf + l_heap_is_wellformed (* l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M *) begin lemma to_tree_order_ptr_in_heap: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "h \ ok (to_tree_order ptr)" shows "ptr |\| object_ptr_kinds h" proof(insert assms(1) assms(4), induct rule: heap_wellformed_induct) case (step parent) then show ?case apply(auto simp add: to_tree_order_def[of parent] map_M_pure_I elim!: bind_is_OK_E3)[1] using get_child_nodes_ptr_in_heap by blast qed lemma to_tree_order_either_ptr_or_in_children: assumes "h \ to_tree_order ptr \\<^sub>r nodes" and "node \ set nodes" and "h \ get_child_nodes ptr \\<^sub>r children" and "node \ ptr" obtains child child_to where "child \ set children" and "h \ to_tree_order (cast child) \\<^sub>r child_to" and "node \ set child_to" proof - obtain treeorders where treeorders: "h \ map_M to_tree_order (map cast children) \\<^sub>r treeorders" using assms apply(auto simp add: to_tree_order_def elim!: bind_returns_result_E)[1] using pure_returns_heap_eq returns_result_eq by fastforce then have "node \ set (concat treeorders)" using assms[simplified to_tree_order_def] by(auto elim!: bind_returns_result_E4 dest: pure_returns_heap_eq) then obtain treeorder where "treeorder \ set treeorders" and node_in_treeorder: "node \ set treeorder" by auto then obtain child where "h \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r treeorder" and "child \ set children" using assms[simplified to_tree_order_def] treeorders by(auto elim!: map_M_pure_E2) then show ?thesis using node_in_treeorder returns_result_eq that by auto qed lemma to_tree_order_ptrs_in_heap: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ to_tree_order ptr \\<^sub>r to" assumes "ptr' \ set to" shows "ptr' |\| object_ptr_kinds h" proof(insert assms(1) assms(4) assms(5), induct ptr arbitrary: to rule: heap_wellformed_induct) case (step parent) have "parent |\| object_ptr_kinds h" using assms(1) assms(2) assms(3) step.prems(1) to_tree_order_ptr_in_heap by blast then obtain children where children: "h \ get_child_nodes parent \\<^sub>r children" by (meson assms(2) assms(3) get_child_nodes_ok is_OK_returns_result_E local.known_ptrs_known_ptr) then show ?case proof (cases "children = []") case True then have "to = [parent]" using step(2) children apply(auto simp add: to_tree_order_def[of parent] map_M_pure_I elim!: bind_returns_result_E2)[1] by (metis list.distinct(1) list.map_disc_iff list.set_cases map_M_pure_E2 returns_result_eq) then show ?thesis using \parent |\| object_ptr_kinds h\ step.prems(2) by auto next case False note f = this then show ?thesis using children step to_tree_order_either_ptr_or_in_children proof (cases "ptr' = parent") case True then show ?thesis using \parent |\| object_ptr_kinds h\ by blast next case False then show ?thesis using children step.hyps to_tree_order_either_ptr_or_in_children by (metis step.prems(1) step.prems(2)) qed qed qed lemma to_tree_order_ok: assumes wellformed: "heap_is_wellformed h" and "ptr |\| object_ptr_kinds h" and "known_ptrs h" and type_wf: "type_wf h" shows "h \ ok (to_tree_order ptr)" proof(insert assms(1) assms(2), induct rule: heap_wellformed_induct) case (step parent) then show ?case using assms(3) type_wf apply(simp add: to_tree_order_def) apply(auto simp add: heap_is_wellformed_def intro!: map_M_ok_I bind_is_OK_pure_I map_M_pure_I)[1] using get_child_nodes_ok known_ptrs_known_ptr apply blast by (simp add: local.heap_is_wellformed_children_in_heap local.to_tree_order_def wellformed) qed lemma to_tree_order_child_subset: assumes "heap_is_wellformed h" and "h \ to_tree_order ptr \\<^sub>r nodes" and "h \ get_child_nodes ptr \\<^sub>r children" and "node \ set children" and "h \ to_tree_order (cast node) \\<^sub>r nodes'" shows "set nodes' \ set nodes" proof fix x assume a1: "x \ set nodes'" moreover obtain treeorders where treeorders: "h \ map_M to_tree_order (map cast children) \\<^sub>r treeorders" using assms(2) assms(3) apply(auto simp add: to_tree_order_def elim!: bind_returns_result_E)[1] using pure_returns_heap_eq returns_result_eq by fastforce then have "nodes' \ set treeorders" using assms(4) assms(5) by(auto elim!: map_M_pure_E dest: returns_result_eq) moreover have "set (concat treeorders) \ set nodes" using treeorders assms(2) assms(3) by(auto simp add: to_tree_order_def elim!: bind_returns_result_E4 dest: pure_returns_heap_eq) ultimately show "x \ set nodes" by auto qed lemma to_tree_order_ptr_in_result: assumes "h \ to_tree_order ptr \\<^sub>r nodes" shows "ptr \ set nodes" using assms apply(simp add: to_tree_order_def) by(auto elim!: bind_returns_result_E2 intro!: map_M_pure_I bind_pure_I) lemma to_tree_order_subset: assumes "heap_is_wellformed h" and "h \ to_tree_order ptr \\<^sub>r nodes" and "node \ set nodes" and "h \ to_tree_order node \\<^sub>r nodes'" and "known_ptrs h" and type_wf: "type_wf h" shows "set nodes' \ set nodes" proof - have "\nodes. h \ to_tree_order ptr \\<^sub>r nodes \ (\node. node \ set nodes \ (\nodes'. h \ to_tree_order node \\<^sub>r nodes' \ set nodes' \ set nodes))" proof(insert assms(1), induct ptr rule: heap_wellformed_induct) case (step parent) then show ?case proof safe fix nodes node nodes' x assume 1: "(\children child. h \ get_child_nodes parent \\<^sub>r children \ child \ set children \ \nodes. h \ to_tree_order (cast child) \\<^sub>r nodes \ (\node. node \ set nodes \ (\nodes'. h \ to_tree_order node \\<^sub>r nodes' \ set nodes' \ set nodes)))" and 2: "h \ to_tree_order parent \\<^sub>r nodes" and 3: "node \ set nodes" and "h \ to_tree_order node \\<^sub>r nodes'" and "x \ set nodes'" have h1: "(\children child nodes node nodes'. h \ get_child_nodes parent \\<^sub>r children \ child \ set children \ h \ to_tree_order (cast child) \\<^sub>r nodes \ (node \ set nodes \ (h \ to_tree_order node \\<^sub>r nodes' \ set nodes' \ set nodes)))" using 1 by blast obtain children where children: "h \ get_child_nodes parent \\<^sub>r children" using 2 by(auto simp add: to_tree_order_def elim!: bind_returns_result_E) then have "set nodes' \ set nodes" proof (cases "children = []") case True then show ?thesis by (metis "2" "3" \h \ to_tree_order node \\<^sub>r nodes'\ children empty_iff list.set(1) subsetI to_tree_order_either_ptr_or_in_children) next case False then show ?thesis proof (cases "node = parent") case True then show ?thesis using "2" \h \ to_tree_order node \\<^sub>r nodes'\ returns_result_eq by fastforce next case False then obtain child nodes_of_child where "child \ set children" and "h \ to_tree_order (cast child) \\<^sub>r nodes_of_child" and "node \ set nodes_of_child" using 2[simplified to_tree_order_def] 3 to_tree_order_either_ptr_or_in_children[where node=node and ptr=parent] children apply(auto elim!: bind_returns_result_E2 intro: map_M_pure_I)[1] using is_OK_returns_result_E 2 a_all_ptrs_in_heap_def assms(1) heap_is_wellformed_def using "3" by blast then have "set nodes' \ set nodes_of_child" using h1 using \h \ to_tree_order node \\<^sub>r nodes'\ children by blast moreover have "set nodes_of_child \ set nodes" using "2" \child \ set children\ \h \ to_tree_order (cast child) \\<^sub>r nodes_of_child\ assms children to_tree_order_child_subset by auto ultimately show ?thesis by blast qed qed then show "x \ set nodes" using \x \ set nodes'\ by blast qed qed then show ?thesis using assms by blast qed lemma to_tree_order_parent: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ to_tree_order ptr \\<^sub>r nodes" assumes "h \ get_parent child \\<^sub>r Some parent" assumes "parent \ set nodes" shows "cast child \ set nodes" proof - obtain nodes' where nodes': "h \ to_tree_order parent \\<^sub>r nodes'" using assms to_tree_order_ok get_parent_parent_in_heap by (meson get_parent_parent_in_heap is_OK_returns_result_E) then have "set nodes' \ set nodes" using to_tree_order_subset assms by blast moreover obtain children where children: "h \ get_child_nodes parent \\<^sub>r children" and child: "child \ set children" using assms get_parent_child_dual by blast then obtain child_to where child_to: "h \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r child_to" by (meson assms(1) assms(2) assms(3) assms(5) is_OK_returns_result_E is_OK_returns_result_I get_parent_ptr_in_heap node_ptr_kinds_commutes to_tree_order_ok) then have "cast child \ set child_to" apply(simp add: to_tree_order_def) by(auto elim!: bind_returns_result_E2 map_M_pure_E dest!: bind_returns_result_E3[rotated, OF children, rotated] intro!: map_M_pure_I) have "cast child \ set nodes'" using nodes' child apply(simp add: to_tree_order_def) apply(auto elim!: bind_returns_result_E2 map_M_pure_E dest!: bind_returns_result_E3[rotated, OF children, rotated] intro!: map_M_pure_I)[1] using child_to \cast child \ set child_to\ returns_result_eq by fastforce ultimately show ?thesis by auto qed lemma to_tree_order_child: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ to_tree_order ptr \\<^sub>r nodes" assumes "h \ get_child_nodes parent \\<^sub>r children" assumes "cast child \ ptr" assumes "child \ set children" assumes "cast child \ set nodes" shows "parent \ set nodes" proof(insert assms(1) assms(4) assms(6) assms(8), induct ptr arbitrary: nodes rule: heap_wellformed_induct) case (step p) have "p |\| object_ptr_kinds h" using \h \ to_tree_order p \\<^sub>r nodes\ to_tree_order_ptr_in_heap using assms(1) assms(2) assms(3) by blast then obtain children where children: "h \ get_child_nodes p \\<^sub>r children" by (meson assms(2) assms(3) get_child_nodes_ok is_OK_returns_result_E local.known_ptrs_known_ptr) then show ?case proof (cases "children = []") case True then show ?thesis using step(2) step(3) step(4) children by(auto simp add: to_tree_order_def[of p] map_M_pure_I elim!: bind_returns_result_E2 dest!: bind_returns_result_E3[rotated, OF children, rotated]) next case False then obtain c child_to where child: "c \ set children" and child_to: "h \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \\<^sub>r child_to" and "cast child \ set child_to" using step(2) children apply(auto simp add: to_tree_order_def[of p] map_M_pure_I elim!: bind_returns_result_E2 dest!: bind_returns_result_E3[rotated, OF children, rotated])[1] by (metis (full_types) assms(1) assms(2) assms(3) get_parent_ptr_in_heap is_OK_returns_result_I l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.child_parent_dual l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms node_ptr_kinds_commutes returns_result_select_result step.prems(1) step.prems(2) step.prems(3) to_tree_order_either_ptr_or_in_children to_tree_order_ok) then have "set child_to \ set nodes" using assms(1) child children step.prems(1) to_tree_order_child_subset by auto show ?thesis proof (cases "c = child") case True then have "parent = p" using step(3) children child assms(5) assms(7) by (meson assms(1) assms(2) assms(3) child_parent_dual option.inject returns_result_eq) then show ?thesis using step.prems(1) to_tree_order_ptr_in_result by blast next case False then show ?thesis using step(1)[OF children child child_to] step(3) step(4) using \set child_to \ set nodes\ using \cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child \ set child_to\ by auto qed qed qed lemma to_tree_order_node_ptrs: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ to_tree_order ptr \\<^sub>r nodes" assumes "ptr' \ ptr" assumes "ptr' \ set nodes" shows "is_node_ptr_kind ptr'" proof(insert assms(1) assms(4) assms(5) assms(6), induct ptr arbitrary: nodes rule: heap_wellformed_induct) case (step p) have "p |\| object_ptr_kinds h" using \h \ to_tree_order p \\<^sub>r nodes\ to_tree_order_ptr_in_heap using assms(1) assms(2) assms(3) by blast then obtain children where children: "h \ get_child_nodes p \\<^sub>r children" by (meson assms(2) assms(3) get_child_nodes_ok is_OK_returns_result_E local.known_ptrs_known_ptr) then show ?case proof (cases "children = []") case True then show ?thesis using step(2) step(3) step(4) children by(auto simp add: to_tree_order_def[of p] map_M_pure_I elim!: bind_returns_result_E2 dest!: bind_returns_result_E3[rotated, OF children, rotated])[1] next case False then obtain c child_to where child: "c \ set children" and child_to: "h \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \\<^sub>r child_to" and "ptr' \ set child_to" using step(2) children apply(auto simp add: to_tree_order_def[of p] map_M_pure_I elim!: bind_returns_result_E2 dest!: bind_returns_result_E3[rotated, OF children, rotated])[1] using step.prems(1) step.prems(2) step.prems(3) to_tree_order_either_ptr_or_in_children by blast then have "set child_to \ set nodes" using assms(1) child children step.prems(1) to_tree_order_child_subset by auto show ?thesis proof (cases "cast c = ptr") case True then show ?thesis using step \ptr' \ set child_to\ assms(5) child child_to children by blast next case False then show ?thesis using \ptr' \ set child_to\ child child_to children is_node_ptr_kind_cast step.hyps by blast qed qed qed lemma to_tree_order_child2: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ to_tree_order ptr \\<^sub>r nodes" assumes "cast child \ ptr" assumes "cast child \ set nodes" obtains parent where "h \ get_parent child \\<^sub>r Some parent" and "parent \ set nodes" proof - assume 1: "(\parent. h \ get_parent child \\<^sub>r Some parent \ parent \ set nodes \ thesis)" show thesis proof(insert assms(1) assms(4) assms(5) assms(6) 1, induct ptr arbitrary: nodes rule: heap_wellformed_induct) case (step p) have "p |\| object_ptr_kinds h" using \h \ to_tree_order p \\<^sub>r nodes\ to_tree_order_ptr_in_heap using assms(1) assms(2) assms(3) by blast then obtain children where children: "h \ get_child_nodes p \\<^sub>r children" by (meson assms(2) assms(3) get_child_nodes_ok is_OK_returns_result_E local.known_ptrs_known_ptr) then show ?case proof (cases "children = []") case True then show ?thesis using step(2) step(3) step(4) children by(auto simp add: to_tree_order_def[of p] map_M_pure_I elim!: bind_returns_result_E2 dest!: bind_returns_result_E3[rotated, OF children, rotated]) next case False then obtain c child_to where child: "c \ set children" and child_to: "h \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \\<^sub>r child_to" and "cast child \ set child_to" using step(2) children apply(auto simp add: to_tree_order_def[of p] map_M_pure_I elim!: bind_returns_result_E2 dest!: bind_returns_result_E3[rotated, OF children, rotated])[1] using step.prems(1) step.prems(2) step.prems(3) to_tree_order_either_ptr_or_in_children by blast then have "set child_to \ set nodes" using assms(1) child children step.prems(1) to_tree_order_child_subset by auto have "cast child |\| object_ptr_kinds h" using assms(1) assms(2) assms(3) assms(4) assms(6) to_tree_order_ptrs_in_heap by blast then obtain parent_opt where parent_opt: "h \ get_parent child \\<^sub>r parent_opt" by (meson assms(2) assms(3) is_OK_returns_result_E get_parent_ok node_ptr_kinds_commutes) then show ?thesis proof (induct parent_opt) case None then show ?case by (metis \cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child \ set child_to\ assms(1) assms(2) assms(3) cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_inject child child_parent_dual child_to children option.distinct(1) returns_result_eq step.hyps) next case (Some option) then show ?case by (meson assms(1) assms(2) assms(3) get_parent_child_dual step.prems(1) step.prems(2) step.prems(3) step.prems(4) to_tree_order_child) qed qed qed qed lemma to_tree_order_parent_child_rel: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ to_tree_order ptr \\<^sub>r to" shows "(ptr, child) \ (parent_child_rel h)\<^sup>* \ child \ set to" proof assume 3: "(ptr, child) \ (parent_child_rel h)\<^sup>*" show "child \ set to" proof (insert 3, induct child rule: heap_wellformed_induct_rev[OF assms(1)]) case (1 child) then show ?case proof (cases "ptr = child") case True then show ?thesis using assms(4) apply(simp add: to_tree_order_def) by(auto simp add: map_M_pure_I elim!: bind_returns_result_E2) next case False obtain child_parent where "(ptr, child_parent) \ (parent_child_rel h)\<^sup>*" and "(child_parent, child) \ (parent_child_rel h)" using \ptr \ child\ by (metis "1.prems" rtranclE) obtain child_node where child_node: "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child_node = child" using \(child_parent, child) \ parent_child_rel h\ node_ptr_casts_commute3 parent_child_rel_node_ptr by blast then have "h \ get_parent child_node \\<^sub>r Some child_parent" using \(child_parent, child) \ (parent_child_rel h)\ by (meson assms(1) assms(2) assms(3) is_OK_returns_result_E l_get_parent_wf.child_parent_dual l_heap_is_wellformed.parent_child_rel_child local.get_child_nodes_ok local.known_ptrs_known_ptr local.l_get_parent_wf_axioms local.l_heap_is_wellformed_axioms local.parent_child_rel_parent_in_heap) then show ?thesis using 1(1) child_node \(ptr, child_parent) \ (parent_child_rel h)\<^sup>*\ using assms(1) assms(2) assms(3) assms(4) to_tree_order_parent by blast qed qed next assume "child \ set to" then show "(ptr, child) \ (parent_child_rel h)\<^sup>*" proof (induct child rule: heap_wellformed_induct_rev[OF assms(1)]) case (1 child) then show ?case proof (cases "ptr = child") case True then show ?thesis by simp next case False then have "\parent. (parent, child) \ (parent_child_rel h)" using 1(2) assms(4) to_tree_order_child2[OF assms(1) assms(2) assms(3) assms(4)] to_tree_order_node_ptrs by (metis assms(1) assms(2) assms(3) node_ptr_casts_commute3 parent_child_rel_parent) then obtain child_node where child_node: "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child_node = child" using node_ptr_casts_commute3 parent_child_rel_node_ptr by blast then obtain child_parent where child_parent: "h \ get_parent child_node \\<^sub>r Some child_parent" using \\parent. (parent, child) \ (parent_child_rel h)\ by (metis "1.prems" False assms(1) assms(2) assms(3) assms(4) to_tree_order_child2) then have "(child_parent, child) \ (parent_child_rel h)" using assms(1) child_node parent_child_rel_parent by blast moreover have "child_parent \ set to" by (metis "1.prems" False assms(1) assms(2) assms(3) assms(4) child_node child_parent get_parent_child_dual to_tree_order_child) then have "(ptr, child_parent) \ (parent_child_rel h)\<^sup>*" using 1 child_node child_parent by blast ultimately show ?thesis by auto qed qed qed end interpretation i_to_tree_order_wf?: l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs to_tree_order known_ptrs get_parent get_parent_locs heap_is_wellformed parent_child_rel get_disconnected_nodes get_disconnected_nodes_locs using instances apply(simp add: l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) done declare l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_to_tree_order_wf = l_heap_is_wellformed_defs + l_type_wf + l_known_ptrs + l_to_tree_order_defs + l_get_parent_defs + l_get_child_nodes_defs + assumes to_tree_order_ok: "heap_is_wellformed h \ ptr |\| object_ptr_kinds h \ known_ptrs h \ type_wf h \ h \ ok (to_tree_order ptr)" assumes to_tree_order_ptrs_in_heap: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ to_tree_order ptr \\<^sub>r to \ ptr' \ set to \ ptr' |\| object_ptr_kinds h" assumes to_tree_order_parent_child_rel: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ to_tree_order ptr \\<^sub>r to \ (ptr, child_ptr) \ (parent_child_rel h)\<^sup>* \ child_ptr \ set to" assumes to_tree_order_child2: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ to_tree_order ptr \\<^sub>r nodes \ cast child \ ptr \ cast child \ set nodes \ (\parent. h \ get_parent child \\<^sub>r Some parent \ parent \ set nodes \ thesis) \ thesis" assumes to_tree_order_node_ptrs: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ to_tree_order ptr \\<^sub>r nodes \ ptr' \ ptr \ ptr' \ set nodes \ is_node_ptr_kind ptr'" assumes to_tree_order_child: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ to_tree_order ptr \\<^sub>r nodes \ h \ get_child_nodes parent \\<^sub>r children \ cast child \ ptr \ child \ set children \ cast child \ set nodes \ parent \ set nodes" assumes to_tree_order_ptr_in_result: "h \ to_tree_order ptr \\<^sub>r nodes \ ptr \ set nodes" assumes to_tree_order_parent: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ to_tree_order ptr \\<^sub>r nodes \ h \ get_parent child \\<^sub>r Some parent \ parent \ set nodes \ cast child \ set nodes" assumes to_tree_order_subset: "heap_is_wellformed h \ h \ to_tree_order ptr \\<^sub>r nodes \ node \ set nodes \ h \ to_tree_order node \\<^sub>r nodes' \ known_ptrs h \ type_wf h \ set nodes' \ set nodes" lemma to_tree_order_wf_is_l_to_tree_order_wf [instances]: "l_to_tree_order_wf heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent get_child_nodes" using instances apply(auto simp add: l_to_tree_order_wf_def l_to_tree_order_wf_axioms_def)[1] using to_tree_order_ok apply blast using to_tree_order_ptrs_in_heap apply blast using to_tree_order_parent_child_rel apply(blast, blast) using to_tree_order_child2 apply blast using to_tree_order_node_ptrs apply blast using to_tree_order_child apply blast using to_tree_order_ptr_in_result apply blast using to_tree_order_parent apply blast using to_tree_order_subset apply blast done subsubsection \get\_root\_node\ locale l_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_to_tree_order_wf begin lemma to_tree_order_get_root_node: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ to_tree_order ptr \\<^sub>r to" assumes "ptr' \ set to" assumes "h \ get_root_node ptr' \\<^sub>r root_ptr" assumes "ptr'' \ set to" shows "h \ get_root_node ptr'' \\<^sub>r root_ptr" proof - obtain ancestors' where ancestors': "h \ get_ancestors ptr' \\<^sub>r ancestors'" by (meson assms(1) assms(2) assms(3) assms(4) assms(5) get_ancestors_ok is_OK_returns_result_E to_tree_order_ptrs_in_heap ) moreover have "ptr \ set ancestors'" using \h \ get_ancestors ptr' \\<^sub>r ancestors'\ using assms(1) assms(2) assms(3) assms(4) assms(5) get_ancestors_parent_child_rel to_tree_order_parent_child_rel by blast ultimately have "h \ get_root_node ptr \\<^sub>r root_ptr" using \h \ get_root_node ptr' \\<^sub>r root_ptr\ using assms(1) assms(2) assms(3) get_ancestors_ptr get_ancestors_same_root_node by blast obtain ancestors'' where ancestors'': "h \ get_ancestors ptr'' \\<^sub>r ancestors''" by (meson assms(1) assms(2) assms(3) assms(4) assms(7) get_ancestors_ok is_OK_returns_result_E to_tree_order_ptrs_in_heap) moreover have "ptr \ set ancestors''" using \h \ get_ancestors ptr'' \\<^sub>r ancestors''\ using assms(1) assms(2) assms(3) assms(4) assms(7) get_ancestors_parent_child_rel to_tree_order_parent_child_rel by blast ultimately show ?thesis using \h \ get_root_node ptr \\<^sub>r root_ptr\ assms(1) assms(2) assms(3) get_ancestors_ptr get_ancestors_same_root_node by blast qed lemma to_tree_order_same_root: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_root_node ptr \\<^sub>r root_ptr" assumes "h \ to_tree_order root_ptr \\<^sub>r to" assumes "ptr' \ set to" shows "h \ get_root_node ptr' \\<^sub>r root_ptr" proof (insert assms(1)(* assms(4) assms(5) *) assms(6), induct ptr' rule: heap_wellformed_induct_rev) case (step child) then show ?case proof (cases "h \ get_root_node child \\<^sub>r child") case True then have "child = root_ptr" using assms(1) assms(2) assms(3) assms(5) step.prems by (metis (no_types, lifting) get_root_node_same_no_parent node_ptr_casts_commute3 option.simps(3) returns_result_eq to_tree_order_child2 to_tree_order_node_ptrs) then show ?thesis using True by blast next case False then obtain child_node parent where "cast child_node = child" and "h \ get_parent child_node \\<^sub>r Some parent" by (metis assms(1) assms(2) assms(3) assms(4) assms(5) local.get_root_node_no_parent local.get_root_node_not_node_same local.get_root_node_same_no_parent local.to_tree_order_child2 local.to_tree_order_ptrs_in_heap node_ptr_casts_commute3 step.prems) then show ?thesis proof (cases "child = root_ptr") case True then have "h \ get_root_node root_ptr \\<^sub>r root_ptr" using assms(4) using \cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child_node = child\ assms(1) assms(2) assms(3) get_root_node_no_parent get_root_node_same_no_parent by blast then show ?thesis using step assms(4) using True by blast next case False then have "parent \ set to" using assms(5) step(2) to_tree_order_child \h \ get_parent child_node \\<^sub>r Some parent\ \cast child_node = child\ by (metis False assms(1) assms(2) assms(3) get_parent_child_dual) then show ?thesis using \cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child_node = child\ \h \ get_parent child_node \\<^sub>r Some parent\ get_root_node_parent_same using step.hyps by blast qed qed qed end interpretation i_to_tree_order_wf_get_root_node_wf?: l_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf known_ptrs heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_parent get_parent_locs get_ancestors get_ancestors_locs get_root_node get_root_node_locs to_tree_order using instances by(simp add: l_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) locale l_to_tree_order_wf_get_root_node_wf = l_type_wf + l_known_ptrs + l_to_tree_order_defs + l_get_root_node_defs + l_heap_is_wellformed_defs + assumes to_tree_order_get_root_node: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ to_tree_order ptr \\<^sub>r to \ ptr' \ set to \ h \ get_root_node ptr' \\<^sub>r root_ptr \ ptr'' \ set to \ h \ get_root_node ptr'' \\<^sub>r root_ptr" assumes to_tree_order_same_root: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_root_node ptr \\<^sub>r root_ptr \ h \ to_tree_order root_ptr \\<^sub>r to \ ptr' \ set to \ h \ get_root_node ptr' \\<^sub>r root_ptr" lemma to_tree_order_wf_get_root_node_wf_is_l_to_tree_order_wf_get_root_node_wf [instances]: "l_to_tree_order_wf_get_root_node_wf type_wf known_ptr known_ptrs to_tree_order get_root_node heap_is_wellformed" using instances apply(auto simp add: l_to_tree_order_wf_get_root_node_wf_def l_to_tree_order_wf_get_root_node_wf_axioms_def)[1] using to_tree_order_get_root_node apply blast using to_tree_order_same_root apply blast done subsection \get\_owner\_document\ locale l_get_owner_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_known_ptrs + l_heap_is_wellformed + l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_ancestors + l_get_ancestors_wf + l_get_parent + l_get_parent_wf + l_get_root_node_wf + l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_owner_document_disconnected_nodes: assumes "heap_is_wellformed h" assumes "h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes" assumes "node_ptr \ set disc_nodes" assumes known_ptrs: "known_ptrs h" assumes type_wf: "type_wf h" shows "h \ get_owner_document (cast node_ptr) \\<^sub>r document_ptr" proof - have 2: "node_ptr |\| node_ptr_kinds h" using assms heap_is_wellformed_disc_nodes_in_heap by blast have 3: "document_ptr |\| document_ptr_kinds h" using assms(2) get_disconnected_nodes_ptr_in_heap by blast have 0: "\!document_ptr\set |h \ document_ptr_kinds_M|\<^sub>r. node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r" by (metis (no_types, lifting) "3" DocumentMonad.ptr_kinds_ptr_kinds_M assms(1) assms(2) assms(3) disjoint_iff_not_equal l_heap_is_wellformed.heap_is_wellformed_one_disc_parent local.get_disconnected_nodes_ok local.l_heap_is_wellformed_axioms returns_result_select_result select_result_I2 type_wf) have "h \ get_parent node_ptr \\<^sub>r None" using heap_is_wellformed_children_disc_nodes_different child_parent_dual assms using "2" disjoint_iff_not_equal local.get_parent_child_dual local.get_parent_ok returns_result_select_result split_option_ex by (metis (no_types, lifting)) then have 4: "h \ get_root_node (cast node_ptr) \\<^sub>r cast node_ptr" using 2 get_root_node_no_parent by blast obtain document_ptrs where document_ptrs: "h \ document_ptr_kinds_M \\<^sub>r document_ptrs" by simp then have "h \ ok (filter_M (\document_ptr. do { disconnected_nodes \ get_disconnected_nodes document_ptr; return (((cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)) \ cast ` set disconnected_nodes) }) document_ptrs)" using assms(1) get_disconnected_nodes_ok type_wf unfolding heap_is_wellformed_def by(auto intro!: bind_is_OK_I2 filter_M_is_OK_I bind_pure_I) then obtain candidates where candidates: "h \ filter_M (\document_ptr. do { disconnected_nodes \ get_disconnected_nodes document_ptr; return (((cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)) \ cast ` set disconnected_nodes) }) document_ptrs \\<^sub>r candidates" by auto have eq: "\document_ptr. document_ptr |\| document_ptr_kinds h \ node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r \ |h \ do { disconnected_nodes \ get_disconnected_nodes document_ptr; return (((cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)) \ cast ` set disconnected_nodes) }|\<^sub>r" apply(auto dest!: get_disconnected_nodes_ok[OF type_wf] intro!: select_result_I[where P=id, simplified] elim!: bind_returns_result_E2)[1] apply(drule select_result_E[where P=id, simplified]) by(auto elim!: bind_returns_result_E2) have filter: "filter (\document_ptr. |h \ do { disconnected_nodes \ get_disconnected_nodes document_ptr; return (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr \ cast ` set disconnected_nodes) }|\<^sub>r) document_ptrs = [document_ptr]" apply(rule filter_ex1) using 0 document_ptrs apply(simp)[1] using eq using local.get_disconnected_nodes_ok apply auto[1] using assms(2) assms(3) apply(auto intro!: intro!: select_result_I[where P=id, simplified] elim!: bind_returns_result_E2)[1] using returns_result_eq apply fastforce using document_ptrs 3 apply(simp) using document_ptrs by simp have "h \ filter_M (\document_ptr. do { disconnected_nodes \ get_disconnected_nodes document_ptr; return (((cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)) \ cast ` set disconnected_nodes) }) document_ptrs \\<^sub>r [document_ptr]" apply(rule filter_M_filter2) using get_disconnected_nodes_ok document_ptrs 3 assms(1) type_wf filter unfolding heap_is_wellformed_def by(auto intro: bind_pure_I bind_is_OK_I2) with 4 document_ptrs have "h \ a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r document_ptr" by(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I split: option.splits)[1] moreover have "known_ptr (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)" using "4" assms(1) known_ptrs type_wf known_ptrs_known_ptr "2" node_ptr_kinds_commutes by blast ultimately show ?thesis using 2 apply(auto simp add: known_ptr_impl get_owner_document_def a_get_owner_document_tups_def)[1] apply(split invoke_splits, (rule conjI | rule impI)+)+ apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl]) apply(drule(1) known_ptr_not_character_data_ptr) apply(drule(1) known_ptr_not_element_ptr) apply(simp add: NodeClass.known_ptr_defs) by(auto split: option.splits intro!: bind_pure_returns_result_I) qed lemma in_disconnected_nodes_no_parent: assumes "heap_is_wellformed h" and "h \ get_parent node_ptr \\<^sub>r None" and "h \ get_owner_document (cast node_ptr) \\<^sub>r owner_document" and "h \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "node_ptr \ set disc_nodes" proof - have 2: "cast node_ptr |\| object_ptr_kinds h" using assms(3) get_owner_document_ptr_in_heap by fast then have 3: "h \ get_root_node (cast node_ptr) \\<^sub>r cast node_ptr" using assms(2) local.get_root_node_no_parent by blast have "\(\parent_ptr. parent_ptr |\| object_ptr_kinds h \ node_ptr \ set |h \ get_child_nodes parent_ptr|\<^sub>r)" apply(auto)[1] using assms(2) child_parent_dual[OF assms(1)] type_wf assms(1) assms(5) get_child_nodes_ok known_ptrs_known_ptr option.simps(3) returns_result_eq returns_result_select_result by (metis (no_types, hide_lams)) moreover have "node_ptr |\| node_ptr_kinds h" using assms(2) get_parent_ptr_in_heap by blast ultimately have 0: "\document_ptr\set |h \ document_ptr_kinds_M|\<^sub>r. node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r" by (metis DocumentMonad.ptr_kinds_ptr_kinds_M assms(1) finite_set_in heap_is_wellformed_children_disc_nodes) then obtain document_ptr where document_ptr: "document_ptr\set |h \ document_ptr_kinds_M|\<^sub>r" and node_ptr_in_disc_nodes: "node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r" by auto then show ?thesis using get_owner_document_disconnected_nodes known_ptrs type_wf assms using DocumentMonad.ptr_kinds_ptr_kinds_M assms(1) assms(3) assms(4) get_disconnected_nodes_ok returns_result_select_result select_result_I2 by (metis (no_types, hide_lams) ) qed lemma get_owner_document_owner_document_in_heap: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_owner_document ptr \\<^sub>r owner_document" shows "owner_document |\| document_ptr_kinds h" using assms apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1] apply(split invoke_split_asm)+ proof - assume "h \ invoke [] ptr () \\<^sub>r owner_document" then show "owner_document |\| document_ptr_kinds h" by (meson invoke_empty is_OK_returns_result_I) next assume "h \ Heap_Error_Monad.bind (check_in_heap ptr) (\_. (local.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r) ptr ()) \\<^sub>r owner_document" then show "owner_document |\| document_ptr_kinds h" by(auto simp add: a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 split: if_splits) next assume 0: "heap_is_wellformed h" and 1: "type_wf h" and 2: "known_ptrs h" and 3: "\ is_element_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 4: "is_character_data_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 5: "h \ Heap_Error_Monad.bind (check_in_heap ptr) (\_. (local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r) ptr ()) \\<^sub>r owner_document" then obtain root where root: "h \ get_root_node ptr \\<^sub>r root" by(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 split: option.splits) then show ?thesis proof (cases "is_document_ptr root") case True then show ?thesis using 4 5 root apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 intro!: filter_M_pure_I bind_pure_I split: option.splits)[1] apply(drule(1) returns_result_eq) apply(auto)[1] using "0" "1" "2" document_ptr_kinds_commutes local.get_root_node_root_in_heap by blast next case False have "known_ptr root" using "0" "1" "2" local.get_root_node_root_in_heap local.known_ptrs_known_ptr root by blast have "root |\| object_ptr_kinds h" using root using "0" "1" "2" local.get_root_node_root_in_heap by blast then have "is_node_ptr_kind root" using False \known_ptr root\ apply(simp add: known_ptr_impl known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs) using is_node_ptr_kind_none by force then have "(\document_ptr \ fset (document_ptr_kinds h). root \ cast ` set |h \ get_disconnected_nodes document_ptr|\<^sub>r)" by (metis (no_types, lifting) "0" "1" "2" \root |\| object_ptr_kinds h\ local.child_parent_dual local.get_child_nodes_ok local.get_root_node_same_no_parent local.heap_is_wellformed_children_disc_nodes local.known_ptrs_known_ptr node_ptr_casts_commute3 node_ptr_inclusion node_ptr_kinds_commutes notin_fset option.distinct(1) returns_result_eq returns_result_select_result root) then obtain some_owner_document where "some_owner_document |\| document_ptr_kinds h" and "root \ cast ` set |h \ get_disconnected_nodes some_owner_document|\<^sub>r" by auto then obtain candidates where candidates: "h \ filter_M (\document_ptr. Heap_Error_Monad.bind (get_disconnected_nodes document_ptr) (\disconnected_nodes. return (root \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ` set disconnected_nodes))) (sorted_list_of_set (fset (document_ptr_kinds h))) \\<^sub>r candidates" by (metis (no_types, lifting) "1" bind_is_OK_I2 bind_pure_I filter_M_is_OK_I finite_fset is_OK_returns_result_E local.get_disconnected_nodes_ok local.get_disconnected_nodes_pure notin_fset return_ok return_pure sorted_list_of_set(1)) then have "some_owner_document \ set candidates" apply(rule filter_M_in_result_if_ok) using \some_owner_document |\| document_ptr_kinds h\ \root \ cast ` set |h \ get_disconnected_nodes some_owner_document|\<^sub>r\ apply(auto intro!: bind_pure_I bind_pure_returns_result_I)[1] apply (simp add: \some_owner_document |\| document_ptr_kinds h\) using "1" \root \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ` set |h \ get_disconnected_nodes some_owner_document|\<^sub>r\ \some_owner_document |\| document_ptr_kinds h\ local.get_disconnected_nodes_ok by auto then have "candidates \ []" by auto then have "owner_document \ set candidates" using 5 root 4 apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 intro!: filter_M_pure_I bind_pure_I split: option.splits)[1] apply (metis candidates list.set_sel(1) returns_result_eq) by (metis \is_node_ptr_kind root\ node_ptr_no_document_ptr_cast returns_result_eq) then show ?thesis using candidates by (meson bind_pure_I bind_returns_result_E2 filter_M_holds_for_result is_OK_returns_result_I local.get_disconnected_nodes_ptr_in_heap local.get_disconnected_nodes_pure return_pure) qed next assume 0: "heap_is_wellformed h" and 1: "type_wf h" and 2: "known_ptrs h" and 3: "is_element_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 4: "h \ Heap_Error_Monad.bind (check_in_heap ptr) (\_. (local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r) ptr ()) \\<^sub>r owner_document" then obtain root where root: "h \ get_root_node ptr \\<^sub>r root" by(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 split: option.splits) then show ?thesis proof (cases "is_document_ptr root") case True then show ?thesis using 3 4 root apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 intro!: filter_M_pure_I bind_pure_I split: option.splits)[1] apply(drule(1) returns_result_eq) apply(auto)[1] using "0" "1" "2" document_ptr_kinds_commutes local.get_root_node_root_in_heap by blast next case False have "known_ptr root" using "0" "1" "2" local.get_root_node_root_in_heap local.known_ptrs_known_ptr root by blast have "root |\| object_ptr_kinds h" using root using "0" "1" "2" local.get_root_node_root_in_heap by blast then have "is_node_ptr_kind root" using False \known_ptr root\ apply(simp add: known_ptr_impl known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs) using is_node_ptr_kind_none by force then have "(\document_ptr \ fset (document_ptr_kinds h). root \ cast ` set |h \ get_disconnected_nodes document_ptr|\<^sub>r)" by (metis (no_types, lifting) "0" "1" "2" \root |\| object_ptr_kinds h\ local.child_parent_dual local.get_child_nodes_ok local.get_root_node_same_no_parent local.heap_is_wellformed_children_disc_nodes local.known_ptrs_known_ptr node_ptr_casts_commute3 node_ptr_inclusion node_ptr_kinds_commutes notin_fset option.distinct(1) returns_result_eq returns_result_select_result root) then obtain some_owner_document where "some_owner_document |\| document_ptr_kinds h" and "root \ cast ` set |h \ get_disconnected_nodes some_owner_document|\<^sub>r" by auto then obtain candidates where candidates: "h \ filter_M (\document_ptr. Heap_Error_Monad.bind (get_disconnected_nodes document_ptr) (\disconnected_nodes. return (root \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ` set disconnected_nodes))) (sorted_list_of_set (fset (document_ptr_kinds h))) \\<^sub>r candidates" by (metis (no_types, lifting) "1" bind_is_OK_I2 bind_pure_I filter_M_is_OK_I finite_fset is_OK_returns_result_E local.get_disconnected_nodes_ok local.get_disconnected_nodes_pure notin_fset return_ok return_pure sorted_list_of_set(1)) then have "some_owner_document \ set candidates" apply(rule filter_M_in_result_if_ok) using \some_owner_document |\| document_ptr_kinds h\ \root \ cast ` set |h \ get_disconnected_nodes some_owner_document|\<^sub>r\ apply(auto intro!: bind_pure_I bind_pure_returns_result_I)[1] apply (simp add: \some_owner_document |\| document_ptr_kinds h\) using "1" \root \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ` set |h \ get_disconnected_nodes some_owner_document|\<^sub>r\ \some_owner_document |\| document_ptr_kinds h\ local.get_disconnected_nodes_ok by auto then have "candidates \ []" by auto then have "owner_document \ set candidates" using 4 root 3 apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 intro!: filter_M_pure_I bind_pure_I split: option.splits)[1] apply (metis candidates list.set_sel(1) returns_result_eq) by (metis \is_node_ptr_kind root\ node_ptr_no_document_ptr_cast returns_result_eq) then show ?thesis using candidates by (meson bind_pure_I bind_returns_result_E2 filter_M_holds_for_result is_OK_returns_result_I local.get_disconnected_nodes_ptr_in_heap local.get_disconnected_nodes_pure return_pure) qed qed lemma get_owner_document_ok: assumes "heap_is_wellformed h" "known_ptrs h" "type_wf h" assumes "ptr |\| object_ptr_kinds h" shows "h \ ok (get_owner_document ptr)" proof - have "known_ptr ptr" using assms(2) assms(4) local.known_ptrs_known_ptr by blast then show ?thesis apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1] apply(split invoke_splits, (rule conjI | rule impI)+)+ apply(auto simp add: known_ptr_impl)[1] using NodeClass.a_known_ptr_def known_ptr_not_character_data_ptr known_ptr_not_document_ptr known_ptr_not_element_ptr apply blast using assms(4) apply(auto simp add: a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_is_OK_pure_I)[1] apply (metis (no_types, lifting) document_ptr_casts_commute3 document_ptr_kinds_commutes is_document_ptr_kind_none option.case_eq_if) using assms(4) apply(auto simp add: a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_is_OK_pure_I)[1] apply (metis (no_types, lifting) assms(1) assms(2) assms(3) is_node_ptr_kind_none local.get_root_node_ok node_ptr_casts_commute3 option.case_eq_if) using assms(4) apply(auto simp add: a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_is_OK_pure_I)[1] apply(auto split: option.splits intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I filter_M_is_OK_I)[1] using assms(3) local.get_disconnected_nodes_ok apply blast apply (simp add: assms(1) assms(2) assms(3) local.get_root_node_ok) using assms(4) apply(auto simp add: a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_is_OK_pure_I)[1] apply(auto split: option.splits intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I filter_M_is_OK_I)[1] apply (simp add: assms(1) assms(2) assms(3) local.get_root_node_ok)[1] apply(auto split: option.splits intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I filter_M_is_OK_I)[1] using assms(3) local.get_disconnected_nodes_ok by blast qed lemma get_owner_document_child_same: assumes "heap_is_wellformed h" "known_ptrs h" "type_wf h" assumes "h \ get_child_nodes ptr \\<^sub>r children" assumes "child \ set children" shows "h \ get_owner_document ptr \\<^sub>r owner_document \ h \ get_owner_document (cast child) \\<^sub>r owner_document" proof - have "ptr |\| object_ptr_kinds h" by (meson assms(4) is_OK_returns_result_I local.get_child_nodes_ptr_in_heap) then have "known_ptr ptr" using assms(2) local.known_ptrs_known_ptr by blast have "cast child |\| object_ptr_kinds h" using assms(1) assms(4) assms(5) local.heap_is_wellformed_children_in_heap node_ptr_kinds_commutes by blast then have "known_ptr (cast child)" using assms(2) local.known_ptrs_known_ptr by blast obtain root where root: "h \ get_root_node ptr \\<^sub>r root" by (meson \ptr |\| object_ptr_kinds h\ assms(1) assms(2) assms(3) is_OK_returns_result_E local.get_root_node_ok) then have "h \ get_root_node (cast child) \\<^sub>r root" using assms(1) assms(2) assms(3) assms(4) assms(5) local.child_parent_dual local.get_root_node_parent_same by blast have "h \ get_owner_document ptr \\<^sub>r owner_document \ h \ a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child () \\<^sub>r owner_document" proof (cases "is_document_ptr ptr") case True then obtain document_ptr where document_ptr: "cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr = ptr" using case_optionE document_ptr_casts_commute by blast then have "root = cast document_ptr" using root by(auto simp add: get_root_node_def get_ancestors_def elim!: bind_returns_result_E2 split: option.splits) then have "h \ a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr () \\<^sub>r owner_document \ h \ a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child () \\<^sub>r owner_document" using document_ptr \h \ get_root_node (cast child) \\<^sub>r root\[simplified \root = cast document_ptr\ document_ptr] apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 dest!: bind_returns_result_E3[rotated, OF \h \ get_root_node (cast child) \\<^sub>r root\[simplified \root = cast document_ptr\ document_ptr], rotated] intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I split: if_splits option.splits)[1] using \ptr |\| object_ptr_kinds h\ document_ptr_kinds_commutes by blast then show ?thesis using \known_ptr ptr\ apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def known_ptr_impl)[1] apply(split invoke_splits, ((rule conjI | rule impI)+)?)+ apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl]) apply(drule(1) known_ptr_not_character_data_ptr) apply(drule(1) known_ptr_not_element_ptr) apply(simp add: NodeClass.known_ptr_defs) using \ptr |\| object_ptr_kinds h\ True by(auto simp add: document_ptr[symmetric] intro!: bind_pure_returns_result_I split: option.splits) next case False then obtain node_ptr where node_ptr: "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr = ptr" using \known_ptr ptr\ by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits) then have "h \ a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r owner_document \ h \ a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child () \\<^sub>r owner_document" using root \h \ get_root_node (cast child) \\<^sub>r root\ unfolding a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def by (meson bind_pure_returns_result_I bind_returns_result_E3 local.get_root_node_pure) then show ?thesis using \known_ptr ptr\ apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def known_ptr_impl)[1] apply(split invoke_splits, ((rule conjI | rule impI)+)?)+ apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl]) apply(drule(1) known_ptr_not_character_data_ptr) apply(drule(1) known_ptr_not_element_ptr) apply(simp add: NodeClass.known_ptr_defs) using \cast child |\| object_ptr_kinds h\ \ptr |\| object_ptr_kinds h\ False apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1] using \cast child |\| object_ptr_kinds h\ \ptr |\| object_ptr_kinds h\ False apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1] using \cast child |\| object_ptr_kinds h\ \ptr |\| object_ptr_kinds h\ False apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1] using \cast child |\| object_ptr_kinds h\ \ptr |\| object_ptr_kinds h\ False apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1] apply(split invoke_splits, ((rule conjI | rule impI)+)?)+ apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1] apply (meson invoke_empty is_OK_returns_result_I) apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1] apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1] by(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1] qed then show ?thesis using \known_ptr (cast child)\ apply(auto simp add: get_owner_document_def[of "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child"] a_get_owner_document_tups_def known_ptr_impl)[1] apply(split invoke_splits, ((rule conjI | rule impI)+)?)+ apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl]) apply(drule(1) known_ptr_not_character_data_ptr) apply(drule(1) known_ptr_not_element_ptr) apply(simp add: NodeClass.known_ptr_defs) using \cast child |\| object_ptr_kinds h\ \ptr |\| object_ptr_kinds h\ apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1] using \cast child |\| object_ptr_kinds h\ \ptr |\| object_ptr_kinds h\ apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1] using \cast child |\| object_ptr_kinds h\ \ptr |\| object_ptr_kinds h\ apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1] using \cast child |\| object_ptr_kinds h\ \ptr |\| object_ptr_kinds h\ apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1] by (smt \cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child |\| object_ptr_kinds h\ cast_document_ptr_not_node_ptr(1) comp_apply invoke_empty invoke_not invoke_returns_result is_OK_returns_result_I node_ptr_casts_commute2 option.sel) qed end locale l_get_owner_document_wf = l_heap_is_wellformed_defs + l_type_wf + l_known_ptrs + l_get_disconnected_nodes_defs + l_get_owner_document_defs + l_get_parent_defs + l_get_child_nodes_defs + assumes get_owner_document_disconnected_nodes: "heap_is_wellformed h \ known_ptrs h \ type_wf h \ h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes \ node_ptr \ set disc_nodes \ h \ get_owner_document (cast node_ptr) \\<^sub>r document_ptr" assumes in_disconnected_nodes_no_parent: "heap_is_wellformed h \ h \ get_parent node_ptr \\<^sub>r None\ h \ get_owner_document (cast node_ptr) \\<^sub>r owner_document \ h \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes \ known_ptrs h \ type_wf h\ node_ptr \ set disc_nodes" assumes get_owner_document_owner_document_in_heap: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_owner_document ptr \\<^sub>r owner_document \ owner_document |\| document_ptr_kinds h" assumes get_owner_document_ok: "heap_is_wellformed h \ known_ptrs h \ type_wf h \ ptr |\| object_ptr_kinds h \ h \ ok (get_owner_document ptr)" assumes get_owner_document_child_same: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_child_nodes ptr \\<^sub>r children \ child \ set children \ h \ get_owner_document ptr \\<^sub>r owner_document \ h \ get_owner_document (cast child) \\<^sub>r owner_document" interpretation i_get_owner_document_wf?: l_get_owner_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr known_ptrs type_wf heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_parent get_parent_locs get_ancestors get_ancestors_locs get_root_node get_root_node_locs get_owner_document by(auto simp add: l_get_owner_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_owner_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_owner_document_wf_is_l_get_owner_document_wf [instances]: "l_get_owner_document_wf heap_is_wellformed type_wf known_ptr known_ptrs get_disconnected_nodes get_owner_document get_parent get_child_nodes" using known_ptrs_is_l_known_ptrs apply(auto simp add: l_get_owner_document_wf_def l_get_owner_document_wf_axioms_def)[1] using get_owner_document_disconnected_nodes apply fast using in_disconnected_nodes_no_parent apply fast using get_owner_document_owner_document_in_heap apply fast using get_owner_document_ok apply fast using get_owner_document_child_same apply (fast, fast) done subsubsection \get\_root\_node\ locale l_get_owner_document_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_root_node_wf + l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_owner_document_wf begin lemma get_root_node_document: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_root_node ptr \\<^sub>r root" assumes "is_document_ptr_kind root" shows "h \ get_owner_document ptr \\<^sub>r the (cast root)" proof - have "ptr |\| object_ptr_kinds h" using assms(4) by (meson is_OK_returns_result_I local.get_root_node_ptr_in_heap) then have "known_ptr ptr" using assms(3) local.known_ptrs_known_ptr by blast { assume "is_document_ptr_kind ptr" then have "ptr = root" using assms(4) by(auto simp add: get_root_node_def get_ancestors_def elim!: bind_returns_result_E2 split: option.splits) then have ?thesis using \is_document_ptr_kind ptr\ \known_ptr ptr\ \ptr |\| object_ptr_kinds h\ apply(auto simp add: known_ptr_impl get_owner_document_def a_get_owner_document_tups_def)[1] apply(split invoke_splits, (rule conjI | rule impI)+)+ apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl]) apply(drule(1) known_ptr_not_character_data_ptr) apply(drule(1) known_ptr_not_element_ptr) apply(simp add: NodeClass.known_ptr_defs) by(auto simp add: a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_pure_returns_result_I split: option.splits) } moreover { assume "is_node_ptr_kind ptr" then have ?thesis using \known_ptr ptr\ \ptr |\| object_ptr_kinds h\ apply(auto simp add: known_ptr_impl get_owner_document_def a_get_owner_document_tups_def)[1] apply(split invoke_splits, (rule conjI | rule impI)+)+ apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl]) apply(drule(1) known_ptr_not_character_data_ptr) apply(drule(1) known_ptr_not_element_ptr) apply(simp add: NodeClass.known_ptr_defs) apply(auto split: option.splits)[1] using \h \ get_root_node ptr \\<^sub>r root\ assms(5) by(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def is_document_ptr_kind_def intro!: bind_pure_returns_result_I split: option.splits)[2] } ultimately show ?thesis using \known_ptr ptr\ by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits) qed lemma get_root_node_same_owner_document: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_root_node ptr \\<^sub>r root" shows "h \ get_owner_document ptr \\<^sub>r owner_document \ h \ get_owner_document root \\<^sub>r owner_document" proof - have "ptr |\| object_ptr_kinds h" by (meson assms(4) is_OK_returns_result_I local.get_root_node_ptr_in_heap) have "root |\| object_ptr_kinds h" using assms(1) assms(2) assms(3) assms(4) local.get_root_node_root_in_heap by blast have "known_ptr ptr" using \ptr |\| object_ptr_kinds h\ assms(3) local.known_ptrs_known_ptr by blast have "known_ptr root" using \root |\| object_ptr_kinds h\ assms(3) local.known_ptrs_known_ptr by blast show ?thesis proof (cases "is_document_ptr_kind ptr") case True then have "ptr = root" using assms(4) apply(auto simp add: get_root_node_def elim!: bind_returns_result_E2)[1] by (metis document_ptr_casts_commute3 last_ConsL local.get_ancestors_not_node node_ptr_no_document_ptr_cast) then show ?thesis by auto next case False then have "is_node_ptr_kind ptr" using \known_ptr ptr\ by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits) then obtain node_ptr where node_ptr: "ptr = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr" by (metis node_ptr_casts_commute3) show ?thesis proof assume "h \ get_owner_document ptr \\<^sub>r owner_document" then have "h \ local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r owner_document" using node_ptr apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1] apply(split invoke_splits)+ apply (meson invoke_empty is_OK_returns_result_I) by(auto elim!: bind_returns_result_E2 split: option.splits) show "h \ get_owner_document root \\<^sub>r owner_document" proof (cases "is_document_ptr_kind root") case True have "is_document_ptr root" using True \known_ptr root\ by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits) have "root = cast owner_document" using True by (smt \h \ get_owner_document ptr \\<^sub>r owner_document\ assms(1) assms(2) assms(3) assms(4) document_ptr_casts_commute3 get_root_node_document returns_result_eq) then show ?thesis apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1] apply(split invoke_splits, (rule conjI | rule impI)+)+ using \is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root\ apply blast using \root |\| object_ptr_kinds h\ by(auto simp add: a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def is_node_ptr_kind_none) next case False then have "is_node_ptr_kind root" using \known_ptr root\ by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits) then obtain root_node_ptr where root_node_ptr: "root = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr" by (metis node_ptr_casts_commute3) then have "h \ local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr () \\<^sub>r owner_document" using \h \ local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r owner_document\ assms(4) apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I split: option.splits)[1] apply (metis assms(1) assms(2) assms(3) local.get_root_node_no_parent local.get_root_node_same_no_parent node_ptr returns_result_eq) using \is_node_ptr_kind root\ node_ptr returns_result_eq by fastforce then show ?thesis apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1] apply(split invoke_splits, (rule conjI | rule impI)+)+ using \is_node_ptr_kind root\ \known_ptr root\ apply(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits)[1] apply(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits)[1] using \root |\| object_ptr_kinds h\ by(auto simp add: root_node_ptr) qed next assume "h \ get_owner_document root \\<^sub>r owner_document" show "h \ get_owner_document ptr \\<^sub>r owner_document" proof (cases "is_document_ptr_kind root") case True have "root = cast owner_document" using \h \ get_owner_document root \\<^sub>r owner_document\ apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1] apply(split invoke_splits)+ apply (meson invoke_empty is_OK_returns_result_I) apply(auto simp add: True a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 split: if_splits)[1] apply (metis True cast_document_ptr_not_node_ptr(2) is_document_ptr_kind_obtains is_node_ptr_kind_none node_ptr_casts_commute3 option.case_eq_if) by (metis True cast_document_ptr_not_node_ptr(1) document_ptr_casts_commute3 is_node_ptr_kind_none node_ptr_casts_commute3 option.case_eq_if) then show ?thesis using assms(1) assms(2) assms(3) assms(4) get_root_node_document by fastforce next case False then have "is_node_ptr_kind root" using \known_ptr root\ by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits) then obtain root_node_ptr where root_node_ptr: "root = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr" by (metis node_ptr_casts_commute3) then have "h \ local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr () \\<^sub>r owner_document" using \h \ get_owner_document root \\<^sub>r owner_document\ apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1] apply(split invoke_splits)+ apply (meson invoke_empty is_OK_returns_result_I) by(auto simp add: is_document_ptr_kind_none elim!: bind_returns_result_E2) then have "h \ local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r owner_document" apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I split: option.splits)[1] using assms(1) assms(2) assms(3) assms(4) local.get_root_node_no_parent local.get_root_node_same_no_parent node_ptr returns_result_eq root_node_ptr by fastforce+ then show ?thesis apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1] apply(split invoke_splits, (rule conjI | rule impI)+)+ using node_ptr \known_ptr ptr\ \ptr |\| object_ptr_kinds h\ by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs intro!: bind_pure_returns_result_I split: option.splits) qed qed qed qed end interpretation get_owner_document_wf_get_root_node_wf?: l_get_owner_document_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_ancestors get_ancestors_locs get_root_node get_root_node_locs heap_is_wellformed parent_child_rel get_disconnected_nodes get_disconnected_nodes_locs get_owner_document by(auto simp add: l_get_owner_document_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_owner_document_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_get_owner_document_wf_get_root_node_wf = l_heap_is_wellformed_defs + l_type_wf + l_known_ptrs + l_get_root_node_defs + l_get_owner_document_defs + assumes get_root_node_document: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_root_node ptr \\<^sub>r root \ is_document_ptr_kind root \ h \ get_owner_document ptr \\<^sub>r the (cast root)" assumes get_root_node_same_owner_document: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_root_node ptr \\<^sub>r root \ h \ get_owner_document ptr \\<^sub>r owner_document \ h \ get_owner_document root \\<^sub>r owner_document" lemma get_owner_document_wf_get_root_node_wf_is_l_get_owner_document_wf_get_root_node_wf [instances]: "l_get_owner_document_wf_get_root_node_wf heap_is_wellformed type_wf known_ptr known_ptrs get_root_node get_owner_document" apply(auto simp add: l_get_owner_document_wf_get_root_node_wf_def l_get_owner_document_wf_get_root_node_wf_axioms_def instances)[1] using get_root_node_document apply blast using get_root_node_same_owner_document apply (blast, blast) done subsection \Preserving heap-wellformedness\ subsection \set\_attribute\ locale l_set_attribute_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_attribute\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_attribute_get_disconnected_nodes + l_set_attribute_get_child_nodes begin lemma set_attribute_preserves_wellformedness: assumes "heap_is_wellformed h" and "h \ set_attribute element_ptr k v \\<^sub>h h'" shows "heap_is_wellformed h'" thm preserves_wellformedness_writes_needed apply(rule preserves_wellformedness_writes_needed[OF assms set_attribute_writes]) using set_attribute_get_child_nodes apply(fast) using set_attribute_get_disconnected_nodes apply(fast) by(auto simp add: all_args_def set_attribute_locs_def) end subsection \remove\_child\ locale l_remove_child_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_heap_is_wellformed + l_set_disconnected_nodes_get_child_nodes begin lemma remove_child_removes_parent: assumes wellformed: "heap_is_wellformed h" and remove_child: "h \ remove_child ptr child \\<^sub>h h2" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "h2 \ get_parent child \\<^sub>r None" proof - obtain children where children: "h \ get_child_nodes ptr \\<^sub>r children" using remove_child remove_child_def by auto then have "child \ set children" using remove_child remove_child_def by(auto elim!: bind_returns_heap_E dest: returns_result_eq split: if_splits) then have h1: "\other_ptr other_children. other_ptr \ ptr \ h \ get_child_nodes other_ptr \\<^sub>r other_children \ child \ set other_children" using assms(1) known_ptrs type_wf child_parent_dual by (meson child_parent_dual children option.inject returns_result_eq) have known_ptr: "known_ptr ptr" using known_ptrs by (meson is_OK_returns_heap_I l_known_ptrs.known_ptrs_known_ptr l_known_ptrs_axioms remove_child remove_child_ptr_in_heap) obtain owner_document disc_nodes h' where owner_document: "h \ get_owner_document (cast child) \\<^sub>r owner_document" and disc_nodes: "h \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes" and h': "h \ set_disconnected_nodes owner_document (child # disc_nodes) \\<^sub>h h'" and h2: "h' \ set_child_nodes ptr (remove1 child children) \\<^sub>h h2" using assms children unfolding remove_child_def apply(auto split: if_splits elim!: bind_returns_heap_E)[1] by (metis (full_types) get_child_nodes_pure get_disconnected_nodes_pure get_owner_document_pure pure_returns_heap_eq returns_result_eq) have "object_ptr_kinds h = object_ptr_kinds h2" using remove_child_writes remove_child unfolding remove_child_locs_def apply(rule writes_small_big) using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by(auto simp add: reflp_def transp_def) then have "|h \ object_ptr_kinds_M|\<^sub>r = |h2 \ object_ptr_kinds_M|\<^sub>r" unfolding object_ptr_kinds_M_defs by simp have "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h'] using set_disconnected_nodes_types_preserved type_wf by(auto simp add: reflp_def transp_def) have "type_wf h2" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF remove_child_writes remove_child] unfolding remove_child_locs_def using set_disconnected_nodes_types_preserved set_child_nodes_types_preserved type_wf apply(auto simp add: reflp_def transp_def)[1] by blast then obtain children' where children': "h2 \ get_child_nodes ptr \\<^sub>r children'" using h2 set_child_nodes_get_child_nodes known_ptr by (metis \object_ptr_kinds h = object_ptr_kinds h2\ children get_child_nodes_ok get_child_nodes_ptr_in_heap is_OK_returns_result_E is_OK_returns_result_I) have "child \ set children'" by (metis (mono_tags, lifting) \type_wf h'\ children children' distinct_remove1_removeAll h2 known_ptr local.heap_is_wellformed_children_distinct local.set_child_nodes_get_child_nodes member_remove remove_code(1) select_result_I2 wellformed) moreover have "\other_ptr other_children. other_ptr \ ptr \ h' \ get_child_nodes other_ptr \\<^sub>r other_children \ child \ set other_children" proof - fix other_ptr other_children assume a1: "other_ptr \ ptr" and a3: "h' \ get_child_nodes other_ptr \\<^sub>r other_children" have "h \ get_child_nodes other_ptr \\<^sub>r other_children" using get_child_nodes_reads set_disconnected_nodes_writes h' a3 apply(rule reads_writes_separate_backwards) using set_disconnected_nodes_get_child_nodes by fast show "child \ set other_children" using \h \ get_child_nodes other_ptr \\<^sub>r other_children\ a1 h1 by blast qed then have "\other_ptr other_children. other_ptr \ ptr \ h2 \ get_child_nodes other_ptr \\<^sub>r other_children \ child \ set other_children" proof - fix other_ptr other_children assume a1: "other_ptr \ ptr" and a3: "h2 \ get_child_nodes other_ptr \\<^sub>r other_children" have "h' \ get_child_nodes other_ptr \\<^sub>r other_children" using get_child_nodes_reads set_child_nodes_writes h2 a3 apply(rule reads_writes_separate_backwards) using set_disconnected_nodes_get_child_nodes a1 set_child_nodes_get_child_nodes_different_pointers by metis then show "child \ set other_children" using \\other_ptr other_children. \other_ptr \ ptr; h' \ get_child_nodes other_ptr \\<^sub>r other_children\ \ child \ set other_children\ a1 by blast qed ultimately have ha: "\other_ptr other_children. h2 \ get_child_nodes other_ptr \\<^sub>r other_children \ child \ set other_children" by (metis (full_types) children' returns_result_eq) moreover obtain ptrs where ptrs: "h2 \ object_ptr_kinds_M \\<^sub>r ptrs" by (simp add: object_ptr_kinds_M_defs) moreover have "\ptr. ptr \ set ptrs \ h2 \ ok (get_child_nodes ptr)" using \type_wf h2\ ptrs get_child_nodes_ok known_ptr using \object_ptr_kinds h = object_ptr_kinds h2\ known_ptrs local.known_ptrs_known_ptr by auto ultimately show "h2 \ get_parent child \\<^sub>r None" apply(auto simp add: get_parent_def intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I)[1] proof - have "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child |\| object_ptr_kinds h" using get_owner_document_ptr_in_heap owner_document by blast then show "h2 \ check_in_heap (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r ()" by (simp add: \object_ptr_kinds h = object_ptr_kinds h2\ check_in_heap_def) next show "(\other_ptr other_children. h2 \ get_child_nodes other_ptr \\<^sub>r other_children \ child \ set other_children) \ ptrs = sorted_list_of_set (fset (object_ptr_kinds h2)) \ (\ptr. ptr |\| object_ptr_kinds h2 \ h2 \ ok get_child_nodes ptr) \ h2 \ filter_M (\ptr. Heap_Error_Monad.bind (get_child_nodes ptr) (\children. return (child \ set children))) (sorted_list_of_set (fset (object_ptr_kinds h2))) \\<^sub>r []" by(auto intro!: filter_M_empty_I bind_pure_I) qed qed end locale l_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_remove_child_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma remove_child_parent_child_rel_subset: assumes "heap_is_wellformed h" and "h \ remove_child ptr child \\<^sub>h h'" and "known_ptrs h" and type_wf: "type_wf h" shows "parent_child_rel h' \ parent_child_rel h" proof (standard, safe) obtain owner_document children_h h2 disconnected_nodes_h where owner_document: "h \ get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r owner_document" and children_h: "h \ get_child_nodes ptr \\<^sub>r children_h" and child_in_children_h: "child \ set children_h" and disconnected_nodes_h: "h \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h" and h2: "h \ set_disconnected_nodes owner_document (child # disconnected_nodes_h) \\<^sub>h h2" and h': "h2 \ set_child_nodes ptr (remove1 child children_h) \\<^sub>h h'" using assms(2) apply(auto simp add: remove_child_def elim!: bind_returns_heap_E dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure] pure_returns_heap_eq[rotated, OF get_child_nodes_pure] split: if_splits)[1] using pure_returns_heap_eq by fastforce have object_ptr_kinds_eq3: "object_ptr_kinds h = object_ptr_kinds h'" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF remove_child_writes assms(2)]) unfolding remove_child_locs_def using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_eq: "\ptrs. h \ object_ptr_kinds_M \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" unfolding object_ptr_kinds_M_defs by simp then have object_ptr_kinds_eq2: "|h \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" using select_result_eq by force then have node_ptr_kinds_eq2: "|h \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by auto then have node_ptr_kinds_eq3: "node_ptr_kinds h = node_ptr_kinds h'" using node_ptr_kinds_M_eq by auto have document_ptr_kinds_eq2: "|h \ document_ptr_kinds_M|\<^sub>r = |h' \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_eq2 document_ptr_kinds_M_eq by auto then have document_ptr_kinds_eq3: "document_ptr_kinds h = document_ptr_kinds h'" using document_ptr_kinds_M_eq by auto have children_eq: "\ptr' children. ptr \ ptr' \ h \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" apply(rule reads_writes_preserved[OF get_child_nodes_reads remove_child_writes assms(2)]) unfolding remove_child_locs_def using set_disconnected_nodes_get_child_nodes set_child_nodes_get_child_nodes_different_pointers by fast then have children_eq2: "\ptr' children. ptr \ ptr' \ |h \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq: "\document_ptr disconnected_nodes. document_ptr \ owner_document \ h \ get_disconnected_nodes document_ptr \\<^sub>r disconnected_nodes = h' \ get_disconnected_nodes document_ptr \\<^sub>r disconnected_nodes" apply(rule reads_writes_preserved[OF get_disconnected_nodes_reads remove_child_writes assms(2)]) unfolding remove_child_locs_def using set_child_nodes_get_disconnected_nodes set_disconnected_nodes_get_disconnected_nodes_different_pointers by (metis (no_types, lifting) Un_iff owner_document select_result_I2) then have disconnected_nodes_eq2: "\document_ptr. document_ptr \ owner_document \ |h \ get_disconnected_nodes document_ptr|\<^sub>r = |h' \ get_disconnected_nodes document_ptr|\<^sub>r" using select_result_eq by force have "h2 \ get_child_nodes ptr \\<^sub>r children_h" apply(rule reads_writes_separate_forwards[OF get_child_nodes_reads set_disconnected_nodes_writes h2 children_h] ) by (simp add: set_disconnected_nodes_get_child_nodes) have "known_ptr ptr" using assms(3) using children_h get_child_nodes_ptr_in_heap local.known_ptrs_known_ptr by blast have "type_wf h2" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h2] using set_disconnected_nodes_types_preserved type_wf by(auto simp add: reflp_def transp_def) then have "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_child_nodes_writes h'] using set_child_nodes_types_preserved by(auto simp add: reflp_def transp_def) have children_h': "h' \ get_child_nodes ptr \\<^sub>r remove1 child children_h" using assms(2) owner_document h2 disconnected_nodes_h children_h apply(auto simp add: remove_child_def split: if_splits)[1] apply(drule bind_returns_heap_E3) apply(auto split: if_splits)[1] apply(simp) apply(auto split: if_splits)[1] apply(drule bind_returns_heap_E3) apply(auto)[1] apply(simp) apply(drule bind_returns_heap_E3) apply(auto)[1] apply(simp) apply(drule bind_returns_heap_E4) apply(auto)[1] apply(simp) using \type_wf h2\ set_child_nodes_get_child_nodes \known_ptr ptr\ h' by blast fix parent child assume a1: "(parent, child) \ parent_child_rel h'" then show "(parent, child) \ parent_child_rel h" proof (cases "parent = ptr") case True then show ?thesis using a1 remove_child_removes_parent[OF assms(1) assms(2)] children_h children_h' get_child_nodes_ptr_in_heap apply(auto simp add: parent_child_rel_def object_ptr_kinds_eq )[1] by (metis notin_set_remove1) next case False then show ?thesis using a1 by(auto simp add: parent_child_rel_def object_ptr_kinds_eq3 children_eq2) qed qed lemma remove_child_heap_is_wellformed_preserved: assumes "heap_is_wellformed h" and "h \ remove_child ptr child \\<^sub>h h'" and "known_ptrs h" and type_wf: "type_wf h" shows "type_wf h'" and "known_ptrs h'" and "heap_is_wellformed h'" proof - obtain owner_document children_h h2 disconnected_nodes_h where owner_document: "h \ get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r owner_document" and children_h: "h \ get_child_nodes ptr \\<^sub>r children_h" and child_in_children_h: "child \ set children_h" and disconnected_nodes_h: "h \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h" and h2: "h \ set_disconnected_nodes owner_document (child # disconnected_nodes_h) \\<^sub>h h2" and h': "h2 \ set_child_nodes ptr (remove1 child children_h) \\<^sub>h h'" using assms(2) apply(auto simp add: remove_child_def elim!: bind_returns_heap_E dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure] pure_returns_heap_eq[rotated, OF get_child_nodes_pure] split: if_splits)[1] using pure_returns_heap_eq by fastforce have object_ptr_kinds_eq3: "object_ptr_kinds h = object_ptr_kinds h'" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF remove_child_writes assms(2)]) unfolding remove_child_locs_def using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_eq: "\ptrs. h \ object_ptr_kinds_M \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" unfolding object_ptr_kinds_M_defs by simp then have object_ptr_kinds_eq2: "|h \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" using select_result_eq by force then have node_ptr_kinds_eq2: "|h \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by auto then have node_ptr_kinds_eq3: "node_ptr_kinds h = node_ptr_kinds h'" using node_ptr_kinds_M_eq by auto have document_ptr_kinds_eq2: "|h \ document_ptr_kinds_M|\<^sub>r = |h' \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_eq2 document_ptr_kinds_M_eq by auto then have document_ptr_kinds_eq3: "document_ptr_kinds h = document_ptr_kinds h'" using document_ptr_kinds_M_eq by auto have children_eq: "\ptr' children. ptr \ ptr' \ h \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" apply(rule reads_writes_preserved[OF get_child_nodes_reads remove_child_writes assms(2)]) unfolding remove_child_locs_def using set_disconnected_nodes_get_child_nodes set_child_nodes_get_child_nodes_different_pointers by fast then have children_eq2: "\ptr' children. ptr \ ptr' \ |h \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq: "\document_ptr disconnected_nodes. document_ptr \ owner_document \ h \ get_disconnected_nodes document_ptr \\<^sub>r disconnected_nodes = h' \ get_disconnected_nodes document_ptr \\<^sub>r disconnected_nodes" apply(rule reads_writes_preserved[OF get_disconnected_nodes_reads remove_child_writes assms(2)]) unfolding remove_child_locs_def using set_child_nodes_get_disconnected_nodes set_disconnected_nodes_get_disconnected_nodes_different_pointers by (metis (no_types, lifting) Un_iff owner_document select_result_I2) then have disconnected_nodes_eq2: "\document_ptr. document_ptr \ owner_document \ |h \ get_disconnected_nodes document_ptr|\<^sub>r = |h' \ get_disconnected_nodes document_ptr|\<^sub>r" using select_result_eq by force have "h2 \ get_child_nodes ptr \\<^sub>r children_h" apply(rule reads_writes_separate_forwards[OF get_child_nodes_reads set_disconnected_nodes_writes h2 children_h] ) by (simp add: set_disconnected_nodes_get_child_nodes) show "known_ptrs h'" using object_ptr_kinds_eq3 known_ptrs_preserved \known_ptrs h\ by blast have "known_ptr ptr" using assms(3) using children_h get_child_nodes_ptr_in_heap local.known_ptrs_known_ptr by blast have "type_wf h2" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h2] using set_disconnected_nodes_types_preserved type_wf by(auto simp add: reflp_def transp_def) then show "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_child_nodes_writes h'] using set_child_nodes_types_preserved by(auto simp add: reflp_def transp_def) have children_h': "h' \ get_child_nodes ptr \\<^sub>r remove1 child children_h" using assms(2) owner_document h2 disconnected_nodes_h children_h apply(auto simp add: remove_child_def split: if_splits)[1] apply(drule bind_returns_heap_E3) apply(auto split: if_splits)[1] apply(simp) apply(auto split: if_splits)[1] apply(drule bind_returns_heap_E3) apply(auto)[1] apply(simp) apply(drule bind_returns_heap_E3) apply(auto)[1] apply(simp) apply(drule bind_returns_heap_E4) apply(auto)[1] apply simp using \type_wf h2\ set_child_nodes_get_child_nodes \known_ptr ptr\ h' by blast have disconnected_nodes_h2: "h2 \ get_disconnected_nodes owner_document \\<^sub>r child # disconnected_nodes_h" using owner_document assms(2) h2 disconnected_nodes_h apply (auto simp add: remove_child_def split: if_splits)[1] apply(drule bind_returns_heap_E2) apply(auto split: if_splits)[1] apply(simp) by(auto simp add: local.set_disconnected_nodes_get_disconnected_nodes split: if_splits) then have disconnected_nodes_h': "h' \ get_disconnected_nodes owner_document \\<^sub>r child # disconnected_nodes_h" apply(rule reads_writes_separate_forwards[OF get_disconnected_nodes_reads set_child_nodes_writes h']) by (simp add: set_child_nodes_get_disconnected_nodes) moreover have "a_acyclic_heap h" using assms(1) by (simp add: heap_is_wellformed_def) have "parent_child_rel h' \ parent_child_rel h" proof (standard, safe) fix parent child assume a1: "(parent, child) \ parent_child_rel h'" then show "(parent, child) \ parent_child_rel h" proof (cases "parent = ptr") case True then show ?thesis using a1 remove_child_removes_parent[OF assms(1) assms(2)] children_h children_h' get_child_nodes_ptr_in_heap apply(auto simp add: parent_child_rel_def object_ptr_kinds_eq )[1] by (metis imageI notin_set_remove1) next case False then show ?thesis using a1 by(auto simp add: parent_child_rel_def object_ptr_kinds_eq3 children_eq2) qed qed then have "a_acyclic_heap h'" using \a_acyclic_heap h\ acyclic_heap_def acyclic_subset by blast moreover have "a_all_ptrs_in_heap h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_all_ptrs_in_heap h'" apply(auto simp add: a_all_ptrs_in_heap_def node_ptr_kinds_eq3 disconnected_nodes_eq)[1] apply (metis (no_types, lifting) \type_wf h'\ assms(2) assms(3) local.get_child_nodes_ok local.known_ptrs_known_ptr local.remove_child_children_subset notin_fset object_ptr_kinds_eq3 returns_result_select_result subset_code(1) type_wf) apply (metis (no_types, lifting) assms(2) disconnected_nodes_eq2 disconnected_nodes_h disconnected_nodes_h' document_ptr_kinds_eq3 finite_set_in local.remove_child_child_in_heap node_ptr_kinds_eq3 select_result_I2 set_ConsD subset_code(1)) done moreover have "a_owner_document_valid h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_owner_document_valid h'" apply(auto simp add: a_owner_document_valid_def object_ptr_kinds_eq3 document_ptr_kinds_eq3 node_ptr_kinds_eq3)[1] proof - fix node_ptr assume 0: "\node_ptr\fset (node_ptr_kinds h'). (\document_ptr. document_ptr |\| document_ptr_kinds h' \ node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r) \ (\parent_ptr. parent_ptr |\| object_ptr_kinds h' \ node_ptr \ set |h \ get_child_nodes parent_ptr|\<^sub>r)" and 1: "node_ptr |\| node_ptr_kinds h'" and 2: "\parent_ptr. parent_ptr |\| object_ptr_kinds h' \ node_ptr \ set |h' \ get_child_nodes parent_ptr|\<^sub>r" then show "\document_ptr. document_ptr |\| document_ptr_kinds h' \ node_ptr \ set |h' \ get_disconnected_nodes document_ptr|\<^sub>r" proof (cases "node_ptr = child") case True show ?thesis apply(rule exI[where x=owner_document]) using children_eq2 disconnected_nodes_eq2 children_h children_h' disconnected_nodes_h' True by (metis (no_types, lifting) get_disconnected_nodes_ptr_in_heap is_OK_returns_result_I list.set_intros(1) select_result_I2) next case False then show ?thesis using 0 1 2 children_eq2 children_h children_h' disconnected_nodes_eq2 disconnected_nodes_h disconnected_nodes_h' apply(auto simp add: children_eq2 disconnected_nodes_eq2 dest!: select_result_I2)[1] by (metis children_eq2 disconnected_nodes_eq2 finite_set_in in_set_remove1 list.set_intros(2)) qed qed moreover { have h0: "a_distinct_lists h" using assms(1) by (simp add: heap_is_wellformed_def) moreover have ha1: "(\x\set |h \ object_ptr_kinds_M|\<^sub>r. set |h \ get_child_nodes x|\<^sub>r) \ (\x\set |h \ document_ptr_kinds_M|\<^sub>r. set |h \ get_disconnected_nodes x|\<^sub>r) = {}" using \a_distinct_lists h\ unfolding a_distinct_lists_def by(auto) have ha2: "ptr |\| object_ptr_kinds h" using children_h get_child_nodes_ptr_in_heap by blast have ha3: "child \ set |h \ get_child_nodes ptr|\<^sub>r" using child_in_children_h children_h by(simp) have child_not_in: "\document_ptr. document_ptr |\| document_ptr_kinds h \ child \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r" using ha1 ha2 ha3 apply(simp) using IntI by fastforce moreover have "distinct |h \ object_ptr_kinds_M|\<^sub>r" apply(rule select_result_I) by(auto simp add: object_ptr_kinds_M_defs) moreover have "distinct |h \ document_ptr_kinds_M|\<^sub>r" apply(rule select_result_I) by(auto simp add: document_ptr_kinds_M_defs) ultimately have "a_distinct_lists h'" proof(simp (no_asm) add: a_distinct_lists_def, safe) assume 1: "a_distinct_lists h" and 3: "distinct |h \ object_ptr_kinds_M|\<^sub>r" assume 1: "a_distinct_lists h" and 3: "distinct |h \ object_ptr_kinds_M|\<^sub>r" have 4: "distinct (concat ((map (\ptr. |h \ get_child_nodes ptr|\<^sub>r) |h \ object_ptr_kinds_M|\<^sub>r)))" using 1 by(auto simp add: a_distinct_lists_def) show "distinct (concat (map (\ptr. |h' \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h')))))" proof(rule distinct_concat_map_I[OF 3[unfolded object_ptr_kinds_eq2], simplified]) fix x assume 5: "x |\| object_ptr_kinds h'" then have 6: "distinct |h \ get_child_nodes x|\<^sub>r" using 4 distinct_concat_map_E object_ptr_kinds_eq2 by fastforce obtain children where children: "h \ get_child_nodes x \\<^sub>r children" and distinct_children: "distinct children" by (metis "5" "6" type_wf assms(3) get_child_nodes_ok local.known_ptrs_known_ptr object_ptr_kinds_eq3 select_result_I) obtain children' where children': "h' \ get_child_nodes x \\<^sub>r children'" using children children_eq children_h' by fastforce then have "distinct children'" proof (cases "ptr = x") case True then show ?thesis using children distinct_children children_h children_h' by (metis children' distinct_remove1 returns_result_eq) next case False then show ?thesis using children distinct_children children_eq[OF False] using children' distinct_lists_children h0 using select_result_I2 by fastforce qed then show "distinct |h' \ get_child_nodes x|\<^sub>r" using children' by(auto simp add: ) next fix x y assume 5: "x |\| object_ptr_kinds h'" and 6: "y |\| object_ptr_kinds h'" and 7: "x \ y" obtain children_x where children_x: "h \ get_child_nodes x \\<^sub>r children_x" by (metis "5" type_wf assms(3) get_child_nodes_ok is_OK_returns_result_E local.known_ptrs_known_ptr object_ptr_kinds_eq3) obtain children_y where children_y: "h \ get_child_nodes y \\<^sub>r children_y" by (metis "6" type_wf assms(3) get_child_nodes_ok is_OK_returns_result_E local.known_ptrs_known_ptr object_ptr_kinds_eq3) obtain children_x' where children_x': "h' \ get_child_nodes x \\<^sub>r children_x'" using children_eq children_h' children_x by fastforce obtain children_y' where children_y': "h' \ get_child_nodes y \\<^sub>r children_y'" using children_eq children_h' children_y by fastforce have "distinct (concat (map (\ptr. |h \ get_child_nodes ptr|\<^sub>r) |h \ object_ptr_kinds_M|\<^sub>r))" using h0 by(auto simp add: a_distinct_lists_def) then have 8: "set children_x \ set children_y = {}" using "7" assms(1) children_x children_y local.heap_is_wellformed_one_parent by blast have "set children_x' \ set children_y' = {}" proof (cases "ptr = x") case True then have "ptr \ y" by(simp add: 7) have "children_x' = remove1 child children_x" using children_h children_h' children_x children_x' True returns_result_eq by fastforce moreover have "children_y' = children_y" using children_y children_y' children_eq[OF \ptr \ y\] by auto ultimately show ?thesis using 8 set_remove1_subset by fastforce next case False then show ?thesis proof (cases "ptr = y") case True have "children_y' = remove1 child children_y" using children_h children_h' children_y children_y' True returns_result_eq by fastforce moreover have "children_x' = children_x" using children_x children_x' children_eq[OF \ptr \ x\] by auto ultimately show ?thesis using 8 set_remove1_subset by fastforce next case False have "children_x' = children_x" using children_x children_x' children_eq[OF \ptr \ x\] by auto moreover have "children_y' = children_y" using children_y children_y' children_eq[OF \ptr \ y\] by auto ultimately show ?thesis using 8 by simp qed qed then show "set |h' \ get_child_nodes x|\<^sub>r \ set |h' \ get_child_nodes y|\<^sub>r = {}" using children_x' children_y' by (metis (no_types, lifting) select_result_I2) qed next assume 2: "distinct |h \ document_ptr_kinds_M|\<^sub>r" then have 4: "distinct (sorted_list_of_set (fset (document_ptr_kinds h')))" by simp have 3: "distinct (concat (map (\document_ptr. |h \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h')))))" using h0 by(simp add: a_distinct_lists_def document_ptr_kinds_eq3) show "distinct (concat (map (\document_ptr. |h' \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h')))))" proof(rule distinct_concat_map_I[OF 4[unfolded document_ptr_kinds_eq3]]) fix x assume 4: "x \ set (sorted_list_of_set (fset (document_ptr_kinds h')))" have 5: "distinct |h \ get_disconnected_nodes x|\<^sub>r" using distinct_lists_disconnected_nodes[OF h0] 4 get_disconnected_nodes_ok by (simp add: type_wf document_ptr_kinds_eq3 select_result_I) show "distinct |h' \ get_disconnected_nodes x|\<^sub>r" proof (cases "x = owner_document") case True have "child \ set |h \ get_disconnected_nodes x|\<^sub>r" using child_not_in document_ptr_kinds_eq2 "4" by fastforce moreover have "|h' \ get_disconnected_nodes x|\<^sub>r = child # |h \ get_disconnected_nodes x|\<^sub>r" using disconnected_nodes_h' disconnected_nodes_h unfolding True by(simp) ultimately show ?thesis using 5 unfolding True by simp next case False show ?thesis using "5" False disconnected_nodes_eq2 by auto qed next fix x y assume 4: "x \ set (sorted_list_of_set (fset (document_ptr_kinds h')))" and 5: "y \ set (sorted_list_of_set (fset (document_ptr_kinds h')))" and "x \ y" obtain disc_nodes_x where disc_nodes_x: "h \ get_disconnected_nodes x \\<^sub>r disc_nodes_x" using 4 get_disconnected_nodes_ok[OF \type_wf h\, of x] document_ptr_kinds_eq2 by auto obtain disc_nodes_y where disc_nodes_y: "h \ get_disconnected_nodes y \\<^sub>r disc_nodes_y" using 5 get_disconnected_nodes_ok[OF \type_wf h\, of y] document_ptr_kinds_eq2 by auto obtain disc_nodes_x' where disc_nodes_x': "h' \ get_disconnected_nodes x \\<^sub>r disc_nodes_x'" using 4 get_disconnected_nodes_ok[OF \type_wf h'\, of x] document_ptr_kinds_eq2 by auto obtain disc_nodes_y' where disc_nodes_y': "h' \ get_disconnected_nodes y \\<^sub>r disc_nodes_y'" using 5 get_disconnected_nodes_ok[OF \type_wf h'\, of y] document_ptr_kinds_eq2 by auto have "distinct (concat (map (\document_ptr. |h \ get_disconnected_nodes document_ptr|\<^sub>r) |h \ document_ptr_kinds_M|\<^sub>r))" using h0 by (simp add: a_distinct_lists_def) then have 6: "set disc_nodes_x \ set disc_nodes_y = {}" using \x \ y\ assms(1) disc_nodes_x disc_nodes_y local.heap_is_wellformed_one_disc_parent by blast have "set disc_nodes_x' \ set disc_nodes_y' = {}" proof (cases "x = owner_document") case True then have "y \ owner_document" using \x \ y\ by simp then have "disc_nodes_y' = disc_nodes_y" using disconnected_nodes_eq[OF \y \ owner_document\] disc_nodes_y disc_nodes_y' by auto have "disc_nodes_x' = child # disc_nodes_x" using disconnected_nodes_h' disc_nodes_x disc_nodes_x' True disconnected_nodes_h returns_result_eq by fastforce have "child \ set disc_nodes_y" using child_not_in disc_nodes_y 5 using document_ptr_kinds_eq2 by fastforce then show ?thesis apply(unfold \disc_nodes_x' = child # disc_nodes_x\ \disc_nodes_y' = disc_nodes_y\) using 6 by auto next case False then show ?thesis proof (cases "y = owner_document") case True then have "disc_nodes_x' = disc_nodes_x" using disconnected_nodes_eq[OF \x \ owner_document\] disc_nodes_x disc_nodes_x' by auto have "disc_nodes_y' = child # disc_nodes_y" using disconnected_nodes_h' disc_nodes_y disc_nodes_y' True disconnected_nodes_h returns_result_eq by fastforce have "child \ set disc_nodes_x" using child_not_in disc_nodes_x 4 using document_ptr_kinds_eq2 by fastforce then show ?thesis apply(unfold \disc_nodes_y' = child # disc_nodes_y\ \disc_nodes_x' = disc_nodes_x\) using 6 by auto next case False have "disc_nodes_x' = disc_nodes_x" using disconnected_nodes_eq[OF \x \ owner_document\] disc_nodes_x disc_nodes_x' by auto have "disc_nodes_y' = disc_nodes_y" using disconnected_nodes_eq[OF \y \ owner_document\] disc_nodes_y disc_nodes_y' by auto then show ?thesis apply(unfold \disc_nodes_y' = disc_nodes_y\ \disc_nodes_x' = disc_nodes_x\) using 6 by auto qed qed then show "set |h' \ get_disconnected_nodes x|\<^sub>r \ set |h' \ get_disconnected_nodes y|\<^sub>r = {}" using disc_nodes_x' disc_nodes_y' by auto qed next fix x xa xb assume 1: "xa \ fset (object_ptr_kinds h')" and 2: "x \ set |h' \ get_child_nodes xa|\<^sub>r" and 3: "xb \ fset (document_ptr_kinds h')" and 4: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" obtain disc_nodes where disc_nodes: "h \ get_disconnected_nodes xb \\<^sub>r disc_nodes" using 3 get_disconnected_nodes_ok[OF \type_wf h\, of xb] document_ptr_kinds_eq2 by auto obtain disc_nodes' where disc_nodes': "h' \ get_disconnected_nodes xb \\<^sub>r disc_nodes'" using 3 get_disconnected_nodes_ok[OF \type_wf h'\, of xb] document_ptr_kinds_eq2 by auto obtain children where children: "h \ get_child_nodes xa \\<^sub>r children" by (metis "1" type_wf assms(3) finite_set_in get_child_nodes_ok is_OK_returns_result_E local.known_ptrs_known_ptr object_ptr_kinds_eq3) obtain children' where children': "h' \ get_child_nodes xa \\<^sub>r children'" using children children_eq children_h' by fastforce have "\x. x \ set |h \ get_child_nodes xa|\<^sub>r \ x \ set |h \ get_disconnected_nodes xb|\<^sub>r \ False" using 1 3 apply(fold \ object_ptr_kinds h = object_ptr_kinds h'\) apply(fold \ document_ptr_kinds h = document_ptr_kinds h'\) using children disc_nodes h0 apply(auto simp add: a_distinct_lists_def)[1] by (metis (no_types, lifting) h0 local.distinct_lists_no_parent select_result_I2) then have 5: "\x. x \ set children \ x \ set disc_nodes \ False" using children disc_nodes by fastforce have 6: "|h' \ get_child_nodes xa|\<^sub>r = children'" using children' by (simp add: ) have 7: "|h' \ get_disconnected_nodes xb|\<^sub>r = disc_nodes'" using disc_nodes' by (simp add: ) have "False" proof (cases "xa = ptr") case True have "distinct children_h" using children_h distinct_lists_children h0 \known_ptr ptr\ by blast have "|h' \ get_child_nodes ptr|\<^sub>r = remove1 child children_h" using children_h' by(simp add: ) have "children = children_h" using True children children_h by auto show ?thesis using disc_nodes' children' 5 2 4 children_h \distinct children_h\ disconnected_nodes_h' apply(auto simp add: 6 7 \xa = ptr\ \|h' \ get_child_nodes ptr|\<^sub>r = remove1 child children_h\ \children = children_h\)[1] by (metis (no_types, lifting) disc_nodes disconnected_nodes_eq2 disconnected_nodes_h select_result_I2 set_ConsD) next case False have "children' = children" using children' children children_eq[OF False[symmetric]] by auto then show ?thesis proof (cases "xb = owner_document") case True then show ?thesis using disc_nodes disconnected_nodes_h disconnected_nodes_h' using "2" "4" "5" "6" "7" False \children' = children\ assms(1) child_in_children_h child_parent_dual children children_h disc_nodes' get_child_nodes_ptr_in_heap list.set_cases list.simps(3) option.simps(1) returns_result_eq set_ConsD by (metis (no_types, hide_lams) assms(3) type_wf) next case False then show ?thesis using "2" "4" "5" "6" "7" \children' = children\ disc_nodes disc_nodes' disconnected_nodes_eq returns_result_eq by metis qed qed then show "x \ {}" by simp qed } ultimately show "heap_is_wellformed h'" using heap_is_wellformed_def by blast qed lemma remove_heap_is_wellformed_preserved: assumes "heap_is_wellformed h" and "h \ remove child \\<^sub>h h'" and "known_ptrs h" and type_wf: "type_wf h" shows "type_wf h'" and "known_ptrs h'" and "heap_is_wellformed h'" using assms by(auto simp add: remove_def intro: remove_child_heap_is_wellformed_preserved elim!: bind_returns_heap_E2 split: option.splits) lemma remove_child_removes_child: assumes wellformed: "heap_is_wellformed h" and remove_child: "h \ remove_child ptr' child \\<^sub>h h'" and children: "h' \ get_child_nodes ptr \\<^sub>r children" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "child \ set children" proof - obtain owner_document children_h h2 disconnected_nodes_h where owner_document: "h \ get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r owner_document" and children_h: "h \ get_child_nodes ptr' \\<^sub>r children_h" and child_in_children_h: "child \ set children_h" and disconnected_nodes_h: "h \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h" and h2: "h \ set_disconnected_nodes owner_document (child # disconnected_nodes_h) \\<^sub>h h2" and h': "h2 \ set_child_nodes ptr' (remove1 child children_h) \\<^sub>h h'" using assms(2) apply(auto simp add: remove_child_def elim!: bind_returns_heap_E dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure] pure_returns_heap_eq[rotated, OF get_child_nodes_pure] split: if_splits)[1] using pure_returns_heap_eq by fastforce have "object_ptr_kinds h = object_ptr_kinds h'" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF remove_child_writes remove_child]) unfolding remove_child_locs_def using set_child_nodes_pointers_preserved set_disconnected_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) moreover have "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF remove_child_writes assms(2)] using set_child_nodes_types_preserved set_disconnected_nodes_types_preserved type_wf unfolding remove_child_locs_def apply(auto simp add: reflp_def transp_def)[1] by blast ultimately show ?thesis using remove_child_removes_parent remove_child_heap_is_wellformed_preserved child_parent_dual by (meson children known_ptrs local.known_ptrs_preserved option.distinct(1) remove_child returns_result_eq type_wf wellformed) qed lemma remove_child_removes_first_child: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_child_nodes ptr \\<^sub>r node_ptr # children" assumes "h \ remove_child ptr node_ptr \\<^sub>h h'" shows "h' \ get_child_nodes ptr \\<^sub>r children" proof - obtain h2 disc_nodes owner_document where "h \ get_owner_document (cast node_ptr) \\<^sub>r owner_document" and "h \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes" and h2: "h \ set_disconnected_nodes owner_document (node_ptr # disc_nodes) \\<^sub>h h2" and "h2 \ set_child_nodes ptr children \\<^sub>h h'" using assms(5) apply(auto simp add: remove_child_def dest!: bind_returns_heap_E3[rotated, OF assms(4) get_child_nodes_pure, rotated])[1] by(auto elim!: bind_returns_heap_E bind_returns_heap_E2[rotated,OF get_owner_document_pure, rotated] bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated]) have "known_ptr ptr" by (meson assms(3) assms(4) is_OK_returns_result_I get_child_nodes_ptr_in_heap known_ptrs_known_ptr) moreover have "h2 \ get_child_nodes ptr \\<^sub>r node_ptr # children" apply(rule reads_writes_separate_forwards[OF get_child_nodes_reads set_disconnected_nodes_writes h2 assms(4)]) using set_disconnected_nodes_get_child_nodes set_child_nodes_get_child_nodes_different_pointers by fast moreover have "type_wf h2" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h2] using \type_wf h\ set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) ultimately show ?thesis using set_child_nodes_get_child_nodes\h2 \ set_child_nodes ptr children \\<^sub>h h'\ by fast qed lemma remove_removes_child: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_child_nodes ptr \\<^sub>r node_ptr # children" assumes "h \ remove node_ptr \\<^sub>h h'" shows "h' \ get_child_nodes ptr \\<^sub>r children" proof - have "h \ get_parent node_ptr \\<^sub>r Some ptr" using child_parent_dual assms by fastforce show ?thesis using assms remove_child_removes_first_child by(auto simp add: remove_def dest!: bind_returns_heap_E3[rotated, OF \h \ get_parent node_ptr \\<^sub>r Some ptr\, rotated] bind_returns_heap_E3[rotated, OF assms(4) get_child_nodes_pure, rotated]) qed lemma remove_for_all_empty_children: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_child_nodes ptr \\<^sub>r children" assumes "h \ forall_M remove children \\<^sub>h h'" shows "h' \ get_child_nodes ptr \\<^sub>r []" using assms proof(induct children arbitrary: h h') case Nil then show ?case by simp next case (Cons a children) have "h \ get_parent a \\<^sub>r Some ptr" using child_parent_dual Cons by fastforce with Cons show ?case proof(auto elim!: bind_returns_heap_E)[1] fix h2 assume 0: "(\h h'. heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_child_nodes ptr \\<^sub>r children \ h \ forall_M remove children \\<^sub>h h' \ h' \ get_child_nodes ptr \\<^sub>r [])" and 1: "heap_is_wellformed h" and 2: "type_wf h" and 3: "known_ptrs h" and 4: "h \ get_child_nodes ptr \\<^sub>r a # children" and 5: "h \ get_parent a \\<^sub>r Some ptr" and 7: "h \ remove a \\<^sub>h h2" and 8: "h2 \ forall_M remove children \\<^sub>h h'" then have "h2 \ get_child_nodes ptr \\<^sub>r children" using remove_removes_child by blast moreover have "heap_is_wellformed h2" using 7 1 2 3 remove_child_heap_is_wellformed_preserved(3) by(auto simp add: remove_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_parent_pure, rotated] split: option.splits) moreover have "type_wf h2" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF remove_writes 7] using \type_wf h\ remove_child_types_preserved by(auto simp add: a_remove_child_locs_def reflp_def transp_def) moreover have "object_ptr_kinds h = object_ptr_kinds h2" using 7 apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF remove_writes]) using remove_child_pointers_preserved by (auto simp add: reflp_def transp_def) then have "known_ptrs h2" using 3 known_ptrs_preserved by blast ultimately show "h' \ get_child_nodes ptr \\<^sub>r []" using 0 8 by fast qed qed end locale l_remove_child_wf2 = l_type_wf + l_known_ptrs + l_remove_child_defs + l_heap_is_wellformed_defs + l_get_child_nodes_defs + l_remove_defs + assumes remove_child_preserves_type_wf: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ remove_child ptr child \\<^sub>h h' \ type_wf h'" assumes remove_child_preserves_known_ptrs: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ remove_child ptr child \\<^sub>h h' \ known_ptrs h'" assumes remove_child_heap_is_wellformed_preserved: "type_wf h \ known_ptrs h \ heap_is_wellformed h \ h \ remove_child ptr child \\<^sub>h h' \ heap_is_wellformed h'" assumes remove_preserves_type_wf: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ remove child \\<^sub>h h' \ type_wf h'" assumes remove_preserves_known_ptrs: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ remove child \\<^sub>h h' \ known_ptrs h'" assumes remove_heap_is_wellformed_preserved: "type_wf h \ known_ptrs h \ heap_is_wellformed h \ h \ remove child \\<^sub>h h' \ heap_is_wellformed h'" assumes remove_child_removes_child: "heap_is_wellformed h \ h \ remove_child ptr' child \\<^sub>h h' \ h' \ get_child_nodes ptr \\<^sub>r children \ known_ptrs h \ type_wf h \ child \ set children" assumes remove_child_removes_first_child: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_child_nodes ptr \\<^sub>r node_ptr # children \ h \ remove_child ptr node_ptr \\<^sub>h h' \ h' \ get_child_nodes ptr \\<^sub>r children" assumes remove_removes_child: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_child_nodes ptr \\<^sub>r node_ptr # children \ h \ remove node_ptr \\<^sub>h h' \ h' \ get_child_nodes ptr \\<^sub>r children" assumes remove_for_all_empty_children: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_child_nodes ptr \\<^sub>r children \ h \ forall_M remove children \\<^sub>h h' \ h' \ get_child_nodes ptr \\<^sub>r []" interpretation i_remove_child_wf2?: l_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs set_child_nodes set_child_nodes_locs get_parent get_parent_locs get_owner_document get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs remove_child remove_child_locs remove type_wf known_ptr known_ptrs heap_is_wellformed parent_child_rel by unfold_locales declare l_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma remove_child_wf2_is_l_remove_child_wf2 [instances]: "l_remove_child_wf2 type_wf known_ptr known_ptrs remove_child heap_is_wellformed get_child_nodes remove" apply(auto simp add: l_remove_child_wf2_def l_remove_child_wf2_axioms_def instances)[1] using remove_child_heap_is_wellformed_preserved apply(fast, fast, fast) using remove_heap_is_wellformed_preserved apply(fast, fast, fast) using remove_child_removes_child apply fast using remove_child_removes_first_child apply fast using remove_removes_child apply fast using remove_for_all_empty_children apply fast done subsection \adopt\_node\ locale l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent_wf + l_get_owner_document_wf + l_remove_child_wf2 + l_heap_is_wellformed begin lemma adopt_node_removes_first_child: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ adopt_node owner_document node \\<^sub>h h'" assumes "h \ get_child_nodes ptr' \\<^sub>r node # children" shows "h' \ get_child_nodes ptr' \\<^sub>r children" proof - obtain old_document parent_opt h2 where old_document: "h \ get_owner_document (cast node) \\<^sub>r old_document" and parent_opt: "h \ get_parent node \\<^sub>r parent_opt" and h2: "h \ (case parent_opt of Some parent \ do { remove_child parent node } | None \ do { return ()}) \\<^sub>h h2" and h': "h2 \ (if owner_document \ old_document then do { old_disc_nodes \ get_disconnected_nodes old_document; set_disconnected_nodes old_document (remove1 node old_disc_nodes); disc_nodes \ get_disconnected_nodes owner_document; set_disconnected_nodes owner_document (node # disc_nodes) } else do { return () }) \\<^sub>h h'" using assms(4) by(auto simp add: adopt_node_def elim!: bind_returns_heap_E dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure] pure_returns_heap_eq[rotated, OF get_parent_pure]) have "h2 \ get_child_nodes ptr' \\<^sub>r children" using h2 remove_child_removes_first_child assms(1) assms(2) assms(3) assms(5) by (metis list.set_intros(1) local.child_parent_dual option.simps(5) parent_opt returns_result_eq) then show ?thesis using h' by(auto elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] dest!: reads_writes_separate_forwards[OF get_child_nodes_reads set_disconnected_nodes_writes] split: if_splits) qed lemma adopt_node_document_in_heap: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "h \ ok (adopt_node owner_document node)" shows "owner_document |\| document_ptr_kinds h" proof - obtain old_document parent_opt h2 h' where old_document: "h \ get_owner_document (cast node) \\<^sub>r old_document" and parent_opt: "h \ get_parent node \\<^sub>r parent_opt" and h2: "h \ (case parent_opt of Some parent \ do { remove_child parent node } | None \ do { return ()}) \\<^sub>h h2" and h': "h2 \ (if owner_document \ old_document then do { old_disc_nodes \ get_disconnected_nodes old_document; set_disconnected_nodes old_document (remove1 node old_disc_nodes); disc_nodes \ get_disconnected_nodes owner_document; set_disconnected_nodes owner_document (node # disc_nodes) } else do { return () }) \\<^sub>h h'" using assms(4) by(auto simp add: adopt_node_def elim!: bind_returns_heap_E dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure] pure_returns_heap_eq[rotated, OF get_parent_pure]) show ?thesis proof (cases "owner_document = old_document") case True then show ?thesis using old_document get_owner_document_owner_document_in_heap assms(1) assms(2) assms(3) by auto next case False then obtain h3 old_disc_nodes disc_nodes where old_disc_nodes: "h2 \ get_disconnected_nodes old_document \\<^sub>r old_disc_nodes" and h3: "h2 \ set_disconnected_nodes old_document (remove1 node old_disc_nodes) \\<^sub>h h3" and old_disc_nodes: "h3 \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes" and h': "h3 \ set_disconnected_nodes owner_document (node # disc_nodes) \\<^sub>h h'" using h' by(auto elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] ) then have "owner_document |\| document_ptr_kinds h3" by (meson is_OK_returns_result_I local.get_disconnected_nodes_ptr_in_heap) moreover have "object_ptr_kinds h = object_ptr_kinds h2" using h2 apply(simp split: option.splits) apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF remove_child_writes]) using remove_child_pointers_preserved by (auto simp add: reflp_def transp_def) moreover have "object_ptr_kinds h2 = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes h3]) using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) ultimately show ?thesis by(auto simp add: document_ptr_kinds_def) qed qed end locale l_adopt_node_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_root_node + l_get_owner_document_wf + l_remove_child_wf2 + l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma adopt_node_removes_child_step: assumes wellformed: "heap_is_wellformed h" and adopt_node: "h \ adopt_node owner_document node_ptr \\<^sub>h h2" and children: "h2 \ get_child_nodes ptr \\<^sub>r children" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "node_ptr \ set children" proof - obtain old_document parent_opt h' where old_document: "h \ get_owner_document (cast node_ptr) \\<^sub>r old_document" and parent_opt: "h \ get_parent node_ptr \\<^sub>r parent_opt" and h': "h \ (case parent_opt of Some parent \ remove_child parent node_ptr | None \ return () ) \\<^sub>h h'" using adopt_node get_parent_pure by(auto simp add: adopt_node_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated] bind_returns_heap_E2[rotated, OF get_parent_pure, rotated] bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] split: if_splits) then have "h' \ get_child_nodes ptr \\<^sub>r children" using adopt_node apply(auto simp add: adopt_node_def dest!: bind_returns_heap_E3[rotated, OF old_document, rotated] bind_returns_heap_E3[rotated, OF parent_opt, rotated] elim!: bind_returns_heap_E4[rotated, OF h', rotated])[1] apply(auto split: if_splits elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated])[1] apply (simp add: set_disconnected_nodes_get_child_nodes children reads_writes_preserved[OF get_child_nodes_reads set_disconnected_nodes_writes]) using children by blast show ?thesis proof(insert parent_opt h', induct parent_opt) case None then show ?case using child_parent_dual wellformed known_ptrs type_wf \h' \ get_child_nodes ptr \\<^sub>r children\ returns_result_eq by fastforce next case (Some option) then show ?case using remove_child_removes_child \h' \ get_child_nodes ptr \\<^sub>r children\ known_ptrs type_wf wellformed by auto qed qed lemma adopt_node_removes_child: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "h \ adopt_node owner_document node_ptr \\<^sub>h h'" shows "\ptr' children'. h' \ get_child_nodes ptr' \\<^sub>r children' \ node_ptr \ set children'" using adopt_node_removes_child_step assms by blast lemma adopt_node_preserves_wellformedness: assumes "heap_is_wellformed h" and "h \ adopt_node document_ptr child \\<^sub>h h'" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "heap_is_wellformed h'" and "known_ptrs h'" and "type_wf h'" proof - obtain old_document parent_opt h2 where old_document: "h \ get_owner_document (cast child) \\<^sub>r old_document" and parent_opt: "h \ get_parent child \\<^sub>r parent_opt" and h2: "h \ (case parent_opt of Some parent \ remove_child parent child | None \ return ()) \\<^sub>h h2" and h': "h2 \ (if document_ptr \ old_document then do { old_disc_nodes \ get_disconnected_nodes old_document; set_disconnected_nodes old_document (remove1 child old_disc_nodes); disc_nodes \ get_disconnected_nodes document_ptr; set_disconnected_nodes document_ptr (child # disc_nodes) } else do { return () }) \\<^sub>h h'" using assms(2) by(auto simp add: adopt_node_def elim!: bind_returns_heap_E dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure] pure_returns_heap_eq[rotated, OF get_parent_pure]) have object_ptr_kinds_h_eq3: "object_ptr_kinds h = object_ptr_kinds h2" using h2 apply(simp split: option.splits) apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF remove_child_writes]) using remove_child_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h: "\ptrs. h \ object_ptr_kinds_M \\<^sub>r ptrs = h2 \ object_ptr_kinds_M \\<^sub>r ptrs" unfolding object_ptr_kinds_M_defs by simp then have object_ptr_kinds_eq_h: "|h \ object_ptr_kinds_M|\<^sub>r = |h2 \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq_h: "|h \ node_ptr_kinds_M|\<^sub>r = |h2 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have wellformed_h2: "heap_is_wellformed h2" using h2 remove_child_heap_is_wellformed_preserved known_ptrs type_wf by (metis (no_types, lifting) assms(1) option.case_eq_if pure_returns_heap_eq return_pure) have "type_wf h2" using h2 remove_child_preserves_type_wf known_ptrs type_wf by (metis (no_types, lifting) assms(1) option.case_eq_if pure_returns_heap_eq return_pure) have "known_ptrs h2" using h2 remove_child_preserves_known_ptrs known_ptrs type_wf by (metis (no_types, lifting) assms(1) option.case_eq_if pure_returns_heap_eq return_pure) have "heap_is_wellformed h' \ known_ptrs h' \ type_wf h'" proof(cases "document_ptr = old_document") case True then show ?thesis using h' wellformed_h2 \type_wf h2\ \known_ptrs h2\ by auto next case False then obtain h3 old_disc_nodes disc_nodes_document_ptr_h3 where docs_neq: "document_ptr \ old_document" and old_disc_nodes: "h2 \ get_disconnected_nodes old_document \\<^sub>r old_disc_nodes" and h3: "h2 \ set_disconnected_nodes old_document (remove1 child old_disc_nodes) \\<^sub>h h3" and disc_nodes_document_ptr_h3: "h3 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_document_ptr_h3" and h': "h3 \ set_disconnected_nodes document_ptr (child # disc_nodes_document_ptr_h3) \\<^sub>h h'" using h' by(auto elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] ) have object_ptr_kinds_h2_eq3: "object_ptr_kinds h2 = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes h3]) using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h2: "\ptrs. h2 \ object_ptr_kinds_M \\<^sub>r ptrs = h3 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_eq_h2: "|h2 \ object_ptr_kinds_M|\<^sub>r = |h3 \ object_ptr_kinds_M|\<^sub>r" by(simp) then have node_ptr_kinds_eq_h2: "|h2 \ node_ptr_kinds_M|\<^sub>r = |h3 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast then have node_ptr_kinds_eq3_h2: "node_ptr_kinds h2 = node_ptr_kinds h3" by auto have document_ptr_kinds_eq2_h2: "|h2 \ document_ptr_kinds_M|\<^sub>r = |h3 \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_eq_h2 document_ptr_kinds_M_eq by auto then have document_ptr_kinds_eq3_h2: "document_ptr_kinds h2 = document_ptr_kinds h3" using object_ptr_kinds_eq_h2 document_ptr_kinds_M_eq by auto have children_eq_h2: "\ptr children. h2 \ get_child_nodes ptr \\<^sub>r children = h3 \ get_child_nodes ptr \\<^sub>r children" using get_child_nodes_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by (simp add: set_disconnected_nodes_get_child_nodes) then have children_eq2_h2: "\ptr. |h2 \ get_child_nodes ptr|\<^sub>r = |h3 \ get_child_nodes ptr|\<^sub>r" using select_result_eq by force have object_ptr_kinds_h3_eq3: "object_ptr_kinds h3 = object_ptr_kinds h'" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes h']) using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h3: "\ptrs. h3 \ object_ptr_kinds_M \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_eq_h3: "|h3 \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" by(simp) then have node_ptr_kinds_eq_h3: "|h3 \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast then have node_ptr_kinds_eq3_h3: "node_ptr_kinds h3 = node_ptr_kinds h'" by auto have document_ptr_kinds_eq2_h3: "|h3 \ document_ptr_kinds_M|\<^sub>r = |h' \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_eq_h3 document_ptr_kinds_M_eq by auto then have document_ptr_kinds_eq3_h3: "document_ptr_kinds h3 = document_ptr_kinds h'" using object_ptr_kinds_eq_h3 document_ptr_kinds_M_eq by auto have children_eq_h3: "\ptr children. h3 \ get_child_nodes ptr \\<^sub>r children = h' \ get_child_nodes ptr \\<^sub>r children" using get_child_nodes_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by (simp add: set_disconnected_nodes_get_child_nodes) then have children_eq2_h3: "\ptr. |h3 \ get_child_nodes ptr|\<^sub>r = |h' \ get_child_nodes ptr|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. old_document \ doc_ptr \ h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by (simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers) then have disconnected_nodes_eq2_h2: "\doc_ptr. old_document \ doc_ptr \ |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force obtain disc_nodes_old_document_h2 where disc_nodes_old_document_h2: "h2 \ get_disconnected_nodes old_document \\<^sub>r disc_nodes_old_document_h2" using old_disc_nodes by blast then have disc_nodes_old_document_h3: "h3 \ get_disconnected_nodes old_document \\<^sub>r remove1 child disc_nodes_old_document_h2" using h3 old_disc_nodes returns_result_eq set_disconnected_nodes_get_disconnected_nodes by fastforce have "distinct disc_nodes_old_document_h2" using disc_nodes_old_document_h2 local.heap_is_wellformed_disconnected_nodes_distinct wellformed_h2 by blast have "type_wf h2" proof (insert h2, induct parent_opt) case None then show ?case using type_wf by simp next case (Some option) then show ?case using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF remove_child_writes] type_wf remove_child_types_preserved by (simp add: reflp_def transp_def) qed then have "type_wf h3" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h3] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) then have "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h'] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) have "known_ptrs h3" using known_ptrs local.known_ptrs_preserved object_ptr_kinds_h2_eq3 object_ptr_kinds_h_eq3 by blast then have "known_ptrs h'" using local.known_ptrs_preserved object_ptr_kinds_h3_eq3 by blast have disconnected_nodes_eq_h3: "\doc_ptr disc_nodes. document_ptr \ doc_ptr \ h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by (simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers) then have disconnected_nodes_eq2_h3: "\doc_ptr. document_ptr \ doc_ptr \ |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have disc_nodes_document_ptr_h2: "h2 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_document_ptr_h3" using disconnected_nodes_eq_h2 docs_neq disc_nodes_document_ptr_h3 by auto have disc_nodes_document_ptr_h': " h' \ get_disconnected_nodes document_ptr \\<^sub>r child # disc_nodes_document_ptr_h3" using h' disc_nodes_document_ptr_h3 using set_disconnected_nodes_get_disconnected_nodes by blast have document_ptr_in_heap: "document_ptr |\| document_ptr_kinds h2" using disc_nodes_document_ptr_h3 document_ptr_kinds_eq2_h2 get_disconnected_nodes_ok assms(1) unfolding heap_is_wellformed_def using disc_nodes_document_ptr_h2 get_disconnected_nodes_ptr_in_heap by blast have old_document_in_heap: "old_document |\| document_ptr_kinds h2" using disc_nodes_old_document_h3 document_ptr_kinds_eq2_h2 get_disconnected_nodes_ok assms(1) unfolding heap_is_wellformed_def using get_disconnected_nodes_ptr_in_heap old_disc_nodes by blast have "child \ set disc_nodes_old_document_h2" proof (insert parent_opt h2, induct parent_opt) case None then have "h = h2" by(auto) moreover have "a_owner_document_valid h" using assms(1) heap_is_wellformed_def by(simp add: heap_is_wellformed_def) ultimately show ?case using old_document disc_nodes_old_document_h2 None(1) child_parent_dual[OF assms(1)] in_disconnected_nodes_no_parent assms(1) known_ptrs type_wf by blast next case (Some option) then show ?case apply(simp split: option.splits) using assms(1) disc_nodes_old_document_h2 old_document remove_child_in_disconnected_nodes known_ptrs by blast qed have "child \ set (remove1 child disc_nodes_old_document_h2)" using disc_nodes_old_document_h3 h3 known_ptrs wellformed_h2 \distinct disc_nodes_old_document_h2\ by auto have "child \ set disc_nodes_document_ptr_h3" proof - have "a_distinct_lists h2" using heap_is_wellformed_def wellformed_h2 by blast then have 0: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>r) |h2 \ document_ptr_kinds_M|\<^sub>r))" by(simp add: a_distinct_lists_def) show ?thesis using distinct_concat_map_E(1)[OF 0] \child \ set disc_nodes_old_document_h2\ disc_nodes_old_document_h2 disc_nodes_document_ptr_h2 by (meson \type_wf h2\ docs_neq known_ptrs local.get_owner_document_disconnected_nodes local.known_ptrs_preserved object_ptr_kinds_h_eq3 returns_result_eq wellformed_h2) qed have child_in_heap: "child |\| node_ptr_kinds h" using get_owner_document_ptr_in_heap[OF is_OK_returns_result_I[OF old_document]] node_ptr_kinds_commutes by blast have "a_acyclic_heap h2" using wellformed_h2 by (simp add: heap_is_wellformed_def) have "parent_child_rel h' \ parent_child_rel h2" proof fix x assume "x \ parent_child_rel h'" then show "x \ parent_child_rel h2" using object_ptr_kinds_h2_eq3 object_ptr_kinds_h3_eq3 children_eq2_h2 children_eq2_h3 mem_Collect_eq object_ptr_kinds_M_eq_h3 select_result_eq split_cong unfolding parent_child_rel_def by(simp) qed then have "a_acyclic_heap h'" using \a_acyclic_heap h2\ acyclic_heap_def acyclic_subset by blast moreover have "a_all_ptrs_in_heap h2" using wellformed_h2 by (simp add: heap_is_wellformed_def) then have "a_all_ptrs_in_heap h3" apply(auto simp add: a_all_ptrs_in_heap_def node_ptr_kinds_eq3_h2 children_eq_h2)[1] apply (simp add: children_eq2_h2 object_ptr_kinds_h2_eq3 subset_code(1)) by (metis (no_types, lifting) \child \ set disc_nodes_old_document_h2\ \type_wf h2\ disc_nodes_old_document_h2 disc_nodes_old_document_h3 disconnected_nodes_eq2_h2 document_ptr_kinds_eq3_h2 in_set_remove1 local.get_disconnected_nodes_ok local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq3_h2 returns_result_select_result select_result_I2 wellformed_h2) then have "a_all_ptrs_in_heap h'" apply(auto simp add: a_all_ptrs_in_heap_def node_ptr_kinds_eq3_h3 children_eq_h3)[1] apply (simp add: children_eq2_h3 object_ptr_kinds_h3_eq3 subset_code(1)) by (metis (no_types, lifting) \child \ set disc_nodes_old_document_h2\ disc_nodes_document_ptr_h' disc_nodes_document_ptr_h2 disc_nodes_old_document_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq3_h3 finite_set_in local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq3_h2 node_ptr_kinds_eq3_h3 select_result_I2 set_ConsD subset_code(1) wellformed_h2) moreover have "a_owner_document_valid h2" using wellformed_h2 by (simp add: heap_is_wellformed_def) then have "a_owner_document_valid h'" apply(simp add: a_owner_document_valid_def node_ptr_kinds_eq_h2 node_ptr_kinds_eq3_h3 object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 children_eq2_h2 children_eq2_h3 ) by (smt disc_nodes_document_ptr_h' disc_nodes_document_ptr_h2 disc_nodes_old_document_h2 disc_nodes_old_document_h3 disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_in_heap document_ptr_kinds_eq3_h2 document_ptr_kinds_eq3_h3 in_set_remove1 list.set_intros(1) node_ptr_kinds_eq3_h2 node_ptr_kinds_eq3_h3 object_ptr_kinds_h2_eq3 object_ptr_kinds_h3_eq3 select_result_I2 set_subset_Cons subset_code(1)) have a_distinct_lists_h2: "a_distinct_lists h2" using wellformed_h2 by (simp add: heap_is_wellformed_def) then have "a_distinct_lists h'" apply(auto simp add: a_distinct_lists_def object_ptr_kinds_eq_h3 object_ptr_kinds_eq_h2 children_eq2_h2 children_eq2_h3)[1] proof - assume 1: "distinct (concat (map (\ptr. |h' \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h')))))" and 2: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h2)))))" and 3: "(\x\fset (object_ptr_kinds h'). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h2). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" show "distinct (concat (map (\document_ptr. |h' \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h')))))" proof(rule distinct_concat_map_I) show "distinct (sorted_list_of_set (fset (document_ptr_kinds h')))" by(auto simp add: document_ptr_kinds_M_def ) next fix x assume a1: "x \ set (sorted_list_of_set (fset (document_ptr_kinds h')))" have 4: "distinct |h2 \ get_disconnected_nodes x|\<^sub>r" using a_distinct_lists_h2 "2" a1 concat_map_all_distinct document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 by fastforce then show "distinct |h' \ get_disconnected_nodes x|\<^sub>r" proof (cases "old_document \ x") case True then show ?thesis proof (cases "document_ptr \ x") case True then show ?thesis using disconnected_nodes_eq2_h2[OF \old_document \ x\] disconnected_nodes_eq2_h3[OF \document_ptr \ x\] 4 by(auto) next case False then show ?thesis using disc_nodes_document_ptr_h3 disc_nodes_document_ptr_h' 4 \child \ set disc_nodes_document_ptr_h3\ by(auto simp add: disconnected_nodes_eq2_h2[OF \old_document \ x\] ) qed next case False then show ?thesis by (metis (no_types, hide_lams) \distinct disc_nodes_old_document_h2\ disc_nodes_old_document_h3 disconnected_nodes_eq2_h3 distinct_remove1 docs_neq select_result_I2) qed next fix x y assume a0: "x \ set (sorted_list_of_set (fset (document_ptr_kinds h')))" and a1: "y \ set (sorted_list_of_set (fset (document_ptr_kinds h')))" and a2: "x \ y" moreover have 5: "set |h2 \ get_disconnected_nodes x|\<^sub>r \ set |h2 \ get_disconnected_nodes y|\<^sub>r = {}" using 2 calculation by (auto simp add: document_ptr_kinds_eq3_h2 document_ptr_kinds_eq3_h3 dest: distinct_concat_map_E(1)) ultimately show "set |h' \ get_disconnected_nodes x|\<^sub>r \ set |h' \ get_disconnected_nodes y|\<^sub>r = {}" proof(cases "old_document = x") case True have "old_document \ y" using \x \ y\ \old_document = x\ by simp have "document_ptr \ x" using docs_neq \old_document = x\ by auto show ?thesis proof(cases "document_ptr = y") case True then show ?thesis using 5 True select_result_I2[OF disc_nodes_document_ptr_h'] select_result_I2[OF disc_nodes_document_ptr_h2] select_result_I2[OF disc_nodes_old_document_h2] select_result_I2[OF disc_nodes_old_document_h3] \old_document = x\ by (metis (no_types, lifting) \child \ set (remove1 child disc_nodes_old_document_h2)\ \document_ptr \ x\ disconnected_nodes_eq2_h3 disjoint_iff_not_equal notin_set_remove1 set_ConsD) next case False then show ?thesis using 5 select_result_I2[OF disc_nodes_document_ptr_h'] select_result_I2[OF disc_nodes_document_ptr_h2] select_result_I2[OF disc_nodes_old_document_h2] select_result_I2[OF disc_nodes_old_document_h3] disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 \old_document = x\ docs_neq \old_document \ y\ by (metis (no_types, lifting) disjoint_iff_not_equal notin_set_remove1) qed next case False then show ?thesis proof(cases "old_document = y") case True then show ?thesis proof(cases "document_ptr = x") case True show ?thesis using 5 select_result_I2[OF disc_nodes_document_ptr_h'] select_result_I2[OF disc_nodes_document_ptr_h2] select_result_I2[OF disc_nodes_old_document_h2] select_result_I2[OF disc_nodes_old_document_h3] \old_document \ x\ \old_document = y\ \document_ptr = x\ apply(simp) by (metis (no_types, lifting) \child \ set (remove1 child disc_nodes_old_document_h2)\ disconnected_nodes_eq2_h3 disjoint_iff_not_equal notin_set_remove1) next case False then show ?thesis using 5 select_result_I2[OF disc_nodes_document_ptr_h'] select_result_I2[OF disc_nodes_document_ptr_h2] select_result_I2[OF disc_nodes_old_document_h2] select_result_I2[OF disc_nodes_old_document_h3] \old_document \ x\ \old_document = y\ \document_ptr \ x\ by (metis (no_types, lifting) disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disjoint_iff_not_equal docs_neq notin_set_remove1) qed next case False have "set |h2 \ get_disconnected_nodes y|\<^sub>r \ set disc_nodes_old_document_h2 = {}" by (metis DocumentMonad.ptr_kinds_M_ok DocumentMonad.ptr_kinds_M_ptr_kinds False \type_wf h2\ a1 disc_nodes_old_document_h2 document_ptr_kinds_M_def document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 l_ptr_kinds_M.ptr_kinds_ptr_kinds_M local.get_disconnected_nodes_ok local.heap_is_wellformed_one_disc_parent returns_result_select_result wellformed_h2) then show ?thesis proof(cases "document_ptr = x") case True then have "document_ptr \ y" using \x \ y\ by auto have "set |h2 \ get_disconnected_nodes y|\<^sub>r \ set disc_nodes_old_document_h2 = {}" using \set |h2 \ get_disconnected_nodes y|\<^sub>r \ set disc_nodes_old_document_h2 = {}\ by blast then show ?thesis using 5 select_result_I2[OF disc_nodes_document_ptr_h'] select_result_I2[OF disc_nodes_document_ptr_h2] select_result_I2[OF disc_nodes_old_document_h2] select_result_I2[OF disc_nodes_old_document_h3] \old_document \ x\ \old_document \ y\ \document_ptr = x\ \document_ptr \ y\ \child \ set disc_nodes_old_document_h2\ disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 \set |h2 \ get_disconnected_nodes y|\<^sub>r \ set disc_nodes_old_document_h2 = {}\ by(auto) next case False then show ?thesis proof(cases "document_ptr = y") case True have f1: "set |h2 \ get_disconnected_nodes x|\<^sub>r \ set disc_nodes_document_ptr_h3 = {}" using 2 a1 document_ptr_in_heap document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 \document_ptr \ x\ select_result_I2[OF disc_nodes_document_ptr_h3, symmetric] disconnected_nodes_eq2_h2[OF docs_neq[symmetric], symmetric] by (simp add: "5" True) moreover have f1: "set |h2 \ get_disconnected_nodes x|\<^sub>r \ set |h2 \ get_disconnected_nodes old_document|\<^sub>r = {}" using 2 a1 old_document_in_heap document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 \old_document \ x\ by (metis (no_types, lifting) a0 distinct_concat_map_E(1) document_ptr_kinds_eq3_h2 document_ptr_kinds_eq3_h3 finite_fset fmember.rep_eq set_sorted_list_of_set) ultimately show ?thesis using 5 select_result_I2[OF disc_nodes_document_ptr_h'] select_result_I2[OF disc_nodes_old_document_h2] \old_document \ x\ \document_ptr \ x\ \document_ptr = y\ \child \ set disc_nodes_old_document_h2\ disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 by auto next case False then show ?thesis using 5 select_result_I2[OF disc_nodes_old_document_h2] \old_document \ x\ \document_ptr \ x\ \document_ptr \ y\ \child \ set disc_nodes_old_document_h2\ disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 by (metis \set |h2 \ get_disconnected_nodes y|\<^sub>r \ set disc_nodes_old_document_h2 = {}\ empty_iff inf.idem) qed qed qed qed qed next fix x xa xb assume 0: "distinct (concat (map (\ptr. |h' \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h')))))" and 1: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h2)))))" and 2: "(\x\fset (object_ptr_kinds h'). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h2). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" and 3: "xa |\| object_ptr_kinds h'" and 4: "x \ set |h' \ get_child_nodes xa|\<^sub>r" and 5: "xb |\| document_ptr_kinds h'" and 6: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" then show False using \child \ set disc_nodes_old_document_h2\ disc_nodes_document_ptr_h' disc_nodes_document_ptr_h2 disc_nodes_old_document_h2 disc_nodes_old_document_h3 disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 old_document_in_heap apply(auto)[1] apply(cases "xb = old_document") proof - assume a1: "xb = old_document" assume a2: "h2 \ get_disconnected_nodes old_document \\<^sub>r disc_nodes_old_document_h2" assume a3: "h3 \ get_disconnected_nodes old_document \\<^sub>r remove1 child disc_nodes_old_document_h2" assume a4: "x \ set |h' \ get_child_nodes xa|\<^sub>r" assume "document_ptr_kinds h2 = document_ptr_kinds h'" assume a5: "(\x\fset (object_ptr_kinds h'). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h'). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" have f6: "old_document |\| document_ptr_kinds h'" using a1 \xb |\| document_ptr_kinds h'\ by blast have f7: "|h2 \ get_disconnected_nodes old_document|\<^sub>r = disc_nodes_old_document_h2" using a2 by simp have "x \ set disc_nodes_old_document_h2" using f6 a3 a1 by (metis (no_types) \type_wf h'\ \x \ set |h' \ get_disconnected_nodes xb|\<^sub>r\ disconnected_nodes_eq_h3 docs_neq get_disconnected_nodes_ok returns_result_eq returns_result_select_result set_remove1_subset subsetCE) then have "set |h' \ get_child_nodes xa|\<^sub>r \ set |h2 \ get_disconnected_nodes xb|\<^sub>r = {}" using f7 f6 a5 a4 \xa |\| object_ptr_kinds h'\ by fastforce then show ?thesis using \x \ set disc_nodes_old_document_h2\ a1 a4 f7 by blast next assume a1: "xb \ old_document" assume a2: "h2 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_document_ptr_h3" assume a3: "h2 \ get_disconnected_nodes old_document \\<^sub>r disc_nodes_old_document_h2" assume a4: "xa |\| object_ptr_kinds h'" assume a5: "h' \ get_disconnected_nodes document_ptr \\<^sub>r child # disc_nodes_document_ptr_h3" assume a6: "old_document |\| document_ptr_kinds h'" assume a7: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" assume a8: "x \ set |h' \ get_child_nodes xa|\<^sub>r" assume a9: "document_ptr_kinds h2 = document_ptr_kinds h'" assume a10: "\doc_ptr. old_document \ doc_ptr \ |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" assume a11: "\doc_ptr. document_ptr \ doc_ptr \ |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" assume a12: "(\x\fset (object_ptr_kinds h'). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h'). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" have f13: "\d. d \ set |h' \ document_ptr_kinds_M|\<^sub>r \ h2 \ ok get_disconnected_nodes d" using a9 \type_wf h2\ get_disconnected_nodes_ok by simp then have f14: "|h2 \ get_disconnected_nodes old_document|\<^sub>r = disc_nodes_old_document_h2" using a6 a3 by simp have "x \ set |h2 \ get_disconnected_nodes xb|\<^sub>r" using a12 a8 a4 \xb |\| document_ptr_kinds h'\ by (meson UN_I disjoint_iff_not_equal fmember.rep_eq) then have "x = child" using f13 a11 a10 a7 a5 a2 a1 by (metis (no_types, lifting) select_result_I2 set_ConsD) then have "child \ set disc_nodes_old_document_h2" using f14 a12 a8 a6 a4 by (metis \type_wf h'\ adopt_node_removes_child assms(1) assms(2) type_wf get_child_nodes_ok known_ptrs local.known_ptrs_known_ptr object_ptr_kinds_h2_eq3 object_ptr_kinds_h3_eq3 object_ptr_kinds_h_eq3 returns_result_select_result) then show ?thesis using \child \ set disc_nodes_old_document_h2\ by fastforce qed qed ultimately show ?thesis using \type_wf h'\ \known_ptrs h'\ \a_owner_document_valid h'\ heap_is_wellformed_def by blast qed then show "heap_is_wellformed h'" and "known_ptrs h'" and "type_wf h'" by auto qed lemma adopt_node_node_in_disconnected_nodes: assumes wellformed: "heap_is_wellformed h" and adopt_node: "h \ adopt_node owner_document node_ptr \\<^sub>h h'" and "h' \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "node_ptr \ set disc_nodes" proof - obtain old_document parent_opt h2 where old_document: "h \ get_owner_document (cast node_ptr) \\<^sub>r old_document" and parent_opt: "h \ get_parent node_ptr \\<^sub>r parent_opt" and h2: "h \ (case parent_opt of Some parent \ remove_child parent node_ptr | None \ return ()) \\<^sub>h h2" and h': "h2 \ (if owner_document \ old_document then do { old_disc_nodes \ get_disconnected_nodes old_document; set_disconnected_nodes old_document (remove1 node_ptr old_disc_nodes); disc_nodes \ get_disconnected_nodes owner_document; set_disconnected_nodes owner_document (node_ptr # disc_nodes) } else do { return () }) \\<^sub>h h'" using assms(2) by(auto simp add: adopt_node_def elim!: bind_returns_heap_E dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure] pure_returns_heap_eq[rotated, OF get_parent_pure]) show ?thesis proof (cases "owner_document = old_document") case True then show ?thesis proof (insert parent_opt h2, induct parent_opt) case None then have "h = h'" using h2 h' by(auto) then show ?case using in_disconnected_nodes_no_parent assms None old_document by blast next case (Some parent) then show ?case using remove_child_in_disconnected_nodes known_ptrs True h' assms(3) old_document by auto qed next case False then show ?thesis using assms(3) h' list.set_intros(1) select_result_I2 set_disconnected_nodes_get_disconnected_nodes apply(auto elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated])[1] proof - fix x and h'a and xb assume a1: "h' \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes" assume a2: "\h document_ptr disc_nodes h'. h \ set_disconnected_nodes document_ptr disc_nodes \\<^sub>h h' \ h' \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes" assume "h'a \ set_disconnected_nodes owner_document (node_ptr # xb) \\<^sub>h h'" then have "node_ptr # xb = disc_nodes" using a2 a1 by (meson returns_result_eq) then show ?thesis by (meson list.set_intros(1)) qed qed qed end interpretation i_adopt_node_wf?: l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_owner_document get_parent get_parent_locs remove_child remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs adopt_node adopt_node_locs known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs set_child_nodes set_child_nodes_locs remove heap_is_wellformed parent_child_rel by(simp add: l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] interpretation i_adopt_node_wf2?: l_adopt_node_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_owner_document get_parent get_parent_locs remove_child remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs adopt_node adopt_node_locs known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs set_child_nodes set_child_nodes_locs remove heap_is_wellformed parent_child_rel get_root_node get_root_node_locs by(simp add: l_adopt_node_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_adopt_node_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_adopt_node_wf = l_heap_is_wellformed + l_known_ptrs + l_type_wf + l_adopt_node_defs + l_get_child_nodes_defs + l_get_disconnected_nodes_defs + assumes adopt_node_preserves_wellformedness: "heap_is_wellformed h \ h \ adopt_node document_ptr child \\<^sub>h h' \ known_ptrs h \ type_wf h \ heap_is_wellformed h'" assumes adopt_node_removes_child: "heap_is_wellformed h \ h \ adopt_node owner_document node_ptr \\<^sub>h h2 \ h2 \ get_child_nodes ptr \\<^sub>r children \ known_ptrs h \ type_wf h \ node_ptr \ set children" assumes adopt_node_node_in_disconnected_nodes: "heap_is_wellformed h \ h \ adopt_node owner_document node_ptr \\<^sub>h h' \ h' \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes \ known_ptrs h \ type_wf h \ node_ptr \ set disc_nodes" assumes adopt_node_removes_first_child: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ adopt_node owner_document node \\<^sub>h h' \ h \ get_child_nodes ptr' \\<^sub>r node # children \ h' \ get_child_nodes ptr' \\<^sub>r children" assumes adopt_node_document_in_heap: "heap_is_wellformed h \ known_ptrs h \ type_wf h \ h \ ok (adopt_node owner_document node) \ owner_document |\| document_ptr_kinds h" lemma adopt_node_wf_is_l_adopt_node_wf [instances]: "l_adopt_node_wf type_wf known_ptr heap_is_wellformed parent_child_rel get_child_nodes get_disconnected_nodes known_ptrs adopt_node" using heap_is_wellformed_is_l_heap_is_wellformed known_ptrs_is_l_known_ptrs apply(auto simp add: l_adopt_node_wf_def l_adopt_node_wf_axioms_def)[1] using adopt_node_preserves_wellformedness apply blast using adopt_node_removes_child apply blast using adopt_node_node_in_disconnected_nodes apply blast using adopt_node_removes_first_child apply blast using adopt_node_document_in_heap apply blast done subsection \insert\_before\ locale l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_adopt_node_wf + l_set_disconnected_nodes_get_child_nodes + l_heap_is_wellformed begin lemma insert_before_removes_child: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "ptr \ ptr'" assumes "h \ insert_before ptr node child \\<^sub>h h'" assumes "h \ get_child_nodes ptr' \\<^sub>r node # children" shows "h' \ get_child_nodes ptr' \\<^sub>r children" proof - obtain owner_document h2 h3 disc_nodes reference_child where "h \ (if Some node = child then a_next_sibling node else return child) \\<^sub>r reference_child" and "h \ get_owner_document ptr \\<^sub>r owner_document" and h2: "h \ adopt_node owner_document node \\<^sub>h h2" and "h2 \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes" and h3: "h2 \ set_disconnected_nodes owner_document (remove1 node disc_nodes) \\<^sub>h h3" and h': "h3 \ a_insert_node ptr node reference_child \\<^sub>h h'" using assms(5) by(auto simp add: insert_before_def a_ensure_pre_insertion_validity_def elim!: bind_returns_heap_E bind_returns_result_E bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated] bind_returns_heap_E2[rotated, OF get_parent_pure, rotated] bind_returns_heap_E2[rotated, OF get_ancestors_pure, rotated] bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated] bind_returns_heap_E2[rotated, OF next_sibling_pure, rotated] bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] split: if_splits option.splits) have "h2 \ get_child_nodes ptr' \\<^sub>r children" using h2 adopt_node_removes_first_child assms(1) assms(2) assms(3) assms(6) by simp then have "h3 \ get_child_nodes ptr' \\<^sub>r children" using h3 by(auto simp add: set_disconnected_nodes_get_child_nodes dest!: reads_writes_separate_forwards[OF get_child_nodes_reads set_disconnected_nodes_writes]) then show ?thesis using h' assms(4) apply(auto simp add: a_insert_node_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated])[1] by(auto simp add: set_child_nodes_get_child_nodes_different_pointers elim!: reads_writes_separate_forwards[OF get_child_nodes_reads set_child_nodes_writes]) qed end locale l_insert_before_wf = l_heap_is_wellformed_defs + l_type_wf + l_known_ptrs + l_insert_before_defs + l_get_child_nodes_defs + assumes insert_before_removes_child: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ ptr \ ptr' \ h \ insert_before ptr node child \\<^sub>h h' \ h \ get_child_nodes ptr' \\<^sub>r node # children \ h' \ get_child_nodes ptr' \\<^sub>r children" interpretation i_insert_before_wf?: l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_parent get_parent_locs get_child_nodes get_child_nodes_locs set_child_nodes set_child_nodes_locs get_ancestors get_ancestors_locs adopt_node adopt_node_locs set_disconnected_nodes set_disconnected_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_owner_document insert_before insert_before_locs append_child type_wf known_ptr known_ptrs heap_is_wellformed parent_child_rel by(simp add: l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma insert_before_wf_is_l_insert_before_wf [instances]: "l_insert_before_wf heap_is_wellformed type_wf known_ptr known_ptrs insert_before get_child_nodes" apply(auto simp add: l_insert_before_wf_def l_insert_before_wf_axioms_def instances)[1] using insert_before_removes_child apply fast done locale l_insert_before_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_child_nodes_get_disconnected_nodes + l_remove_child + l_get_root_node_wf + l_set_disconnected_nodes_get_disconnected_nodes_wf + l_set_disconnected_nodes_get_ancestors + l_get_ancestors_wf + l_get_owner_document + l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma insert_before_heap_is_wellformed_preserved: assumes wellformed: "heap_is_wellformed h" and insert_before: "h \ insert_before ptr node child \\<^sub>h h'" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "heap_is_wellformed h'" and "type_wf h'" and "known_ptrs h'" proof - obtain ancestors reference_child owner_document h2 h3 disconnected_nodes_h2 where ancestors: "h \ get_ancestors ptr \\<^sub>r ancestors" and node_not_in_ancestors: "cast node \ set ancestors" and reference_child: "h \ (if Some node = child then a_next_sibling node else return child) \\<^sub>r reference_child" and owner_document: "h \ get_owner_document ptr \\<^sub>r owner_document" and h2: "h \ adopt_node owner_document node \\<^sub>h h2" and disconnected_nodes_h2: "h2 \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h2" and h3: "h2 \ set_disconnected_nodes owner_document (remove1 node disconnected_nodes_h2) \\<^sub>h h3" and h': "h3 \ a_insert_node ptr node reference_child \\<^sub>h h'" using assms(2) by(auto simp add: insert_before_def a_ensure_pre_insertion_validity_def elim!: bind_returns_heap_E bind_returns_result_E bind_returns_heap_E2[rotated, OF get_parent_pure, rotated] bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated] bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] bind_returns_heap_E2[rotated, OF get_ancestors_pure, rotated] bind_returns_heap_E2[rotated, OF next_sibling_pure, rotated] bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated] split: if_splits option.splits) have "known_ptr ptr" by (meson get_owner_document_ptr_in_heap is_OK_returns_result_I known_ptrs l_known_ptrs.known_ptrs_known_ptr l_known_ptrs_axioms owner_document) have "type_wf h2" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF adopt_node_writes h2] using type_wf adopt_node_types_preserved by(auto simp add: a_remove_child_locs_def reflp_def transp_def) then have "type_wf h3" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h3] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) then show "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF insert_node_writes h'] using set_child_nodes_types_preserved by(auto simp add: reflp_def transp_def) have object_ptr_kinds_M_eq3_h: "object_ptr_kinds h = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF adopt_node_writes h2]) using adopt_node_pointers_preserved apply blast by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h: "\ptrs. h \ object_ptr_kinds_M \\<^sub>r ptrs = h2 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs ) then have object_ptr_kinds_M_eq2_h: "|h \ object_ptr_kinds_M|\<^sub>r = |h2 \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h: "|h \ node_ptr_kinds_M|\<^sub>r = |h2 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have "known_ptrs h2" using known_ptrs object_ptr_kinds_M_eq3_h known_ptrs_preserved by blast have wellformed_h2: "heap_is_wellformed h2" using adopt_node_preserves_wellformedness[OF wellformed h2] known_ptrs type_wf . have object_ptr_kinds_M_eq3_h2: "object_ptr_kinds h2 = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes h3]) unfolding a_remove_child_locs_def using set_disconnected_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h2: "\ptrs. h2 \ object_ptr_kinds_M \\<^sub>r ptrs = h3 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_M_eq2_h2: "|h2 \ object_ptr_kinds_M|\<^sub>r = |h3 \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h2: "|h2 \ node_ptr_kinds_M|\<^sub>r = |h3 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have document_ptr_kinds_eq2_h2: "|h2 \ document_ptr_kinds_M|\<^sub>r = |h3 \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_M_eq2_h2 document_ptr_kinds_M_eq by auto have "known_ptrs h3" using object_ptr_kinds_M_eq3_h2 known_ptrs_preserved \known_ptrs h2\ by blast have object_ptr_kinds_M_eq3_h': "object_ptr_kinds h3 = object_ptr_kinds h'" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF insert_node_writes h']) unfolding a_remove_child_locs_def using set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h3: "\ptrs. h3 \ object_ptr_kinds_M \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_M_eq2_h3: "|h3 \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h3: "|h3 \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have document_ptr_kinds_eq2_h3: "|h3 \ document_ptr_kinds_M|\<^sub>r = |h' \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_M_eq2_h3 document_ptr_kinds_M_eq by auto show "known_ptrs h'" using object_ptr_kinds_M_eq3_h' known_ptrs_preserved \known_ptrs h3\ by blast have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. owner_document \ doc_ptr \ h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by (auto simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers) then have disconnected_nodes_eq2_h2: "\doc_ptr. doc_ptr \ owner_document \ |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have disconnected_nodes_h3: "h3 \ get_disconnected_nodes owner_document \\<^sub>r remove1 node disconnected_nodes_h2" using h3 set_disconnected_nodes_get_disconnected_nodes by blast have disconnected_nodes_eq_h3: "\doc_ptr disc_nodes. h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads insert_node_writes h' apply(rule reads_writes_preserved) using set_child_nodes_get_disconnected_nodes by fast then have disconnected_nodes_eq2_h3: "\doc_ptr. |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. h2 \ get_child_nodes ptr' \\<^sub>r children = h3 \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by (auto simp add: set_disconnected_nodes_get_child_nodes) then have children_eq2_h2: "\ptr'. |h2 \ get_child_nodes ptr'|\<^sub>r = |h3 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have children_eq_h3: "\ptr' children. ptr \ ptr' \ h3 \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads insert_node_writes h' apply(rule reads_writes_preserved) by (auto simp add: set_child_nodes_get_child_nodes_different_pointers) then have children_eq2_h3: "\ptr'. ptr \ ptr' \ |h3 \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force obtain children_h3 where children_h3: "h3 \ get_child_nodes ptr \\<^sub>r children_h3" using h' a_insert_node_def by auto have children_h': "h' \ get_child_nodes ptr \\<^sub>r insert_before_list node reference_child children_h3" using h' \type_wf h3\ \known_ptr ptr\ by(auto simp add: a_insert_node_def elim!: bind_returns_heap_E2 dest!: set_child_nodes_get_child_nodes returns_result_eq[OF children_h3]) have ptr_in_heap: "ptr |\| object_ptr_kinds h3" using children_h3 get_child_nodes_ptr_in_heap by blast have node_in_heap: "node |\| node_ptr_kinds h" using h2 adopt_node_child_in_heap by fast have child_not_in_any_children: "\p children. h2 \ get_child_nodes p \\<^sub>r children \ node \ set children" using wellformed h2 adopt_node_removes_child \type_wf h\ \known_ptrs h\ by auto have "node \ set disconnected_nodes_h2" using disconnected_nodes_h2 h2 adopt_node_node_in_disconnected_nodes assms(1) \type_wf h\ \known_ptrs h\ by blast have node_not_in_disconnected_nodes: "\d. d |\| document_ptr_kinds h3 \ node \ set |h3 \ get_disconnected_nodes d|\<^sub>r" proof - fix d assume "d |\| document_ptr_kinds h3" show "node \ set |h3 \ get_disconnected_nodes d|\<^sub>r" proof (cases "d = owner_document") case True then show ?thesis using disconnected_nodes_h2 wellformed_h2 h3 remove_from_disconnected_nodes_removes wellformed_h2 \d |\| document_ptr_kinds h3\ disconnected_nodes_h3 by fastforce next case False then have "set |h2 \ get_disconnected_nodes d|\<^sub>r \ set |h2 \ get_disconnected_nodes owner_document|\<^sub>r = {}" using distinct_concat_map_E(1) wellformed_h2 by (metis (no_types, lifting) \d |\| document_ptr_kinds h3\ \type_wf h2\ disconnected_nodes_h2 document_ptr_kinds_M_def document_ptr_kinds_eq2_h2 l_ptr_kinds_M.ptr_kinds_ptr_kinds_M local.get_disconnected_nodes_ok local.heap_is_wellformed_one_disc_parent returns_result_select_result select_result_I2) then show ?thesis using disconnected_nodes_eq2_h2[OF False] \node \ set disconnected_nodes_h2\ disconnected_nodes_h2 by fastforce qed qed have "cast node \ ptr" using ancestors node_not_in_ancestors get_ancestors_ptr by fast obtain ancestors_h2 where ancestors_h2: "h2 \ get_ancestors ptr \\<^sub>r ancestors_h2" using get_ancestors_ok object_ptr_kinds_M_eq2_h2 \known_ptrs h2\ \type_wf h2\ by (metis is_OK_returns_result_E object_ptr_kinds_M_eq3_h2 ptr_in_heap wellformed_h2) have ancestors_h3: "h3 \ get_ancestors ptr \\<^sub>r ancestors_h2" using get_ancestors_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_separate_forwards) using \heap_is_wellformed h2\ ancestors_h2 by (auto simp add: set_disconnected_nodes_get_ancestors) have node_not_in_ancestors_h2: "cast node \ set ancestors_h2" apply(rule get_ancestors_remains_not_in_ancestors[OF assms(1) wellformed_h2 ancestors ancestors_h2]) using adopt_node_children_subset using h2 \known_ptrs h\ \ type_wf h\ apply(blast) using node_not_in_ancestors apply(blast) using object_ptr_kinds_M_eq3_h apply(blast) using \known_ptrs h\ apply(blast) using \type_wf h\ apply(blast) using \type_wf h2\ by blast moreover have "a_acyclic_heap h'" proof - have "acyclic (parent_child_rel h2)" using wellformed_h2 by (simp add: heap_is_wellformed_def acyclic_heap_def) then have "acyclic (parent_child_rel h3)" by(auto simp add: parent_child_rel_def object_ptr_kinds_M_eq3_h2 children_eq2_h2) moreover have "cast node \ {x. (x, ptr) \ (parent_child_rel h2)\<^sup>*}" using get_ancestors_parent_child_rel node_not_in_ancestors_h2 \known_ptrs h2\ \type_wf h2\ using ancestors_h2 wellformed_h2 by blast then have "cast node \ {x. (x, ptr) \ (parent_child_rel h3)\<^sup>*}" by(auto simp add: parent_child_rel_def object_ptr_kinds_M_eq3_h2 children_eq2_h2) moreover have "parent_child_rel h' = insert (ptr, cast node) ((parent_child_rel h3))" using children_h3 children_h' ptr_in_heap apply(auto simp add: parent_child_rel_def object_ptr_kinds_M_eq3_h' children_eq2_h3 insert_before_list_node_in_set)[1] apply (metis (no_types, lifting) children_eq2_h3 insert_before_list_in_set select_result_I2) by (metis (no_types, lifting) children_eq2_h3 imageI insert_before_list_in_set select_result_I2) ultimately show ?thesis by(auto simp add: acyclic_heap_def) qed moreover have "a_all_ptrs_in_heap h2" using wellformed_h2 by (simp add: heap_is_wellformed_def) have "a_all_ptrs_in_heap h'" proof - have "a_all_ptrs_in_heap h3" using \a_all_ptrs_in_heap h2\ apply(auto simp add: a_all_ptrs_in_heap_def object_ptr_kinds_M_eq2_h2 node_ptr_kinds_eq2_h2 children_eq_h2)[1] using disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3 using node_ptr_kinds_eq2_h2 apply auto[1] apply (metis \known_ptrs h2\ \type_wf h3\ children_eq_h2 local.get_child_nodes_ok local.heap_is_wellformed_children_in_heap local.known_ptrs_known_ptr object_ptr_kinds_M_eq3_h2 returns_result_select_result wellformed_h2) by (metis (no_types, lifting) disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3 document_ptr_kinds_commutes finite_set_in node_ptr_kinds_commutes object_ptr_kinds_M_eq3_h2 select_result_I2 set_remove1_subset subsetD) have "set children_h3 \ set |h' \ node_ptr_kinds_M|\<^sub>r" using children_h3 \a_all_ptrs_in_heap h3\ apply(auto simp add: a_all_ptrs_in_heap_def node_ptr_kinds_eq2_h3)[1] by (metis children_eq_h2 l_heap_is_wellformed.heap_is_wellformed_children_in_heap local.l_heap_is_wellformed_axioms node_ptr_kinds_commutes object_ptr_kinds_M_eq3_h' object_ptr_kinds_M_eq3_h2 wellformed_h2) then have "set (insert_before_list node reference_child children_h3) \ set |h' \ node_ptr_kinds_M|\<^sub>r" using node_in_heap apply(auto simp add: node_ptr_kinds_eq2_h node_ptr_kinds_eq2_h2 node_ptr_kinds_eq2_h3)[1] by (metis (no_types, hide_lams) contra_subsetD finite_set_in insert_before_list_in_set node_ptr_kinds_commutes object_ptr_kinds_M_eq3_h object_ptr_kinds_M_eq3_h' object_ptr_kinds_M_eq3_h2) then show ?thesis using \a_all_ptrs_in_heap h3\ apply(auto simp add: object_ptr_kinds_M_eq3_h' a_all_ptrs_in_heap_def node_ptr_kinds_def node_ptr_kinds_eq2_h3 disconnected_nodes_eq_h3)[1] using children_eq_h3 children_h' apply (metis (no_types, lifting) children_eq2_h3 finite_set_in select_result_I2 subsetD) by (metis (no_types) \type_wf h'\ disconnected_nodes_eq2_h3 disconnected_nodes_eq_h3 finite_set_in is_OK_returns_result_I local.get_disconnected_nodes_ok local.get_disconnected_nodes_ptr_in_heap returns_result_select_result subsetD) qed moreover have "a_distinct_lists h2" using wellformed_h2 by (simp add: heap_is_wellformed_def) then have "a_distinct_lists h3" proof(auto simp add: a_distinct_lists_def object_ptr_kinds_M_eq2_h2 document_ptr_kinds_eq2_h2 children_eq2_h2 intro!: distinct_concat_map_I)[1] fix x assume 1: "x |\| document_ptr_kinds h3" and 2: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" show "distinct |h3 \ get_disconnected_nodes x|\<^sub>r" using distinct_concat_map_E(2)[OF 2] select_result_I2[OF disconnected_nodes_h3] disconnected_nodes_eq2_h2 select_result_I2[OF disconnected_nodes_h2] 1 by (metis (full_types) distinct_remove1 finite_fset fmember.rep_eq set_sorted_list_of_set) next fix x y xa assume 1: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" and 2: "x |\| document_ptr_kinds h3" and 3: "y |\| document_ptr_kinds h3" and 4: "x \ y" and 5: "xa \ set |h3 \ get_disconnected_nodes x|\<^sub>r" and 6: "xa \ set |h3 \ get_disconnected_nodes y|\<^sub>r" show False proof (cases "x = owner_document") case True then have "y \ owner_document" using 4 by simp show ?thesis using distinct_concat_map_E(1)[OF 1] using 2 3 4 5 6 select_result_I2[OF disconnected_nodes_h3] select_result_I2[OF disconnected_nodes_h2] apply(auto simp add: True disconnected_nodes_eq2_h2[OF \y \ owner_document\])[1] by (metis (no_types, hide_lams) disconnected_nodes_eq2_h2 disjoint_iff_not_equal notin_set_remove1) next case False then show ?thesis proof (cases "y = owner_document") case True then show ?thesis using distinct_concat_map_E(1)[OF 1] using 2 3 4 5 6 select_result_I2[OF disconnected_nodes_h3] select_result_I2[OF disconnected_nodes_h2] apply(auto simp add: True disconnected_nodes_eq2_h2[OF \x \ owner_document\])[1] by (metis (no_types, hide_lams) disconnected_nodes_eq2_h2 disjoint_iff_not_equal notin_set_remove1) next case False then show ?thesis using distinct_concat_map_E(1)[OF 1, simplified, OF 2 3 4] 5 6 using disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3 disjoint_iff_not_equal finite_fset fmember.rep_eq notin_set_remove1 select_result_I2 set_sorted_list_of_set by (metis (no_types, lifting)) qed qed next fix x xa xb assume 1: "(\x\fset (object_ptr_kinds h3). set |h3 \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h3). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" and 2: "xa |\| object_ptr_kinds h3" and 3: "x \ set |h3 \ get_child_nodes xa|\<^sub>r" and 4: "xb |\| document_ptr_kinds h3" and 5: "x \ set |h3 \ get_disconnected_nodes xb|\<^sub>r" have 6: "set |h3 \ get_child_nodes xa|\<^sub>r \ set |h2 \ get_disconnected_nodes xb|\<^sub>r = {}" using 1 2 4 by (metis \type_wf h2\ children_eq2_h2 document_ptr_kinds_commutes known_ptrs local.get_child_nodes_ok local.get_disconnected_nodes_ok local.heap_is_wellformed_children_disc_nodes_different local.known_ptrs_known_ptr object_ptr_kinds_M_eq3_h object_ptr_kinds_M_eq3_h2 returns_result_select_result wellformed_h2) show False proof (cases "xb = owner_document") case True then show ?thesis using select_result_I2[OF disconnected_nodes_h3,folded select_result_I2[OF disconnected_nodes_h2]] by (metis (no_types, lifting) "3" "5" "6" disjoint_iff_not_equal notin_set_remove1) next case False show ?thesis using 2 3 4 5 6 unfolding disconnected_nodes_eq2_h2[OF False] by auto qed qed then have "a_distinct_lists h'" proof(auto simp add: a_distinct_lists_def document_ptr_kinds_eq2_h3 object_ptr_kinds_M_eq2_h3 disconnected_nodes_eq2_h3 intro!: distinct_concat_map_I)[1] fix x assume 1: "distinct (concat (map (\ptr. |h3 \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h')))))" and 2: "x |\| object_ptr_kinds h'" have 3: "\p. p |\| object_ptr_kinds h' \ distinct |h3 \ get_child_nodes p|\<^sub>r" using 1 by (auto elim: distinct_concat_map_E) show "distinct |h' \ get_child_nodes x|\<^sub>r" proof(cases "ptr = x") case True show ?thesis using 3[OF 2] children_h3 children_h' by(auto simp add: True insert_before_list_distinct dest: child_not_in_any_children[unfolded children_eq_h2]) next case False show ?thesis using children_eq2_h3[OF False] 3[OF 2] by auto qed next fix x y xa assume 1: "distinct (concat (map (\ptr. |h3 \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h')))))" and 2: "x |\| object_ptr_kinds h'" and 3: "y |\| object_ptr_kinds h'" and 4: "x \ y" and 5: "xa \ set |h' \ get_child_nodes x|\<^sub>r" and 6: "xa \ set |h' \ get_child_nodes y|\<^sub>r" have 7:"set |h3 \ get_child_nodes x|\<^sub>r \ set |h3 \ get_child_nodes y|\<^sub>r = {}" using distinct_concat_map_E(1)[OF 1] 2 3 4 by auto show False proof (cases "ptr = x") case True then have "ptr \ y" using 4 by simp then show ?thesis using children_h3 children_h' child_not_in_any_children[unfolded children_eq_h2] 5 6 apply(auto simp add: True children_eq2_h3[OF \ptr \ y\])[1] by (metis (no_types, hide_lams) "3" "7" \type_wf h3\ children_eq2_h3 disjoint_iff_not_equal get_child_nodes_ok insert_before_list_in_set known_ptrs local.known_ptrs_known_ptr object_ptr_kinds_M_eq3_h object_ptr_kinds_M_eq3_h' object_ptr_kinds_M_eq3_h2 returns_result_select_result select_result_I2) next case False then show ?thesis proof (cases "ptr = y") case True then show ?thesis using children_h3 children_h' child_not_in_any_children[unfolded children_eq_h2] 5 6 apply(auto simp add: True children_eq2_h3[OF \ptr \ x\])[1] by (metis (no_types, hide_lams) "2" "4" "7" IntI \known_ptrs h3\ \type_wf h'\ children_eq_h3 empty_iff insert_before_list_in_set local.get_child_nodes_ok local.known_ptrs_known_ptr object_ptr_kinds_M_eq3_h' returns_result_select_result select_result_I2) next case False then show ?thesis using children_eq2_h3[OF \ptr \ x\] children_eq2_h3[OF \ptr \ y\] 5 6 7 by auto qed qed next fix x xa xb assume 1: " (\x\fset (object_ptr_kinds h'). set |h3 \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h'). set |h' \ get_disconnected_nodes x|\<^sub>r) = {} " and 2: "xa |\| object_ptr_kinds h'" and 3: "x \ set |h' \ get_child_nodes xa|\<^sub>r" and 4: "xb |\| document_ptr_kinds h'" and 5: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" have 6: "set |h3 \ get_child_nodes xa|\<^sub>r \ set |h' \ get_disconnected_nodes xb|\<^sub>r = {}" using 1 2 3 4 5 proof - have "\h d. \ type_wf h \ d |\| document_ptr_kinds h \ h \ ok get_disconnected_nodes d" using local.get_disconnected_nodes_ok by satx then have "h' \ ok get_disconnected_nodes xb" using "4" \type_wf h'\ by fastforce then have f1: "h3 \ get_disconnected_nodes xb \\<^sub>r |h' \ get_disconnected_nodes xb|\<^sub>r" by (simp add: disconnected_nodes_eq_h3) have "xa |\| object_ptr_kinds h3" using "2" object_ptr_kinds_M_eq3_h' by blast then show ?thesis using f1 \local.a_distinct_lists h3\ local.distinct_lists_no_parent by fastforce qed show False proof (cases "ptr = xa") case True show ?thesis using 6 node_not_in_disconnected_nodes 3 4 5 select_result_I2[OF children_h'] select_result_I2[OF children_h3] True disconnected_nodes_eq2_h3 by (metis (no_types, lifting) "2" DocumentMonad.ptr_kinds_ptr_kinds_M \a_distinct_lists h3\ \type_wf h'\ disconnected_nodes_eq_h3 distinct_lists_no_parent document_ptr_kinds_eq2_h3 get_disconnected_nodes_ok insert_before_list_in_set object_ptr_kinds_M_eq3_h' returns_result_select_result) next case False then show ?thesis using 1 2 3 4 5 children_eq2_h3[OF False] by fastforce qed qed moreover have "a_owner_document_valid h2" using wellformed_h2 by (simp add: heap_is_wellformed_def) then have "a_owner_document_valid h'" apply(auto simp add: a_owner_document_valid_def object_ptr_kinds_M_eq2_h2 object_ptr_kinds_M_eq2_h3 node_ptr_kinds_eq2_h2 node_ptr_kinds_eq2_h3 document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 children_eq2_h2)[1] apply(auto simp add: document_ptr_kinds_eq2_h2[simplified] document_ptr_kinds_eq2_h3[simplified] object_ptr_kinds_M_eq2_h2[simplified] object_ptr_kinds_M_eq2_h3[simplified] node_ptr_kinds_eq2_h2[simplified] node_ptr_kinds_eq2_h3[simplified])[1] apply(auto simp add: disconnected_nodes_eq2_h3[symmetric])[1] by (smt children_eq2_h3 children_h' children_h3 disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3 finite_set_in in_set_remove1 insert_before_list_in_set object_ptr_kinds_M_eq3_h' ptr_in_heap select_result_I2) ultimately show "heap_is_wellformed h'" by (simp add: heap_is_wellformed_def) qed end locale l_insert_before_wf2 = l_type_wf + l_known_ptrs + l_insert_before_defs + l_heap_is_wellformed_defs + l_get_child_nodes_defs + l_remove_defs + assumes insert_before_preserves_type_wf: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ insert_before ptr child ref \\<^sub>h h' \ type_wf h'" assumes insert_before_preserves_known_ptrs: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ insert_before ptr child ref \\<^sub>h h' \ known_ptrs h'" assumes insert_before_heap_is_wellformed_preserved: "type_wf h \ known_ptrs h \ heap_is_wellformed h \ h \ insert_before ptr child ref \\<^sub>h h' \ heap_is_wellformed h'" interpretation i_insert_before_wf2?: l_insert_before_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_parent get_parent_locs get_child_nodes get_child_nodes_locs set_child_nodes set_child_nodes_locs get_ancestors get_ancestors_locs adopt_node adopt_node_locs set_disconnected_nodes set_disconnected_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_owner_document insert_before insert_before_locs append_child type_wf known_ptr known_ptrs heap_is_wellformed parent_child_rel remove_child remove_child_locs get_root_node get_root_node_locs by(simp add: l_insert_before_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_insert_before_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma insert_before_wf2_is_l_insert_before_wf2 [instances]: "l_insert_before_wf2 type_wf known_ptr known_ptrs insert_before heap_is_wellformed" apply(auto simp add: l_insert_before_wf2_def l_insert_before_wf2_axioms_def instances)[1] using insert_before_heap_is_wellformed_preserved apply(fast, fast, fast) done locale l_append_child_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_insert_before_wf + l_insert_before_wf2 + l_get_child_nodes begin lemma append_child_heap_is_wellformed_preserved: assumes wellformed: "heap_is_wellformed h" and append_child: "h \ append_child ptr node \\<^sub>h h'" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "heap_is_wellformed h'" and "type_wf h'" and "known_ptrs h'" using assms by(auto simp add: append_child_def intro: insert_before_preserves_type_wf insert_before_preserves_known_ptrs insert_before_heap_is_wellformed_preserved) lemma append_child_children: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_child_nodes ptr \\<^sub>r xs" assumes "h \ append_child ptr node \\<^sub>h h'" assumes "node \ set xs" shows "h' \ get_child_nodes ptr \\<^sub>r xs @ [node]" proof - obtain ancestors owner_document h2 h3 disconnected_nodes_h2 where ancestors: "h \ get_ancestors ptr \\<^sub>r ancestors" and node_not_in_ancestors: "cast node \ set ancestors" and owner_document: "h \ get_owner_document ptr \\<^sub>r owner_document" and h2: "h \ adopt_node owner_document node \\<^sub>h h2" and disconnected_nodes_h2: "h2 \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h2" and h3: "h2 \ set_disconnected_nodes owner_document (remove1 node disconnected_nodes_h2) \\<^sub>h h3" and h': "h3 \ a_insert_node ptr node None \\<^sub>h h'" using assms(5) by(auto simp add: append_child_def insert_before_def a_ensure_pre_insertion_validity_def elim!: bind_returns_heap_E bind_returns_result_E bind_returns_heap_E2[rotated, OF get_parent_pure, rotated] bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated] bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] bind_returns_heap_E2[rotated, OF get_ancestors_pure, rotated] bind_returns_heap_E2[rotated, OF next_sibling_pure, rotated] bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated] split: if_splits option.splits) have "\parent. |h \ get_parent node|\<^sub>r = Some parent \ parent \ ptr" using assms(1) assms(4) assms(6) by (metis (no_types, lifting) assms(2) assms(3) h2 is_OK_returns_heap_I is_OK_returns_result_E local.adopt_node_child_in_heap local.get_parent_child_dual local.get_parent_ok select_result_I2) have "h2 \ get_child_nodes ptr \\<^sub>r xs" using get_child_nodes_reads adopt_node_writes h2 assms(4) apply(rule reads_writes_separate_forwards) using \\parent. |h \ get_parent node|\<^sub>r = Some parent \ parent \ ptr\ apply(auto simp add: adopt_node_locs_def remove_child_locs_def)[1] by (meson local.set_child_nodes_get_child_nodes_different_pointers) have "h3 \ get_child_nodes ptr \\<^sub>r xs" using get_child_nodes_reads set_disconnected_nodes_writes h3 \h2 \ get_child_nodes ptr \\<^sub>r xs\ apply(rule reads_writes_separate_forwards) by(auto) have "ptr |\| object_ptr_kinds h" by (meson ancestors is_OK_returns_result_I local.get_ancestors_ptr_in_heap) then have "known_ptr ptr" using assms(3) using local.known_ptrs_known_ptr by blast have "type_wf h2" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF adopt_node_writes h2] using adopt_node_types_preserved \type_wf h\ by(auto simp add: adopt_node_locs_def remove_child_locs_def reflp_def transp_def split: if_splits) then have "type_wf h3" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h3] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) show "h' \ get_child_nodes ptr \\<^sub>r xs@[node]" using h' apply(auto simp add: a_insert_node_def dest!: bind_returns_heap_E3[rotated, OF \h3 \ get_child_nodes ptr \\<^sub>r xs\ get_child_nodes_pure, rotated])[1] using \type_wf h3\ set_child_nodes_get_child_nodes \known_ptr ptr\ by metis qed lemma append_child_for_all_on_children: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_child_nodes ptr \\<^sub>r xs" assumes "h \ forall_M (append_child ptr) nodes \\<^sub>h h'" assumes "set nodes \ set xs = {}" assumes "distinct nodes" shows "h' \ get_child_nodes ptr \\<^sub>r xs@nodes" using assms apply(induct nodes arbitrary: h xs) apply(simp) proof(auto elim!: bind_returns_heap_E)[1]fix a nodes h xs h'a assume 0: "(\h xs. heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_child_nodes ptr \\<^sub>r xs \ h \ forall_M (append_child ptr) nodes \\<^sub>h h' \ set nodes \ set xs = {} \ h' \ get_child_nodes ptr \\<^sub>r xs @ nodes)" and 1: "heap_is_wellformed h" and 2: "type_wf h" and 3: "known_ptrs h" and 4: "h \ get_child_nodes ptr \\<^sub>r xs" and 5: "h \ append_child ptr a \\<^sub>r ()" and 6: "h \ append_child ptr a \\<^sub>h h'a" and 7: "h'a \ forall_M (append_child ptr) nodes \\<^sub>h h'" and 8: "a \ set xs" and 9: "set nodes \ set xs = {}" and 10: "a \ set nodes" and 11: "distinct nodes" then have "h'a \ get_child_nodes ptr \\<^sub>r xs @ [a]" using append_child_children 6 using "1" "2" "3" "4" "8" by blast moreover have "heap_is_wellformed h'a" and "type_wf h'a" and "known_ptrs h'a" using insert_before_heap_is_wellformed_preserved insert_before_preserves_known_ptrs insert_before_preserves_type_wf 1 2 3 6 append_child_def by metis+ moreover have "set nodes \ set (xs @ [a]) = {}" using 9 10 by auto ultimately show "h' \ get_child_nodes ptr \\<^sub>r xs @ a # nodes" using 0 7 by fastforce qed lemma append_child_for_all_on_no_children: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_child_nodes ptr \\<^sub>r []" assumes "h \ forall_M (append_child ptr) nodes \\<^sub>h h'" assumes "distinct nodes" shows "h' \ get_child_nodes ptr \\<^sub>r nodes" using assms append_child_for_all_on_children by force end locale l_append_child_wf = l_type_wf + l_known_ptrs + l_append_child_defs + l_heap_is_wellformed_defs + assumes append_child_preserves_type_wf: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ append_child ptr child \\<^sub>h h' \ type_wf h'" assumes append_child_preserves_known_ptrs: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ append_child ptr child \\<^sub>h h' \ known_ptrs h'" assumes append_child_heap_is_wellformed_preserved: "type_wf h \ known_ptrs h \ heap_is_wellformed h \ h \ append_child ptr child \\<^sub>h h' \ heap_is_wellformed h'" interpretation i_append_child_wf?: l_append_child_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_owner_document get_parent get_parent_locs remove_child remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs adopt_node adopt_node_locs known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs set_child_nodes set_child_nodes_locs remove get_ancestors get_ancestors_locs insert_before insert_before_locs append_child heap_is_wellformed parent_child_rel by(auto simp add: l_append_child_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) lemma append_child_wf_is_l_append_child_wf [instances]: "l_append_child_wf type_wf known_ptr known_ptrs append_child heap_is_wellformed" apply(auto simp add: l_append_child_wf_def l_append_child_wf_axioms_def instances)[1] using append_child_heap_is_wellformed_preserved by fast+ subsection \create\_element\ locale l_create_element_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed parent_child_rel + l_new_element_get_disconnected_nodes get_disconnected_nodes get_disconnected_nodes_locs + l_set_tag_name_get_disconnected_nodes type_wf set_tag_name set_tag_name_locs get_disconnected_nodes get_disconnected_nodes_locs + l_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs set_tag_name set_tag_name_locs type_wf create_element known_ptr + l_new_element_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs + l_set_tag_name_get_child_nodes type_wf set_tag_name set_tag_name_locs known_ptr get_child_nodes get_child_nodes_locs + l_set_disconnected_nodes_get_child_nodes set_disconnected_nodes set_disconnected_nodes_locs get_child_nodes get_child_nodes_locs + l_set_disconnected_nodes type_wf set_disconnected_nodes set_disconnected_nodes_locs + l_set_disconnected_nodes_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs + l_new_element type_wf + l_known_ptrs known_ptr known_ptrs for known_ptr :: "(_::linorder) object_ptr \ bool" and known_ptrs :: "(_) heap \ bool" and type_wf :: "(_) heap \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and set_tag_name :: "(_) element_ptr \ char list \ ((_) heap, exception, unit) prog" and set_tag_name_locs :: "(_) element_ptr \ ((_) heap, exception, unit) prog set" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" and create_element :: "(_) document_ptr \ char list \ ((_) heap, exception, (_) element_ptr) prog" begin lemma create_element_preserves_wellformedness: assumes "heap_is_wellformed h" and "h \ create_element document_ptr tag \\<^sub>h h'" and "type_wf h" and "known_ptrs h" shows "heap_is_wellformed h'" and "type_wf h'" and "known_ptrs h'" proof - obtain new_element_ptr h2 h3 disc_nodes_h3 where new_element_ptr: "h \ new_element \\<^sub>r new_element_ptr" and h2: "h \ new_element \\<^sub>h h2" and h3: "h2 \ set_tag_name new_element_ptr tag \\<^sub>h h3" and disc_nodes_h3: "h3 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" and h': "h3 \ set_disconnected_nodes document_ptr (cast new_element_ptr # disc_nodes_h3) \\<^sub>h h'" using assms(2) by(auto simp add: create_element_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] ) then have "h \ create_element document_ptr tag \\<^sub>r new_element_ptr" apply(auto simp add: create_element_def intro!: bind_returns_result_I)[1] apply (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust) apply (metis is_OK_returns_heap_E is_OK_returns_result_I local.get_disconnected_nodes_pure pure_returns_heap_eq) by (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust) have "new_element_ptr \ set |h \ element_ptr_kinds_M|\<^sub>r" using new_element_ptr ElementMonad.ptr_kinds_ptr_kinds_M h2 using new_element_ptr_not_in_heap by blast then have "cast new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r" by simp then have "cast new_element_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r" by simp have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\| {|cast new_element_ptr|}" using new_element_new_ptr h2 new_element_ptr by blast then have node_ptr_kinds_eq_h: "node_ptr_kinds h2 = node_ptr_kinds h |\| {|cast new_element_ptr|}" apply(simp add: node_ptr_kinds_def) by force then have element_ptr_kinds_eq_h: "element_ptr_kinds h2 = element_ptr_kinds h |\| {|new_element_ptr|}" apply(simp add: element_ptr_kinds_def) by force have character_data_ptr_kinds_eq_h: "character_data_ptr_kinds h2 = character_data_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: node_ptr_kinds_def character_data_ptr_kinds_def) have document_ptr_kinds_eq_h: "document_ptr_kinds h2 = document_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: document_ptr_kinds_def) have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_tag_name_writes h3]) using set_tag_name_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2" using object_ptr_kinds_eq_h2 by(auto simp add: node_ptr_kinds_def) have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_disconnected_nodes_writes h']) using set_disconnected_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3" using object_ptr_kinds_eq_h3 by(auto simp add: node_ptr_kinds_def) have "known_ptr (cast new_element_ptr)" using \h \ create_element document_ptr tag \\<^sub>r new_element_ptr\ local.create_element_known_ptr by blast then have "known_ptrs h2" using known_ptrs_new_ptr object_ptr_kinds_eq_h \known_ptrs h\ h2 by blast then have "known_ptrs h3" using known_ptrs_preserved object_ptr_kinds_eq_h2 by blast then show "known_ptrs h'" using known_ptrs_preserved object_ptr_kinds_eq_h3 by blast have "document_ptr |\| document_ptr_kinds h" using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2 get_disconnected_nodes_ptr_in_heap \type_wf h\ document_ptr_kinds_def by (metis is_OK_returns_result_I) have children_eq_h: "\(ptr'::(_) object_ptr) children. ptr' \ cast new_element_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads h2 get_child_nodes_new_element[rotated, OF new_element_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have children_eq2_h: "\ptr'. ptr' \ cast new_element_ptr \ |h \ get_child_nodes ptr'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have "h2 \ get_child_nodes (cast new_element_ptr) \\<^sub>r []" using new_element_ptr h2 new_element_ptr_in_heap[OF h2 new_element_ptr] new_element_is_element_ptr[OF new_element_ptr] new_element_no_child_nodes by blast have disconnected_nodes_eq_h: "\doc_ptr disc_nodes. h \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads h2 get_disconnected_nodes_new_element[OF new_element_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have disconnected_nodes_eq2_h: "\doc_ptr. |h \ get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. h2 \ get_child_nodes ptr' \\<^sub>r children = h3 \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads set_tag_name_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_tag_name_get_child_nodes) then have children_eq2_h2: "\ptr'. |h2 \ get_child_nodes ptr'|\<^sub>r = |h3 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_tag_name_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_tag_name_get_disconnected_nodes) then have disconnected_nodes_eq2_h2: "\doc_ptr. |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have "type_wf h2" using \type_wf h\ new_element_types_preserved h2 by blast then have "type_wf h3" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_tag_name_writes h3] using set_tag_name_types_preserved by(auto simp add: reflp_def transp_def) then show "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h'] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) have children_eq_h3: "\ptr' children. h3 \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_disconnected_nodes_get_child_nodes) then have children_eq2_h3: "\ptr'. |h3 \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h3: "\doc_ptr disc_nodes. document_ptr \ doc_ptr \ h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers) then have disconnected_nodes_eq2_h3: "\doc_ptr. document_ptr \ doc_ptr \ |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have disc_nodes_document_ptr_h2: "h2 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" using disconnected_nodes_eq_h2 disc_nodes_h3 by auto then have disc_nodes_document_ptr_h: "h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" using disconnected_nodes_eq_h by auto then have "cast new_element_ptr \ set disc_nodes_h3" using \heap_is_wellformed h\ using \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ a_all_ptrs_in_heap_def heap_is_wellformed_def using NodeMonad.ptr_kinds_ptr_kinds_M local.heap_is_wellformed_disc_nodes_in_heap by blast have "acyclic (parent_child_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def acyclic_heap_def) also have "parent_child_rel h = parent_child_rel h2" proof(auto simp add: parent_child_rel_def)[1] fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h2" by (simp add: object_ptr_kinds_eq_h) next fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "x \ set |h2 \ get_child_nodes a|\<^sub>r" by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2_h) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h" using object_ptr_kinds_eq_h \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ by(auto) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "x \ set |h \ get_child_nodes a|\<^sub>r" by (metis (no_types, lifting) \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ children_eq2_h empty_iff empty_set image_eqI select_result_I2) qed also have "\ = parent_child_rel h3" by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2) also have "\ = parent_child_rel h'" by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3) finally have "a_acyclic_heap h'" by (simp add: acyclic_heap_def) have "a_all_ptrs_in_heap h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) then have "a_all_ptrs_in_heap h2" apply(auto simp add: a_all_ptrs_in_heap_def)[1] apply (metis \known_ptrs h2\ \parent_child_rel h = parent_child_rel h2\ \type_wf h2\ assms(1) assms(3) funion_iff local.get_child_nodes_ok local.known_ptrs_known_ptr local.parent_child_rel_child_in_heap local.parent_child_rel_child_nodes2 node_ptr_kinds_commutes node_ptr_kinds_eq_h returns_result_select_result) by (metis assms(1) assms(3) disconnected_nodes_eq2_h document_ptr_kinds_eq_h funion_iff local.get_disconnected_nodes_ok local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq_h returns_result_select_result) then have "a_all_ptrs_in_heap h3" by (simp add: children_eq2_h2 disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 local.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 object_ptr_kinds_eq_h2) then have "a_all_ptrs_in_heap h'" by (smt \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ children_eq2_h3 disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 finite_set_in h' is_OK_returns_result_I l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes local.a_all_ptrs_in_heap_def local.get_child_nodes_ptr_in_heap local.l_set_disconnected_nodes_get_disconnected_nodes_axioms node_ptr_kinds_commutes object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 select_result_I2 set_ConsD subset_code(1)) have "\p. p |\| object_ptr_kinds h \ cast new_element_ptr \ set |h \ get_child_nodes p|\<^sub>r" using \heap_is_wellformed h\ \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ heap_is_wellformed_children_in_heap by (meson NodeMonad.ptr_kinds_ptr_kinds_M a_all_ptrs_in_heap_def assms(3) assms(4) fset_mp fset_of_list_elem get_child_nodes_ok known_ptrs_known_ptr returns_result_select_result) then have "\p. p |\| object_ptr_kinds h2 \ cast new_element_ptr \ set |h2 \ get_child_nodes p|\<^sub>r" using children_eq2_h apply(auto simp add: object_ptr_kinds_eq_h)[1] using \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ apply auto[1] by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\) then have "\p. p |\| object_ptr_kinds h3 \ cast new_element_ptr \ set |h3 \ get_child_nodes p|\<^sub>r" using object_ptr_kinds_eq_h2 children_eq2_h2 by auto then have new_element_ptr_not_in_any_children: "\p. p |\| object_ptr_kinds h' \ cast new_element_ptr \ set |h' \ get_child_nodes p|\<^sub>r" using object_ptr_kinds_eq_h3 children_eq2_h3 by auto have "a_distinct_lists h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) then have "a_distinct_lists h2" using \h2 \ get_child_nodes (cast new_element_ptr) \\<^sub>r []\ apply(auto simp add: a_distinct_lists_def object_ptr_kinds_eq_h document_ptr_kinds_eq_h disconnected_nodes_eq2_h intro!: distinct_concat_map_I)[1] - apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert) + apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert_remove) apply(case_tac "x=cast new_element_ptr") apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply (metis IntI assms(1) assms(3) assms(4) empty_iff local.get_child_nodes_ok local.heap_is_wellformed_one_parent local.known_ptrs_known_ptr returns_result_select_result) apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] by (metis \local.a_distinct_lists h\ \type_wf h2\ disconnected_nodes_eq_h document_ptr_kinds_eq_h local.distinct_lists_no_parent local.get_disconnected_nodes_ok returns_result_select_result) then have "a_distinct_lists h3" by(auto simp add: a_distinct_lists_def disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 children_eq2_h2 object_ptr_kinds_eq_h2) then have "a_distinct_lists h'" proof(auto simp add: a_distinct_lists_def disconnected_nodes_eq2_h3 children_eq2_h3 object_ptr_kinds_eq_h3 document_ptr_kinds_eq_h3 intro!: distinct_concat_map_I)[1] fix x assume "distinct (concat (map (\document_ptr. |h3 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" and "x |\| document_ptr_kinds h3" then show "distinct |h' \ get_disconnected_nodes x|\<^sub>r" using document_ptr_kinds_eq_h3 disconnected_nodes_eq_h3 h' set_disconnected_nodes_get_disconnected_nodes by (metis (no_types, lifting) \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set disc_nodes_h3\ \a_distinct_lists h3\ \type_wf h'\ disc_nodes_h3 distinct.simps(2) distinct_lists_disconnected_nodes get_disconnected_nodes_ok returns_result_eq returns_result_select_result) next fix x y xa assume "distinct (concat (map (\document_ptr. |h3 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" and "x |\| document_ptr_kinds h3" and "y |\| document_ptr_kinds h3" and "x \ y" and "xa \ set |h' \ get_disconnected_nodes x|\<^sub>r" and "xa \ set |h' \ get_disconnected_nodes y|\<^sub>r" moreover have "set |h3 \ get_disconnected_nodes x|\<^sub>r \ set |h3 \ get_disconnected_nodes y|\<^sub>r = {}" using calculation by(auto dest: distinct_concat_map_E(1)) ultimately show "False" apply(-) apply(cases "x = document_ptr") apply (smt NodeMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ \local.a_all_ptrs_in_heap h\ disc_nodes_h3 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 finite_set_in h' l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms select_result_I2 set_ConsD subsetD) by (smt NodeMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ \local.a_all_ptrs_in_heap h\ disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 finite_set_in h' l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms select_result_I2 set_ConsD subsetD) next fix x xa xb assume 2: "(\x\fset (object_ptr_kinds h3). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h3). set |h3 \ get_disconnected_nodes x|\<^sub>r) = {}" and 3: "xa |\| object_ptr_kinds h3" and 4: "x \ set |h' \ get_child_nodes xa|\<^sub>r" and 5: "xb |\| document_ptr_kinds h3" and 6: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" show "False" using disc_nodes_document_ptr_h disconnected_nodes_eq2_h3 apply - apply(cases "xb = document_ptr") apply (metis (no_types, hide_lams) "3" "4" "6" \\p. p |\| object_ptr_kinds h3 \ cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h3 \ get_child_nodes p|\<^sub>r\ \a_distinct_lists h3\ children_eq2_h3 disc_nodes_h3 distinct_lists_no_parent h' select_result_I2 set_ConsD set_disconnected_nodes_get_disconnected_nodes) by (metis "3" "4" "5" "6" \a_distinct_lists h3\ \type_wf h3\ children_eq2_h3 distinct_lists_no_parent get_disconnected_nodes_ok returns_result_select_result) qed have "a_owner_document_valid h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) then have "a_owner_document_valid h'" using disc_nodes_h3 \document_ptr |\| document_ptr_kinds h\ apply(auto simp add: a_owner_document_valid_def)[1] apply(auto simp add: object_ptr_kinds_eq_h object_ptr_kinds_eq_h3 )[1] apply(auto simp add: object_ptr_kinds_eq_h2)[1] apply(auto simp add: document_ptr_kinds_eq_h document_ptr_kinds_eq_h3 )[1] apply(auto simp add: document_ptr_kinds_eq_h2)[1] apply(auto simp add: node_ptr_kinds_eq_h node_ptr_kinds_eq_h3 )[1] apply(auto simp add: node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h )[1] apply(auto simp add: children_eq2_h2[symmetric] children_eq2_h3[symmetric] disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3)[1] apply (metis (no_types, lifting) document_ptr_kinds_eq_h h' list.set_intros(1) local.set_disconnected_nodes_get_disconnected_nodes select_result_I2) apply(simp add: object_ptr_kinds_eq_h) by(metis (no_types, lifting) NodeMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ children_eq2_h children_eq2_h2 children_eq2_h3 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h finite_set_in h' l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes list.set_intros(2) local.l_set_disconnected_nodes_get_disconnected_nodes_axioms node_ptr_kinds_commutes select_result_I2) show "heap_is_wellformed h'" using \a_acyclic_heap h'\ \a_all_ptrs_in_heap h'\ \a_distinct_lists h'\ \a_owner_document_valid h'\ by(simp add: heap_is_wellformed_def) qed end interpretation i_create_element_wf?: l_create_element_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr known_ptrs type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed parent_child_rel set_tag_name set_tag_name_locs set_disconnected_nodes set_disconnected_nodes_locs create_element using instances by(auto simp add: l_create_element_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_create_element_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsection \create\_character\_data\ locale l_create_character_data_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed parent_child_rel + l_new_character_data_get_disconnected_nodes get_disconnected_nodes get_disconnected_nodes_locs + l_set_val_get_disconnected_nodes type_wf set_val set_val_locs get_disconnected_nodes get_disconnected_nodes_locs + l_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs set_val set_val_locs type_wf create_character_data known_ptr + l_new_character_data_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs + l_set_val_get_child_nodes type_wf set_val set_val_locs known_ptr get_child_nodes get_child_nodes_locs + l_set_disconnected_nodes_get_child_nodes set_disconnected_nodes set_disconnected_nodes_locs get_child_nodes get_child_nodes_locs + l_set_disconnected_nodes type_wf set_disconnected_nodes set_disconnected_nodes_locs + l_set_disconnected_nodes_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs + l_new_character_data type_wf + l_known_ptrs known_ptr known_ptrs for known_ptr :: "(_::linorder) object_ptr \ bool" and type_wf :: "(_) heap \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and set_val :: "(_) character_data_ptr \ char list \ ((_) heap, exception, unit) prog" and set_val_locs :: "(_) character_data_ptr \ ((_) heap, exception, unit) prog set" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" and create_character_data :: "(_) document_ptr \ char list \ ((_) heap, exception, (_) character_data_ptr) prog" and known_ptrs :: "(_) heap \ bool" begin lemma create_character_data_preserves_wellformedness: assumes "heap_is_wellformed h" and "h \ create_character_data document_ptr text \\<^sub>h h'" and "type_wf h" and "known_ptrs h" shows "heap_is_wellformed h'" and "type_wf h'" and "known_ptrs h'" proof - obtain new_character_data_ptr h2 h3 disc_nodes_h3 where new_character_data_ptr: "h \ new_character_data \\<^sub>r new_character_data_ptr" and h2: "h \ new_character_data \\<^sub>h h2" and h3: "h2 \ set_val new_character_data_ptr text \\<^sub>h h3" and disc_nodes_h3: "h3 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" and h': "h3 \ set_disconnected_nodes document_ptr (cast new_character_data_ptr # disc_nodes_h3) \\<^sub>h h'" using assms(2) by(auto simp add: create_character_data_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] ) then have "h \ create_character_data document_ptr text \\<^sub>r new_character_data_ptr" apply(auto simp add: create_character_data_def intro!: bind_returns_result_I)[1] apply (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust) apply (metis is_OK_returns_heap_E is_OK_returns_result_I local.get_disconnected_nodes_pure pure_returns_heap_eq) by (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust) have "new_character_data_ptr \ set |h \ character_data_ptr_kinds_M|\<^sub>r" using new_character_data_ptr CharacterDataMonad.ptr_kinds_ptr_kinds_M h2 using new_character_data_ptr_not_in_heap by blast then have "cast new_character_data_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r" by simp then have "cast new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r" by simp have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\| {|cast new_character_data_ptr|}" using new_character_data_new_ptr h2 new_character_data_ptr by blast then have node_ptr_kinds_eq_h: "node_ptr_kinds h2 = node_ptr_kinds h |\| {|cast new_character_data_ptr|}" apply(simp add: node_ptr_kinds_def) by force then have character_data_ptr_kinds_eq_h: "character_data_ptr_kinds h2 = character_data_ptr_kinds h |\| {|new_character_data_ptr|}" apply(simp add: character_data_ptr_kinds_def) by force have element_ptr_kinds_eq_h: "element_ptr_kinds h2 = element_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: node_ptr_kinds_def element_ptr_kinds_def) have document_ptr_kinds_eq_h: "document_ptr_kinds h2 = document_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: document_ptr_kinds_def) have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_val_writes h3]) using set_val_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2" using object_ptr_kinds_eq_h2 by(auto simp add: node_ptr_kinds_def) have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_disconnected_nodes_writes h']) using set_disconnected_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3" using object_ptr_kinds_eq_h3 by(auto simp add: node_ptr_kinds_def) have "known_ptr (cast new_character_data_ptr)" using \h \ create_character_data document_ptr text \\<^sub>r new_character_data_ptr\ local.create_character_data_known_ptr by blast then have "known_ptrs h2" using known_ptrs_new_ptr object_ptr_kinds_eq_h \known_ptrs h\ h2 by blast then have "known_ptrs h3" using known_ptrs_preserved object_ptr_kinds_eq_h2 by blast then show "known_ptrs h'" using known_ptrs_preserved object_ptr_kinds_eq_h3 by blast have "document_ptr |\| document_ptr_kinds h" using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2 get_disconnected_nodes_ptr_in_heap \type_wf h\ document_ptr_kinds_def by (metis is_OK_returns_result_I) have children_eq_h: "\(ptr'::(_) object_ptr) children. ptr' \ cast new_character_data_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads h2 get_child_nodes_new_character_data[rotated, OF new_character_data_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have children_eq2_h: "\ptr'. ptr' \ cast new_character_data_ptr \ |h \ get_child_nodes ptr'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\| {|cast new_character_data_ptr|}" using new_character_data_new_ptr h2 new_character_data_ptr by blast then have node_ptr_kinds_eq_h: "node_ptr_kinds h2 = node_ptr_kinds h |\| {|cast new_character_data_ptr|}" apply(simp add: node_ptr_kinds_def) by force then have character_data_ptr_kinds_eq_h: "character_data_ptr_kinds h2 = character_data_ptr_kinds h |\| {|new_character_data_ptr|}" apply(simp add: character_data_ptr_kinds_def) by force have element_ptr_kinds_eq_h: "element_ptr_kinds h2 = element_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: node_ptr_kinds_def element_ptr_kinds_def) have document_ptr_kinds_eq_h: "document_ptr_kinds h2 = document_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: document_ptr_kinds_def) have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_val_writes h3]) using set_val_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2" using object_ptr_kinds_eq_h2 by(auto simp add: node_ptr_kinds_def) have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_disconnected_nodes_writes h']) using set_disconnected_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3" using object_ptr_kinds_eq_h3 by(auto simp add: node_ptr_kinds_def) have "document_ptr |\| document_ptr_kinds h" using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2 get_disconnected_nodes_ptr_in_heap \type_wf h\ document_ptr_kinds_def by (metis is_OK_returns_result_I) have children_eq_h: "\(ptr'::(_) object_ptr) children. ptr' \ cast new_character_data_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads h2 get_child_nodes_new_character_data[rotated, OF new_character_data_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have children_eq2_h: "\ptr'. ptr' \ cast new_character_data_ptr \ |h \ get_child_nodes ptr'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have "h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []" using new_character_data_ptr h2 new_character_data_ptr_in_heap[OF h2 new_character_data_ptr] new_character_data_is_character_data_ptr[OF new_character_data_ptr] new_character_data_no_child_nodes by blast have disconnected_nodes_eq_h: "\doc_ptr disc_nodes. h \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads h2 get_disconnected_nodes_new_character_data[OF new_character_data_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have disconnected_nodes_eq2_h: "\doc_ptr. |h \ get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. h2 \ get_child_nodes ptr' \\<^sub>r children = h3 \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads set_val_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_val_get_child_nodes) then have children_eq2_h2: "\ptr'. |h2 \ get_child_nodes ptr'|\<^sub>r = |h3 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_val_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_val_get_disconnected_nodes) then have disconnected_nodes_eq2_h2: "\doc_ptr. |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have "type_wf h2" using \type_wf h\ new_character_data_types_preserved h2 by blast then have "type_wf h3" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_val_writes h3] using set_val_types_preserved by(auto simp add: reflp_def transp_def) then show "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h'] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) have children_eq_h3: "\ptr' children. h3 \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_disconnected_nodes_get_child_nodes) then have children_eq2_h3: " \ptr'. |h3 \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h3: "\doc_ptr disc_nodes. document_ptr \ doc_ptr \ h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers) then have disconnected_nodes_eq2_h3: "\doc_ptr. document_ptr \ doc_ptr \ |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have disc_nodes_document_ptr_h2: "h2 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" using disconnected_nodes_eq_h2 disc_nodes_h3 by auto then have disc_nodes_document_ptr_h: "h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" using disconnected_nodes_eq_h by auto then have "cast new_character_data_ptr \ set disc_nodes_h3" using \heap_is_wellformed h\ using \cast new_character_data_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ a_all_ptrs_in_heap_def heap_is_wellformed_def using NodeMonad.ptr_kinds_ptr_kinds_M local.heap_is_wellformed_disc_nodes_in_heap by blast have "acyclic (parent_child_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def acyclic_heap_def) also have "parent_child_rel h = parent_child_rel h2" proof(auto simp add: parent_child_rel_def)[1] fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h2" by (simp add: object_ptr_kinds_eq_h) next fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "x \ set |h2 \ get_child_nodes a|\<^sub>r" by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2_h) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h" using object_ptr_kinds_eq_h \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ by(auto) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "x \ set |h \ get_child_nodes a|\<^sub>r" by (metis (no_types, lifting) \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ children_eq2_h empty_iff empty_set image_eqI select_result_I2) qed also have "\ = parent_child_rel h3" by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2) also have "\ = parent_child_rel h'" by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3) finally have "a_acyclic_heap h'" by (simp add: acyclic_heap_def) have "a_all_ptrs_in_heap h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) then have "a_all_ptrs_in_heap h2" apply(auto simp add: a_all_ptrs_in_heap_def)[1] using node_ptr_kinds_eq_h \cast new_character_data_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ apply (metis (no_types, lifting) NodeMonad.ptr_kinds_ptr_kinds_M \parent_child_rel h = parent_child_rel h2\ children_eq2_h finite_set_in finsert_iff funion_finsert_right local.parent_child_rel_child local.parent_child_rel_parent_in_heap node_ptr_kinds_commutes object_ptr_kinds_eq_h select_result_I2 subsetD sup_bot.right_neutral) by (metis assms(1) assms(3) disconnected_nodes_eq2_h document_ptr_kinds_eq_h funionI1 local.get_disconnected_nodes_ok local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq_h returns_result_select_result) then have "a_all_ptrs_in_heap h3" by (simp add: children_eq2_h2 disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 local.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 object_ptr_kinds_eq_h2) then have "a_all_ptrs_in_heap h'" by (smt character_data_ptr_kinds_commutes children_eq2_h3 disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 finite_set_in h' h2 local.a_all_ptrs_in_heap_def local.set_disconnected_nodes_get_disconnected_nodes new_character_data_ptr new_character_data_ptr_in_heap node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h3 object_ptr_kinds_eq_h3 select_result_I2 set_ConsD subset_code(1)) have "\p. p |\| object_ptr_kinds h \ cast new_character_data_ptr \ set |h \ get_child_nodes p|\<^sub>r" using \heap_is_wellformed h\ \cast new_character_data_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ heap_is_wellformed_children_in_heap by (meson NodeMonad.ptr_kinds_ptr_kinds_M a_all_ptrs_in_heap_def assms(3) assms(4) fset_mp fset_of_list_elem get_child_nodes_ok known_ptrs_known_ptr returns_result_select_result) then have "\p. p |\| object_ptr_kinds h2 \ cast new_character_data_ptr \ set |h2 \ get_child_nodes p|\<^sub>r" using children_eq2_h apply(auto simp add: object_ptr_kinds_eq_h)[1] using \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ apply auto[1] by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\) then have "\p. p |\| object_ptr_kinds h3 \ cast new_character_data_ptr \ set |h3 \ get_child_nodes p|\<^sub>r" using object_ptr_kinds_eq_h2 children_eq2_h2 by auto then have new_character_data_ptr_not_in_any_children: "\p. p |\| object_ptr_kinds h' \ cast new_character_data_ptr \ set |h' \ get_child_nodes p|\<^sub>r" using object_ptr_kinds_eq_h3 children_eq2_h3 by auto have "a_distinct_lists h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) then have "a_distinct_lists h2" using \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ apply(auto simp add: a_distinct_lists_def object_ptr_kinds_eq_h document_ptr_kinds_eq_h disconnected_nodes_eq2_h intro!: distinct_concat_map_I)[1] - apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert) + apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert_remove) apply(case_tac "x=cast new_character_data_ptr") apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply (metis IntI assms(1) assms(3) assms(4) empty_iff local.get_child_nodes_ok local.heap_is_wellformed_one_parent local.known_ptrs_known_ptr returns_result_select_result) apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] by (metis \local.a_distinct_lists h\ \type_wf h2\ disconnected_nodes_eq_h document_ptr_kinds_eq_h local.distinct_lists_no_parent local.get_disconnected_nodes_ok returns_result_select_result) then have "a_distinct_lists h3" by(auto simp add: a_distinct_lists_def disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 children_eq2_h2 object_ptr_kinds_eq_h2)[1] then have "a_distinct_lists h'" proof(auto simp add: a_distinct_lists_def disconnected_nodes_eq2_h3 children_eq2_h3 object_ptr_kinds_eq_h3 document_ptr_kinds_eq_h3 intro!: distinct_concat_map_I)[1] fix x assume "distinct (concat (map (\document_ptr. |h3 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" and "x |\| document_ptr_kinds h3" then show "distinct |h' \ get_disconnected_nodes x|\<^sub>r" using document_ptr_kinds_eq_h3 disconnected_nodes_eq_h3 h' set_disconnected_nodes_get_disconnected_nodes by (metis (no_types, lifting) \cast new_character_data_ptr \ set disc_nodes_h3\ \a_distinct_lists h3\ \type_wf h'\ disc_nodes_h3 distinct.simps(2) distinct_lists_disconnected_nodes get_disconnected_nodes_ok returns_result_eq returns_result_select_result) next fix x y xa assume "distinct (concat (map (\document_ptr. |h3 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" and "x |\| document_ptr_kinds h3" and "y |\| document_ptr_kinds h3" and "x \ y" and "xa \ set |h' \ get_disconnected_nodes x|\<^sub>r" and "xa \ set |h' \ get_disconnected_nodes y|\<^sub>r" moreover have "set |h3 \ get_disconnected_nodes x|\<^sub>r \ set |h3 \ get_disconnected_nodes y|\<^sub>r = {}" using calculation by(auto dest: distinct_concat_map_E(1)) ultimately show "False" by (smt NodeMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_character_data_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ \local.a_all_ptrs_in_heap h\ disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 finite_set_in h' l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms select_result_I2 set_ConsD subsetD) next fix x xa xb assume 2: "(\x\fset (object_ptr_kinds h3). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h3). set |h3 \ get_disconnected_nodes x|\<^sub>r) = {}" and 3: "xa |\| object_ptr_kinds h3" and 4: "x \ set |h' \ get_child_nodes xa|\<^sub>r" and 5: "xb |\| document_ptr_kinds h3" and 6: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" show "False" using disc_nodes_document_ptr_h disconnected_nodes_eq2_h3 apply(cases "xb = document_ptr") apply (metis (no_types, hide_lams) "3" "4" "6" \\p. p |\| object_ptr_kinds h3 \ cast new_character_data_ptr \ set |h3 \ get_child_nodes p|\<^sub>r\ \a_distinct_lists h3\ children_eq2_h3 disc_nodes_h3 distinct_lists_no_parent h' select_result_I2 set_ConsD set_disconnected_nodes_get_disconnected_nodes) by (metis "3" "4" "5" "6" \a_distinct_lists h3\ \type_wf h3\ children_eq2_h3 distinct_lists_no_parent get_disconnected_nodes_ok returns_result_select_result) qed have "a_owner_document_valid h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) then have "a_owner_document_valid h'" using disc_nodes_h3 \document_ptr |\| document_ptr_kinds h\ apply(simp add: a_owner_document_valid_def) apply(simp add: object_ptr_kinds_eq_h object_ptr_kinds_eq_h3 ) apply(simp add: object_ptr_kinds_eq_h2) apply(simp add: document_ptr_kinds_eq_h document_ptr_kinds_eq_h3 ) apply(simp add: document_ptr_kinds_eq_h2) apply(simp add: node_ptr_kinds_eq_h node_ptr_kinds_eq_h3 ) apply(simp add: node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h ) apply(auto simp add: children_eq2_h2[symmetric] children_eq2_h3[symmetric] disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3)[1] apply (metis (no_types, lifting) document_ptr_kinds_eq_h h' list.set_intros(1) local.set_disconnected_nodes_get_disconnected_nodes select_result_I2) apply(simp add: object_ptr_kinds_eq_h) by (metis (mono_tags, lifting) \cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2_h disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h finite_set_in h' l_ptr_kinds_M.ptr_kinds_ptr_kinds_M l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes list.set_intros(2) local.l_set_disconnected_nodes_get_disconnected_nodes_axioms object_ptr_kinds_M_def select_result_I2) show "heap_is_wellformed h'" using \a_acyclic_heap h'\ \a_all_ptrs_in_heap h'\ \a_distinct_lists h'\ \a_owner_document_valid h'\ by(simp add: heap_is_wellformed_def) qed end interpretation i_create_character_data_wf?: l_create_character_data_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed parent_child_rel set_val set_val_locs set_disconnected_nodes set_disconnected_nodes_locs create_character_data known_ptrs using instances by (auto simp add: l_create_character_data_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_create_character_data_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsection \create\_document\ locale l_create_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed parent_child_rel + l_new_document_get_disconnected_nodes get_disconnected_nodes get_disconnected_nodes_locs + l_create_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M create_document + l_new_document_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs + l_new_document type_wf + l_known_ptrs known_ptr known_ptrs for known_ptr :: "(_::linorder) object_ptr \ bool" and type_wf :: "(_) heap \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and set_val :: "(_) character_data_ptr \ char list \ ((_) heap, exception, unit) prog" and set_val_locs :: "(_) character_data_ptr \ ((_) heap, exception, unit) prog set" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" and create_document :: "((_) heap, exception, (_) document_ptr) prog" and known_ptrs :: "(_) heap \ bool" begin lemma create_document_preserves_wellformedness: assumes "heap_is_wellformed h" and "h \ create_document \\<^sub>h h'" and "type_wf h" and "known_ptrs h" shows "heap_is_wellformed h'" proof - obtain new_document_ptr where new_document_ptr: "h \ new_document \\<^sub>r new_document_ptr" and h': "h \ new_document \\<^sub>h h'" using assms(2) apply(simp add: create_document_def) using new_document_ok by blast have "new_document_ptr \ set |h \ document_ptr_kinds_M|\<^sub>r" using new_document_ptr DocumentMonad.ptr_kinds_ptr_kinds_M using new_document_ptr_not_in_heap h' by blast then have "cast new_document_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r" by simp have "new_document_ptr |\| document_ptr_kinds h" using new_document_ptr DocumentMonad.ptr_kinds_ptr_kinds_M using new_document_ptr_not_in_heap h' by blast then have "cast new_document_ptr |\| object_ptr_kinds h" by simp have object_ptr_kinds_eq: "object_ptr_kinds h' = object_ptr_kinds h |\| {|cast new_document_ptr|}" using new_document_new_ptr h' new_document_ptr by blast then have node_ptr_kinds_eq: "node_ptr_kinds h' = node_ptr_kinds h" apply(simp add: node_ptr_kinds_def) by force then have character_data_ptr_kinds_eq_h: "character_data_ptr_kinds h' = character_data_ptr_kinds h" by(simp add: character_data_ptr_kinds_def) have element_ptr_kinds_eq_h: "element_ptr_kinds h' = element_ptr_kinds h" using object_ptr_kinds_eq by(auto simp add: node_ptr_kinds_def element_ptr_kinds_def) have document_ptr_kinds_eq_h: "document_ptr_kinds h' = document_ptr_kinds h |\| {|new_document_ptr|}" using object_ptr_kinds_eq apply(auto simp add: document_ptr_kinds_def)[1] by (metis (no_types, lifting) document_ptr_kinds_commutes document_ptr_kinds_def finsertI1 fset.map_comp) have children_eq: "\(ptr'::(_) object_ptr) children. ptr' \ cast new_document_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads h' get_child_nodes_new_document[rotated, OF new_document_ptr h'] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have children_eq2: "\ptr'. ptr' \ cast new_document_ptr \ |h \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have "h' \ get_child_nodes (cast new_document_ptr) \\<^sub>r []" using new_document_ptr h' new_document_ptr_in_heap[OF h' new_document_ptr] new_document_is_document_ptr[OF new_document_ptr] new_document_no_child_nodes by blast have disconnected_nodes_eq_h: "\doc_ptr disc_nodes. doc_ptr \ new_document_ptr \ h \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads h' get_disconnected_nodes_new_document_different_pointers new_document_ptr apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by (metis(full_types) \\thesis. (\new_document_ptr. \h \ new_document \\<^sub>r new_document_ptr; h \ new_document \\<^sub>h h'\ \ thesis) \ thesis\ local.get_disconnected_nodes_new_document_different_pointers new_document_ptr)+ then have disconnected_nodes_eq2_h: "\doc_ptr. doc_ptr \ new_document_ptr \ |h \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have "h' \ get_disconnected_nodes new_document_ptr \\<^sub>r []" using h' local.new_document_no_disconnected_nodes new_document_ptr by blast have "type_wf h'" using \type_wf h\ new_document_types_preserved h' by blast have "acyclic (parent_child_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def acyclic_heap_def) also have "parent_child_rel h = parent_child_rel h'" proof(auto simp add: parent_child_rel_def)[1] fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h'" by (simp add: object_ptr_kinds_eq) next fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "x \ set |h' \ get_child_nodes a|\<^sub>r" by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast new_document_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2) next fix a x assume 0: "a |\| object_ptr_kinds h'" and 1: "x \ set |h' \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h" using object_ptr_kinds_eq \h' \ get_child_nodes (cast new_document_ptr) \\<^sub>r []\ by(auto) next fix a x assume 0: "a |\| object_ptr_kinds h'" and 1: "x \ set |h' \ get_child_nodes a|\<^sub>r" then show "x \ set |h \ get_child_nodes a|\<^sub>r" by (metis (no_types, lifting) \h' \ get_child_nodes (cast new_document_ptr) \\<^sub>r []\ children_eq2 empty_iff empty_set image_eqI select_result_I2) qed finally have "a_acyclic_heap h'" by (simp add: acyclic_heap_def) have "a_all_ptrs_in_heap h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) then have "a_all_ptrs_in_heap h'" apply(auto simp add: a_all_ptrs_in_heap_def)[1] using ObjectMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ \parent_child_rel h = parent_child_rel h'\ assms(1) children_eq fset_of_list_elem local.heap_is_wellformed_children_in_heap local.parent_child_rel_child local.parent_child_rel_parent_in_heap node_ptr_kinds_eq apply (metis (no_types, lifting) \h' \ get_child_nodes (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr) \\<^sub>r []\ children_eq2 finite_set_in finsert_iff funion_finsert_right object_ptr_kinds_eq select_result_I2 subsetD sup_bot.right_neutral) by (metis (no_types, lifting) \cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr |\| object_ptr_kinds h\ \h' \ get_child_nodes (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr) \\<^sub>r []\ \h' \ get_disconnected_nodes new_document_ptr \\<^sub>r []\ \parent_child_rel h = parent_child_rel h'\ \type_wf h'\ assms(1) disconnected_nodes_eq_h local.get_disconnected_nodes_ok local.heap_is_wellformed_disc_nodes_in_heap local.parent_child_rel_child local.parent_child_rel_parent_in_heap node_ptr_kinds_eq returns_result_select_result select_result_I2) have "a_distinct_lists h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) then have "a_distinct_lists h'" using \h' \ get_disconnected_nodes new_document_ptr \\<^sub>r []\ \h' \ get_child_nodes (cast new_document_ptr) \\<^sub>r []\ apply(auto simp add: children_eq2[symmetric] a_distinct_lists_def insort_split object_ptr_kinds_eq document_ptr_kinds_eq_h disconnected_nodes_eq2_h intro!: distinct_concat_map_I)[1] - apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert) + apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert_remove) apply(auto simp add: dest: distinct_concat_map_E)[1] apply(auto simp add: dest: distinct_concat_map_E)[1] using \new_document_ptr |\| document_ptr_kinds h\ apply(auto simp add: distinct_insort dest: distinct_concat_map_E)[1] using disconnected_nodes_eq_h apply (metis assms(1) assms(3) disconnected_nodes_eq2_h local.get_disconnected_nodes_ok local.heap_is_wellformed_disconnected_nodes_distinct returns_result_select_result) proof - fix x :: "(_) document_ptr" and y :: "(_) document_ptr" and xa :: "(_) node_ptr" assume a1: "x \ y" assume a2: "x |\| document_ptr_kinds h" assume a3: "x \ new_document_ptr" assume a4: "y |\| document_ptr_kinds h" assume a5: "y \ new_document_ptr" assume a6: "distinct (concat (map (\document_ptr. |h \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h)))))" assume a7: "xa \ set |h' \ get_disconnected_nodes x|\<^sub>r" assume a8: "xa \ set |h' \ get_disconnected_nodes y|\<^sub>r" have f9: "xa \ set |h \ get_disconnected_nodes x|\<^sub>r" using a7 a3 disconnected_nodes_eq2_h by presburger have f10: "xa \ set |h \ get_disconnected_nodes y|\<^sub>r" using a8 a5 disconnected_nodes_eq2_h by presburger have f11: "y \ set (sorted_list_of_set (fset (document_ptr_kinds h)))" using a4 by simp have "x \ set (sorted_list_of_set (fset (document_ptr_kinds h)))" using a2 by simp then show False using f11 f10 f9 a6 a1 by (meson disjoint_iff_not_equal distinct_concat_map_E(1)) next fix x xa xb assume 0: "h' \ get_disconnected_nodes new_document_ptr \\<^sub>r []" and 1: "h' \ get_child_nodes (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr) \\<^sub>r []" and 2: "distinct (concat (map (\ptr. |h \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h)))))" and 3: "distinct (concat (map (\document_ptr. |h \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h)))))" and 4: "(\x\fset (object_ptr_kinds h). set |h \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h). set |h \ get_disconnected_nodes x|\<^sub>r) = {}" and 5: "x \ set |h \ get_child_nodes xa|\<^sub>r" and 6: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" and 7: "xa |\| object_ptr_kinds h" and 8: "xa \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr" and 9: "xb |\| document_ptr_kinds h" and 10: "xb \ new_document_ptr" then show "False" by (metis \local.a_distinct_lists h\ assms(3) disconnected_nodes_eq2_h local.distinct_lists_no_parent local.get_disconnected_nodes_ok returns_result_select_result) qed have "a_owner_document_valid h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) then have "a_owner_document_valid h'" apply(auto simp add: a_owner_document_valid_def)[1] by (metis \cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr |\| object_ptr_kinds h\ children_eq2 disconnected_nodes_eq2_h document_ptr_kinds_commutes finite_set_in funion_iff node_ptr_kinds_eq object_ptr_kinds_eq) show "heap_is_wellformed h'" using \a_acyclic_heap h'\ \a_all_ptrs_in_heap h'\ \a_distinct_lists h'\ \a_owner_document_valid h'\ by(simp add: heap_is_wellformed_def) qed end interpretation i_create_document_wf?: l_create_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed parent_child_rel set_val set_val_locs set_disconnected_nodes set_disconnected_nodes_locs create_document known_ptrs using instances by (auto simp add: l_create_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_create_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] end diff --git a/thys/Featherweight_OCL/collection_types/UML_Set.thy b/thys/Featherweight_OCL/collection_types/UML_Set.thy --- a/thys/Featherweight_OCL/collection_types/UML_Set.thy +++ b/thys/Featherweight_OCL/collection_types/UML_Set.thy @@ -1,3170 +1,3172 @@ (***************************************************************************** * Featherweight-OCL --- A Formal Semantics for UML-OCL Version OCL 2.5 * for the OMG Standard. * http://www.brucker.ch/projects/hol-testgen/ * * UML_Set.thy --- Library definitions. * This file is part of HOL-TestGen. * * Copyright (c) 2012-2015 Université Paris-Saclay, Univ. Paris-Sud, France * 2013-2015 IRT SystemX, 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. ******************************************************************************) theory UML_Set imports "../basic_types/UML_Void" "../basic_types/UML_Boolean" "../basic_types/UML_Integer" "../basic_types/UML_String" "../basic_types/UML_Real" begin no_notation None ("\") section\Collection Type Set: Operations \label{formal-set}\ subsection\As a Motivation for the (infinite) Type Construction: Type-Extensions as Sets \label{sec:type-extensions}\ text\Our notion of typed set goes beyond the usual notion of a finite executable set and is powerful enough to capture \emph{the extension of a type} in UML and OCL. This means we can have in Featherweight OCL Sets containing all possible elements of a type, not only those (finite) ones representable in a state. This holds for base types as well as class types, although the notion for class-types --- involving object id's not occurring in a state --- requires some care. In a world with @{term invalid} and @{term null}, there are two notions extensions possible: \begin{enumerate} \item the set of all \emph{defined} values of a type @{term T} (for which we will introduce the constant @{term T}) \item the set of all \emph{valid} values of a type @{term T}, so including @{term null} (for which we will introduce the constant @{term T\<^sub>n\<^sub>u\<^sub>l\<^sub>l}). \end{enumerate} \ text\We define the set extensions for the base type @{term Integer} as follows:\ definition Integer :: "('\,Integer\<^sub>b\<^sub>a\<^sub>s\<^sub>e) Set" where "Integer \ (\ \. (Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e o Some o Some) ((Some o Some) ` (UNIV::int set)))" definition Integer\<^sub>n\<^sub>u\<^sub>l\<^sub>l :: "('\,Integer\<^sub>b\<^sub>a\<^sub>s\<^sub>e) Set" where "Integer\<^sub>n\<^sub>u\<^sub>l\<^sub>l \ (\ \. (Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e o Some o Some) (Some ` (UNIV::int option set)))" lemma Integer_defined : "\ Integer = true" apply(rule ext, auto simp: Integer_def defined_def false_def true_def bot_fun_def null_fun_def null_option_def) by(simp_all add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_option_def) lemma Integer\<^sub>n\<^sub>u\<^sub>l\<^sub>l_defined : "\ Integer\<^sub>n\<^sub>u\<^sub>l\<^sub>l = true" apply(rule ext, auto simp: Integer\<^sub>n\<^sub>u\<^sub>l\<^sub>l_def defined_def false_def true_def bot_fun_def null_fun_def null_option_def) by(simp_all add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_option_def) text\This allows the theorems: \\ \ \ x \ \ \ (Integer->includes\<^sub>S\<^sub>e\<^sub>t(x))\ \\ \ \ x \ \ \ Integer \ (Integer->including\<^sub>S\<^sub>e\<^sub>t(x))\ and \\ \ \ x \ \ \ (Integer\<^sub>n\<^sub>u\<^sub>l\<^sub>l->includes\<^sub>S\<^sub>e\<^sub>t(x))\ \\ \ \ x \ \ \ Integer\<^sub>n\<^sub>u\<^sub>l\<^sub>l \ (Integer\<^sub>n\<^sub>u\<^sub>l\<^sub>l->including\<^sub>S\<^sub>e\<^sub>t(x))\ which characterize the infiniteness of these sets by a recursive property on these sets. \ text\In the same spirit, we proceed similarly for the remaining base types:\ definition Void\<^sub>n\<^sub>u\<^sub>l\<^sub>l :: "('\,Void\<^sub>b\<^sub>a\<^sub>s\<^sub>e) Set" where "Void\<^sub>n\<^sub>u\<^sub>l\<^sub>l \ (\ \. (Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e o Some o Some) {Abs_Void\<^sub>b\<^sub>a\<^sub>s\<^sub>e (Some None)})" definition Void\<^sub>e\<^sub>m\<^sub>p\<^sub>t\<^sub>y :: "('\,Void\<^sub>b\<^sub>a\<^sub>s\<^sub>e) Set" where "Void\<^sub>e\<^sub>m\<^sub>p\<^sub>t\<^sub>y \ (\ \. (Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e o Some o Some) {})" lemma Void\<^sub>n\<^sub>u\<^sub>l\<^sub>l_defined : "\ Void\<^sub>n\<^sub>u\<^sub>l\<^sub>l = true" apply(rule ext, auto simp: Void\<^sub>n\<^sub>u\<^sub>l\<^sub>l_def defined_def false_def true_def bot_fun_def null_fun_def null_option_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def) by((subst (asm) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject, auto simp add: bot_option_def null_option_def bot_Void_def), (subst (asm) Abs_Void\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject, auto simp add: bot_option_def null_option_def))+ lemma Void\<^sub>e\<^sub>m\<^sub>p\<^sub>t\<^sub>y_defined : "\ Void\<^sub>e\<^sub>m\<^sub>p\<^sub>t\<^sub>y = true" apply(rule ext, auto simp: Void\<^sub>e\<^sub>m\<^sub>p\<^sub>t\<^sub>y_def defined_def false_def true_def bot_fun_def null_fun_def null_option_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def) by((subst (asm) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject, auto simp add: bot_option_def null_option_def bot_Void_def))+ lemma assumes "\ \ \ (V :: ('\,Void\<^sub>b\<^sub>a\<^sub>s\<^sub>e) Set)" shows "\ \ V \ Void\<^sub>n\<^sub>u\<^sub>l\<^sub>l \ \ \ V \ Void\<^sub>e\<^sub>m\<^sub>p\<^sub>t\<^sub>y" proof - have A:"\x y. x \ {} \ \y. y\ x" by (metis all_not_in_conv) show "?thesis" apply(case_tac "V \") proof - fix y show "V \ = Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e y \ y \ {X. X = \ \ X = null \ (\x\\\X\\. x \ \)} \ \ \ V \ Void\<^sub>n\<^sub>u\<^sub>l\<^sub>l \ \ \ V \ Void\<^sub>e\<^sub>m\<^sub>p\<^sub>t\<^sub>y" apply(insert assms, case_tac y, simp add: bot_option_def, simp add: bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def foundation16) apply(simp add: bot_option_def null_option_def) apply(erule disjE, metis OclValid_def defined_def foundation2 null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_fun_def true_def) proof - fix a show "V \ = Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \a\ \ \x\\a\. x \ \ \ \ \ V \ Void\<^sub>n\<^sub>u\<^sub>l\<^sub>l \ \ \ V \ Void\<^sub>e\<^sub>m\<^sub>p\<^sub>t\<^sub>y" apply(case_tac a, simp, insert assms, metis OclValid_def foundation16 null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def true_def) apply(simp) proof - fix aa show " V \ = Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\aa\\ \ \x\aa. x \ \ \ \ \ V \ Void\<^sub>n\<^sub>u\<^sub>l\<^sub>l \ \ \ V \ Void\<^sub>e\<^sub>m\<^sub>p\<^sub>t\<^sub>y" apply(case_tac "aa = {}", rule disjI2, insert assms, simp add: Void\<^sub>e\<^sub>m\<^sub>p\<^sub>t\<^sub>y_def OclValid_def StrongEq_def true_def, rule disjI1) apply(subgoal_tac "aa = {Abs_Void\<^sub>b\<^sub>a\<^sub>s\<^sub>e \None\}", simp add: StrongEq_def OclValid_def true_def Void\<^sub>n\<^sub>u\<^sub>l\<^sub>l_def) apply(drule A, erule exE) proof - fix y show "V \ = Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\aa\\ \ \x\aa. x \ \ \ \ \ \ V \ y \ aa \ aa = {Abs_Void\<^sub>b\<^sub>a\<^sub>s\<^sub>e \None\}" apply(rule equalityI, rule subsetI, simp) proof - fix x show " V \ = Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\aa\\ \ \x\aa. x \ \ \ \ \ \ V \ y \ aa \ x \ aa \ x = Abs_Void\<^sub>b\<^sub>a\<^sub>s\<^sub>e \None\" apply(case_tac x, simp) by (metis bot_Void_def bot_option_def null_option_def) apply_end(simp_all) apply_end(erule ballE[where x = y], simp_all) apply_end(case_tac y, simp add: bot_option_def null_option_def OclValid_def defined_def split: if_split_asm, simp add: false_def true_def) qed (erule disjE, simp add: bot_Void_def, simp) qed qed qed qed qed definition Boolean :: "('\,Boolean\<^sub>b\<^sub>a\<^sub>s\<^sub>e) Set" where "Boolean \ (\ \. (Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e o Some o Some) ((Some o Some) ` (UNIV::bool set)))" definition Boolean\<^sub>n\<^sub>u\<^sub>l\<^sub>l :: "('\,Boolean\<^sub>b\<^sub>a\<^sub>s\<^sub>e) Set" where "Boolean\<^sub>n\<^sub>u\<^sub>l\<^sub>l \ (\ \. (Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e o Some o Some) (Some ` (UNIV::bool option set)))" lemma Boolean_defined : "\ Boolean = true" apply(rule ext, auto simp: Boolean_def defined_def false_def true_def bot_fun_def null_fun_def null_option_def) by(simp_all add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_option_def) lemma Boolean\<^sub>n\<^sub>u\<^sub>l\<^sub>l_defined : "\ Boolean\<^sub>n\<^sub>u\<^sub>l\<^sub>l = true" apply(rule ext, auto simp: Boolean\<^sub>n\<^sub>u\<^sub>l\<^sub>l_def defined_def false_def true_def bot_fun_def null_fun_def null_option_def) by(simp_all add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_option_def) definition String :: "('\,String\<^sub>b\<^sub>a\<^sub>s\<^sub>e) Set" where "String \ (\ \. (Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e o Some o Some) ((Some o Some) ` (UNIV::string set)))" definition String\<^sub>n\<^sub>u\<^sub>l\<^sub>l :: "('\,String\<^sub>b\<^sub>a\<^sub>s\<^sub>e) Set" where "String\<^sub>n\<^sub>u\<^sub>l\<^sub>l \ (\ \. (Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e o Some o Some) (Some ` (UNIV::string option set)))" lemma String_defined : "\ String = true" apply(rule ext, auto simp: String_def defined_def false_def true_def bot_fun_def null_fun_def null_option_def) by(simp_all add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_option_def) lemma String\<^sub>n\<^sub>u\<^sub>l\<^sub>l_defined : "\ String\<^sub>n\<^sub>u\<^sub>l\<^sub>l = true" apply(rule ext, auto simp: String\<^sub>n\<^sub>u\<^sub>l\<^sub>l_def defined_def false_def true_def bot_fun_def null_fun_def null_option_def) by(simp_all add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_option_def) definition Real :: "('\,Real\<^sub>b\<^sub>a\<^sub>s\<^sub>e) Set" where "Real \ (\ \. (Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e o Some o Some) ((Some o Some) ` (UNIV::real set)))" definition Real\<^sub>n\<^sub>u\<^sub>l\<^sub>l :: "('\,Real\<^sub>b\<^sub>a\<^sub>s\<^sub>e) Set" where "Real\<^sub>n\<^sub>u\<^sub>l\<^sub>l \ (\ \. (Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e o Some o Some) (Some ` (UNIV::real option set)))" lemma Real_defined : "\ Real = true" apply(rule ext, auto simp: Real_def defined_def false_def true_def bot_fun_def null_fun_def null_option_def) by(simp_all add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_option_def) lemma Real\<^sub>n\<^sub>u\<^sub>l\<^sub>l_defined : "\ Real\<^sub>n\<^sub>u\<^sub>l\<^sub>l = true" apply(rule ext, auto simp: Real\<^sub>n\<^sub>u\<^sub>l\<^sub>l_def defined_def false_def true_def bot_fun_def null_fun_def null_option_def) by(simp_all add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_option_def) subsection\Basic Properties of the Set Type\ text\Every element in a defined set is valid.\ lemma Set_inv_lemma: "\ \ (\ X) \ \x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. x \ bot" apply(insert Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e [of "X \"], simp) apply(auto simp: OclValid_def defined_def false_def true_def cp_def bot_fun_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_fun_def split:if_split_asm) apply(erule contrapos_pp [of "Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \) = bot"]) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject[symmetric], rule Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e, simp) apply(simp add: Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_option_def) apply(erule contrapos_pp [of "Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \) = null"]) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject[symmetric], rule Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e, simp) apply(simp add: Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse null_option_def) by (simp add: bot_option_def) lemma Set_inv_lemma' : assumes x_def : "\ \ \ X" and e_mem : "e \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\" shows "\ \ \ (\_. e)" apply(rule Set_inv_lemma[OF x_def, THEN ballE[where x = e]]) apply(simp add: foundation18') by(simp add: e_mem) lemma abs_rep_simp' : assumes S_all_def : "\ \ \ S" shows "Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\\\ = S \" proof - have discr_eq_false_true : "\\. (false \ = true \) = False" by(simp add: false_def true_def) show ?thesis apply(insert S_all_def, simp add: OclValid_def defined_def) apply(rule mp[OF Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_induct[where P = "\S. (if S = \ \ \ S = null \ then false \ else true \) = true \ \ Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e S\\\\ = S"]], rename_tac S') apply(simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse discr_eq_false_true) apply(case_tac S') apply(simp add: bot_fun_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def)+ apply(rename_tac S'', case_tac S'') apply(simp add: null_fun_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def)+ done qed lemma S_lift' : assumes S_all_def : "(\ :: '\ st) \ \ S" shows "\S'. (\a (_::'\ st). a) ` \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ = (\a (_::'\ st). \a\) ` S'" apply(rule_tac x = "(\a. \a\) ` \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\" in exI) apply(simp only: image_comp) apply(simp add: comp_def) apply(rule image_cong, fast) (* *) apply(drule Set_inv_lemma'[OF S_all_def]) by(case_tac x, (simp add: bot_option_def foundation18')+) lemma invalid_set_OclNot_defined [simp,code_unfold]:"\(invalid::('\,'\::null) Set) = false" by simp lemma null_set_OclNot_defined [simp,code_unfold]:"\(null::('\,'\::null) Set) = false" by(simp add: defined_def null_fun_def) lemma invalid_set_valid [simp,code_unfold]:"\(invalid::('\,'\::null) Set) = false" by simp lemma null_set_valid [simp,code_unfold]:"\(null::('\,'\::null) Set) = true" apply(simp add: valid_def null_fun_def bot_fun_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject,simp_all add: null_option_def bot_option_def) done text\... which means that we can have a type \('\,('\,('\) Integer) Set) Set\ corresponding exactly to Set(Set(Integer)) in OCL notation. Note that the parameter \'\\ still refers to the object universe; making the OCL semantics entirely parametric in the object universe makes it possible to study (and prove) its properties independently from a concrete class diagram.\ subsection\Definition: Strict Equality \label{sec:set-strict-equality}\ text\After the part of foundational operations on sets, we detail here equality on sets. Strong equality is inherited from the OCL core, but we have to consider the case of the strict equality. We decide to overload strict equality in the same way we do for other value's in OCL:\ overloading StrictRefEq \ "StrictRefEq :: [('\,'\::null)Set,('\,'\::null)Set] \ ('\)Boolean" begin definition StrictRefEq\<^sub>S\<^sub>e\<^sub>t : "(x::('\,'\::null)Set) \ y \ \ \. if (\ x) \ = true \ \ (\ y) \ = true \ then (x \ y)\ else invalid \" end text\One might object here that for the case of objects, this is an empty definition. The answer is no, we will restrain later on states and objects such that any object has its oid stored inside the object (so the ref, under which an object can be referenced in the store will represented in the object itself). For such well-formed stores that satisfy this invariant (the WFF-invariant), the referential equality and the strong equality---and therefore the strict equality on sets in the sense above---coincides.\ text\Property proof in terms of @{term "profile_bin\<^sub>S\<^sub>t\<^sub>r\<^sub>o\<^sub>n\<^sub>g\<^sub>E\<^sub>q_\<^sub>v_\<^sub>v"}\ interpretation StrictRefEq\<^sub>S\<^sub>e\<^sub>t : profile_bin\<^sub>S\<^sub>t\<^sub>r\<^sub>o\<^sub>n\<^sub>g\<^sub>E\<^sub>q_\<^sub>v_\<^sub>v "\ x y. (x::('\,'\::null)Set) \ y" by unfold_locales (auto simp: StrictRefEq\<^sub>S\<^sub>e\<^sub>t) subsection\Constants: mtSet\ definition mtSet::"('\,'\::null) Set" ("Set{}") where "Set{} \ (\ \. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\{}::'\ set\\ )" lemma mtSet_defined[simp,code_unfold]:"\(Set{}) = true" apply(rule ext, auto simp: mtSet_def defined_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_fun_def null_fun_def) by(simp_all add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_option_def) lemma mtSet_valid[simp,code_unfold]:"\(Set{}) = true" apply(rule ext,auto simp: mtSet_def valid_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_fun_def null_fun_def) by(simp_all add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_option_def) lemma mtSet_rep_set: "\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (Set{} \)\\ = {}" apply(simp add: mtSet_def, subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse) by(simp add: bot_option_def)+ lemma [simp,code_unfold]: "const Set{}" by(simp add: const_def mtSet_def) text\Note that the collection types in OCL allow for null to be included; however, there is the null-collection into which inclusion yields invalid.\ subsection\Definition: Including\ definition OclIncluding :: "[('\,'\::null) Set,('\,'\) val] \ ('\,'\) Set" where "OclIncluding x y = (\ \. if (\ x) \ = true \ \ (\ y) \ = true \ then Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \)\\ \ {y \} \\ else invalid \ )" notation OclIncluding ("_->including\<^sub>S\<^sub>e\<^sub>t'(_')") interpretation OclIncluding : profile_bin\<^sub>d_\<^sub>v OclIncluding "\x y. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\ \ {y}\\" proof - have A : "None \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(simp add: bot_option_def) have B : "\None\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(simp add: null_option_def bot_option_def) have C : "\x y. x \ \ \ x \ null \ y \ \ \ \\insert y \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(auto intro!:Set_inv_lemma[simplified OclValid_def defined_def false_def true_def null_fun_def bot_fun_def]) show "profile_bin\<^sub>d_\<^sub>v OclIncluding (\x y. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\ \ {y}\\)" apply unfold_locales apply(auto simp:OclIncluding_def bot_option_def null_option_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def) apply(erule_tac Q="Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\insert y \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\\\ = Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e None" in contrapos_pp) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject[OF C A]) apply(simp_all add: null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_option_def) apply(erule_tac Q="Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\insert y \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\\\ = Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \None\" in contrapos_pp) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject[OF C B]) apply(simp_all add: null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_option_def) done qed syntax "_OclFinset" :: "args => ('\,'a::null) Set" ("Set{(_)}") translations "Set{x, xs}" == "CONST OclIncluding (Set{xs}) x" "Set{x}" == "CONST OclIncluding (Set{}) x " subsection\Definition: Excluding\ definition OclExcluding :: "[('\,'\::null) Set,('\,'\) val] \ ('\,'\) Set" where "OclExcluding x y = (\ \. if (\ x) \ = true \ \ (\ y) \ = true \ then Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \)\\ - {y \} \\ else \ )" notation OclExcluding ("_->excluding\<^sub>S\<^sub>e\<^sub>t'(_')") lemma OclExcluding_inv: "(x:: Set('b::{null})) \ \ \ x \ null \ y \ \ \ \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\ - {y}\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" proof - fix X :: "'a state \ 'a state \ Set('b)" fix \ show "x \ \ \ x \ null \ y \ \ \ ?thesis" when "x = X \" by(simp add: that Set_inv_lemma[simplified OclValid_def defined_def null_fun_def bot_fun_def, of X \]) qed simp_all interpretation OclExcluding : profile_bin\<^sub>d_\<^sub>v OclExcluding "\x y. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\ - {y}\\" proof - have A : "None \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(simp add: bot_option_def) have B : "\None\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(simp add: null_option_def bot_option_def) show "profile_bin\<^sub>d_\<^sub>v OclExcluding (\x y. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x:: Set('b))\\ - {y}\\)" apply unfold_locales apply(auto simp:OclExcluding_def bot_option_def null_option_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def invalid_def) apply(erule_tac Q="Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\ - {y}\\ = Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e None" in contrapos_pp) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject[OF OclExcluding_inv A]) apply(simp_all add: null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_option_def) apply(erule_tac Q="Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\ - {y}\\ = Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \None\" in contrapos_pp) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject[OF OclExcluding_inv B]) apply(simp_all add: null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_option_def) done qed subsection\Definition: Includes\ definition OclIncludes :: "[('\,'\::null) Set,('\,'\) val] \ '\ Boolean" where "OclIncludes x y = (\ \. if (\ x) \ = true \ \ (\ y) \ = true \ then \\(y \) \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \)\\ \\ else \ )" notation OclIncludes ("_->includes\<^sub>S\<^sub>e\<^sub>t'(_')" (*[66,65]65*)) interpretation OclIncludes : profile_bin\<^sub>d_\<^sub>v OclIncludes "\x y. \\y \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\\\" by(unfold_locales, auto simp:OclIncludes_def bot_option_def null_option_def invalid_def) subsection\Definition: Excludes\ definition OclExcludes :: "[('\,'\::null) Set,('\,'\) val] \ '\ Boolean" where "OclExcludes x y = (not(OclIncludes x y))" notation OclExcludes ("_->excludes\<^sub>S\<^sub>e\<^sub>t'(_')" (*[66,65]65*)) text\The case of the size definition is somewhat special, we admit explicitly in Featherweight OCL the possibility of infinite sets. For the size definition, this requires an extra condition that assures that the cardinality of the set is actually a defined integer.\ interpretation OclExcludes : profile_bin\<^sub>d_\<^sub>v OclExcludes "\x y. \\y \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\\\" by(unfold_locales, auto simp:OclExcludes_def OclIncludes_def OclNot_def bot_option_def null_option_def invalid_def) subsection\Definition: Size\ definition OclSize :: "('\,'\::null)Set \ '\ Integer" where "OclSize x = (\ \. if (\ x) \ = true \ \ finite(\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \)\\) then \\ int(card \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \)\\) \\ else \ )" notation (* standard ascii syntax *) OclSize ("_->size\<^sub>S\<^sub>e\<^sub>t'(')" (*[66]*)) text\The following definition follows the requirement of the standard to treat null as neutral element of sets. It is a well-documented exception from the general strictness rule and the rule that the distinguished argument self should be non-null.\ (*TODO Locale - Equivalent*) subsection\Definition: IsEmpty\ definition OclIsEmpty :: "('\,'\::null) Set \ '\ Boolean" where "OclIsEmpty x = ((\ x and not (\ x)) or ((OclSize x) \ \))" notation OclIsEmpty ("_->isEmpty\<^sub>S\<^sub>e\<^sub>t'(')" (*[66]*)) (*TODO Locale - Equivalent*) subsection\Definition: NotEmpty\ definition OclNotEmpty :: "('\,'\::null) Set \ '\ Boolean" where "OclNotEmpty x = not(OclIsEmpty x)" notation OclNotEmpty ("_->notEmpty\<^sub>S\<^sub>e\<^sub>t'(')" (*[66]*)) (*TODO Locale - Equivalent*) subsection\Definition: Any\ (* Slight breach of naming convention in order to avoid naming conflict on constant.*) definition OclANY :: "[('\,'\::null) Set] \ ('\,'\) val" where "OclANY x = (\ \. if (\ x) \ = true \ then if (\ x and OclNotEmpty x) \ = true \ then SOME y. y \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \)\\ else null \ else \ )" notation OclANY ("_->any\<^sub>S\<^sub>e\<^sub>t'(')") (*TODO Locale - Equivalent*) (* actually, this definition covers only: X->any\<^sub>S\<^sub>e\<^sub>t(true) of the standard, which foresees a (totally correct) high-level definition source->any\<^sub>S\<^sub>e\<^sub>t(iterator | body) = source->select(iterator | body)->asSequence()->first(). Since we don't have sequences, we have to go for a direct---restricted---definition. *) subsection\Definition: Forall\ text\The definition of OclForall mimics the one of @{term "OclAnd"}: OclForall is not a strict operation.\ definition OclForall :: "[('\,'\::null)Set,('\,'\)val\('\)Boolean] \ '\ Boolean" where "OclForall S P = (\ \. if (\ S) \ = true \ then if (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\. P(\ _. x) \ = false \) then false \ else if (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\. P(\ _. x) \ = invalid \) then invalid \ else if (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\. P(\ _. x) \ = null \) then null \ else true \ else \)" syntax "_OclForallSet" :: "[('\,'\::null) Set,id,('\)Boolean] \ '\ Boolean" ("(_)->forAll\<^sub>S\<^sub>e\<^sub>t'(_|_')") translations "X->forAll\<^sub>S\<^sub>e\<^sub>t(x | P)" == "CONST UML_Set.OclForall X (%x. P)" (*TODO Locale - Equivalent*) subsection\Definition: Exists\ text\Like OclForall, OclExists is also not strict.\ definition OclExists :: "[('\,'\::null) Set,('\,'\)val\('\)Boolean] \ '\ Boolean" where "OclExists S P = not(UML_Set.OclForall S (\ X. not (P X)))" syntax "_OclExistSet" :: "[('\,'\::null) Set,id,('\)Boolean] \ '\ Boolean" ("(_)->exists\<^sub>S\<^sub>e\<^sub>t'(_|_')") translations "X->exists\<^sub>S\<^sub>e\<^sub>t(x | P)" == "CONST UML_Set.OclExists X (%x. P)" (*TODO Locale - Equivalent*) subsection\Definition: Iterate\ definition OclIterate :: "[('\,'\::null) Set,('\,'\::null)val, ('\,'\)val\('\,'\)val\('\,'\)val] \ ('\,'\)val" where "OclIterate S A F = (\ \. if (\ S) \ = true \ \ (\ A) \ = true \ \ finite\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ then (Finite_Set.fold (F) (A) ((\a \. a) ` \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\))\ else \)" syntax "_OclIterateSet" :: "[('\,'\::null) Set, idt, idt, '\, '\] => ('\,'\)val" ("_ ->iterate\<^sub>S\<^sub>e\<^sub>t'(_;_=_ | _')" (*[71,100,70]50*)) translations "X->iterate\<^sub>S\<^sub>e\<^sub>t(a; x = A | P)" == "CONST OclIterate X A (%a. (% x. P))" (*TODO Locale - Equivalent*) subsection\Definition: Select\ definition OclSelect :: "[('\,'\::null)Set,('\,'\)val\('\)Boolean] \ ('\,'\)Set" where "OclSelect S P = (\\. if (\ S) \ = true \ then if (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\. P(\ _. x) \ = invalid \) then invalid \ else Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\{x\\\ Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\. P (\_. x) \ \ false \}\\ else invalid \)" syntax "_OclSelectSet" :: "[('\,'\::null) Set,id,('\)Boolean] \ '\ Boolean" ("(_)->select\<^sub>S\<^sub>e\<^sub>t'(_|_')") translations "X->select\<^sub>S\<^sub>e\<^sub>t(x | P)" == "CONST OclSelect X (% x. P)" (*TODO Locale - Equivalent*) subsection\Definition: Reject\ definition OclReject :: "[('\,'\::null)Set,('\,'\)val\('\)Boolean] \ ('\,'\::null)Set" where "OclReject S P = OclSelect S (not o P)" syntax "_OclRejectSet" :: "[('\,'\::null) Set,id,('\)Boolean] \ '\ Boolean" ("(_)->reject\<^sub>S\<^sub>e\<^sub>t'(_|_')") translations "X->reject\<^sub>S\<^sub>e\<^sub>t(x | P)" == "CONST OclReject X (% x. P)" (*TODO Locale - Equivalent*) subsection\Definition: IncludesAll\ definition OclIncludesAll :: "[('\,'\::null) Set,('\,'\) Set] \ '\ Boolean" where "OclIncludesAll x y = (\ \. if (\ x) \ = true \ \ (\ y) \ = true \ then \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (y \)\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \)\\ \\ else \ )" notation OclIncludesAll ("_->includesAll\<^sub>S\<^sub>e\<^sub>t'(_')" (*[66,65]65*)) interpretation OclIncludesAll : profile_bin\<^sub>d_\<^sub>d OclIncludesAll "\x y. \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e y\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\\\" by(unfold_locales, auto simp:OclIncludesAll_def bot_option_def null_option_def invalid_def) subsection\Definition: ExcludesAll\ definition OclExcludesAll :: "[('\,'\::null) Set,('\,'\) Set] \ '\ Boolean" where "OclExcludesAll x y = (\ \. if (\ x) \ = true \ \ (\ y) \ = true \ then \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (y \)\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \)\\ = {} \\ else \ )" notation OclExcludesAll ("_->excludesAll\<^sub>S\<^sub>e\<^sub>t'(_')" (*[66,65]65*)) interpretation OclExcludesAll : profile_bin\<^sub>d_\<^sub>d OclExcludesAll "\x y. \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e y\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\ = {}\\" by(unfold_locales, auto simp:OclExcludesAll_def bot_option_def null_option_def invalid_def) subsection\Definition: Union\ definition OclUnion :: "[('\,'\::null) Set,('\,'\) Set] \ ('\,'\) Set" where "OclUnion x y = (\ \. if (\ x) \ = true \ \ (\ y) \ = true \ then Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (y \)\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \)\\ \\ else \ )" notation OclUnion ("_->union\<^sub>S\<^sub>e\<^sub>t'(_')" (*[66,65]65*)) lemma OclUnion_inv: "(x:: Set('b::{null})) \ \ \ x \ null \ y \ \ \ y \ null \ \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e y\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" proof - fix X Y :: "'a state \ 'a state \ Set('b)" fix \ show "x \ \ \ x \ null \ y \ \ \ y \ null \ ?thesis" when "x = X \" "y = Y \" by(auto simp: that, insert Set_inv_lemma[simplified OclValid_def defined_def null_fun_def bot_fun_def, of Y \] Set_inv_lemma[simplified OclValid_def defined_def null_fun_def bot_fun_def, of X \], auto) qed simp_all interpretation OclUnion : profile_bin\<^sub>d_\<^sub>d OclUnion "\x y. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e y\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\\\" proof - have A : "None \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(simp add: bot_option_def) have B : "\None\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(simp add: null_option_def bot_option_def) show "profile_bin\<^sub>d_\<^sub>d OclUnion (\x y. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e y\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\\\)" apply unfold_locales apply(auto simp:OclUnion_def bot_option_def null_option_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def invalid_def) apply(erule_tac Q="Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e y\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\\\ = Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e None" in contrapos_pp) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject[OF OclUnion_inv A]) apply(simp_all add: null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_option_def) apply(erule_tac Q="Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e y\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\\\ = Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \None\" in contrapos_pp) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject[OF OclUnion_inv B]) apply(simp_all add: null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_option_def) done qed subsection\Definition: Intersection\ definition OclIntersection :: "[('\,'\::null) Set,('\,'\) Set] \ ('\,'\) Set" where "OclIntersection x y = (\ \. if (\ x) \ = true \ \ (\ y) \ = true \ then Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (y \)\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \)\\\\ else \ )" notation OclIntersection("_->intersection\<^sub>S\<^sub>e\<^sub>t'(_')" (*[71,70]70*)) lemma OclIntersection_inv: "(x:: Set('b::{null})) \ \ \ x \ null \ y \ \ \ y \ null \ \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e y\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" proof - fix X Y :: "'a state \ 'a state \ Set('b)" fix \ show "x \ \ \ x \ null \ y \ \ \ y \ null \ ?thesis" when "x = X \" "y = Y \" by(auto simp: that, insert Set_inv_lemma[simplified OclValid_def defined_def null_fun_def bot_fun_def, of Y \] Set_inv_lemma[simplified OclValid_def defined_def null_fun_def bot_fun_def, of X \], auto) qed simp_all interpretation OclIntersection : profile_bin\<^sub>d_\<^sub>d OclIntersection "\x y. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e y\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\\\" proof - have A : "None \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(simp add: bot_option_def) have B : "\None\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(simp add: null_option_def bot_option_def) show "profile_bin\<^sub>d_\<^sub>d OclIntersection (\x y. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e y\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\\\)" apply unfold_locales apply(auto simp:OclIntersection_def bot_option_def null_option_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def invalid_def) apply(erule_tac Q="Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e y\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\\\ = Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e None" in contrapos_pp) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject[OF OclIntersection_inv A]) apply(simp_all add: null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_option_def) apply(erule_tac Q="Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e y\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\\\ = Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \None\" in contrapos_pp) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject[OF OclIntersection_inv B]) apply(simp_all add: null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_option_def) done qed subsection\Definition (future operators)\ consts (* abstract set collection operations *) OclCount :: "[('\,'\::null) Set,('\,'\) Set] \ '\ Integer" OclSum :: " ('\,'\::null) Set \ '\ Integer" notation OclCount ("_->count\<^sub>S\<^sub>e\<^sub>t'(_')" (*[66,65]65*)) notation OclSum ("_->sum\<^sub>S\<^sub>e\<^sub>t'(')" (*[66]*)) subsection\Logical Properties\ text\OclIncluding\ lemma OclIncluding_valid_args_valid: "(\ \ \(X->including\<^sub>S\<^sub>e\<^sub>t(x))) = ((\ \(\ X)) \ (\ \(\ x)))" by (metis (hide_lams, no_types) OclIncluding.def_valid_then_def OclIncluding.defined_args_valid) lemma OclIncluding_valid_args_valid''[simp,code_unfold]: "\(X->including\<^sub>S\<^sub>e\<^sub>t(x)) = ((\ X) and (\ x))" by (simp add: OclIncluding.def_valid_then_def) text\etc. etc.\ text_raw\\isatagafp\ text\OclExcluding\ lemma OclExcluding_valid_args_valid: "(\ \ \(X->excluding\<^sub>S\<^sub>e\<^sub>t(x))) = ((\ \(\ X)) \ (\ \(\ x)))" by (metis OclExcluding.def_valid_then_def OclExcluding.defined_args_valid) lemma OclExcluding_valid_args_valid''[simp,code_unfold]: "\(X->excluding\<^sub>S\<^sub>e\<^sub>t(x)) = ((\ X) and (\ x))" by (simp add: OclExcluding.def_valid_then_def) text\OclIncludes\ lemma OclIncludes_valid_args_valid: "(\ \ \(X->includes\<^sub>S\<^sub>e\<^sub>t(x))) = ((\ \(\ X)) \ (\ \(\ x)))" by (simp add: OclIncludes.def_valid_then_def foundation10') lemma OclIncludes_valid_args_valid''[simp,code_unfold]: "\(X->includes\<^sub>S\<^sub>e\<^sub>t(x)) = ((\ X) and (\ x))" by (simp add: OclIncludes.def_valid_then_def) text\OclExcludes\ lemma OclExcludes_valid_args_valid: "(\ \ \(X->excludes\<^sub>S\<^sub>e\<^sub>t(x))) = ((\ \(\ X)) \ (\ \(\ x)))" by (simp add: OclExcludes.def_valid_then_def foundation10') lemma OclExcludes_valid_args_valid''[simp,code_unfold]: "\(X->excludes\<^sub>S\<^sub>e\<^sub>t(x)) = ((\ X) and (\ x))" by (simp add: OclExcludes.def_valid_then_def) text\OclSize\ lemma OclSize_defined_args_valid: "\ \ \ (X->size\<^sub>S\<^sub>e\<^sub>t()) \ \ \ \ X" by(auto simp: OclSize_def OclValid_def true_def valid_def false_def StrongEq_def defined_def invalid_def bot_fun_def null_fun_def split: bool.split_asm HOL.if_split_asm option.split) lemma OclSize_infinite: assumes non_finite:"\ \ not(\(S->size\<^sub>S\<^sub>e\<^sub>t()))" shows "(\ \ not(\(S))) \ \ finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\" apply(insert non_finite, simp) apply(rule impI) apply(simp add: OclSize_def OclValid_def defined_def) apply(case_tac "finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\", simp_all add:null_fun_def null_option_def bot_fun_def bot_option_def) done lemma "\ \ \ X \ \ finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ \ \ \ \ \ (X->size\<^sub>S\<^sub>e\<^sub>t())" by(simp add: OclSize_def OclValid_def defined_def bot_fun_def false_def true_def) lemma size_defined: assumes X_finite: "\\. finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\" shows "\ (X->size\<^sub>S\<^sub>e\<^sub>t()) = \ X" apply(rule ext, simp add: cp_defined[of "X->size\<^sub>S\<^sub>e\<^sub>t()"] OclSize_def) apply(simp add: defined_def bot_option_def bot_fun_def null_option_def null_fun_def X_finite) done lemma size_defined': assumes X_finite: "finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\" shows "(\ \ \ (X->size\<^sub>S\<^sub>e\<^sub>t())) = (\ \ \ X)" apply(simp add: cp_defined[of "X->size\<^sub>S\<^sub>e\<^sub>t()"] OclSize_def OclValid_def) apply(simp add: defined_def bot_option_def bot_fun_def null_option_def null_fun_def X_finite) done text\OclIsEmpty\ lemma OclIsEmpty_defined_args_valid:"\ \ \ (X->isEmpty\<^sub>S\<^sub>e\<^sub>t()) \ \ \ \ X" apply(auto simp: OclIsEmpty_def OclValid_def defined_def valid_def false_def true_def bot_fun_def null_fun_def OclAnd_def OclOr_def OclNot_def split: if_split_asm) apply(case_tac "(X->size\<^sub>S\<^sub>e\<^sub>t() \ \) \", simp add: bot_option_def, simp, rename_tac x) apply(case_tac x, simp add: null_option_def bot_option_def, simp) apply(simp add: OclSize_def StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r valid_def) by (metis (hide_lams, no_types) bot_fun_def OclValid_def defined_def foundation2 invalid_def) lemma "\ \ \ (null->isEmpty\<^sub>S\<^sub>e\<^sub>t())" by(auto simp: OclIsEmpty_def OclValid_def defined_def valid_def false_def true_def bot_fun_def null_fun_def OclAnd_def OclOr_def OclNot_def null_is_valid split: if_split_asm) lemma OclIsEmpty_infinite: "\ \ \ X \ \ finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ \ \ \ \ \ (X->isEmpty\<^sub>S\<^sub>e\<^sub>t())" apply(auto simp: OclIsEmpty_def OclValid_def defined_def valid_def false_def true_def bot_fun_def null_fun_def OclAnd_def OclOr_def OclNot_def split: if_split_asm) apply(case_tac "(X->size\<^sub>S\<^sub>e\<^sub>t() \ \) \", simp add: bot_option_def, simp, rename_tac x) apply(case_tac x, simp add: null_option_def bot_option_def, simp) by(simp add: OclSize_def StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r valid_def bot_fun_def false_def true_def invalid_def) text\OclNotEmpty\ lemma OclNotEmpty_defined_args_valid:"\ \ \ (X->notEmpty\<^sub>S\<^sub>e\<^sub>t()) \ \ \ \ X" by (metis (hide_lams, no_types) OclNotEmpty_def OclNot_defargs OclNot_not foundation6 foundation9 OclIsEmpty_defined_args_valid) lemma "\ \ \ (null->notEmpty\<^sub>S\<^sub>e\<^sub>t())" by (metis (hide_lams, no_types) OclNotEmpty_def OclAnd_false1 OclAnd_idem OclIsEmpty_def OclNot3 OclNot4 OclOr_def defined2 defined4 transform1 valid2) lemma OclNotEmpty_infinite: "\ \ \ X \ \ finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ \ \ \ \ \ (X->notEmpty\<^sub>S\<^sub>e\<^sub>t())" apply(simp add: OclNotEmpty_def) apply(drule OclIsEmpty_infinite, simp) by (metis OclNot_defargs OclNot_not foundation6 foundation9) lemma OclNotEmpty_has_elt : "\ \ \ X \ \ \ X->notEmpty\<^sub>S\<^sub>e\<^sub>t() \ \e. e \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\" apply(simp add: OclNotEmpty_def OclIsEmpty_def deMorgan1 deMorgan2, drule foundation5) apply(subst (asm) (2) OclNot_def, simp add: OclValid_def StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r StrongEq_def split: if_split_asm) prefer 2 apply(simp add: invalid_def bot_option_def true_def) apply(simp add: OclSize_def valid_def split: if_split_asm, simp_all add: false_def true_def bot_option_def bot_fun_def OclInt0_def) by (metis equals0I) text\OclANY\ lemma OclANY_defined_args_valid: "\ \ \ (X->any\<^sub>S\<^sub>e\<^sub>t()) \ \ \ \ X" by(auto simp: OclANY_def OclValid_def true_def valid_def false_def StrongEq_def defined_def invalid_def bot_fun_def null_fun_def OclAnd_def split: bool.split_asm HOL.if_split_asm option.split) lemma "\ \ \ X \ \ \ X->isEmpty\<^sub>S\<^sub>e\<^sub>t() \ \ \ \ \ (X->any\<^sub>S\<^sub>e\<^sub>t())" apply(simp add: OclANY_def OclValid_def) apply(subst cp_defined, subst cp_OclAnd, simp add: OclNotEmpty_def, subst (1 2) cp_OclNot, simp add: cp_OclNot[symmetric] cp_OclAnd[symmetric] cp_defined[symmetric], simp add: false_def true_def) by(drule foundation20[simplified OclValid_def true_def], simp) lemma OclANY_valid_args_valid: "(\ \ \(X->any\<^sub>S\<^sub>e\<^sub>t())) = (\ \ \ X)" proof - have A: "(\ \ \(X->any\<^sub>S\<^sub>e\<^sub>t())) \ ((\ \(\ X)))" by(auto simp: OclANY_def OclValid_def true_def valid_def false_def StrongEq_def defined_def invalid_def bot_fun_def null_fun_def split: bool.split_asm HOL.if_split_asm option.split) have B: "(\ \(\ X)) \ (\ \ \(X->any\<^sub>S\<^sub>e\<^sub>t()))" apply(auto simp: OclANY_def OclValid_def true_def false_def StrongEq_def defined_def invalid_def valid_def bot_fun_def null_fun_def bot_option_def null_option_def null_is_valid OclAnd_def split: bool.split_asm HOL.if_split_asm option.split) apply(frule Set_inv_lemma[OF foundation16[THEN iffD2], OF conjI], simp) apply(subgoal_tac "(\ X) \ = true \") prefer 2 apply (metis (hide_lams, no_types) OclValid_def foundation16) apply(simp add: true_def, drule OclNotEmpty_has_elt[simplified OclValid_def true_def], simp) by(erule exE, insert someI2[where Q = "\x. x \ \" and P = "\y. y \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\"], simp) show ?thesis by(auto dest:A intro:B) qed lemma OclANY_valid_args_valid''[simp,code_unfold]: "\(X->any\<^sub>S\<^sub>e\<^sub>t()) = (\ X)" by(auto intro!: OclANY_valid_args_valid transform2_rev) (* and higher order ones : forall, exists, iterate, select, reject... *) text_raw\\endisatagafp\ subsection\Execution Laws with Invalid or Null or Infinite Set as Argument\ text\OclIncluding\ (* properties already generated by the corresponding locale *) text\OclExcluding\ (* properties already generated by the corresponding locale *) text\OclIncludes\ (* properties already generated by the corresponding locale *) text\OclExcludes\ (* properties already generated by the corresponding locale *) text\OclSize\ lemma OclSize_invalid[simp,code_unfold]:"(invalid->size\<^sub>S\<^sub>e\<^sub>t()) = invalid" by(simp add: bot_fun_def OclSize_def invalid_def defined_def valid_def false_def true_def) lemma OclSize_null[simp,code_unfold]:"(null->size\<^sub>S\<^sub>e\<^sub>t()) = invalid" by(rule ext, simp add: bot_fun_def null_fun_def null_is_valid OclSize_def invalid_def defined_def valid_def false_def true_def) text\OclIsEmpty\ lemma OclIsEmpty_invalid[simp,code_unfold]:"(invalid->isEmpty\<^sub>S\<^sub>e\<^sub>t()) = invalid" by(simp add: OclIsEmpty_def) lemma OclIsEmpty_null[simp,code_unfold]:"(null->isEmpty\<^sub>S\<^sub>e\<^sub>t()) = true" by(simp add: OclIsEmpty_def) text\OclNotEmpty\ lemma OclNotEmpty_invalid[simp,code_unfold]:"(invalid->notEmpty\<^sub>S\<^sub>e\<^sub>t()) = invalid" by(simp add: OclNotEmpty_def) lemma OclNotEmpty_null[simp,code_unfold]:"(null->notEmpty\<^sub>S\<^sub>e\<^sub>t()) = false" by(simp add: OclNotEmpty_def) text\OclANY\ lemma OclANY_invalid[simp,code_unfold]:"(invalid->any\<^sub>S\<^sub>e\<^sub>t()) = invalid" by(simp add: bot_fun_def OclANY_def invalid_def defined_def valid_def false_def true_def) lemma OclANY_null[simp,code_unfold]:"(null->any\<^sub>S\<^sub>e\<^sub>t()) = null" by(simp add: OclANY_def false_def true_def) text\OclForall\ lemma OclForall_invalid[simp,code_unfold]:"invalid->forAll\<^sub>S\<^sub>e\<^sub>t(a| P a) = invalid" by(simp add: bot_fun_def invalid_def OclForall_def defined_def valid_def false_def true_def) lemma OclForall_null[simp,code_unfold]:"null->forAll\<^sub>S\<^sub>e\<^sub>t(a | P a) = invalid" by(simp add: bot_fun_def invalid_def OclForall_def defined_def valid_def false_def true_def) text\OclExists\ lemma OclExists_invalid[simp,code_unfold]:"invalid->exists\<^sub>S\<^sub>e\<^sub>t(a| P a) = invalid" by(simp add: OclExists_def) lemma OclExists_null[simp,code_unfold]:"null->exists\<^sub>S\<^sub>e\<^sub>t(a | P a) = invalid" by(simp add: OclExists_def) text\OclIterate\ lemma OclIterate_invalid[simp,code_unfold]:"invalid->iterate\<^sub>S\<^sub>e\<^sub>t(a; x = A | P a x) = invalid" by(simp add: bot_fun_def invalid_def OclIterate_def defined_def valid_def false_def true_def) lemma OclIterate_null[simp,code_unfold]:"null->iterate\<^sub>S\<^sub>e\<^sub>t(a; x = A | P a x) = invalid" by(simp add: bot_fun_def invalid_def OclIterate_def defined_def valid_def false_def true_def) lemma OclIterate_invalid_args[simp,code_unfold]:"S->iterate\<^sub>S\<^sub>e\<^sub>t(a; x = invalid | P a x) = invalid" by(simp add: bot_fun_def invalid_def OclIterate_def defined_def valid_def false_def true_def) text\An open question is this ...\ lemma (*OclIterate_null_args[simp,code_unfold]:*) "S->iterate\<^sub>S\<^sub>e\<^sub>t(a; x = null | P a x) = invalid" oops (* In the definition above, this does not hold in general. And I believe, this is how it should be ... *) lemma OclIterate_infinite: assumes non_finite: "\ \ not(\(S->size\<^sub>S\<^sub>e\<^sub>t()))" shows "(OclIterate S A F) \ = invalid \" apply(insert non_finite [THEN OclSize_infinite]) apply(subst (asm) foundation9, simp) by(metis OclIterate_def OclValid_def invalid_def) text\OclSelect\ lemma OclSelect_invalid[simp,code_unfold]:"invalid->select\<^sub>S\<^sub>e\<^sub>t(a | P a) = invalid" by(simp add: bot_fun_def invalid_def OclSelect_def defined_def valid_def false_def true_def) lemma OclSelect_null[simp,code_unfold]:"null->select\<^sub>S\<^sub>e\<^sub>t(a | P a) = invalid" by(simp add: bot_fun_def invalid_def OclSelect_def defined_def valid_def false_def true_def) text\OclReject\ lemma OclReject_invalid[simp,code_unfold]:"invalid->reject\<^sub>S\<^sub>e\<^sub>t(a | P a) = invalid" by(simp add: OclReject_def) lemma OclReject_null[simp,code_unfold]:"null->reject\<^sub>S\<^sub>e\<^sub>t(a | P a) = invalid" by(simp add: OclReject_def) text_raw\\isatagafp\ subsubsection\Context Passing\ lemma cp_OclIncludes1: "(X->includes\<^sub>S\<^sub>e\<^sub>t(x)) \ = (X->includes\<^sub>S\<^sub>e\<^sub>t(\ _. x \)) \" by(auto simp: OclIncludes_def StrongEq_def invalid_def cp_defined[symmetric] cp_valid[symmetric]) lemma cp_OclSize: "X->size\<^sub>S\<^sub>e\<^sub>t() \ = ((\_. X \)->size\<^sub>S\<^sub>e\<^sub>t()) \" by(simp add: OclSize_def cp_defined[symmetric]) lemma cp_OclIsEmpty: "X->isEmpty\<^sub>S\<^sub>e\<^sub>t() \ = ((\_. X \)->isEmpty\<^sub>S\<^sub>e\<^sub>t()) \" apply(simp only: OclIsEmpty_def) apply(subst (2) cp_OclOr, subst cp_OclAnd, subst cp_OclNot, subst StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.cp0) by(simp add: cp_defined[symmetric] cp_valid[symmetric] StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.cp0[symmetric] cp_OclSize[symmetric] cp_OclNot[symmetric] cp_OclAnd[symmetric] cp_OclOr[symmetric]) lemma cp_OclNotEmpty: "X->notEmpty\<^sub>S\<^sub>e\<^sub>t() \ = ((\_. X \)->notEmpty\<^sub>S\<^sub>e\<^sub>t()) \" apply(simp only: OclNotEmpty_def) apply(subst (2) cp_OclNot) by(simp add: cp_OclNot[symmetric] cp_OclIsEmpty[symmetric]) lemma cp_OclANY: "X->any\<^sub>S\<^sub>e\<^sub>t() \ = ((\_. X \)->any\<^sub>S\<^sub>e\<^sub>t()) \" apply(simp only: OclANY_def) apply(subst (2) cp_OclAnd) by(simp only: cp_OclAnd[symmetric] cp_defined[symmetric] cp_valid[symmetric] cp_OclNotEmpty[symmetric]) lemma cp_OclForall: "(S->forAll\<^sub>S\<^sub>e\<^sub>t(x | P x)) \ = ((\ _. S \)->forAll\<^sub>S\<^sub>e\<^sub>t(x | P (\ _. x \))) \" by(simp add: OclForall_def cp_defined[symmetric]) (* first-order version !*) lemma cp_OclForall1 [simp,intro!]: "cp S \ cp (\X. ((S X)->forAll\<^sub>S\<^sub>e\<^sub>t(x | P x)))" apply(simp add: cp_def) apply(erule exE, rule exI, intro allI) apply(erule_tac x=X in allE) by(subst cp_OclForall, simp) lemma (*cp_OclForall2 [simp,intro!]:*) "cp (\X St x. P (\\. x) X St) \ cp S \ cp (\X. (S X)->forAll\<^sub>S\<^sub>e\<^sub>t(x|P x X)) " apply(simp only: cp_def) oops lemma (*cp_OclForall:*) "cp S \ (\ x. cp(P x)) \ cp(\X. ((S X)->forAll\<^sub>S\<^sub>e\<^sub>t(x | P x X)))" oops lemma cp_OclExists: "(S->exists\<^sub>S\<^sub>e\<^sub>t(x | P x)) \ = ((\ _. S \)->exists\<^sub>S\<^sub>e\<^sub>t(x | P (\ _. x \))) \" by(simp add: OclExists_def OclNot_def, subst cp_OclForall, simp) (* first-order version !*) lemma cp_OclExists1 [simp,intro!]: "cp S \ cp (\X. ((S X)->exists\<^sub>S\<^sub>e\<^sub>t(x | P x)))" apply(simp add: cp_def) apply(erule exE, rule exI, intro allI) apply(erule_tac x=X in allE) by(subst cp_OclExists,simp) lemma cp_OclIterate: "(X->iterate\<^sub>S\<^sub>e\<^sub>t(a; x = A | P a x)) \ = ((\ _. X \)->iterate\<^sub>S\<^sub>e\<^sub>t(a; x = A | P a x)) \" by(simp add: OclIterate_def cp_defined[symmetric]) lemma cp_OclSelect: "(X->select\<^sub>S\<^sub>e\<^sub>t(a | P a)) \ = ((\ _. X \)->select\<^sub>S\<^sub>e\<^sub>t(a | P a)) \" by(simp add: OclSelect_def cp_defined[symmetric]) lemma cp_OclReject: "(X->reject\<^sub>S\<^sub>e\<^sub>t(a | P a)) \ = ((\ _. X \)->reject\<^sub>S\<^sub>e\<^sub>t(a | P a)) \" by(simp add: OclReject_def, subst cp_OclSelect, simp) lemmas cp_intro''\<^sub>S\<^sub>e\<^sub>t[intro!,simp,code_unfold] = cp_OclSize [THEN allI[THEN allI[THEN cpI1], of "OclSize"]] cp_OclIsEmpty [THEN allI[THEN allI[THEN cpI1], of "OclIsEmpty"]] cp_OclNotEmpty [THEN allI[THEN allI[THEN cpI1], of "OclNotEmpty"]] cp_OclANY [THEN allI[THEN allI[THEN cpI1], of "OclANY"]] subsubsection\Const\ lemma const_OclIncluding[simp,code_unfold] : assumes const_x : "const x" and const_S : "const S" shows "const (S->including\<^sub>S\<^sub>e\<^sub>t(x))" proof - have A:"\\ \'. \ (\ \ \ x) \ (S->including\<^sub>S\<^sub>e\<^sub>t(x) \) = (S->including\<^sub>S\<^sub>e\<^sub>t(x) \')" apply(simp add: foundation18) apply(erule const_subst[OF const_x const_invalid],simp_all) by(rule const_charn[OF const_invalid]) have B: "\ \ \'. \ (\ \ \ S) \ (S->including\<^sub>S\<^sub>e\<^sub>t(x) \) = (S->including\<^sub>S\<^sub>e\<^sub>t(x) \')" apply(simp add: foundation16', elim disjE) apply(erule const_subst[OF const_S const_invalid],simp_all) apply(rule const_charn[OF const_invalid]) apply(erule const_subst[OF const_S const_null],simp_all) by(rule const_charn[OF const_invalid]) show ?thesis apply(simp only: const_def,intro allI, rename_tac \ \') apply(case_tac "\ (\ \ \ x)", simp add: A) apply(case_tac "\ (\ \ \ S)", simp_all add: B) apply(frule_tac \'1= \' in const_OclValid2[OF const_x, THEN iffD1]) apply(frule_tac \'1= \' in const_OclValid1[OF const_S, THEN iffD1]) apply(simp add: OclIncluding_def OclValid_def) apply(subst const_charn[OF const_x]) apply(subst const_charn[OF const_S]) by simp qed text_raw\\endisatagafp\ subsection\General Algebraic Execution Rules\ subsubsection\Execution Rules on Including\ lemma OclIncluding_finite_rep_set : assumes X_def : "\ \ \ X" and x_val : "\ \ \ x" shows "finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X->including\<^sub>S\<^sub>e\<^sub>t(x) \)\\ = finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\" proof - have C : "\\insert (x \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(insert X_def x_val, frule Set_inv_lemma, simp add: foundation18 invalid_def) show "?thesis" by(insert X_def x_val, auto simp: OclIncluding_def Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse[OF C] dest: foundation13[THEN iffD2, THEN foundation22[THEN iffD1]]) qed lemma OclIncluding_rep_set: assumes S_def: "\ \ \ S" shows "\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S->including\<^sub>S\<^sub>e\<^sub>t(\_. \\x\\) \)\\ = insert \\x\\ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\" apply(simp add: OclIncluding_def S_def[simplified OclValid_def]) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def null_option_def) apply(insert Set_inv_lemma[OF S_def], metis bot_option_def not_Some_eq) by(simp) lemma OclIncluding_notempty_rep_set: assumes X_def: "\ \ \ X" and a_val: "\ \ \ a" shows "\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X->including\<^sub>S\<^sub>e\<^sub>t(a) \)\\ \ {}" apply(simp add: OclIncluding_def X_def[simplified OclValid_def] a_val[simplified OclValid_def]) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def null_option_def) apply(insert Set_inv_lemma[OF X_def], metis a_val foundation18') by(simp) lemma OclIncluding_includes0: assumes "\ \ X->includes\<^sub>S\<^sub>e\<^sub>t(x)" shows "X->including\<^sub>S\<^sub>e\<^sub>t(x) \ = X \" proof - have includes_def: "\ \ X->includes\<^sub>S\<^sub>e\<^sub>t(x) \ \ \ \ X" by (metis bot_fun_def OclIncludes_def OclValid_def defined3 foundation16) have includes_val: "\ \ X->includes\<^sub>S\<^sub>e\<^sub>t(x) \ \ \ \ x" using foundation5 foundation6 by fastforce show ?thesis apply(insert includes_def[OF assms] includes_val[OF assms] assms, simp add: OclIncluding_def OclIncludes_def OclValid_def true_def) apply(drule insert_absorb, simp, subst abs_rep_simp') by(simp_all add: OclValid_def true_def) qed lemma OclIncluding_includes: assumes "\ \ X->includes\<^sub>S\<^sub>e\<^sub>t(x)" shows "\ \ X->including\<^sub>S\<^sub>e\<^sub>t(x) \ X" by(simp add: StrongEq_def OclValid_def true_def OclIncluding_includes0[OF assms]) lemma OclIncluding_commute0 : assumes S_def : "\ \ \ S" and i_val : "\ \ \ i" and j_val : "\ \ \ j" shows "\ \ ((S :: ('\, 'a::null) Set)->including\<^sub>S\<^sub>e\<^sub>t(i)->including\<^sub>S\<^sub>e\<^sub>t(j) \ (S->including\<^sub>S\<^sub>e\<^sub>t(j)->including\<^sub>S\<^sub>e\<^sub>t(i)))" proof - have A : "\\insert (i \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(insert S_def i_val, frule Set_inv_lemma, simp add: foundation18 invalid_def) have B : "\\insert (j \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(insert S_def j_val, frule Set_inv_lemma, simp add: foundation18 invalid_def) have G1 : "Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\insert (i \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\\\ \ Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e None" by(insert A, simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def null_option_def) have G2 : "Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\insert (i \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\\\ \ Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \None\" by(insert A, simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def null_option_def) have G3 : "Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\insert (j \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\\\ \ Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e None" by(insert B, simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def null_option_def) have G4 : "Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\insert (j \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\\\ \ Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \None\" by(insert B, simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def null_option_def) have * : "(\ (\_. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\insert (i \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\\\)) \ = \\True\\" by(auto simp: OclValid_def false_def defined_def null_fun_def true_def bot_fun_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def S_def i_val G1 G2) have ** : "(\ (\_. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\insert (j \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\\\)) \ = \\True\\" by(auto simp: OclValid_def false_def defined_def null_fun_def true_def bot_fun_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def S_def i_val G3 G4) have *** : "Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\insert(j \)\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e(Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\insert(i \)\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e(S \)\\\\)\\\\ = Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\insert(i \)\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e(Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\insert(j \)\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e(S \)\\\\)\\\\" by(simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse[OF A] Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse[OF B] Set.insert_commute) show ?thesis apply(simp add: OclIncluding_def S_def[simplified OclValid_def] i_val[simplified OclValid_def] j_val[simplified OclValid_def] true_def OclValid_def StrongEq_def) apply(subst cp_defined, simp add: S_def[simplified OclValid_def] i_val[simplified OclValid_def] j_val[simplified OclValid_def] true_def *) apply(subst cp_defined, simp add: S_def[simplified OclValid_def] i_val[simplified OclValid_def] j_val[simplified OclValid_def] true_def ** ***) apply(subst cp_defined, simp add: S_def[simplified OclValid_def] i_val[simplified OclValid_def] j_val[simplified OclValid_def] true_def *) apply(subst cp_defined, simp add: S_def[simplified OclValid_def] i_val[simplified OclValid_def] j_val[simplified OclValid_def] true_def * ) apply(subst cp_defined, simp add: S_def[simplified OclValid_def] i_val[simplified OclValid_def] j_val[simplified OclValid_def] true_def * **) done qed lemma OclIncluding_commute[simp,code_unfold]: "((S :: ('\, 'a::null) Set)->including\<^sub>S\<^sub>e\<^sub>t(i)->including\<^sub>S\<^sub>e\<^sub>t(j) = (S->including\<^sub>S\<^sub>e\<^sub>t(j)->including\<^sub>S\<^sub>e\<^sub>t(i)))" proof - have A: "\ \. \ \ (i \ invalid) \ (S->including\<^sub>S\<^sub>e\<^sub>t(i)->including\<^sub>S\<^sub>e\<^sub>t(j)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have A': "\ \. \ \ (i \ invalid) \ (S->including\<^sub>S\<^sub>e\<^sub>t(j)->including\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have B:"\ \. \ \ (j \ invalid) \ (S->including\<^sub>S\<^sub>e\<^sub>t(i)->including\<^sub>S\<^sub>e\<^sub>t(j)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have B':"\ \. \ \ (j \ invalid) \ (S->including\<^sub>S\<^sub>e\<^sub>t(j)->including\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have C: "\ \. \ \ (S \ invalid) \ (S->including\<^sub>S\<^sub>e\<^sub>t(i)->including\<^sub>S\<^sub>e\<^sub>t(j)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have C': "\ \. \ \ (S \ invalid) \ (S->including\<^sub>S\<^sub>e\<^sub>t(j)->including\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have D: "\ \. \ \ (S \ null) \ (S->including\<^sub>S\<^sub>e\<^sub>t(i)->including\<^sub>S\<^sub>e\<^sub>t(j)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have D': "\ \. \ \ (S \ null) \ (S->including\<^sub>S\<^sub>e\<^sub>t(j)->including\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) show ?thesis apply(rule ext, rename_tac \) apply(case_tac "\ \ (\ i)") apply(case_tac "\ \ (\ j)") apply(case_tac "\ \ (\ S)") apply(simp only: OclIncluding_commute0[THEN foundation22[THEN iffD1]]) apply(simp add: foundation16', elim disjE) apply(simp add: C[OF foundation22[THEN iffD2]] C'[OF foundation22[THEN iffD2]]) apply(simp add: D[OF foundation22[THEN iffD2]] D'[OF foundation22[THEN iffD2]]) apply(simp add:foundation18 B[OF foundation22[THEN iffD2]] B'[OF foundation22[THEN iffD2]]) apply(simp add:foundation18 A[OF foundation22[THEN iffD2]] A'[OF foundation22[THEN iffD2]]) done qed subsubsection\Execution Rules on Excluding\ lemma OclExcluding_finite_rep_set : assumes X_def : "\ \ \ X" and x_val : "\ \ \ x" shows "finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X->excluding\<^sub>S\<^sub>e\<^sub>t(x) \)\\ = finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\" proof - have C : "\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ - {x \}\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" apply(insert X_def x_val, frule Set_inv_lemma) apply(simp add: foundation18 invalid_def) done show "?thesis" by(insert X_def x_val, auto simp: OclExcluding_def Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse[OF C] dest: foundation13[THEN iffD2, THEN foundation22[THEN iffD1]]) qed lemma OclExcluding_rep_set: assumes S_def: "\ \ \ S" shows "\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S->excluding\<^sub>S\<^sub>e\<^sub>t(\_. \\x\\) \)\\ = \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ - {\\x\\}" apply(simp add: OclExcluding_def S_def[simplified OclValid_def]) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def null_option_def) apply(insert Set_inv_lemma[OF S_def], metis Diff_iff bot_option_def not_None_eq) by(simp) lemma OclExcluding_excludes0: assumes "\ \ X->excludes\<^sub>S\<^sub>e\<^sub>t(x)" shows "X->excluding\<^sub>S\<^sub>e\<^sub>t(x) \ = X \" proof - have excludes_def: "\ \ X->excludes\<^sub>S\<^sub>e\<^sub>t(x) \ \ \ \ X" by (metis OclExcludes.def_valid_then_def OclExcludes_valid_args_valid'' foundation10' foundation6) have excludes_val: "\ \ X->excludes\<^sub>S\<^sub>e\<^sub>t(x) \ \ \ \ x" by (metis OclExcludes.def_valid_then_def OclExcludes_valid_args_valid'' foundation10' foundation6) show ?thesis apply(insert excludes_def[OF assms] excludes_val[OF assms] assms, simp add: OclExcluding_def OclExcludes_def OclIncludes_def OclNot_def OclValid_def true_def) by (metis (hide_lams, no_types) abs_rep_simp' assms excludes_def) qed lemma OclExcluding_excludes: assumes "\ \ X->excludes\<^sub>S\<^sub>e\<^sub>t(x)" shows "\ \ X->excluding\<^sub>S\<^sub>e\<^sub>t(x) \ X" by(simp add: StrongEq_def OclValid_def true_def OclExcluding_excludes0[OF assms]) lemma OclExcluding_charn0[simp]: assumes val_x:"\ \ (\ x)" shows "\ \ ((Set{}->excluding\<^sub>S\<^sub>e\<^sub>t(x)) \ Set{})" proof - have A : "\None\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(simp add: null_option_def bot_option_def) have B : "\\{}\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(simp add: mtSet_def) show ?thesis using val_x apply(auto simp: OclValid_def OclIncludes_def OclNot_def false_def true_def StrongEq_def OclExcluding_def mtSet_def defined_def bot_fun_def null_fun_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def) apply(auto simp: mtSet_def Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e.Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e.Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject[OF B A]) done qed lemma OclExcluding_commute0 : assumes S_def : "\ \ \ S" and i_val : "\ \ \ i" and j_val : "\ \ \ j" shows "\ \ ((S :: ('\, 'a::null) Set)->excluding\<^sub>S\<^sub>e\<^sub>t(i)->excluding\<^sub>S\<^sub>e\<^sub>t(j) \ (S->excluding\<^sub>S\<^sub>e\<^sub>t(j)->excluding\<^sub>S\<^sub>e\<^sub>t(i)))" proof - have A : "\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ - {i \}\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(insert S_def i_val, frule Set_inv_lemma, simp add: foundation18 invalid_def) have B : "\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ - {j \}\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(insert S_def j_val, frule Set_inv_lemma, simp add: foundation18 invalid_def) have G1 : "Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ - {i \}\\ \ Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e None" by(insert A, simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def null_option_def) have G2 : "Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ - {i \}\\ \ Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \None\" by(insert A, simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def null_option_def) have G3 : "Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ - {j \}\\ \ Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e None" by(insert B, simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def null_option_def) have G4 : "Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ - {j \}\\ \ Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \None\" by(insert B, simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def null_option_def) have * : "(\ (\_. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ - {i \}\\)) \ = \\True\\" by(auto simp: OclValid_def false_def defined_def null_fun_def true_def bot_fun_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def S_def i_val G1 G2) have ** : "(\ (\_. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ - {j \}\\)) \ = \\True\\" by(auto simp: OclValid_def false_def defined_def null_fun_def true_def bot_fun_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def S_def i_val G3 G4) have *** : "Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e(Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e(S \)\\-{i \}\\)\\-{j \}\\ = Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e(Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e(S \)\\-{j \}\\)\\-{i \}\\" apply(simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse[OF A] Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse[OF B]) by (metis Diff_insert2 insert_commute) show ?thesis apply(simp add: OclExcluding_def S_def[simplified OclValid_def] i_val[simplified OclValid_def] j_val[simplified OclValid_def] true_def OclValid_def StrongEq_def) apply(subst cp_defined, simp add: S_def[simplified OclValid_def] i_val[simplified OclValid_def] j_val[simplified OclValid_def] true_def *) apply(subst cp_defined, simp add: S_def[simplified OclValid_def] i_val[simplified OclValid_def] j_val[simplified OclValid_def] true_def ** ***) apply(subst cp_defined, simp add: S_def[simplified OclValid_def] i_val[simplified OclValid_def] j_val[simplified OclValid_def] true_def *) apply(subst cp_defined, simp add: S_def[simplified OclValid_def] i_val[simplified OclValid_def] j_val[simplified OclValid_def] true_def * ) apply(subst cp_defined, simp add: S_def[simplified OclValid_def] i_val[simplified OclValid_def] j_val[simplified OclValid_def] true_def * **) done qed lemma OclExcluding_commute[simp,code_unfold]: "((S :: ('\, 'a::null) Set)->excluding\<^sub>S\<^sub>e\<^sub>t(i)->excluding\<^sub>S\<^sub>e\<^sub>t(j) = (S->excluding\<^sub>S\<^sub>e\<^sub>t(j)->excluding\<^sub>S\<^sub>e\<^sub>t(i)))" proof - have A: "\ \. \ \ i \ invalid \ (S->excluding\<^sub>S\<^sub>e\<^sub>t(i)->excluding\<^sub>S\<^sub>e\<^sub>t(j)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have A': "\ \. \ \ i \ invalid \ (S->excluding\<^sub>S\<^sub>e\<^sub>t(j)->excluding\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have B:"\ \. \ \ j \ invalid \ (S->excluding\<^sub>S\<^sub>e\<^sub>t(i)->excluding\<^sub>S\<^sub>e\<^sub>t(j)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have B':"\ \. \ \ j \ invalid \ (S->excluding\<^sub>S\<^sub>e\<^sub>t(j)->excluding\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have C: "\ \. \ \ S \ invalid \ (S->excluding\<^sub>S\<^sub>e\<^sub>t(i)->excluding\<^sub>S\<^sub>e\<^sub>t(j)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have C': "\ \. \ \ S \ invalid \ (S->excluding\<^sub>S\<^sub>e\<^sub>t(j)->excluding\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have D: "\ \. \ \ S \ null \ (S->excluding\<^sub>S\<^sub>e\<^sub>t(i)->excluding\<^sub>S\<^sub>e\<^sub>t(j)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have D': "\ \. \ \ S \ null \ (S->excluding\<^sub>S\<^sub>e\<^sub>t(j)->excluding\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) show ?thesis apply(rule ext, rename_tac \) apply(case_tac "\ \ (\ i)") apply(case_tac "\ \ (\ j)") apply(case_tac "\ \ (\ S)") apply(simp only: OclExcluding_commute0[THEN foundation22[THEN iffD1]]) apply(simp add: foundation16', elim disjE) apply(simp add: C[OF foundation22[THEN iffD2]] C'[OF foundation22[THEN iffD2]]) apply(simp add: D[OF foundation22[THEN iffD2]] D'[OF foundation22[THEN iffD2]]) apply(simp add:foundation18 B[OF foundation22[THEN iffD2]] B'[OF foundation22[THEN iffD2]]) apply(simp add:foundation18 A[OF foundation22[THEN iffD2]] A'[OF foundation22[THEN iffD2]]) done qed lemma OclExcluding_charn0_exec[simp,code_unfold]: "(Set{}->excluding\<^sub>S\<^sub>e\<^sub>t(x)) = (if (\ x) then Set{} else invalid endif)" proof - have A: "\ \. (Set{}->excluding\<^sub>S\<^sub>e\<^sub>t(invalid)) \ = (if (\ invalid) then Set{} else invalid endif) \" by simp have B: "\ \ x. \ \ (\ x) \ (Set{}->excluding\<^sub>S\<^sub>e\<^sub>t(x)) \ = (if (\ x) then Set{} else invalid endif) \" by(simp add: OclExcluding_charn0[THEN foundation22[THEN iffD1]]) show ?thesis apply(rule ext, rename_tac \) apply(case_tac "\ \ (\ x)") apply(simp add: B) apply(simp add: foundation18) apply(subst OclExcluding.cp0, simp) apply(simp add: cp_OclIf[symmetric] OclExcluding.cp0[symmetric] cp_valid[symmetric] A) done qed lemma OclExcluding_charn1: assumes def_X:"\ \ (\ X)" and val_x:"\ \ (\ x)" and val_y:"\ \ (\ y)" and neq :"\ \ not(x \ y)" shows "\ \ ((X->including\<^sub>S\<^sub>e\<^sub>t(x))->excluding\<^sub>S\<^sub>e\<^sub>t(y)) \ ((X->excluding\<^sub>S\<^sub>e\<^sub>t(y))->including\<^sub>S\<^sub>e\<^sub>t(x))" proof - have C : "\\insert (x \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(insert def_X val_x, frule Set_inv_lemma, simp add: foundation18 invalid_def) have D : "\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ - {y \}\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(insert def_X val_x, frule Set_inv_lemma, simp add: foundation18 invalid_def) have E : "x \ \ y \" by(insert neq, auto simp: OclValid_def bot_fun_def OclIncluding_def OclIncludes_def false_def true_def defined_def valid_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_fun_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def StrongEq_def OclNot_def) have G1 : "Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\insert (x \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\\\ \ Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e None" by(insert C, simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def null_option_def) have G2 : "Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\insert (x \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\\\ \ Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \None\" by(insert C, simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def null_option_def) have G : "(\ (\_. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\insert (x \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\\\)) \ = true \" by(auto simp: OclValid_def false_def true_def defined_def bot_fun_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_fun_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def G1 G2) have H1 : "Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ - {y \}\\ \ Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e None" by(insert D, simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def null_option_def) have H2 : "Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ - {y \}\\ \ Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \None\" by(insert D, simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def null_option_def) have H : "(\ (\_. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ - {y \}\\)) \ = true \" by(auto simp: OclValid_def false_def true_def defined_def bot_fun_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_fun_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def H1 H2) have Z : "insert (x \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ - {y \} = insert (x \) (\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ - {y \})" by(auto simp: E) show ?thesis apply(insert def_X[THEN foundation13[THEN iffD2]] val_x[THEN foundation13[THEN iffD2]] val_y[THEN foundation13[THEN iffD2]]) apply(simp add: foundation22 OclIncluding_def OclExcluding_def def_X[THEN foundation16[THEN iffD1]]) apply(subst cp_defined, simp)+ apply(simp add: G H Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse[OF C] Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse[OF D] Z) done qed lemma OclExcluding_charn2: assumes def_X:"\ \ (\ X)" and val_x:"\ \ (\ x)" shows "\ \ (((X->including\<^sub>S\<^sub>e\<^sub>t(x))->excluding\<^sub>S\<^sub>e\<^sub>t(x)) \ (X->excluding\<^sub>S\<^sub>e\<^sub>t(x)))" proof - have C : "\\insert (x \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(insert def_X val_x, frule Set_inv_lemma, simp add: foundation18 invalid_def) have G1 : "Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\insert (x \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\\\ \ Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e None" by(insert C, simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def null_option_def) have G2 : "Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\insert (x \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\\\ \ Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \None\" by(insert C, simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def null_option_def) show ?thesis apply(insert def_X[THEN foundation16[THEN iffD1]] val_x[THEN foundation18[THEN iffD1]]) apply(auto simp: OclValid_def bot_fun_def OclIncluding_def OclIncludes_def false_def true_def invalid_def defined_def valid_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_fun_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def StrongEq_def) apply(subst OclExcluding.cp0) apply(auto simp:OclExcluding_def) apply(simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse[OF C]) apply(simp_all add: false_def true_def defined_def valid_def null_fun_def bot_fun_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def split: bool.split_asm HOL.if_split_asm option.split) apply(auto simp: G1 G2) done qed theorem OclExcluding_charn3: "((X->including\<^sub>S\<^sub>e\<^sub>t(x))->excluding\<^sub>S\<^sub>e\<^sub>t(x)) = (X->excluding\<^sub>S\<^sub>e\<^sub>t(x))" proof - have A1 : "\\. \ \ (X \ invalid) \ (X->including\<^sub>S\<^sub>e\<^sub>t(x)->excluding\<^sub>S\<^sub>e\<^sub>t(x)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have A1': "\\. \ \ (X \ invalid) \ (X->excluding\<^sub>S\<^sub>e\<^sub>t(x)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have A2 : "\\. \ \ (X \ null) \ (X->including\<^sub>S\<^sub>e\<^sub>t(x)->excluding\<^sub>S\<^sub>e\<^sub>t(x)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have A2': "\\. \ \ (X \ null) \ (X->excluding\<^sub>S\<^sub>e\<^sub>t(x)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have A3 : "\\. \ \ (x \ invalid) \ (X->including\<^sub>S\<^sub>e\<^sub>t(x)->excluding\<^sub>S\<^sub>e\<^sub>t(x)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have A3': "\\. \ \ (x \ invalid) \ (X->excluding\<^sub>S\<^sub>e\<^sub>t(x)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) show ?thesis apply(rule ext, rename_tac "\") apply(case_tac "\ \ (\ x)") apply(case_tac "\ \ (\ X)") apply(simp only: OclExcluding_charn2[THEN foundation22[THEN iffD1]]) apply(simp add: foundation16', elim disjE) apply(simp add: A1[OF foundation22[THEN iffD2]] A1'[OF foundation22[THEN iffD2]]) apply(simp add: A2[OF foundation22[THEN iffD2]] A2'[OF foundation22[THEN iffD2]]) apply(simp add:foundation18 A3[OF foundation22[THEN iffD2]] A3'[OF foundation22[THEN iffD2]]) done qed text\One would like a generic theorem of the form: \begin{isar}[mathescape] lemma OclExcluding_charn_exec: "(X->including$_{Set}$(x::('$\mathfrak{A}$,'a::null)val)->excluding$_{Set}$(y)) = (if \ X then if x \ y then X->excluding$_{Set}$(y) else X->excluding$_{Set}$(y)->including$_{Set}$(x) endif else invalid endif)" \end{isar} Unfortunately, this does not hold in general, since referential equality is an overloaded concept and has to be defined for each type individually. Consequently, it is only valid for concrete type instances for Boolean, Integer, and Sets thereof... \ text\The computational law \emph{OclExcluding-charn-exec} becomes generic since it uses strict equality which in itself is generic. It is possible to prove the following generic theorem and instantiate it later (using properties that link the polymorphic logical strong equality with the concrete instance of strict quality).\ lemma OclExcluding_charn_exec: assumes strict1: "(invalid \ y) = invalid" and strict2: "(x \ invalid) = invalid" and StrictRefEq_valid_args_valid: "\ (x::('\,'a::null)val) y \. (\ \ \ (x \ y)) = ((\ \ (\ x)) \ (\ \ \ y))" and cp_StrictRefEq: "\ (X::('\,'a::null)val) Y \. (X \ Y) \ = ((\_. X \) \ (\_. Y \)) \" and StrictRefEq_vs_StrongEq: "\ (x::('\,'a::null)val) y \. \ \ \ x \ \ \ \ y \ (\ \ ((x \ y) \ (x \ y)))" shows "(X->including\<^sub>S\<^sub>e\<^sub>t(x::('\,'a::null)val)->excluding\<^sub>S\<^sub>e\<^sub>t(y)) = (if \ X then if x \ y then X->excluding\<^sub>S\<^sub>e\<^sub>t(y) else X->excluding\<^sub>S\<^sub>e\<^sub>t(y)->including\<^sub>S\<^sub>e\<^sub>t(x) endif else invalid endif)" proof - (* Lifting theorems, largely analogous OclIncludes_execute_generic, with the same problems wrt. strict equality. *) have A1: "\\. \ \ (X \ invalid) \ (X->including\<^sub>S\<^sub>e\<^sub>t(x)->includes\<^sub>S\<^sub>e\<^sub>t(y)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have B1: "\\. \ \ (X \ null) \ (X->including\<^sub>S\<^sub>e\<^sub>t(x)->includes\<^sub>S\<^sub>e\<^sub>t(y)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have A2: "\\. \ \ (X \ invalid) \ X->including\<^sub>S\<^sub>e\<^sub>t(x)->excluding\<^sub>S\<^sub>e\<^sub>t(y) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have B2: "\\. \ \ (X \ null) \ X->including\<^sub>S\<^sub>e\<^sub>t(x)->excluding\<^sub>S\<^sub>e\<^sub>t(y) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) note [simp] = cp_StrictRefEq [THEN allI[THEN allI[THEN allI[THEN cpI2]], of "StrictRefEq"]] have C: "\\. \ \ (x \ invalid) \ (X->including\<^sub>S\<^sub>e\<^sub>t(x)->excluding\<^sub>S\<^sub>e\<^sub>t(y)) \ = (if x \ y then X->excluding\<^sub>S\<^sub>e\<^sub>t(y) else X->excluding\<^sub>S\<^sub>e\<^sub>t(y)->including\<^sub>S\<^sub>e\<^sub>t(x) endif) \" apply(rule foundation22[THEN iffD1]) apply(erule StrongEq_L_subst2_rev,simp,simp) by(simp add: strict1) have D: "\\. \ \ (y \ invalid) \ (X->including\<^sub>S\<^sub>e\<^sub>t(x)->excluding\<^sub>S\<^sub>e\<^sub>t(y)) \ = (if x \ y then X->excluding\<^sub>S\<^sub>e\<^sub>t(y) else X->excluding\<^sub>S\<^sub>e\<^sub>t(y)->including\<^sub>S\<^sub>e\<^sub>t(x) endif) \" apply(rule foundation22[THEN iffD1]) apply(erule StrongEq_L_subst2_rev,simp,simp) by (simp add: strict2) have E: "\\. \ \ \ x \ \ \ \ y \ (if x \ y then X->excluding\<^sub>S\<^sub>e\<^sub>t(y) else X->excluding\<^sub>S\<^sub>e\<^sub>t(y)->including\<^sub>S\<^sub>e\<^sub>t(x) endif) \ = (if x \ y then X->excluding\<^sub>S\<^sub>e\<^sub>t(y) else X->excluding\<^sub>S\<^sub>e\<^sub>t(y)->including\<^sub>S\<^sub>e\<^sub>t(x) endif) \" apply(subst cp_OclIf) apply(subst StrictRefEq_vs_StrongEq[THEN foundation22[THEN iffD1]]) by(simp_all add: cp_OclIf[symmetric]) have F: "\\. \ \ \ X \ \ \ \ x \ \ \ (x \ y) \ (X->including\<^sub>S\<^sub>e\<^sub>t(x)->excluding\<^sub>S\<^sub>e\<^sub>t(y) \) = (X->excluding\<^sub>S\<^sub>e\<^sub>t(y) \)" apply(drule StrongEq_L_sym) apply(rule foundation22[THEN iffD1]) apply(erule StrongEq_L_subst2_rev,simp) by(simp add: OclExcluding_charn2) show ?thesis apply(rule ext, rename_tac "\") apply(case_tac "\ (\ \ (\ X))", simp add:defined_split,elim disjE A1 B1 A2 B2) apply(case_tac "\ (\ \ (\ x))", simp add:foundation18 foundation22[symmetric], drule StrongEq_L_sym) apply(simp add: foundation22 C) apply(case_tac "\ (\ \ (\ y))", simp add:foundation18 foundation22[symmetric], drule StrongEq_L_sym, simp add: foundation22 D, simp) apply(subst E,simp_all) apply(case_tac "\ \ not (x \ y)") apply(simp add: OclExcluding_charn1[simplified foundation22] OclExcluding_charn2[simplified foundation22]) apply(simp add: foundation9 F) done qed (* Hack to work around OF-Bug *) schematic_goal OclExcluding_charn_exec\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r[simp,code_unfold]: "?X" by(rule OclExcluding_charn_exec[OF StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.strict1 StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.strict2 StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.defined_args_valid StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.cp0 StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.StrictRefEq_vs_StrongEq], simp_all) schematic_goal OclExcluding_charn_exec\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n[simp,code_unfold]: "?X" by(rule OclExcluding_charn_exec[OF StrictRefEq\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n.strict1 StrictRefEq\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n.strict2 StrictRefEq\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n.defined_args_valid StrictRefEq\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n.cp0 StrictRefEq\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n.StrictRefEq_vs_StrongEq], simp_all) schematic_goal OclExcluding_charn_exec\<^sub>S\<^sub>e\<^sub>t[simp,code_unfold]: "?X" by(rule OclExcluding_charn_exec[OF StrictRefEq\<^sub>S\<^sub>e\<^sub>t.strict1 StrictRefEq\<^sub>S\<^sub>e\<^sub>t.strict2 StrictRefEq\<^sub>S\<^sub>e\<^sub>t.defined_args_valid StrictRefEq\<^sub>S\<^sub>e\<^sub>t.cp0 StrictRefEq\<^sub>S\<^sub>e\<^sub>t.StrictRefEq_vs_StrongEq], simp_all) subsubsection\Execution Rules on Includes\ lemma OclIncludes_charn0[simp]: assumes val_x:"\ \ (\ x)" shows "\ \ not(Set{}->includes\<^sub>S\<^sub>e\<^sub>t(x))" using val_x apply(auto simp: OclValid_def OclIncludes_def OclNot_def false_def true_def) apply(auto simp: mtSet_def Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e.Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse) done lemma OclIncludes_charn0'[simp,code_unfold]: "Set{}->includes\<^sub>S\<^sub>e\<^sub>t(x) = (if \ x then false else invalid endif)" proof - have A: "\ \. (Set{}->includes\<^sub>S\<^sub>e\<^sub>t(invalid)) \ = (if (\ invalid) then false else invalid endif) \" by simp have B: "\ \ x. \ \ (\ x) \ (Set{}->includes\<^sub>S\<^sub>e\<^sub>t(x)) \ = (if \ x then false else invalid endif) \" apply(frule OclIncludes_charn0, simp add: OclValid_def) apply(rule foundation21[THEN fun_cong, simplified StrongEq_def,simplified, THEN iffD1, of _ _ "false"]) by simp show ?thesis apply(rule ext, rename_tac \) apply(case_tac "\ \ (\ x)") apply(simp_all add: B foundation18) apply(subst OclIncludes.cp0, simp add: OclIncludes.cp0[symmetric] A) done qed lemma OclIncludes_charn1: assumes def_X:"\ \ (\ X)" assumes val_x:"\ \ (\ x)" shows "\ \ (X->including\<^sub>S\<^sub>e\<^sub>t(x)->includes\<^sub>S\<^sub>e\<^sub>t(x))" proof - have C : "\\insert (x \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(insert def_X val_x, frule Set_inv_lemma, simp add: foundation18 invalid_def) show ?thesis apply(subst OclIncludes_def, simp add: foundation10[simplified OclValid_def] OclValid_def def_X[simplified OclValid_def] val_x[simplified OclValid_def]) apply(simp add: OclIncluding_def def_X[simplified OclValid_def] val_x[simplified OclValid_def] Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse[OF C] true_def) done qed lemma OclIncludes_charn2: assumes def_X:"\ \ (\ X)" and val_x:"\ \ (\ x)" and val_y:"\ \ (\ y)" and neq :"\ \ not(x \ y)" shows "\ \ (X->including\<^sub>S\<^sub>e\<^sub>t(x)->includes\<^sub>S\<^sub>e\<^sub>t(y)) \ (X->includes\<^sub>S\<^sub>e\<^sub>t(y))" proof - have C : "\\insert (x \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(insert def_X val_x, frule Set_inv_lemma, simp add: foundation18 invalid_def) show ?thesis apply(subst OclIncludes_def, simp add: def_X[simplified OclValid_def] val_x[simplified OclValid_def] val_y[simplified OclValid_def] foundation10[simplified OclValid_def] OclValid_def StrongEq_def) apply(simp add: OclIncluding_def OclIncludes_def def_X[simplified OclValid_def] val_x[simplified OclValid_def] val_y[simplified OclValid_def] Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse[OF C] true_def) by(metis foundation22 foundation6 foundation9 neq) qed text\Here is again a generic theorem similar as above.\ lemma OclIncludes_execute_generic: assumes strict1: "(invalid \ y) = invalid" and strict2: "(x \ invalid) = invalid" and cp_StrictRefEq: "\ (X::('\,'a::null)val) Y \. (X \ Y) \ = ((\_. X \) \ (\_. Y \)) \" and StrictRefEq_vs_StrongEq: "\ (x::('\,'a::null)val) y \. \ \ \ x \ \ \ \ y \ (\ \ ((x \ y) \ (x \ y)))" shows "(X->including\<^sub>S\<^sub>e\<^sub>t(x::('\,'a::null)val)->includes\<^sub>S\<^sub>e\<^sub>t(y)) = (if \ X then if x \ y then true else X->includes\<^sub>S\<^sub>e\<^sub>t(y) endif else invalid endif)" proof - have A: "\\. \ \ (X \ invalid) \ (X->including\<^sub>S\<^sub>e\<^sub>t(x)->includes\<^sub>S\<^sub>e\<^sub>t(y)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev,simp,simp) have B: "\\. \ \ (X \ null) \ (X->including\<^sub>S\<^sub>e\<^sub>t(x)->includes\<^sub>S\<^sub>e\<^sub>t(y)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev,simp,simp) note [simp] = cp_StrictRefEq [THEN allI[THEN allI[THEN allI[THEN cpI2]], of "StrictRefEq"]] have C: "\\. \ \ (x \ invalid) \ (X->including\<^sub>S\<^sub>e\<^sub>t(x)->includes\<^sub>S\<^sub>e\<^sub>t(y)) \ = (if x \ y then true else X->includes\<^sub>S\<^sub>e\<^sub>t(y) endif) \" apply(rule foundation22[THEN iffD1]) apply(erule StrongEq_L_subst2_rev,simp,simp) by (simp add: strict1) have D:"\\. \ \ (y \ invalid) \ (X->including\<^sub>S\<^sub>e\<^sub>t(x)->includes\<^sub>S\<^sub>e\<^sub>t(y)) \ = (if x \ y then true else X->includes\<^sub>S\<^sub>e\<^sub>t(y) endif) \" apply(rule foundation22[THEN iffD1]) apply(erule StrongEq_L_subst2_rev,simp,simp) by (simp add: strict2) have E: "\\. \ \ \ x \ \ \ \ y \ (if x \ y then true else X->includes\<^sub>S\<^sub>e\<^sub>t(y) endif) \ = (if x \ y then true else X->includes\<^sub>S\<^sub>e\<^sub>t(y) endif) \" apply(subst cp_OclIf) apply(subst StrictRefEq_vs_StrongEq[THEN foundation22[THEN iffD1]]) by(simp_all add: cp_OclIf[symmetric]) have F: "\\. \ \ (x \ y) \ (X->including\<^sub>S\<^sub>e\<^sub>t(x)->includes\<^sub>S\<^sub>e\<^sub>t(y)) \ = (X->including\<^sub>S\<^sub>e\<^sub>t(x)->includes\<^sub>S\<^sub>e\<^sub>t(x)) \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev,simp, simp) show ?thesis apply(rule ext, rename_tac "\") apply(case_tac "\ (\ \ (\ X))", simp add:defined_split,elim disjE A B) apply(case_tac "\ (\ \ (\ x))", simp add:foundation18 foundation22[symmetric], drule StrongEq_L_sym) apply(simp add: foundation22 C) apply(case_tac "\ (\ \ (\ y))", simp add:foundation18 foundation22[symmetric], drule StrongEq_L_sym, simp add: foundation22 D, simp) apply(subst E,simp_all) apply(case_tac "\ \ not(x \ y)") apply(simp add: OclIncludes_charn2[simplified foundation22]) apply(simp add: foundation9 F OclIncludes_charn1[THEN foundation13[THEN iffD2], THEN foundation22[THEN iffD1]]) done qed (* Hack to work around OF-Bug *) schematic_goal OclIncludes_execute\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r[simp,code_unfold]: "?X" by(rule OclIncludes_execute_generic[OF StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.strict1 StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.strict2 StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.cp0 StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.StrictRefEq_vs_StrongEq], simp_all) schematic_goal OclIncludes_execute\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n[simp,code_unfold]: "?X" by(rule OclIncludes_execute_generic[OF StrictRefEq\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n.strict1 StrictRefEq\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n.strict2 StrictRefEq\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n.cp0 StrictRefEq\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n.StrictRefEq_vs_StrongEq], simp_all) schematic_goal OclIncludes_execute\<^sub>S\<^sub>e\<^sub>t[simp,code_unfold]: "?X" by(rule OclIncludes_execute_generic[OF StrictRefEq\<^sub>S\<^sub>e\<^sub>t.strict1 StrictRefEq\<^sub>S\<^sub>e\<^sub>t.strict2 StrictRefEq\<^sub>S\<^sub>e\<^sub>t.cp0 StrictRefEq\<^sub>S\<^sub>e\<^sub>t.StrictRefEq_vs_StrongEq], simp_all) lemma OclIncludes_including_generic : assumes OclIncludes_execute_generic [simp] : "\X x y. (X->including\<^sub>S\<^sub>e\<^sub>t(x::('\,'a::null)val)->includes\<^sub>S\<^sub>e\<^sub>t(y)) = (if \ X then if x \ y then true else X->includes\<^sub>S\<^sub>e\<^sub>t(y) endif else invalid endif)" and StrictRefEq_strict'' : "\x y. \ ((x::('\,'a::null)val) \ y) = (\(x) and \(y))" and a_val : "\ \ \ a" and x_val : "\ \ \ x" and S_incl : "\ \ (S)->includes\<^sub>S\<^sub>e\<^sub>t((x::('\,'a::null)val))" shows "\ \ S->including\<^sub>S\<^sub>e\<^sub>t((a::('\,'a::null)val))->includes\<^sub>S\<^sub>e\<^sub>t(x)" proof - have discr_eq_bot1_true : "\\. (\ \ = true \) = False" by (metis bot_fun_def foundation1 foundation18' valid3) have discr_eq_bot2_true : "\\. (\ = true \) = False" by (metis bot_fun_def discr_eq_bot1_true) have discr_neq_invalid_true : "\\. (invalid \ \ true \) = True" by (metis discr_eq_bot2_true invalid_def) have discr_eq_invalid_true : "\\. (invalid \ = true \) = False" by (metis bot_option_def invalid_def option.simps(2) true_def) show ?thesis apply(simp) apply(subgoal_tac "\ \ \ S") prefer 2 apply(insert S_incl[simplified OclIncludes_def], simp add: OclValid_def) apply(metis discr_eq_bot2_true) apply(simp add: cp_OclIf[of "\ S"] OclValid_def OclIf_def x_val[simplified OclValid_def] discr_neq_invalid_true discr_eq_invalid_true) by (metis OclValid_def S_incl StrictRefEq_strict'' a_val foundation10 foundation6 x_val) qed lemmas OclIncludes_including\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r = OclIncludes_including_generic[OF OclIncludes_execute\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.def_homo] subsubsection\Execution Rules on Excludes\ lemma OclExcludes_charn1: assumes def_X:"\ \ (\ X)" assumes val_x:"\ \ (\ x)" shows "\ \ (X->excluding\<^sub>S\<^sub>e\<^sub>t(x)->excludes\<^sub>S\<^sub>e\<^sub>t(x))" proof - let ?OclSet = "\S. \\S\\ \ {X. X = \ \ X = null \ (\x\\\X\\. x \ \)}" have diff_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e : "?OclSet (\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ - {x \})" apply(simp, (rule disjI2)+) by (metis (hide_lams, no_types) Diff_iff Set_inv_lemma def_X) show ?thesis apply(subst OclExcludes_def, simp add: foundation10[simplified OclValid_def] OclValid_def def_X[simplified OclValid_def] val_x[simplified OclValid_def]) apply(subst OclIncludes_def, simp add: OclNot_def) apply(simp add: OclExcluding_def def_X[simplified OclValid_def] val_x[simplified OclValid_def] Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse[OF diff_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e] true_def) by(simp add: OclAnd_def def_X[simplified OclValid_def] val_x[simplified OclValid_def] true_def) qed subsubsection\Execution Rules on Size\ lemma [simp,code_unfold]: "Set{} ->size\<^sub>S\<^sub>e\<^sub>t() = \" apply(rule ext) apply(simp add: defined_def mtSet_def OclSize_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_fun_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_fun_def) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject, simp_all add: bot_option_def null_option_def) + by(simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse bot_option_def null_option_def OclInt0_def) lemma OclSize_including_exec[simp,code_unfold]: "((X ->including\<^sub>S\<^sub>e\<^sub>t(x)) ->size\<^sub>S\<^sub>e\<^sub>t()) = (if \ X and \ x then X ->size\<^sub>S\<^sub>e\<^sub>t() +\<^sub>i\<^sub>n\<^sub>t if X ->includes\<^sub>S\<^sub>e\<^sub>t(x) then \ else \ endif else invalid endif)" proof - have valid_inject_true : "\\ P. (\ P) \ \ true \ \ (\ P) \ = false \" apply(simp add: valid_def true_def false_def bot_fun_def bot_option_def null_fun_def null_option_def) by (case_tac "P \ = \", simp_all add: true_def) have defined_inject_true : "\\ P. (\ P) \ \ true \ \ (\ P) \ = false \" apply(simp add: defined_def true_def false_def bot_fun_def bot_option_def null_fun_def null_option_def) by (case_tac " P \ = \ \ P \ = null", simp_all add: true_def) show ?thesis apply(rule ext, rename_tac \) proof - fix \ have includes_notin: "\ \ \ X->includes\<^sub>S\<^sub>e\<^sub>t(x) \ (\ X) \ = true \ \ (\ x) \ = true \ \ x \ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\" by(simp add: OclIncludes_def OclValid_def true_def) have includes_def: "\ \ X->includes\<^sub>S\<^sub>e\<^sub>t(x) \ \ \ \ X" by (metis bot_fun_def OclIncludes_def OclValid_def defined3 foundation16) have includes_val: "\ \ X->includes\<^sub>S\<^sub>e\<^sub>t(x) \ \ \ \ x" using foundation5 foundation6 by fastforce have ins_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e: "\ \ \ X \ \ \ \ x \ \\insert (x \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\\\ \ {X. X = \ \ X = null \ (\x\\\X\\. x \ \)}" apply(simp add: bot_option_def null_option_def) by (metis (hide_lams, no_types) Set_inv_lemma foundation18' foundation5) have m : "\\. (\_. \) = (\_. invalid \)" by(rule ext, simp add:invalid_def) show "X->including\<^sub>S\<^sub>e\<^sub>t(x)->size\<^sub>S\<^sub>e\<^sub>t() \ = (if \ X and \ x then X->size\<^sub>S\<^sub>e\<^sub>t() +\<^sub>i\<^sub>n\<^sub>t if X->includes\<^sub>S\<^sub>e\<^sub>t(x) then \ else \ endif else invalid endif) \" apply(case_tac "\ \ \ X and \ x", simp) apply(subst OclAdd\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.cp0) apply(case_tac "\ \ X->includes\<^sub>S\<^sub>e\<^sub>t(x)", simp add: OclAdd\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.cp0[symmetric]) apply(case_tac "\ \ ((\ (X->size\<^sub>S\<^sub>e\<^sub>t())) and not (\ (X->size\<^sub>S\<^sub>e\<^sub>t())))", simp) apply(drule foundation5[where P = "\ X->size\<^sub>S\<^sub>e\<^sub>t()"], erule conjE) apply(drule OclSize_infinite) apply(frule includes_def, drule includes_val, simp) apply(subst OclSize_def, subst OclIncluding_finite_rep_set, assumption+) apply (metis (hide_lams, no_types) invalid_def) apply(subst OclIf_false', metis (hide_lams, no_types) defined5 defined6 defined_and_I defined_not_I foundation1 foundation9) apply(subst cp_OclSize, simp add: OclIncluding_includes0 cp_OclSize[symmetric]) (* *) apply(subst OclIf_false', subst foundation9, auto, simp add: OclSize_def) apply(drule foundation5) apply(subst (1 2) OclIncluding_finite_rep_set, fast+) apply(subst (1 2) cp_OclAnd, subst (1 2) OclAdd\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.cp0, simp) apply(rule conjI) apply(simp add: OclIncluding_def) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse[OF ins_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e], fast+) apply(subst (asm) (2 3) OclValid_def, simp add: OclAdd\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r_def OclInt1_def) apply(rule impI) apply(drule Finite_Set.card.insert[where x = "x \"]) apply(rule includes_notin, simp, simp) apply (metis Suc_eq_plus1 of_nat_1 of_nat_add) apply(subst (1 2) m[of \], simp only: OclAdd\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.cp0[symmetric],simp, simp add:invalid_def) apply(subst OclIncluding_finite_rep_set, fast+, simp add: OclValid_def) (* *) apply(subst OclIf_false', metis (hide_lams, no_types) defined6 foundation1 foundation9 OclExcluding_valid_args_valid'') by (metis cp_OclSize foundation18' OclIncluding_valid_args_valid'' invalid_def OclSize_invalid) qed qed subsubsection\Execution Rules on IsEmpty\ lemma [simp,code_unfold]: "Set{}->isEmpty\<^sub>S\<^sub>e\<^sub>t() = true" by(simp add: OclIsEmpty_def) lemma OclIsEmpty_including [simp]: assumes X_def: "\ \ \ X" and X_finite: "finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\" and a_val: "\ \ \ a" shows "X->including\<^sub>S\<^sub>e\<^sub>t(a)->isEmpty\<^sub>S\<^sub>e\<^sub>t() \ = false \" proof - have A1 : "\\ X. X \ = true \ \ X \ = false \ \ (X and not X) \ = false \" by (metis (no_types) OclAnd_false1 OclAnd_idem OclImplies_def OclNot3 OclNot_not OclOr_false1 cp_OclAnd cp_OclNot deMorgan1 deMorgan2) have defined_inject_true : "\\ P. (\ P) \ \ true \ \ (\ P) \ = false \" apply(simp add: defined_def true_def false_def bot_fun_def bot_option_def null_fun_def null_option_def) by (case_tac " P \ = \ \ P \ = null", simp_all add: true_def) have B : "\X \. \ \ \ X \ X \ \ \ \ \ (X \ \) \ = false \" apply(simp add: foundation22[symmetric] foundation14 foundation9) apply(erule StrongEq_L_subst4_rev[THEN iffD2, OF StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.StrictRefEq_vs_StrongEq]) by(simp_all) show ?thesis apply(simp add: OclIsEmpty_def del: OclSize_including_exec) apply(subst cp_OclOr, subst A1) apply (metis OclExcludes.def_homo defined_inject_true) apply(simp add: cp_OclOr[symmetric] del: OclSize_including_exec) apply(rule B, rule foundation20, metis OclIncluding.def_homo OclIncluding_finite_rep_set X_def X_finite a_val foundation10' size_defined') apply(simp add: OclSize_def OclIncluding_finite_rep_set[OF X_def a_val] X_finite OclInt0_def) by (metis OclValid_def X_def a_val foundation10 foundation6 OclIncluding_notempty_rep_set[OF X_def a_val]) qed subsubsection\Execution Rules on NotEmpty\ lemma [simp,code_unfold]: "Set{}->notEmpty\<^sub>S\<^sub>e\<^sub>t() = false" by(simp add: OclNotEmpty_def) lemma OclNotEmpty_including [simp,code_unfold]: assumes X_def: "\ \ \ X" and X_finite: "finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\" and a_val: "\ \ \ a" shows "X->including\<^sub>S\<^sub>e\<^sub>t(a)->notEmpty\<^sub>S\<^sub>e\<^sub>t() \ = true \" apply(simp add: OclNotEmpty_def) apply(subst cp_OclNot, subst OclIsEmpty_including, simp_all add: assms) by (metis OclNot4 cp_OclNot) subsubsection\Execution Rules on Any\ lemma [simp,code_unfold]: "Set{}->any\<^sub>S\<^sub>e\<^sub>t() = null" by(rule ext, simp add: OclANY_def, simp add: false_def true_def) lemma OclANY_singleton_exec[simp,code_unfold]: "(Set{}->including\<^sub>S\<^sub>e\<^sub>t(a))->any\<^sub>S\<^sub>e\<^sub>t() = a" apply(rule ext, rename_tac \, simp add: mtSet_def OclANY_def) apply(case_tac "\ \ \ a") apply(simp add: OclValid_def mtSet_defined[simplified mtSet_def] mtSet_valid[simplified mtSet_def] mtSet_rep_set[simplified mtSet_def]) apply(subst (1 2) cp_OclAnd, subst (1 2) OclNotEmpty_including[where X = "Set{}", simplified mtSet_def]) apply(simp add: mtSet_defined[simplified mtSet_def]) apply(metis (hide_lams, no_types) finite.emptyI mtSet_def mtSet_rep_set) apply(simp add: OclValid_def) apply(simp add: OclIncluding_def) apply(rule conjI) apply(subst (1 2) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def null_option_def) apply(simp, metis OclValid_def foundation18') apply(simp) apply(simp add: mtSet_defined[simplified mtSet_def]) (* *) apply(subgoal_tac "a \ = \") prefer 2 apply(simp add: OclValid_def valid_def bot_fun_def split: if_split_asm) apply(simp) apply(subst (1 2 3 4) cp_OclAnd, simp add: mtSet_defined[simplified mtSet_def] valid_def bot_fun_def) by(simp add: cp_OclAnd[symmetric], rule impI, simp add: false_def true_def) subsubsection\Execution Rules on Forall\ lemma OclForall_mtSet_exec[simp,code_unfold] :"((Set{})->forAll\<^sub>S\<^sub>e\<^sub>t(z| P(z))) = true" apply(simp add: OclForall_def) apply(subst mtSet_def)+ apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp_all add: true_def)+ done text\The following rule is a main theorem of our approach: From a denotational definition that assures consistency, but may be --- as in the case of the @{term "X->forAll\<^sub>S\<^sub>e\<^sub>t(x | P x)"} --- dauntingly complex, we derive operational rules that can serve as a gold-standard for operational execution, since they may be evaluated in whatever situation and according to whatever strategy. In the case of @{term "X->forAll\<^sub>S\<^sub>e\<^sub>t(x | P x)"}, the operational rule gives immediately a way to evaluation in any finite (in terms of conventional OCL: denotable) set, although the rule also holds for the infinite case: @{term "Integer\<^sub>n\<^sub>u\<^sub>l\<^sub>l ->forAll\<^sub>S\<^sub>e\<^sub>t(x | (Integer\<^sub>n\<^sub>u\<^sub>l\<^sub>l ->forAll\<^sub>S\<^sub>e\<^sub>t(y | x +\<^sub>i\<^sub>n\<^sub>t y \ y +\<^sub>i\<^sub>n\<^sub>t x)))"} or even: @{term "Integer ->forAll\<^sub>S\<^sub>e\<^sub>t(x | (Integer ->forAll\<^sub>S\<^sub>e\<^sub>t(y | x +\<^sub>i\<^sub>n\<^sub>t y \ y +\<^sub>i\<^sub>n\<^sub>t x)))"} are valid OCL statements in any context $\tau$. \ theorem OclForall_including_exec[simp,code_unfold] : assumes cp0 : "cp P" shows "((S->including\<^sub>S\<^sub>e\<^sub>t(x))->forAll\<^sub>S\<^sub>e\<^sub>t(z | P(z))) = (if \ S and \ x then P x and (S->forAll\<^sub>S\<^sub>e\<^sub>t(z | P(z))) else invalid endif)" proof - have cp: "\\. P x \ = P (\_. x \) \" by(insert cp0, auto simp: cp_def) have cp_eq : "\\ v. (P x \ = v) = (P (\_. x \) \ = v)" by(subst cp, simp) have cp_OclNot_eq : "\\ v. (P x \ \ v) = (P (\_. x \) \ \ v)" by(subst cp, simp) have insert_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e : "\\. (\ \(\ S)) \ (\ \(\ x)) \ \\insert (x \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(frule Set_inv_lemma, simp add: foundation18 invalid_def) have forall_including_invert : "\\ f. (f x \ = f (\ _. x \) \) \ \ \ (\ S and \ x) \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S->including\<^sub>S\<^sub>e\<^sub>t(x) \)\\. f (\_. x) \) = (f x \ \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\. f (\_. x) \))" apply(drule foundation5, simp add: OclIncluding_def) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse) apply(rule insert_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e, fast+) by(simp add: OclValid_def) have exists_including_invert : "\\ f. (f x \ = f (\ _. x \) \) \ \ \ (\ S and \ x) \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S->including\<^sub>S\<^sub>e\<^sub>t(x) \)\\. f (\_. x) \) = (f x \ \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\. f (\_. x) \))" apply(subst arg_cong[where f = "\x. \x", OF forall_including_invert[where f = "\x \. \ (f x \)"], simplified]) by simp_all have contradict_Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e: "\\ S f. \x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e S\\. f (\_. x) \ \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e S\\. \ (f (\_. x) \)) = False" by(case_tac "(\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e S\\. \ (f (\_. x) \)) = True", simp_all) have bot_invalid : "\ = invalid" by(rule ext, simp add: invalid_def bot_fun_def) have bot_invalid2 : "\\. \ = invalid \" by(simp add: invalid_def) have C1 : "\\. P x \ = false \ \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\. P (\_. x) \ = false \) \ \ \ (\ S and \ x) \ false \ = (P x and OclForall S P) \" apply(simp add: cp_OclAnd[of "P x"]) apply(elim disjE, simp) apply(simp only: cp_OclAnd[symmetric], simp) apply(subgoal_tac "OclForall S P \ = false \") apply(simp only: cp_OclAnd[symmetric], simp) apply(simp add: OclForall_def) apply(fold OclValid_def, simp add: foundation10') done have C2 : "\\. \ \ (\ S and \ x) \ P x \ = null \ \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\. P (\_. x) \ = null \) \ P x \ = invalid \ \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\. P (\_. x) \ = invalid \) \ \x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S->including\<^sub>S\<^sub>e\<^sub>t(x) \)\\. P (\_. x) \ \ false \ \ invalid \ = (P x and OclForall S P) \" apply(subgoal_tac "(\ S)\ = true \") prefer 2 apply(simp add: foundation10', simp add: OclValid_def) apply(drule forall_including_invert[of "\ x \. P x \ \ false \", OF cp_OclNot_eq, THEN iffD1]) apply(assumption) apply(simp add: cp_OclAnd[of "P x"],elim disjE, simp_all) apply(simp add: invalid_def null_fun_def null_option_def bot_fun_def bot_option_def) apply(subgoal_tac "OclForall S P \ = invalid \") apply(simp only:cp_OclAnd[symmetric],simp,simp add:invalid_def bot_fun_def) apply(unfold OclForall_def, simp add: invalid_def false_def bot_fun_def,simp) apply(simp add:cp_OclAnd[symmetric],simp) apply(erule conjE) apply(subgoal_tac "(P x \ = invalid \) \ (P x \ = null \) \ (P x \ = true \) \ (P x \ = false \)") prefer 2 apply(rule bool_split_0) apply(elim disjE, simp_all) apply(simp only:cp_OclAnd[symmetric],simp)+ done have A : "\\. \ \ (\ S and \ x) \ OclForall (S->including\<^sub>S\<^sub>e\<^sub>t(x)) P \ = (P x and OclForall S P) \" proof - fix \ assume 0 : "\ \ (\ S and \ x)" let ?S = "\ocl. P x \ \ ocl \ \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\. P (\_. x) \ \ ocl \)" let ?S' = "\ocl. \x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S->including\<^sub>S\<^sub>e\<^sub>t(x) \)\\. P (\_. x) \ \ ocl \" let ?assms_1 = "?S' null" let ?assms_2 = "?S' invalid" let ?assms_3 = "?S' false" have 4 : "?assms_3 \ ?S false" apply(subst forall_including_invert[of "\ x \. P x \ \ false \",symmetric]) by(simp_all add: cp_OclNot_eq 0) have 5 : "?assms_2 \ ?S invalid" apply(subst forall_including_invert[of "\ x \. P x \ \ invalid \",symmetric]) by(simp_all add: cp_OclNot_eq 0) have 6 : "?assms_1 \ ?S null" apply(subst forall_including_invert[of "\ x \. P x \ \ null \",symmetric]) by(simp_all add: cp_OclNot_eq 0) have 7 : "(\ S) \ = true \" by(insert 0, simp add: foundation10', simp add: OclValid_def) show "?thesis \" apply(subst OclForall_def) apply(simp add: cp_OclAnd[THEN sym] OclValid_def contradict_Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e) apply(intro conjI impI,fold OclValid_def) apply(simp_all add: exists_including_invert[where f = "\ x \. P x \ = null \", OF cp_eq]) apply(simp_all add: exists_including_invert[where f = "\ x \. P x \ = invalid \", OF cp_eq]) apply(simp_all add: exists_including_invert[where f = "\ x \. P x \ = false \", OF cp_eq]) proof - assume 1 : "P x \ = null \ \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\. P (\_. x) \ = null \)" and 2 : ?assms_2 and 3 : ?assms_3 show "null \ = (P x and OclForall S P) \" proof - note 4 = 4[OF 3] note 5 = 5[OF 2] have 6 : "P x \ = null \ \ P x \ = true \" by(metis 4 5 bool_split_0) show ?thesis apply(insert 6, elim disjE) apply(subst cp_OclAnd) apply(simp add: OclForall_def 7 4[THEN conjunct2] 5[THEN conjunct2]) apply(simp_all add:cp_OclAnd[symmetric]) apply(subst cp_OclAnd, simp_all add:cp_OclAnd[symmetric] OclForall_def) apply(simp add:4[THEN conjunct2] 5[THEN conjunct2] 0[simplified OclValid_def] 7) apply(insert 1, elim disjE, auto) done qed next assume 1 : ?assms_1 and 2 : "P x \ = invalid \ \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\. P (\_. x) \ = invalid \)" and 3 : ?assms_3 show "invalid \ = (P x and OclForall S P) \" proof - note 4 = 4[OF 3] note 6 = 6[OF 1] have 5 : "P x \ = invalid \ \ P x \ = true \" by(metis 4 6 bool_split_0) show ?thesis apply(insert 5, elim disjE) apply(subst cp_OclAnd) apply(simp add: OclForall_def 4[THEN conjunct2] 6[THEN conjunct2] 7) apply(simp_all add:cp_OclAnd[symmetric]) apply(subst cp_OclAnd, simp_all add:cp_OclAnd[symmetric] OclForall_def) apply(insert 2, elim disjE, simp add: invalid_def true_def bot_option_def) apply(simp add: 0[simplified OclValid_def] 4[THEN conjunct2] 6[THEN conjunct2] 7) by(auto) qed next assume 1 : ?assms_1 and 2 : ?assms_2 and 3 : ?assms_3 show "true \ = (P x and OclForall S P) \" proof - note 4 = 4[OF 3] note 5 = 5[OF 2] note 6 = 6[OF 1] have 8 : "P x \ = true \" by(metis 4 5 6 bool_split_0) show ?thesis apply(subst cp_OclAnd, simp add: 8 cp_OclAnd[symmetric]) by(simp add: OclForall_def 4 5 6 7) qed qed ( simp add: 0 | rule C1, simp+ | rule C2, simp add: 0 )+ qed have B : "\\. \ (\ \ (\ S and \ x)) \ OclForall (S->including\<^sub>S\<^sub>e\<^sub>t(x)) P \ = invalid \" apply(rule foundation22[THEN iffD1]) apply(simp only: foundation10' de_Morgan_conj foundation18'', elim disjE) apply(simp add: defined_split, elim disjE) apply(erule StrongEq_L_subst2_rev, simp+)+ done show ?thesis apply(rule ext, rename_tac \) apply(simp add: OclIf_def) apply(simp add: cp_defined[of "\ S and \ x"] cp_defined[THEN sym]) apply(intro conjI impI) by(auto intro!: A B simp: OclValid_def) qed subsubsection\Execution Rules on Exists\ lemma OclExists_mtSet_exec[simp,code_unfold] : "((Set{})->exists\<^sub>S\<^sub>e\<^sub>t(z | P(z))) = false" by(simp add: OclExists_def) lemma OclExists_including_exec[simp,code_unfold] : assumes cp: "cp P" shows "((S->including\<^sub>S\<^sub>e\<^sub>t(x))->exists\<^sub>S\<^sub>e\<^sub>t(z | P(z))) = (if \ S and \ x then P x or (S->exists\<^sub>S\<^sub>e\<^sub>t(z | P(z))) else invalid endif)" by(simp add: OclExists_def OclOr_def cp OclNot_inject) subsubsection\Execution Rules on Iterate\ lemma OclIterate_empty[simp,code_unfold]: "((Set{})->iterate\<^sub>S\<^sub>e\<^sub>t(a; x = A | P a x)) = A" proof - have C : "\ \. (\ (\\. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\{}\\)) \ = true \" by (metis (no_types) defined_def mtSet_def mtSet_defined null_fun_def) show ?thesis apply(simp add: OclIterate_def mtSet_def Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse valid_def C) apply(rule ext, rename_tac \) apply(case_tac "A \ = \ \", simp_all, simp add:true_def false_def bot_fun_def) apply(simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse) done qed text\In particular, this does hold for A = null.\ lemma OclIterate_including: assumes S_finite: "\ \ \(S->size\<^sub>S\<^sub>e\<^sub>t())" and F_valid_arg: "(\ A) \ = (\ (F a A)) \" and F_commute: "comp_fun_commute F" and F_cp: "\ x y \. F x y \ = F (\ _. x \) y \" shows "((S->including\<^sub>S\<^sub>e\<^sub>t(a))->iterate\<^sub>S\<^sub>e\<^sub>t(a; x = A | F a x)) \ = ((S->excluding\<^sub>S\<^sub>e\<^sub>t(a))->iterate\<^sub>S\<^sub>e\<^sub>t(a; x = F a A | F a x)) \" proof - have insert_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e : "\\. (\ \(\ S)) \ (\ \(\ a)) \ \\insert (a \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(frule Set_inv_lemma, simp add: foundation18 invalid_def) have insert_defined : "\\. (\ \(\ S)) \ (\ \(\ a)) \ (\ (\_. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\insert (a \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\\\)) \ = true \" apply(subst defined_def) apply(simp add: bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_fun_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_fun_def) by(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject, rule insert_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e, simp_all add: null_option_def bot_option_def)+ have remove_finite : "finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ \ finite ((\a \. a) ` (\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ - {a \}))" by(simp) have remove_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e : "\\. (\ \(\ S)) \ (\ \(\ a)) \ \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ - {a \}\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(frule Set_inv_lemma, simp add: foundation18 invalid_def) have remove_defined : "\\. (\ \(\ S)) \ (\ \(\ a)) \ (\ (\_. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ - {a \}\\)) \ = true \" apply(subst defined_def) apply(simp add: bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_fun_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_fun_def) by(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject, rule remove_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e, simp_all add: null_option_def bot_option_def)+ have abs_rep: "\x. \\x\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)} \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\x\\)\\ = x" by(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp_all) have inject : "inj (\a \. a)" by(rule inj_fun, simp) + interpret F_commute: comp_fun_commute "F" + by (fact F_commute) show ?thesis apply(subst (1 2) cp_OclIterate, subst OclIncluding_def, subst OclExcluding_def) apply(case_tac "\ ((\ S) \ = true \ \ (\ a) \ = true \)", simp add: invalid_def) apply(subgoal_tac "OclIterate (\_. \) A F \ = OclIterate (\_. \) (F a A) F \", simp) apply(rule conjI, blast+) apply(simp add: OclIterate_def defined_def bot_option_def bot_fun_def false_def true_def) apply(simp add: OclIterate_def) apply((subst abs_rep[OF insert_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e[simplified OclValid_def], of \], simp_all)+, (subst abs_rep[OF remove_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e[simplified OclValid_def], of \], simp_all)+, (subst insert_defined, simp_all add: OclValid_def)+, (subst remove_defined, simp_all add: OclValid_def)+) apply(case_tac "\ ((\ A) \ = true \)", (simp add: F_valid_arg)+) apply(rule impI, - subst Finite_Set.comp_fun_commute.fold_fun_left_comm[symmetric, OF F_commute], + subst F_commute.fold_fun_left_comm[symmetric], rule remove_finite, simp) apply(subst image_set_diff[OF inject], simp) apply(subgoal_tac "Finite_Set.fold F A (insert (\\'. a \) ((\a \. a) ` \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\)) \ = F (\\'. a \) (Finite_Set.fold F A ((\a \. a) ` \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ - {\\'. a \})) \") apply(subst F_cp, simp) - by(subst Finite_Set.comp_fun_commute.fold_insert_remove[OF F_commute], simp+) + by(subst F_commute.fold_insert_remove, simp+) qed subsubsection\Execution Rules on Select\ lemma OclSelect_mtSet_exec[simp,code_unfold]: "OclSelect mtSet P = mtSet" apply(rule ext, rename_tac \) apply(simp add: OclSelect_def mtSet_def defined_def false_def true_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_fun_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_fun_def) by(( subst (1 2 3 4 5) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse | subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject), (simp add: null_option_def bot_option_def)+)+ definition "OclSelect_body :: _ \ _ \ _ \ ('\, 'a option option) Set \ (\P x acc. if P x \ false then acc else acc->including\<^sub>S\<^sub>e\<^sub>t(x) endif)" theorem OclSelect_including_exec[simp,code_unfold]: assumes P_cp : "cp P" shows "OclSelect (X->including\<^sub>S\<^sub>e\<^sub>t(y)) P = OclSelect_body P y (OclSelect (X->excluding\<^sub>S\<^sub>e\<^sub>t(y)) P)" (is "_ = ?select") proof - have P_cp: "\x \. P x \ = P (\_. x \) \" by(insert P_cp, auto simp: cp_def) have ex_including : "\f X y \. \ \ \ X \ \ \ \ y \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X->including\<^sub>S\<^sub>e\<^sub>t(y) \)\\. f (P (\_. x)) \) = (f (P (\_. y \)) \ \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. f (P (\_. x)) \))" apply(simp add: OclIncluding_def OclValid_def) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp, (rule disjI2)+) by (metis (hide_lams, no_types) OclValid_def Set_inv_lemma foundation18',simp) have al_including : "\f X y \. \ \ \ X \ \ \ \ y \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X->including\<^sub>S\<^sub>e\<^sub>t(y) \)\\. f (P (\_. x)) \) = (f (P (\_. y \)) \ \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. f (P (\_. x)) \))" apply(simp add: OclIncluding_def OclValid_def) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp, (rule disjI2)+) by (metis (hide_lams, no_types) OclValid_def Set_inv_lemma foundation18', simp) have ex_excluding1 : "\f X y \. \ \ \ X \ \ \ \ y \ \ (f (P (\_. y \)) \) \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. f (P (\_. x)) \) = (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X->excluding\<^sub>S\<^sub>e\<^sub>t(y) \)\\. f (P (\_. x)) \)" apply(simp add: OclExcluding_def OclValid_def) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp, (rule disjI2)+) by (metis (no_types) Diff_iff OclValid_def Set_inv_lemma) auto have al_excluding1 : "\f X y \. \ \ \ X \ \ \ \ y \ f (P (\_. y \)) \ \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. f (P (\_. x)) \) = (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X->excluding\<^sub>S\<^sub>e\<^sub>t(y) \)\\. f (P (\_. x)) \)" apply(simp add: OclExcluding_def OclValid_def) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp, (rule disjI2)+) by (metis (no_types) Diff_iff OclValid_def Set_inv_lemma) auto have in_including : "\f X y \. \ \ \ X \ \ \ \ y \ {x \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X->including\<^sub>S\<^sub>e\<^sub>t(y) \)\\. f (P (\_. x) \)} = (let s = {x \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. f (P (\_. x) \)} in if f (P (\_. y \) \) then insert (y \) s else s)" apply(simp add: OclIncluding_def OclValid_def) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp, (rule disjI2)+) apply (metis (hide_lams, no_types) OclValid_def Set_inv_lemma foundation18') by(simp add: Let_def, auto) let ?OclSet = "\S. \\S\\ \ {X. X = \ \ X = null \ (\x\\\X\\. x \ \)}" have diff_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e : "\\. (\ X) \ = true \ \ ?OclSet (\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ - {y \})" apply(simp, (rule disjI2)+) by (metis (mono_tags) Diff_iff OclValid_def Set_inv_lemma) have ins_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e : "\\. (\ X) \ = true \ \ (\ y) \ = true \ \ ?OclSet (insert (y \) {x \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. P (\_. x) \ \ false \})" apply(simp, (rule disjI2)+) by (metis (hide_lams, no_types) OclValid_def Set_inv_lemma foundation18') have ins_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e' : "\\. (\ X) \ = true \ \ (\ y) \ = true \ \ ?OclSet (insert (y \) {x \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. x \ y \ \ P (\_. x) \ \ false \})" apply(simp, (rule disjI2)+) by (metis (hide_lams, no_types) OclValid_def Set_inv_lemma foundation18') have ins_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e'' : "\\. (\ X) \ = true \ \ ?OclSet {x \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. P (\_. x) \ \ false \}" apply(simp, (rule disjI2)+) by (metis (hide_lams, no_types) OclValid_def Set_inv_lemma) have ins_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e''' : "\\. (\ X) \ = true \ \ ?OclSet {x \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. x \ y \ \ P (\_. x) \ \ false \}" apply(simp, (rule disjI2)+) by(metis (hide_lams, no_types) OclValid_def Set_inv_lemma) have if_same : "\a b c d \. \ \ \ a \ b \ = d \ \ c \ = d \ \ (if a then b else c endif) \ = d \" by(simp add: OclIf_def OclValid_def) have invert_including : "\P y \. P \ = \ \ P->including\<^sub>S\<^sub>e\<^sub>t(y) \ = \" by (metis (hide_lams, no_types) foundation16[THEN iffD1] foundation18' OclIncluding_valid_args_valid) have exclude_defined : "\\. \ \ \ X \ (\(\_. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\{x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. x \ y \ \ P (\_. x) \\false \}\\)) \ = true \" apply(subst defined_def, simp add: false_def true_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_fun_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_fun_def) by(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject[OF ins_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e'''[simplified false_def]], (simp add: OclValid_def bot_option_def null_option_def)+)+ have if_eq : "\x A B \. \ \ \ x \ \ \ ((if x \ false then A else B endif) \ (if x \ false then A else B endif))" apply(simp add: StrictRefEq\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n OclValid_def) apply(subst (2) StrongEq_def) by(subst cp_OclIf, simp add: cp_OclIf[symmetric] true_def) have OclSelect_body_bot: "\\. \ \ \ X \ \ \ \ y \ P y \ \ \ \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. P (\_. x) \ = \) \ \ = ?select \" apply(drule ex_excluding1[where X2 = X and y2 = y and f2 = "\x \. x \ = \"], (simp add: P_cp[symmetric])+) apply(subgoal_tac "\ \ (\ \ ?select)", simp add: OclValid_def StrongEq_def true_def bot_fun_def) apply(simp add: OclSelect_body_def) apply(subst StrongEq_L_subst3[OF _ if_eq], simp, metis foundation18') apply(simp add: OclValid_def, subst StrongEq_def, subst true_def, simp) apply(subgoal_tac "\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X->excluding\<^sub>S\<^sub>e\<^sub>t(y) \)\\. P (\_. x) \ = \ \") prefer 2 apply (metis bot_fun_def ) apply(subst if_same[where d5 = "\"]) apply (metis defined7 transform1) apply(simp add: OclSelect_def bot_option_def bot_fun_def invalid_def) apply(subst invert_including) by(simp add: OclSelect_def bot_option_def bot_fun_def invalid_def)+ have d_and_v_inject : "\\ X y. (\ X and \ y) \ \ true \ \ (\ X and \ y) \ = false \" apply(fold OclValid_def, subst foundation22[symmetric]) apply(auto simp:foundation10' defined_split) apply(erule StrongEq_L_subst2_rev,simp,simp) apply(erule StrongEq_L_subst2_rev,simp,simp) by(erule foundation7'[THEN iffD2, THEN foundation15[THEN iffD2, THEN StrongEq_L_subst2_rev]],simp,simp) have OclSelect_body_bot': "\\. (\ X and \ y) \ \ true \ \ \ = ?select \" apply(drule d_and_v_inject) apply(simp add: OclSelect_def OclSelect_body_def) apply(subst cp_OclIf, subst OclIncluding.cp0, simp add: false_def true_def) apply(subst cp_OclIf[symmetric], subst OclIncluding.cp0[symmetric]) by (metis (lifting, no_types) OclIf_def foundation18 foundation18' invert_including) have conj_split2 : "\a b c \. ((a \ false) \ = false \ \ b) \ ((a \ false) \ = true \ \ c) \ (a \ \ false \ \ b) \ (a \ = false \ \ c)" by (metis OclValid_def defined7 foundation14 foundation22 foundation9) have defined_inject_true : "\\ P. (\ P) \ \ true \ \ (\ P) \ = false \" apply(simp add: defined_def true_def false_def bot_fun_def bot_option_def null_fun_def null_option_def) by (case_tac " P \ = \ \ P \ = null", simp_all add: true_def) have cp_OclSelect_body : "\\. ?select \ = OclSelect_body P y (\_.(OclSelect (X->excluding\<^sub>S\<^sub>e\<^sub>t(y))P)\)\" apply(simp add: OclSelect_body_def) by(subst (1 2) cp_OclIf, subst (1 2) OclIncluding.cp0, blast) have OclSelect_body_strict1 : "OclSelect_body P y invalid = invalid" by(rule ext, simp add: OclSelect_body_def OclIf_def) have bool_invalid: "\(x::('\)Boolean) y \. \ (\ \ \ x) \ \ \ ((x \ y) \ invalid)" by(simp add: StrictRefEq\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n OclValid_def StrongEq_def true_def) have conj_comm : "\p q r. (p \ q \ r) = ((p \ q) \ r)" by blast have inv_bot : "\\. invalid \ = \ \" by (metis bot_fun_def invalid_def) have inv_bot' : "\\. invalid \ = \" by (simp add: invalid_def) show ?thesis apply(rule ext, rename_tac \) apply(subst OclSelect_def) apply(case_tac "(\ (X->including\<^sub>S\<^sub>e\<^sub>t(y))) \ = true \", simp) apply(( subst ex_including | subst in_including), metis OclValid_def foundation5, metis OclValid_def foundation5)+ apply(simp add: Let_def inv_bot) apply(subst (2 4 7 9) bot_fun_def) apply(subst (4) false_def, subst (4) bot_fun_def, simp add: bot_option_def P_cp[symmetric]) (* *) apply(case_tac "\ (\ \ (\ P y))") apply(subgoal_tac "P y \ \ false \") prefer 2 apply (metis (hide_lams, no_types) foundation1 foundation18' valid4) apply(simp) (* *) apply(subst conj_comm, rule conjI) apply(drule_tac y11 = false in bool_invalid) apply(simp only: OclSelect_body_def, metis OclIf_def OclValid_def defined_def foundation2 foundation22 bot_fun_def invalid_def) (* *) apply(drule foundation5[simplified OclValid_def], subst al_including[simplified OclValid_def], simp, simp) apply(simp add: P_cp[symmetric]) apply (metis bot_fun_def foundation18') apply(simp add: foundation18' bot_fun_def OclSelect_body_bot OclSelect_body_bot') (* *) apply(subst (1 2) al_including, metis OclValid_def foundation5, metis OclValid_def foundation5) apply(simp add: P_cp[symmetric], subst (4) false_def, subst (4) bot_option_def, simp) apply(simp add: OclSelect_def[simplified inv_bot'] OclSelect_body_def StrictRefEq\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n) apply(subst (1 2 3 4) cp_OclIf, subst (1 2 3 4) foundation18'[THEN iffD2, simplified OclValid_def], simp, simp only: cp_OclIf[symmetric] refl if_True) apply(subst (1 2) OclIncluding.cp0, rule conj_split2, simp add: cp_OclIf[symmetric]) apply(subst (1 2 3 4 5 6 7 8) cp_OclIf[symmetric], simp) apply(( subst ex_excluding1[symmetric] | subst al_excluding1[symmetric] ), metis OclValid_def foundation5, metis OclValid_def foundation5, simp add: P_cp[symmetric] bot_fun_def)+ apply(simp add: bot_fun_def) apply(subst (1 2) invert_including, simp+) (* *) apply(rule conjI, blast) apply(intro impI conjI) apply(subst OclExcluding_def) apply(drule foundation5[simplified OclValid_def], simp) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse[OF diff_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e], fast) apply(simp add: OclIncluding_def cp_valid[symmetric]) apply((erule conjE)+, frule exclude_defined[simplified OclValid_def], simp) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse[OF ins_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e'''], simp+) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject[OF ins_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e ins_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e'], fast+) (* *) apply(simp add: OclExcluding_def) apply(simp add: foundation10[simplified OclValid_def]) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse[OF diff_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e], simp+) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject[OF ins_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e'' ins_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e'''], simp+) apply(subgoal_tac "P (\_. y \) \ = false \") prefer 2 apply(subst P_cp[symmetric], metis OclValid_def foundation22) apply(rule equalityI) apply(rule subsetI, simp, metis) apply(rule subsetI, simp) (* *) apply(drule defined_inject_true) apply(subgoal_tac "\ (\ \ \ X) \ \ (\ \ \ y)") prefer 2 apply (metis OclIncluding.def_homo OclIncluding_valid_args_valid OclIncluding_valid_args_valid'' OclValid_def foundation18 valid1) apply(subst cp_OclSelect_body, subst cp_OclSelect, subst OclExcluding_def) apply(simp add: OclValid_def false_def true_def, rule conjI, blast) apply(simp add: OclSelect_invalid[simplified invalid_def] OclSelect_body_strict1[simplified invalid_def] inv_bot') done qed subsubsection\Execution Rules on Reject\ lemma OclReject_mtSet_exec[simp,code_unfold]: "OclReject mtSet P = mtSet" by(simp add: OclReject_def) lemma OclReject_including_exec[simp,code_unfold]: assumes P_cp : "cp P" shows "OclReject (X->including\<^sub>S\<^sub>e\<^sub>t(y)) P = OclSelect_body (not o P) y (OclReject (X->excluding\<^sub>S\<^sub>e\<^sub>t(y)) P)" apply(simp add: OclReject_def comp_def, rule OclSelect_including_exec) by (metis assms cp_intro'(5)) subsubsection\Execution Rules Combining Previous Operators\ text\OclIncluding\ (* logical level : *) lemma OclIncluding_idem0 : assumes "\ \ \ S" and "\ \ \ i" shows "\ \ (S->including\<^sub>S\<^sub>e\<^sub>t(i)->including\<^sub>S\<^sub>e\<^sub>t(i) \ (S->including\<^sub>S\<^sub>e\<^sub>t(i)))" by(simp add: OclIncluding_includes OclIncludes_charn1 assms) (* Pure algebraic level *) theorem OclIncluding_idem[simp,code_unfold]: "((S :: ('\,'a::null)Set)->including\<^sub>S\<^sub>e\<^sub>t(i)->including\<^sub>S\<^sub>e\<^sub>t(i) = (S->including\<^sub>S\<^sub>e\<^sub>t(i)))" proof - have A: "\ \. \ \ (i \ invalid) \ (S->including\<^sub>S\<^sub>e\<^sub>t(i)->including\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have A':"\ \. \ \ (i \ invalid) \ (S->including\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have C: "\ \. \ \ (S \ invalid) \ (S->including\<^sub>S\<^sub>e\<^sub>t(i)->including\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have C': "\ \. \ \ (S \ invalid) \ (S->including\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have D: "\ \. \ \ (S \ null) \ (S->including\<^sub>S\<^sub>e\<^sub>t(i)->including\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have D': "\ \. \ \ (S \ null) \ (S->including\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) show ?thesis apply(rule ext, rename_tac \) apply(case_tac "\ \ (\ i)") apply(case_tac "\ \ (\ S)") apply(simp only: OclIncluding_idem0[THEN foundation22[THEN iffD1]]) apply(simp add: foundation16', elim disjE) apply(simp add: C[OF foundation22[THEN iffD2]] C'[OF foundation22[THEN iffD2]]) apply(simp add: D[OF foundation22[THEN iffD2]] D'[OF foundation22[THEN iffD2]]) apply(simp add:foundation18 A[OF foundation22[THEN iffD2]] A'[OF foundation22[THEN iffD2]]) done qed text\OclExcluding\ (* logical level : *) lemma OclExcluding_idem0 : assumes "\ \ \ S" and "\ \ \ i" shows "\ \ (S->excluding\<^sub>S\<^sub>e\<^sub>t(i)->excluding\<^sub>S\<^sub>e\<^sub>t(i) \ (S->excluding\<^sub>S\<^sub>e\<^sub>t(i)))" by(simp add: OclExcluding_excludes OclExcludes_charn1 assms) (* Pure algebraic level *) theorem OclExcluding_idem[simp,code_unfold]: "((S->excluding\<^sub>S\<^sub>e\<^sub>t(i))->excluding\<^sub>S\<^sub>e\<^sub>t(i)) = (S->excluding\<^sub>S\<^sub>e\<^sub>t(i))" proof - have A: "\ \. \ \ (i \ invalid) \ (S->excluding\<^sub>S\<^sub>e\<^sub>t(i)->excluding\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have A':"\ \. \ \ (i \ invalid) \ (S->excluding\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have C: "\ \. \ \ (S \ invalid) \ (S->excluding\<^sub>S\<^sub>e\<^sub>t(i)->excluding\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have C': "\ \. \ \ (S \ invalid) \ (S->excluding\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have D: "\ \. \ \ (S \ null) \ (S->excluding\<^sub>S\<^sub>e\<^sub>t(i)->excluding\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have D': "\ \. \ \ (S \ null) \ (S->excluding\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) show ?thesis apply(rule ext, rename_tac \) apply(case_tac "\ \ (\ i)") apply(case_tac "\ \ (\ S)") apply(simp only: OclExcluding_idem0[THEN foundation22[THEN iffD1]]) apply(simp add: foundation16', elim disjE) apply(simp add: C[OF foundation22[THEN iffD2]] C'[OF foundation22[THEN iffD2]]) apply(simp add: D[OF foundation22[THEN iffD2]] D'[OF foundation22[THEN iffD2]]) apply(simp add:foundation18 A[OF foundation22[THEN iffD2]] A'[OF foundation22[THEN iffD2]]) done qed text\OclIncludes\ lemma OclIncludes_any[simp,code_unfold]: "X->includes\<^sub>S\<^sub>e\<^sub>t(X->any\<^sub>S\<^sub>e\<^sub>t()) = (if \ X then if \ (X->size\<^sub>S\<^sub>e\<^sub>t()) then not(X->isEmpty\<^sub>S\<^sub>e\<^sub>t()) else X->includes\<^sub>S\<^sub>e\<^sub>t(null) endif else invalid endif)" proof - have defined_inject_true : "\\ P. (\ P) \ \ true \ \ (\ P) \ = false \" apply(simp add: defined_def true_def false_def bot_fun_def bot_option_def null_fun_def null_option_def) by (case_tac " P \ = \ \ P \ = null", simp_all add: true_def) have valid_inject_true : "\\ P. (\ P) \ \ true \ \ (\ P) \ = false \" apply(simp add: valid_def true_def false_def bot_fun_def bot_option_def null_fun_def null_option_def) by (case_tac "P \ = \", simp_all add: true_def) have notempty': "\\ X. \ \ \ X \ finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ \ not (X->isEmpty\<^sub>S\<^sub>e\<^sub>t()) \ \ true \ \ X \ = Set{} \" apply(case_tac "X \", rename_tac X', simp add: mtSet_def Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject) apply(erule disjE, metis (hide_lams, mono_tags) bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_option_def foundation16) apply(erule disjE, metis (hide_lams, no_types) bot_option_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_option_def foundation16[THEN iffD1]) apply(case_tac X', simp, metis (hide_lams, no_types) bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def foundation16[THEN iffD1]) apply(rename_tac X'', case_tac X'', simp) apply (metis (hide_lams, no_types) foundation16[THEN iffD1] null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def) apply(simp add: OclIsEmpty_def OclSize_def) apply(subst (asm) cp_OclNot, subst (asm) cp_OclOr, subst (asm) StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.cp0, subst (asm) cp_OclAnd, subst (asm) cp_OclNot) apply(simp only: OclValid_def foundation20[simplified OclValid_def] cp_OclNot[symmetric] cp_OclAnd[symmetric] cp_OclOr[symmetric]) apply(simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse split: if_split_asm) by(simp add: true_def OclInt0_def OclNot_def StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r StrongEq_def) have B: "\X \. \ finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ \ (\ (X->size\<^sub>S\<^sub>e\<^sub>t())) \ = false \" apply(subst cp_defined) apply(simp add: OclSize_def) by (metis bot_fun_def defined_def) show ?thesis apply(rule ext, rename_tac \, simp only: OclIncludes_def OclANY_def) apply(subst cp_OclIf, subst (2) cp_valid) apply(case_tac "(\ X) \ = true \", simp only: foundation20[simplified OclValid_def] cp_OclIf[symmetric], simp, subst (1 2) cp_OclAnd, simp add: cp_OclAnd[symmetric]) apply(case_tac "finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\") apply(frule size_defined'[THEN iffD2, simplified OclValid_def], assumption) apply(subst (1 2 3 4) cp_OclIf, simp) apply(subst (1 2 3 4) cp_OclIf[symmetric], simp) apply(case_tac "(X->notEmpty\<^sub>S\<^sub>e\<^sub>t()) \ = true \", simp) apply(frule OclNotEmpty_has_elt[simplified OclValid_def], simp) apply(simp add: OclNotEmpty_def cp_OclIf[symmetric]) apply(subgoal_tac "(SOME y. y \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\) \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\", simp add: true_def) apply(metis OclValid_def Set_inv_lemma foundation18' null_option_def true_def) apply(rule someI_ex, simp) apply(simp add: OclNotEmpty_def cp_valid[symmetric]) apply(subgoal_tac "\ (null \ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\)", simp) apply(subst OclIsEmpty_def, simp add: OclSize_def) apply(subst cp_OclNot, subst cp_OclOr, subst StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.cp0, subst cp_OclAnd, subst cp_OclNot, simp add: OclValid_def foundation20[simplified OclValid_def] cp_OclNot[symmetric] cp_OclAnd[symmetric] cp_OclOr[symmetric]) apply(frule notempty'[simplified OclValid_def], (simp add: mtSet_def Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse OclInt0_def false_def)+) apply(drule notempty'[simplified OclValid_def], simp, simp) apply (metis (hide_lams, no_types) empty_iff mtSet_rep_set) (* *) apply(frule B) apply(subst (1 2 3 4) cp_OclIf, simp) apply(subst (1 2 3 4) cp_OclIf[symmetric], simp) apply(case_tac "(X->notEmpty\<^sub>S\<^sub>e\<^sub>t()) \ = true \", simp) apply(frule OclNotEmpty_has_elt[simplified OclValid_def], simp) apply(simp add: OclNotEmpty_def OclIsEmpty_def) apply(subgoal_tac "X->size\<^sub>S\<^sub>e\<^sub>t() \ = \") prefer 2 apply (metis (hide_lams, no_types) OclSize_def) apply(subst (asm) cp_OclNot, subst (asm) cp_OclOr, subst (asm) StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.cp0, subst (asm) cp_OclAnd, subst (asm) cp_OclNot) apply(simp add: OclValid_def foundation20[simplified OclValid_def] cp_OclNot[symmetric] cp_OclAnd[symmetric] cp_OclOr[symmetric]) apply(simp add: OclNot_def StrongEq_def StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r valid_def false_def true_def bot_option_def bot_fun_def invalid_def) apply (metis bot_fun_def null_fun_def null_is_valid valid_def) by(drule defined_inject_true, simp add: false_def true_def OclIf_false[simplified false_def] invalid_def) qed text\OclSize\ lemma [simp,code_unfold]: "\ (Set{} ->size\<^sub>S\<^sub>e\<^sub>t()) = true" by simp lemma [simp,code_unfold]: "\ ((X ->including\<^sub>S\<^sub>e\<^sub>t(x)) ->size\<^sub>S\<^sub>e\<^sub>t()) = (\(X->size\<^sub>S\<^sub>e\<^sub>t()) and \(x))" proof - have defined_inject_true : "\\ P. (\ P) \ \ true \ \ (\ P) \ = false \" apply(simp add: defined_def true_def false_def bot_fun_def bot_option_def null_fun_def null_option_def) by (case_tac " P \ = \ \ P \ = null", simp_all add: true_def) have valid_inject_true : "\\ P. (\ P) \ \ true \ \ (\ P) \ = false \" apply(simp add: valid_def true_def false_def bot_fun_def bot_option_def null_fun_def null_option_def) by (case_tac "P \ = \", simp_all add: true_def) have OclIncluding_finite_rep_set : "\\. (\ X and \ x) \ = true \ \ finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X->including\<^sub>S\<^sub>e\<^sub>t(x) \)\\ = finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\" apply(rule OclIncluding_finite_rep_set) by(metis OclValid_def foundation5)+ have card_including_exec : "\\. (\ (\_. \\int (card \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X->including\<^sub>S\<^sub>e\<^sub>t(x) \)\\)\\)) \ = (\ (\_. \\int (card \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\)\\)) \" by(simp add: defined_def bot_fun_def bot_option_def null_fun_def null_option_def) show ?thesis apply(rule ext, rename_tac \) apply(case_tac "(\ (X->including\<^sub>S\<^sub>e\<^sub>t(x)->size\<^sub>S\<^sub>e\<^sub>t())) \ = true \", simp del: OclSize_including_exec) apply(subst cp_OclAnd, subst cp_defined, simp only: cp_defined[of "X->including\<^sub>S\<^sub>e\<^sub>t(x)->size\<^sub>S\<^sub>e\<^sub>t()"], simp add: OclSize_def) apply(case_tac "((\ X and \ x) \ = true \ \ finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X->including\<^sub>S\<^sub>e\<^sub>t(x) \)\\)", simp) apply(erule conjE, simp add: OclIncluding_finite_rep_set[simplified OclValid_def] card_including_exec cp_OclAnd[of "\ X" "\ x"] cp_OclAnd[of "true", THEN sym]) apply(subgoal_tac "(\ X) \ = true \ \ (\ x) \ = true \", simp) apply(rule foundation5[of _ "\ X" "\ x", simplified OclValid_def], simp only: cp_OclAnd[THEN sym]) apply(simp, simp add: defined_def true_def false_def bot_fun_def bot_option_def) apply(drule defined_inject_true[of "X->including\<^sub>S\<^sub>e\<^sub>t(x)->size\<^sub>S\<^sub>e\<^sub>t()"], simp del: OclSize_including_exec, simp only: cp_OclAnd[of "\ (X->size\<^sub>S\<^sub>e\<^sub>t())" "\ x"], simp add: cp_defined[of "X->including\<^sub>S\<^sub>e\<^sub>t(x)->size\<^sub>S\<^sub>e\<^sub>t()" ] cp_defined[of "X->size\<^sub>S\<^sub>e\<^sub>t()" ] del: OclSize_including_exec, simp add: OclSize_def card_including_exec del: OclSize_including_exec) apply(case_tac "(\ X and \ x) \ = true \ \ finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\", simp add: OclIncluding_finite_rep_set[simplified OclValid_def] card_including_exec, simp only: cp_OclAnd[THEN sym], simp add: defined_def bot_fun_def) apply(split if_split_asm) apply(simp add: OclIncluding_finite_rep_set[simplified OclValid_def] card_including_exec)+ apply(simp only: cp_OclAnd[THEN sym], simp, rule impI, erule conjE) apply(case_tac "(\ x) \ = true \", simp add: cp_OclAnd[of "\ X" "\ x"]) by(drule valid_inject_true[of "x"], simp add: cp_OclAnd[of _ "\ x"]) qed lemma [simp,code_unfold]: "\ ((X ->excluding\<^sub>S\<^sub>e\<^sub>t(x)) ->size\<^sub>S\<^sub>e\<^sub>t()) = (\(X->size\<^sub>S\<^sub>e\<^sub>t()) and \(x))" proof - have defined_inject_true : "\\ P. (\ P) \ \ true \ \ (\ P) \ = false \" apply(simp add: defined_def true_def false_def bot_fun_def bot_option_def null_fun_def null_option_def) by (case_tac " P \ = \ \ P \ = null", simp_all add: true_def) have valid_inject_true : "\\ P. (\ P) \ \ true \ \ (\ P) \ = false \" apply(simp add: valid_def true_def false_def bot_fun_def bot_option_def null_fun_def null_option_def) by (case_tac "P \ = \", simp_all add: true_def) have OclExcluding_finite_rep_set : "\\. (\ X and \ x) \ = true \ \ finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X->excluding\<^sub>S\<^sub>e\<^sub>t(x) \)\\ = finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\" apply(rule OclExcluding_finite_rep_set) by(metis OclValid_def foundation5)+ have card_excluding_exec : "\\. (\ (\_. \\int (card \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X->excluding\<^sub>S\<^sub>e\<^sub>t(x) \)\\)\\)) \ = (\ (\_. \\int (card \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\)\\)) \" by(simp add: defined_def bot_fun_def bot_option_def null_fun_def null_option_def) show ?thesis apply(rule ext, rename_tac \) apply(case_tac "(\ (X->excluding\<^sub>S\<^sub>e\<^sub>t(x)->size\<^sub>S\<^sub>e\<^sub>t())) \ = true \", simp) apply(subst cp_OclAnd, subst cp_defined, simp only: cp_defined[of "X->excluding\<^sub>S\<^sub>e\<^sub>t(x)->size\<^sub>S\<^sub>e\<^sub>t()"], simp add: OclSize_def) apply(case_tac "((\ X and \ x) \ = true \ \ finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X->excluding\<^sub>S\<^sub>e\<^sub>t(x) \)\\)", simp) apply(erule conjE, simp add: OclExcluding_finite_rep_set[simplified OclValid_def] card_excluding_exec cp_OclAnd[of "\ X" "\ x"] cp_OclAnd[of "true", THEN sym]) apply(subgoal_tac "(\ X) \ = true \ \ (\ x) \ = true \", simp) apply(rule foundation5[of _ "\ X" "\ x", simplified OclValid_def], simp only: cp_OclAnd[THEN sym]) apply(simp, simp add: defined_def true_def false_def bot_fun_def bot_option_def) apply(drule defined_inject_true[of "X->excluding\<^sub>S\<^sub>e\<^sub>t(x)->size\<^sub>S\<^sub>e\<^sub>t()"], simp, simp only: cp_OclAnd[of "\ (X->size\<^sub>S\<^sub>e\<^sub>t())" "\ x"], simp add: cp_defined[of "X->excluding\<^sub>S\<^sub>e\<^sub>t(x)->size\<^sub>S\<^sub>e\<^sub>t()" ] cp_defined[of "X->size\<^sub>S\<^sub>e\<^sub>t()" ], simp add: OclSize_def card_excluding_exec) apply(case_tac "(\ X and \ x) \ = true \ \ finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\", simp add: OclExcluding_finite_rep_set[simplified OclValid_def] card_excluding_exec, simp only: cp_OclAnd[THEN sym], simp add: defined_def bot_fun_def) apply(split if_split_asm) apply(simp add: OclExcluding_finite_rep_set[simplified OclValid_def] card_excluding_exec)+ apply(simp only: cp_OclAnd[THEN sym], simp, rule impI, erule conjE) apply(case_tac "(\ x) \ = true \", simp add: cp_OclAnd[of "\ X" "\ x"]) by(drule valid_inject_true[of "x"], simp add: cp_OclAnd[of _ "\ x"]) qed lemma [simp]: assumes X_finite: "\\. finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\" shows "\ ((X ->including\<^sub>S\<^sub>e\<^sub>t(x)) ->size\<^sub>S\<^sub>e\<^sub>t()) = (\(X) and \(x))" by(simp add: size_defined[OF X_finite] del: OclSize_including_exec) text\OclForall\ lemma OclForall_rep_set_false: assumes "\ \ \ X" shows "(OclForall X P \ = false \) = (\x \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. P (\\. x) \ = false \)" by(insert assms, simp add: OclForall_def OclValid_def false_def true_def invalid_def bot_fun_def bot_option_def null_fun_def null_option_def) lemma OclForall_rep_set_true: assumes "\ \ \ X" shows "(\ \ OclForall X P) = (\x \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. \ \ P (\\. x))" proof - have destruct_ocl : "\x \. x = true \ \ x = false \ \ x = null \ \ x = \ \" apply(case_tac x) apply (metis bot_Boolean_def) apply(rename_tac x', case_tac x') apply (metis null_Boolean_def) apply(rename_tac x'', case_tac x'') apply (metis (full_types) true_def) by (metis (full_types) false_def) have disjE4 : "\ P1 P2 P3 P4 R. (P1 \ P2 \ P3 \ P4) \ (P1 \ R) \ (P2 \ R) \ (P3 \ R) \ (P4 \ R) \ R" by metis show ?thesis apply(simp add: OclForall_def OclValid_def true_def false_def invalid_def bot_fun_def bot_option_def null_fun_def null_option_def split: if_split_asm) apply(rule conjI, rule impI) apply (metis drop.simps option.distinct(1) invalid_def) apply(rule impI, rule conjI, rule impI) apply (metis option.distinct(1)) apply(rule impI, rule conjI, rule impI) apply (metis drop.simps) apply(intro conjI impI ballI) proof - fix x show "\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. P (\_. x) \ \ \None\ \ \x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. \y. P (\_. x) \ = \y\ \ \x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. P (\_. x) \ \ \\False\\ \ x \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ \ P (\\. x) \ = \\True\\" apply(erule_tac x = x in ballE)+ by(rule disjE4[OF destruct_ocl[of "P (\\. x) \"]], (simp add: true_def false_def null_fun_def null_option_def bot_fun_def bot_option_def)+) qed(simp add: assms[simplified OclValid_def true_def])+ qed lemma OclForall_includes : assumes x_def : "\ \ \ x" and y_def : "\ \ \ y" shows "(\ \ OclForall x (OclIncludes y)) = (\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \)\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (y \)\\)" apply(simp add: OclForall_rep_set_true[OF x_def], simp add: OclIncludes_def OclValid_def y_def[simplified OclValid_def]) apply(insert Set_inv_lemma[OF x_def], simp add: valid_def false_def true_def bot_fun_def) by(rule iffI, simp add: subsetI, simp add: subsetD) lemma OclForall_not_includes : assumes x_def : "\ \ \ x" and y_def : "\ \ \ y" shows "(OclForall x (OclIncludes y) \ = false \) = (\ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \)\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (y \)\\)" apply(simp add: OclForall_rep_set_false[OF x_def], simp add: OclIncludes_def OclValid_def y_def[simplified OclValid_def]) apply(insert Set_inv_lemma[OF x_def], simp add: valid_def false_def true_def bot_fun_def) by(rule iffI, metis rev_subsetD, metis subsetI) lemma OclForall_iterate: assumes S_finite: "finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\" shows "S->forAll\<^sub>S\<^sub>e\<^sub>t(x | P x) \ = (S->iterate\<^sub>S\<^sub>e\<^sub>t(x; acc = true | acc and P x)) \" proof - - have and_comm : "comp_fun_commute (\x acc. acc and P x)" + interpret and_comm: comp_fun_commute "(\x acc. acc and P x)" apply(simp add: comp_fun_commute_def comp_def) - by (metis OclAnd_assoc OclAnd_commute) + by (metis OclAnd_assoc OclAnd_commute) have ex_insert : "\x F P. (\x\insert x F. P x) = (P x \ (\x\F. P x))" by (metis insert_iff) have destruct_ocl : "\x \. x = true \ \ x = false \ \ x = null \ \ x = \ \" apply(case_tac x) apply (metis bot_Boolean_def) apply(rename_tac x', case_tac x') apply (metis null_Boolean_def) apply(rename_tac x'', case_tac x'') apply (metis (full_types) true_def) by (metis (full_types) false_def) have disjE4 : "\ P1 P2 P3 P4 R. (P1 \ P2 \ P3 \ P4) \ (P1 \ R) \ (P2 \ R) \ (P3 \ R) \ (P4 \ R) \ R" by metis let ?P_eq = "\x b \. P (\_. x) \ = b \" let ?P = "\set b \. \x\set. ?P_eq x b \" let ?if = "\f b c. if f b \ then b \ else c" let ?forall = "\P. ?if P false (?if P invalid (?if P null (true \)))" show ?thesis apply(simp only: OclForall_def OclIterate_def) apply(case_tac "\ \ \ S", simp only: OclValid_def) apply(subgoal_tac "let set = \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ in ?forall (?P set) = Finite_Set.fold (\x acc. acc and P x) true ((\a \. a) ` set) \", simp only: Let_def, simp add: S_finite, simp only: Let_def) apply(case_tac "\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ = {}", simp) apply(rule finite_ne_induct[OF S_finite], simp) (* *) apply(simp only: image_insert) - apply(subst comp_fun_commute.fold_insert[OF and_comm], simp) + apply(subst and_comm.fold_insert, simp) apply (metis empty_iff image_empty) apply(simp add: invalid_def) apply (metis bot_fun_def destruct_ocl null_fun_def) (* *) apply(simp only: image_insert) - apply(subst comp_fun_commute.fold_insert[OF and_comm], simp) + apply(subst and_comm.fold_insert, simp) apply (metis (mono_tags) imageE) (* *) apply(subst cp_OclAnd) apply(drule sym, drule sym, simp only:, drule sym, simp only:) apply(simp only: ex_insert) apply(subgoal_tac "\x. x\F") prefer 2 apply(metis all_not_in_conv) proof - fix x F show "(\ S) \ = true \ \ \x. x \ F \ ?forall (\b \. ?P_eq x b \ \ ?P F b \) = ((\_. ?forall (?P F)) and (\_. P (\\. x) \)) \" apply(rule disjE4[OF destruct_ocl[where x1 = "P (\\. x) \"]]) apply(simp_all add: true_def false_def invalid_def OclAnd_def null_fun_def null_option_def bot_fun_def bot_option_def) by (metis (lifting) option.distinct(1))+ qed(simp add: OclValid_def)+ qed lemma OclForall_cong: assumes "\x. x \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ \ \ \ P (\\. x) \ \ \ Q (\\. x)" assumes P: "\ \ OclForall X P" shows "\ \ OclForall X Q" proof - have def_X: "\ \ \ X" by(insert P, simp add: OclForall_def OclValid_def bot_option_def true_def split: if_split_asm) show ?thesis apply(insert P) apply(subst (asm) OclForall_rep_set_true[OF def_X], subst OclForall_rep_set_true[OF def_X]) by (simp add: assms) qed lemma OclForall_cong': assumes "\x. x \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ \ \ \ P (\\. x) \ \ \ Q (\\. x) \ \ \ R (\\. x)" assumes P: "\ \ OclForall X P" assumes Q: "\ \ OclForall X Q" shows "\ \ OclForall X R" proof - have def_X: "\ \ \ X" by(insert P, simp add: OclForall_def OclValid_def bot_option_def true_def split: if_split_asm) show ?thesis apply(insert P Q) apply(subst (asm) (1 2) OclForall_rep_set_true[OF def_X], subst OclForall_rep_set_true[OF def_X]) by (simp add: assms) qed text\Strict Equality\ lemma StrictRefEq\<^sub>S\<^sub>e\<^sub>t_defined : assumes x_def: "\ \ \ x" assumes y_def: "\ \ \ y" shows "((x::('\,'\::null)Set) \ y) \ = (x->forAll\<^sub>S\<^sub>e\<^sub>t(z| y->includes\<^sub>S\<^sub>e\<^sub>t(z)) and (y->forAll\<^sub>S\<^sub>e\<^sub>t(z| x->includes\<^sub>S\<^sub>e\<^sub>t(z)))) \" proof - have rep_set_inj : "\\. (\ x) \ = true \ \ (\ y) \ = true \ \ x \ \ y \ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (y \)\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \)\\" apply(simp add: defined_def) apply(split if_split_asm, simp add: false_def true_def)+ apply(simp add: null_fun_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_fun_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def) apply(case_tac "x \", rename_tac x') apply(case_tac x', simp_all, rename_tac x'') apply(case_tac x'', simp_all) apply(case_tac "y \", rename_tac y') apply(case_tac y', simp_all, rename_tac y'') apply(case_tac y'', simp_all) apply(simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse) by(blast) show ?thesis apply(simp add: StrictRefEq\<^sub>S\<^sub>e\<^sub>t StrongEq_def foundation20[OF x_def, simplified OclValid_def] foundation20[OF y_def, simplified OclValid_def]) apply(subgoal_tac "\\x \ = y \\\ = true \ \ \\x \ = y \\\ = false \") prefer 2 apply(simp add: false_def true_def) (* *) apply(erule disjE) apply(simp add: true_def) apply(subgoal_tac "(\ \ OclForall x (OclIncludes y)) \ (\ \ OclForall y (OclIncludes x))") apply(subst cp_OclAnd, simp add: true_def OclValid_def) apply(simp add: OclForall_includes[OF x_def y_def] OclForall_includes[OF y_def x_def]) (* *) apply(simp) apply(subgoal_tac "OclForall x (OclIncludes y) \ = false \ \ OclForall y (OclIncludes x) \ = false \") apply(subst cp_OclAnd, metis OclAnd_false1 OclAnd_false2 cp_OclAnd) apply(simp only: OclForall_not_includes[OF x_def y_def, simplified OclValid_def] OclForall_not_includes[OF y_def x_def, simplified OclValid_def], simp add: false_def) by (metis OclValid_def rep_set_inj subset_antisym x_def y_def) qed lemma StrictRefEq\<^sub>S\<^sub>e\<^sub>t_exec[simp,code_unfold] : "((x::('\,'\::null)Set) \ y) = (if \ x then (if \ y then ((x->forAll\<^sub>S\<^sub>e\<^sub>t(z| y->includes\<^sub>S\<^sub>e\<^sub>t(z)) and (y->forAll\<^sub>S\<^sub>e\<^sub>t(z| x->includes\<^sub>S\<^sub>e\<^sub>t(z))))) else if \ y then false \ \\x'->includes = null\\ else invalid endif endif) else if \ x \ \\null = ???\\ then if \ y then not(\ y) else invalid endif else invalid endif endif)" proof - have defined_inject_true : "\\ P. (\ (\ \ \ P)) = ((\ P) \ = false \)" by (metis bot_fun_def OclValid_def defined_def foundation16 null_fun_def) have valid_inject_true : "\\ P. (\ (\ \ \ P)) = ((\ P) \ = false \)" by (metis bot_fun_def OclIf_true' OclIncludes_charn0 OclIncludes_charn0' OclValid_def valid_def foundation6 foundation9) show ?thesis apply(rule ext, rename_tac \) (* *) apply(simp add: OclIf_def defined_inject_true[simplified OclValid_def] valid_inject_true[simplified OclValid_def], subst false_def, subst true_def, simp) apply(subst (1 2) cp_OclNot, simp, simp add: cp_OclNot[symmetric]) apply(simp add: StrictRefEq\<^sub>S\<^sub>e\<^sub>t_defined[simplified OclValid_def]) by(simp add: StrictRefEq\<^sub>S\<^sub>e\<^sub>t StrongEq_def false_def true_def valid_def defined_def) qed lemma StrictRefEq\<^sub>S\<^sub>e\<^sub>t_L_subst1 : "cp P \ \ \ \ x \ \ \ \ y \ \ \ \ P x \ \ \ \ P y \ \ \ (x::('\,'\::null)Set) \ y \ \ \ (P x ::('\,'\::null)Set) \ P y" apply(simp only: StrictRefEq\<^sub>S\<^sub>e\<^sub>t OclValid_def) apply(split if_split_asm) apply(simp add: StrongEq_L_subst1[simplified OclValid_def]) by (simp add: invalid_def bot_option_def true_def) lemma OclIncluding_cong' : shows "\ \ \ s \ \ \ \ t \ \ \ \ x \ \ \ ((s::('\,'a::null)Set) \ t) \ \ \ (s->including\<^sub>S\<^sub>e\<^sub>t(x) \ (t->including\<^sub>S\<^sub>e\<^sub>t(x)))" proof - have cp: "cp (\s. (s->including\<^sub>S\<^sub>e\<^sub>t(x)))" apply(simp add: cp_def, subst OclIncluding.cp0) by (rule_tac x = "(\xab ab. ((\_. xab)->including\<^sub>S\<^sub>e\<^sub>t(\_. x ab)) ab)" in exI, simp) show "\ \ \ s \ \ \ \ t \ \ \ \ x \ \ \ (s \ t) \ ?thesis" apply(rule_tac P = "\s. (s->including\<^sub>S\<^sub>e\<^sub>t(x))" in StrictRefEq\<^sub>S\<^sub>e\<^sub>t_L_subst1) apply(rule cp) apply(simp add: foundation20) apply(simp add: foundation20) apply (simp add: foundation10 foundation6)+ done qed lemma OclIncluding_cong : "\(s::('\,'a::null)Set) t x y \. \ \ \ t \ \ \ \ y \ \ \ s \ t \ x = y \ \ \ s->including\<^sub>S\<^sub>e\<^sub>t(x) \ (t->including\<^sub>S\<^sub>e\<^sub>t(y))" apply(simp only:) apply(rule OclIncluding_cong', simp_all only:) by(auto simp: OclValid_def OclIf_def invalid_def bot_option_def OclNot_def split : if_split_asm) (* < *) lemma const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_empty : "const X \ const (X \ Set{})" apply(rule StrictRefEq\<^sub>S\<^sub>e\<^sub>t.const, assumption) by(simp) lemma const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including : "const a \ const S \ const X \ const (X \ S->including\<^sub>S\<^sub>e\<^sub>t(a))" apply(rule StrictRefEq\<^sub>S\<^sub>e\<^sub>t.const, assumption) by(rule const_OclIncluding) subsection\Test Statements\ Assert "(\ \ (Set{\_. \\x\\} \ Set{\_. \\x\\}))" Assert "(\ \ (Set{\_. \x\} \ Set{\_. \x\}))" instantiation Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e :: (equal)equal begin definition "HOL.equal k l \ (k::('a::equal)Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e) = l" instance by standard (rule equal_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def) end lemma equal_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_code [code]: "HOL.equal k (l::('a::{equal,null})Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e) \ Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e k = Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e l" by (auto simp add: equal Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e.Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject) Assert "\ \ (Set{} \ Set{})" Assert "\ \ (Set{\,\} \ Set{}->including\<^sub>S\<^sub>e\<^sub>t(\)->including\<^sub>S\<^sub>e\<^sub>t(\))" Assert "\ \ (Set{\,invalid,\} \ invalid)" Assert "\ \ (Set{\,\}->including\<^sub>S\<^sub>e\<^sub>t(null) \ Set{null,\,\})" Assert "\ \ (Set{\,\}->including\<^sub>S\<^sub>e\<^sub>t(null) \ Set{\,\,null})" (* Assert "\ (\ \ (Set{\,\,\} \ Set{\,\}))" Assert "\ (\ \ (Set{\,\} \ Set{\,\}))" *) (* > *) end diff --git a/thys/FinFun/FinFun.thy b/thys/FinFun/FinFun.thy --- a/thys/FinFun/FinFun.thy +++ b/thys/FinFun/FinFun.thy @@ -1,1582 +1,1582 @@ (* Author: Andreas Lochbihler, Uni Karlsruhe *) section \Almost everywhere constant functions\ theory FinFun imports "HOL-Library.Cardinality" begin text \ This theory defines functions which are constant except for finitely many points (FinFun) and introduces a type finfin along with a number of operators for them. The code generator is set up such that such functions can be represented as data in the generated code and all operators are executable. For details, see Formalising FinFuns - Generating Code for Functions as Data by A. Lochbihler in TPHOLs 2009. \ subsection \The \map_default\ operation\ definition map_default :: "'b \ ('a \ 'b) \ 'a \ 'b" where "map_default b f a \ case f a of None \ b | Some b' \ b'" lemma map_default_delete [simp]: "map_default b (f(a := None)) = (map_default b f)(a := b)" by(simp add: map_default_def fun_eq_iff) lemma map_default_insert: "map_default b (f(a \ b')) = (map_default b f)(a := b')" by(simp add: map_default_def fun_eq_iff) lemma map_default_empty [simp]: "map_default b Map.empty = (\a. b)" by(simp add: fun_eq_iff map_default_def) lemma map_default_inject: fixes g g' :: "'a \ 'b" assumes infin_eq: "\ finite (UNIV :: 'a set) \ b = b'" and fin: "finite (dom g)" and b: "b \ ran g" and fin': "finite (dom g')" and b': "b' \ ran g'" and eq': "map_default b g = map_default b' g'" shows "b = b'" "g = g'" proof - from infin_eq show bb': "b = b'" proof assume infin: "\ finite (UNIV :: 'a set)" from fin fin' have "finite (dom g \ dom g')" by auto with infin have "UNIV - (dom g \ dom g') \ {}" by(auto dest: finite_subset) then obtain a where a: "a \ dom g \ dom g'" by auto hence "map_default b g a = b" "map_default b' g' a = b'" by(auto simp add: map_default_def) with eq' show "b = b'" by simp qed show "g = g'" proof fix x show "g x = g' x" proof(cases "g x") case None hence "map_default b g x = b" by(simp add: map_default_def) with bb' eq' have "map_default b' g' x = b'" by simp with b' have "g' x = None" by(simp add: map_default_def ran_def split: option.split_asm) with None show ?thesis by simp next case (Some c) with b have cb: "c \ b" by(auto simp add: ran_def) moreover from Some have "map_default b g x = c" by(simp add: map_default_def) with eq' have "map_default b' g' x = c" by simp ultimately have "g' x = Some c" using b' bb' by(auto simp add: map_default_def split: option.splits) with Some show ?thesis by simp qed qed qed subsection \The finfun type\ definition "finfun = {f::'a\'b. \b. finite {a. f a \ b}}" typedef ('a,'b) finfun ("(_ \f /_)" [22, 21] 21) = "finfun :: ('a => 'b) set" morphisms finfun_apply Abs_finfun proof - have "\f. finite {x. f x \ undefined}" proof show "finite {x. (\y. undefined) x \ undefined}" by auto qed then show ?thesis unfolding finfun_def by auto qed type_notation finfun ("(_ \f /_)" [22, 21] 21) setup_lifting type_definition_finfun lemma fun_upd_finfun: "y(a := b) \ finfun \ y \ finfun" proof - { fix b' have "finite {a'. (y(a := b)) a' \ b'} = finite {a'. y a' \ b'}" proof(cases "b = b'") case True hence "{a'. (y(a := b)) a' \ b'} = {a'. y a' \ b'} - {a}" by auto thus ?thesis by simp next case False hence "{a'. (y(a := b)) a' \ b'} = insert a {a'. y a' \ b'}" by auto thus ?thesis by simp qed } thus ?thesis unfolding finfun_def by blast qed lemma const_finfun: "(\x. a) \ finfun" by(auto simp add: finfun_def) lemma finfun_left_compose: assumes "y \ finfun" shows "g \ y \ finfun" proof - from assms obtain b where "finite {a. y a \ b}" unfolding finfun_def by blast hence "finite {c. g (y c) \ g b}" proof(induct "{a. y a \ b}" arbitrary: y) case empty hence "y = (\a. b)" by(auto) thus ?case by(simp) next case (insert x F) note IH = \\y. F = {a. y a \ b} \ finite {c. g (y c) \ g b}\ from \insert x F = {a. y a \ b}\ \x \ F\ have F: "F = {a. (y(x := b)) a \ b}" by(auto) show ?case proof(cases "g (y x) = g b") case True hence "{c. g ((y(x := b)) c) \ g b} = {c. g (y c) \ g b}" by auto with IH[OF F] show ?thesis by simp next case False hence "{c. g (y c) \ g b} = insert x {c. g ((y(x := b)) c) \ g b}" by auto with IH[OF F] show ?thesis by(simp) qed qed thus ?thesis unfolding finfun_def by auto qed lemma assumes "y \ finfun" shows fst_finfun: "fst \ y \ finfun" and snd_finfun: "snd \ y \ finfun" proof - from assms obtain b c where bc: "finite {a. y a \ (b, c)}" unfolding finfun_def by auto have "{a. fst (y a) \ b} \ {a. y a \ (b, c)}" and "{a. snd (y a) \ c} \ {a. y a \ (b, c)}" by auto hence "finite {a. fst (y a) \ b}" and "finite {a. snd (y a) \ c}" using bc by(auto intro: finite_subset) thus "fst \ y \ finfun" "snd \ y \ finfun" unfolding finfun_def by auto qed lemma map_of_finfun: "map_of xs \ finfun" unfolding finfun_def by(induct xs)(auto simp add: Collect_neg_eq Collect_conj_eq Collect_imp_eq intro: finite_subset) lemma Diag_finfun: "(\x. (f x, g x)) \ finfun \ f \ finfun \ g \ finfun" by(auto intro: finite_subset simp add: Collect_neg_eq Collect_imp_eq Collect_conj_eq finfun_def) lemma finfun_right_compose: assumes g: "g \ finfun" and inj: "inj f" shows "g o f \ finfun" proof - from g obtain b where b: "finite {a. g a \ b}" unfolding finfun_def by blast moreover have "f ` {a. g (f a) \ b} \ {a. g a \ b}" by auto moreover from inj have "inj_on f {a. g (f a) \ b}" by(rule subset_inj_on) blast ultimately have "finite {a. g (f a) \ b}" by(blast intro: finite_imageD[where f=f] finite_subset) thus ?thesis unfolding finfun_def by auto qed lemma finfun_curry: assumes fin: "f \ finfun" shows "curry f \ finfun" "curry f a \ finfun" proof - from fin obtain c where c: "finite {ab. f ab \ c}" unfolding finfun_def by blast moreover have "{a. \b. f (a, b) \ c} = fst ` {ab. f ab \ c}" by(force) hence "{a. curry f a \ (\b. c)} = fst ` {ab. f ab \ c}" by(auto simp add: curry_def fun_eq_iff) ultimately have "finite {a. curry f a \ (\b. c)}" by simp thus "curry f \ finfun" unfolding finfun_def by blast have "snd ` {ab. f ab \ c} = {b. \a. f (a, b) \ c}" by(force) hence "{b. f (a, b) \ c} \ snd ` {ab. f ab \ c}" by auto hence "finite {b. f (a, b) \ c}" by(rule finite_subset)(rule finite_imageI[OF c]) thus "curry f a \ finfun" unfolding finfun_def by auto qed bundle finfun begin lemmas [simp] = fst_finfun snd_finfun Abs_finfun_inverse finfun_apply_inverse Abs_finfun_inject finfun_apply_inject Diag_finfun finfun_curry lemmas [iff] = const_finfun fun_upd_finfun finfun_apply map_of_finfun lemmas [intro] = finfun_left_compose fst_finfun snd_finfun end lemma Abs_finfun_inject_finite: fixes x y :: "'a \ 'b" assumes fin: "finite (UNIV :: 'a set)" shows "Abs_finfun x = Abs_finfun y \ x = y" proof assume "Abs_finfun x = Abs_finfun y" moreover have "x \ finfun" "y \ finfun" unfolding finfun_def by(auto intro: finite_subset[OF _ fin]) ultimately show "x = y" by(simp add: Abs_finfun_inject) qed simp lemma Abs_finfun_inject_finite_class: fixes x y :: "('a :: finite) \ 'b" shows "Abs_finfun x = Abs_finfun y \ x = y" using finite_UNIV by(simp add: Abs_finfun_inject_finite) lemma Abs_finfun_inj_finite: assumes fin: "finite (UNIV :: 'a set)" shows "inj (Abs_finfun :: ('a \ 'b) \ 'a \f 'b)" proof(rule inj_onI) fix x y :: "'a \ 'b" assume "Abs_finfun x = Abs_finfun y" moreover have "x \ finfun" "y \ finfun" unfolding finfun_def by(auto intro: finite_subset[OF _ fin]) ultimately show "x = y" by(simp add: Abs_finfun_inject) qed lemma Abs_finfun_inverse_finite: fixes x :: "'a \ 'b" assumes fin: "finite (UNIV :: 'a set)" shows "finfun_apply (Abs_finfun x) = x" including finfun proof - from fin have "x \ finfun" by(auto simp add: finfun_def intro: finite_subset) thus ?thesis by simp qed lemma Abs_finfun_inverse_finite_class: fixes x :: "('a :: finite) \ 'b" shows "finfun_apply (Abs_finfun x) = x" using finite_UNIV by(simp add: Abs_finfun_inverse_finite) lemma finfun_eq_finite_UNIV: "finite (UNIV :: 'a set) \ (finfun :: ('a \ 'b) set) = UNIV" unfolding finfun_def by(auto intro: finite_subset) lemma finfun_finite_UNIV_class: "finfun = (UNIV :: ('a :: finite \ 'b) set)" by(simp add: finfun_eq_finite_UNIV) lemma map_default_in_finfun: assumes fin: "finite (dom f)" shows "map_default b f \ finfun" unfolding finfun_def proof(intro CollectI exI) from fin show "finite {a. map_default b f a \ b}" by(auto simp add: map_default_def dom_def Collect_conj_eq split: option.splits) qed lemma finfun_cases_map_default: obtains b g where "f = Abs_finfun (map_default b g)" "finite (dom g)" "b \ ran g" proof - obtain y where f: "f = Abs_finfun y" and y: "y \ finfun" by(cases f) from y obtain b where b: "finite {a. y a \ b}" unfolding finfun_def by auto let ?g = "(\a. if y a = b then None else Some (y a))" have "map_default b ?g = y" by(simp add: fun_eq_iff map_default_def) with f have "f = Abs_finfun (map_default b ?g)" by simp moreover from b have "finite (dom ?g)" by(auto simp add: dom_def) moreover have "b \ ran ?g" by(auto simp add: ran_def) ultimately show ?thesis by(rule that) qed subsection \Kernel functions for type @{typ "'a \f 'b"}\ lift_definition finfun_const :: "'b \ 'a \f 'b" ("K$/ _" [0] 1) is "\ b x. b" by (rule const_finfun) lift_definition finfun_update :: "'a \f 'b \ 'a \ 'b \ 'a \f 'b" ("_'(_ $:= _')" [1000,0,0] 1000) is "fun_upd" by (simp add: fun_upd_finfun) lemma finfun_update_twist: "a \ a' \ f(a $:= b)(a' $:= b') = f(a' $:= b')(a $:= b)" by transfer (simp add: fun_upd_twist) lemma finfun_update_twice [simp]: "f(a $:= b)(a $:= b') = f(a $:= b')" by transfer simp lemma finfun_update_const_same: "(K$ b)(a $:= b) = (K$ b)" by transfer (simp add: fun_eq_iff) subsection \Code generator setup\ definition finfun_update_code :: "'a \f 'b \ 'a \ 'b \ 'a \f 'b" where [simp, code del]: "finfun_update_code = finfun_update" code_datatype finfun_const finfun_update_code lemma finfun_update_const_code [code]: "(K$ b)(a $:= b') = (if b = b' then (K$ b) else finfun_update_code (K$ b) a b')" by(simp add: finfun_update_const_same) lemma finfun_update_update_code [code]: "(finfun_update_code f a b)(a' $:= b') = (if a = a' then f(a $:= b') else finfun_update_code (f(a' $:= b')) a b)" by(simp add: finfun_update_twist) subsection \Setup for quickcheck\ quickcheck_generator finfun constructors: finfun_update_code, "finfun_const :: 'b \ 'a \f 'b" subsection \\finfun_update\ as instance of \comp_fun_commute\\ interpretation finfun_update: comp_fun_commute "\a f. f(a :: 'a $:= b')" including finfun proof fix a a' :: 'a show "(\f. f(a $:= b')) \ (\f. f(a' $:= b')) = (\f. f(a' $:= b')) \ (\f. f(a $:= b'))" proof fix b have "(finfun_apply b)(a := b', a' := b') = (finfun_apply b)(a' := b', a := b')" by(cases "a = a'")(auto simp add: fun_upd_twist) then have "b(a $:= b')(a' $:= b') = b(a' $:= b')(a $:= b')" by(auto simp add: finfun_update_def fun_upd_twist) then show "((\f. f(a $:= b')) \ (\f. f(a' $:= b'))) b = ((\f. f(a' $:= b')) \ (\f. f(a $:= b'))) b" by (simp add: fun_eq_iff) qed qed lemma fold_finfun_update_finite_univ: assumes fin: "finite (UNIV :: 'a set)" shows "Finite_Set.fold (\a f. f(a $:= b')) (K$ b) (UNIV :: 'a set) = (K$ b')" proof - { fix A :: "'a set" from fin have "finite A" by(auto intro: finite_subset) hence "Finite_Set.fold (\a f. f(a $:= b')) (K$ b) A = Abs_finfun (\a. if a \ A then b' else b)" proof(induct) case (insert x F) have "(\a. if a = x then b' else (if a \ F then b' else b)) = (\a. if a = x \ a \ F then b' else b)" by(auto) with insert show ?case by(simp add: finfun_const_def fun_upd_def)(simp add: finfun_update_def Abs_finfun_inverse_finite[OF fin] fun_upd_def) qed(simp add: finfun_const_def) } thus ?thesis by(simp add: finfun_const_def) qed subsection \Default value for FinFuns\ definition finfun_default_aux :: "('a \ 'b) \ 'b" where [code del]: "finfun_default_aux f = (if finite (UNIV :: 'a set) then undefined else THE b. finite {a. f a \ b})" lemma finfun_default_aux_infinite: fixes f :: "'a \ 'b" assumes infin: "\ finite (UNIV :: 'a set)" and fin: "finite {a. f a \ b}" shows "finfun_default_aux f = b" proof - let ?B = "{a. f a \ b}" from fin have "(THE b. finite {a. f a \ b}) = b" proof(rule the_equality) fix b' assume "finite {a. f a \ b'}" (is "finite ?B'") with infin fin have "UNIV - (?B' \ ?B) \ {}" by(auto dest: finite_subset) then obtain a where a: "a \ ?B' \ ?B" by auto thus "b' = b" by auto qed thus ?thesis using infin by(simp add: finfun_default_aux_def) qed lemma finite_finfun_default_aux: fixes f :: "'a \ 'b" assumes fin: "f \ finfun" shows "finite {a. f a \ finfun_default_aux f}" proof(cases "finite (UNIV :: 'a set)") case True thus ?thesis using fin by(auto simp add: finfun_def finfun_default_aux_def intro: finite_subset) next case False from fin obtain b where b: "finite {a. f a \ b}" (is "finite ?B") unfolding finfun_def by blast with False show ?thesis by(simp add: finfun_default_aux_infinite) qed lemma finfun_default_aux_update_const: fixes f :: "'a \ 'b" assumes fin: "f \ finfun" shows "finfun_default_aux (f(a := b)) = finfun_default_aux f" proof(cases "finite (UNIV :: 'a set)") case False from fin obtain b' where b': "finite {a. f a \ b'}" unfolding finfun_def by blast hence "finite {a'. (f(a := b)) a' \ b'}" proof(cases "b = b' \ f a \ b'") case True hence "{a. f a \ b'} = insert a {a'. (f(a := b)) a' \ b'}" by auto thus ?thesis using b' by simp next case False moreover { assume "b \ b'" hence "{a'. (f(a := b)) a' \ b'} = insert a {a. f a \ b'}" by auto hence ?thesis using b' by simp } moreover { assume "b = b'" "f a = b'" hence "{a'. (f(a := b)) a' \ b'} = {a. f a \ b'}" by auto hence ?thesis using b' by simp } ultimately show ?thesis by blast qed with False b' show ?thesis by(auto simp del: fun_upd_apply simp add: finfun_default_aux_infinite) next case True thus ?thesis by(simp add: finfun_default_aux_def) qed lift_definition finfun_default :: "'a \f 'b \ 'b" is "finfun_default_aux" . lemma finite_finfun_default: "finite {a. finfun_apply f a \ finfun_default f}" by transfer (erule finite_finfun_default_aux) lemma finfun_default_const: "finfun_default ((K$ b) :: 'a \f 'b) = (if finite (UNIV :: 'a set) then undefined else b)" by(transfer)(auto simp add: finfun_default_aux_infinite finfun_default_aux_def) lemma finfun_default_update_const: "finfun_default (f(a $:= b)) = finfun_default f" by transfer (simp add: finfun_default_aux_update_const) lemma finfun_default_const_code [code]: "finfun_default ((K$ c) :: 'a :: card_UNIV \f 'b) = (if CARD('a) = 0 then c else undefined)" by(simp add: finfun_default_const) lemma finfun_default_update_code [code]: "finfun_default (finfun_update_code f a b) = finfun_default f" by(simp add: finfun_default_update_const) subsection \Recursion combinator and well-formedness conditions\ definition finfun_rec :: "('b \ 'c) \ ('a \ 'b \ 'c \ 'c) \ ('a \f 'b) \ 'c" where [code del]: "finfun_rec cnst upd f \ let b = finfun_default f; g = THE g. f = Abs_finfun (map_default b g) \ finite (dom g) \ b \ ran g in Finite_Set.fold (\a. upd a (map_default b g a)) (cnst b) (dom g)" locale finfun_rec_wf_aux = fixes cnst :: "'b \ 'c" and upd :: "'a \ 'b \ 'c \ 'c" assumes upd_const_same: "upd a b (cnst b) = cnst b" and upd_commute: "a \ a' \ upd a b (upd a' b' c) = upd a' b' (upd a b c)" and upd_idemp: "b \ b' \ upd a b'' (upd a b' (cnst b)) = upd a b'' (cnst b)" begin lemma upd_left_comm: "comp_fun_commute (\a. upd a (f a))" by(unfold_locales)(auto intro: upd_commute simp add: fun_eq_iff) lemma upd_upd_twice: "upd a b'' (upd a b' (cnst b)) = upd a b'' (cnst b)" by(cases "b \ b'")(auto simp add: fun_upd_def upd_const_same upd_idemp) lemma map_default_update_const: assumes fin: "finite (dom f)" and anf: "a \ dom f" and fg: "f \\<^sub>m g" shows "upd a d (Finite_Set.fold (\a. upd a (map_default d g a)) (cnst d) (dom f)) = Finite_Set.fold (\a. upd a (map_default d g a)) (cnst d) (dom f)" proof - let ?upd = "\a. upd a (map_default d g a)" let ?fr = "\A. Finite_Set.fold ?upd (cnst d) A" interpret gwf: comp_fun_commute "?upd" by(rule upd_left_comm) from fin anf fg show ?thesis proof(induct "dom f" arbitrary: f) case empty from \{} = dom f\ have "f = Map.empty" by(auto simp add: dom_def) thus ?case by(simp add: finfun_const_def upd_const_same) next case (insert a' A) note IH = \\f. \ A = dom f; a \ dom f; f \\<^sub>m g \ \ upd a d (?fr (dom f)) = ?fr (dom f)\ note fin = \finite A\ note anf = \a \ dom f\ note a'nA = \a' \ A\ note domf = \insert a' A = dom f\ note fg = \f \\<^sub>m g\ from domf obtain b where b: "f a' = Some b" by auto let ?f' = "f(a' := None)" have "upd a d (?fr (insert a' A)) = upd a d (upd a' (map_default d g a') (?fr A))" by(subst gwf.fold_insert[OF fin a'nA]) rule also from b fg have "g a' = f a'" by(auto simp add: map_le_def intro: domI dest: bspec) hence ga': "map_default d g a' = map_default d f a'" by(simp add: map_default_def) also from anf domf have "a \ a'" by auto note upd_commute[OF this] also from domf a'nA anf fg have "a \ dom ?f'" "?f' \\<^sub>m g" and A: "A = dom ?f'" by(auto simp add: ran_def map_le_def) note A also note IH[OF A \a \ dom ?f'\ \?f' \\<^sub>m g\] also have "upd a' (map_default d f a') (?fr (dom (f(a' := None)))) = ?fr (dom f)" unfolding domf[symmetric] gwf.fold_insert[OF fin a'nA] ga' unfolding A .. also have "insert a' (dom ?f') = dom f" using domf by auto finally show ?case . qed qed lemma map_default_update_twice: assumes fin: "finite (dom f)" and anf: "a \ dom f" and fg: "f \\<^sub>m g" shows "upd a d'' (upd a d' (Finite_Set.fold (\a. upd a (map_default d g a)) (cnst d) (dom f))) = upd a d'' (Finite_Set.fold (\a. upd a (map_default d g a)) (cnst d) (dom f))" proof - let ?upd = "\a. upd a (map_default d g a)" let ?fr = "\A. Finite_Set.fold ?upd (cnst d) A" interpret gwf: comp_fun_commute "?upd" by(rule upd_left_comm) from fin anf fg show ?thesis proof(induct "dom f" arbitrary: f) case empty from \{} = dom f\ have "f = Map.empty" by(auto simp add: dom_def) thus ?case by(auto simp add: finfun_const_def finfun_update_def upd_upd_twice) next case (insert a' A) note IH = \\f. \A = dom f; a \ dom f; f \\<^sub>m g\ \ upd a d'' (upd a d' (?fr (dom f))) = upd a d'' (?fr (dom f))\ note fin = \finite A\ note anf = \a \ dom f\ note a'nA = \a' \ A\ note domf = \insert a' A = dom f\ note fg = \f \\<^sub>m g\ from domf obtain b where b: "f a' = Some b" by auto let ?f' = "f(a' := None)" let ?b' = "case f a' of None \ d | Some b \ b" from domf have "upd a d'' (upd a d' (?fr (dom f))) = upd a d'' (upd a d' (?fr (insert a' A)))" by simp also note gwf.fold_insert[OF fin a'nA] also from b fg have "g a' = f a'" by(auto simp add: map_le_def intro: domI dest: bspec) hence ga': "map_default d g a' = map_default d f a'" by(simp add: map_default_def) also from anf domf have ana': "a \ a'" by auto note upd_commute[OF this] also note upd_commute[OF ana'] also from domf a'nA anf fg have "a \ dom ?f'" "?f' \\<^sub>m g" and A: "A = dom ?f'" by(auto simp add: ran_def map_le_def) note A also note IH[OF A \a \ dom ?f'\ \?f' \\<^sub>m g\] also note upd_commute[OF ana'[symmetric]] also note ga'[symmetric] also note A[symmetric] also note gwf.fold_insert[symmetric, OF fin a'nA] also note domf finally show ?case . qed qed lemma map_default_eq_id [simp]: "map_default d ((\a. Some (f a)) |` {a. f a \ d}) = f" by(auto simp add: map_default_def restrict_map_def) lemma finite_rec_cong1: assumes f: "comp_fun_commute f" and g: "comp_fun_commute g" and fin: "finite A" and eq: "\a. a \ A \ f a = g a" shows "Finite_Set.fold f z A = Finite_Set.fold g z A" proof - interpret f: comp_fun_commute f by(rule f) interpret g: comp_fun_commute g by(rule g) { fix B assume BsubA: "B \ A" with fin have "finite B" by(blast intro: finite_subset) hence "B \ A \ Finite_Set.fold f z B = Finite_Set.fold g z B" proof(induct) case empty thus ?case by simp next case (insert a B) note finB = \finite B\ note anB = \a \ B\ note sub = \insert a B \ A\ note IH = \B \ A \ Finite_Set.fold f z B = Finite_Set.fold g z B\ from sub anB have BpsubA: "B \ A" and BsubA: "B \ A" and aA: "a \ A" by auto from IH[OF BsubA] eq[OF aA] finB anB show ?case by(auto) qed with BsubA have "Finite_Set.fold f z B = Finite_Set.fold g z B" by blast } thus ?thesis by blast qed lemma finfun_rec_upd [simp]: "finfun_rec cnst upd (f(a' $:= b')) = upd a' b' (finfun_rec cnst upd f)" including finfun proof - obtain b where b: "b = finfun_default f" by auto let ?the = "\f g. f = Abs_finfun (map_default b g) \ finite (dom g) \ b \ ran g" obtain g where g: "g = The (?the f)" by blast obtain y where f: "f = Abs_finfun y" and y: "y \ finfun" by (cases f) from f y b have bfin: "finite {a. y a \ b}" by(simp add: finfun_default_def finite_finfun_default_aux) let ?g = "(\a. Some (y a)) |` {a. y a \ b}" from bfin have fing: "finite (dom ?g)" by auto have bran: "b \ ran ?g" by(auto simp add: ran_def restrict_map_def) have yg: "y = map_default b ?g" by simp have gg: "g = ?g" unfolding g proof(rule the_equality) from f y bfin show "?the f ?g" by(auto)(simp add: restrict_map_def ran_def split: if_split_asm) next fix g' assume "?the f g'" hence fin': "finite (dom g')" and ran': "b \ ran g'" and eq: "Abs_finfun (map_default b ?g) = Abs_finfun (map_default b g')" using f yg by auto from fin' fing have "map_default b ?g \ finfun" "map_default b g' \ finfun" by(blast intro: map_default_in_finfun)+ with eq have "map_default b ?g = map_default b g'" by simp with fing bran fin' ran' show "g' = ?g" by(rule map_default_inject[OF disjI2[OF refl], THEN sym]) qed show ?thesis proof(cases "b' = b") case True note b'b = True let ?g' = "(\a. Some ((y(a' := b)) a)) |` {a. (y(a' := b)) a \ b}" from bfin b'b have fing': "finite (dom ?g')" by(auto simp add: Collect_conj_eq Collect_imp_eq intro: finite_subset) have brang': "b \ ran ?g'" by(auto simp add: ran_def restrict_map_def) let ?b' = "\a. case ?g' a of None \ b | Some b \ b" let ?b = "map_default b ?g" from upd_left_comm upd_left_comm fing' have "Finite_Set.fold (\a. upd a (?b' a)) (cnst b) (dom ?g') = Finite_Set.fold (\a. upd a (?b a)) (cnst b) (dom ?g')" by(rule finite_rec_cong1)(auto simp add: restrict_map_def b'b b map_default_def) also interpret gwf: comp_fun_commute "\a. upd a (?b a)" by(rule upd_left_comm) have "Finite_Set.fold (\a. upd a (?b a)) (cnst b) (dom ?g') = upd a' b' (Finite_Set.fold (\a. upd a (?b a)) (cnst b) (dom ?g))" proof(cases "y a' = b") case True with b'b have g': "?g' = ?g" by(auto simp add: restrict_map_def) from True have a'ndomg: "a' \ dom ?g" by auto from f b'b b show ?thesis unfolding g' by(subst map_default_update_const[OF fing a'ndomg map_le_refl, symmetric]) simp next case False hence domg: "dom ?g = insert a' (dom ?g')" by auto from False b'b have a'ndomg': "a' \ dom ?g'" by auto have "Finite_Set.fold (\a. upd a (?b a)) (cnst b) (insert a' (dom ?g')) = upd a' (?b a') (Finite_Set.fold (\a. upd a (?b a)) (cnst b) (dom ?g'))" using fing' a'ndomg' unfolding b'b by(rule gwf.fold_insert) hence "upd a' b (Finite_Set.fold (\a. upd a (?b a)) (cnst b) (insert a' (dom ?g'))) = upd a' b (upd a' (?b a') (Finite_Set.fold (\a. upd a (?b a)) (cnst b) (dom ?g')))" by simp also from b'b have g'leg: "?g' \\<^sub>m ?g" by(auto simp add: restrict_map_def map_le_def) note map_default_update_twice[OF fing' a'ndomg' this, of b "?b a'" b] also note map_default_update_const[OF fing' a'ndomg' g'leg, of b] finally show ?thesis unfolding b'b domg[unfolded b'b] by(rule sym) qed also have "The (?the (f(a' $:= b'))) = ?g'" proof(rule the_equality) from f y b b'b brang' fing' show "?the (f(a' $:= b')) ?g'" by(auto simp del: fun_upd_apply simp add: finfun_update_def) next fix g' assume "?the (f(a' $:= b')) g'" hence fin': "finite (dom g')" and ran': "b \ ran g'" and eq: "f(a' $:= b') = Abs_finfun (map_default b g')" by(auto simp del: fun_upd_apply) from fin' fing' have "map_default b g' \ finfun" "map_default b ?g' \ finfun" by(blast intro: map_default_in_finfun)+ with eq f b'b b have "map_default b ?g' = map_default b g'" by(simp del: fun_upd_apply add: finfun_update_def) with fing' brang' fin' ran' show "g' = ?g'" by(rule map_default_inject[OF disjI2[OF refl], THEN sym]) qed ultimately show ?thesis unfolding finfun_rec_def Let_def b gg[unfolded g b] using bfin b'b b by(simp only: finfun_default_update_const map_default_def) next case False note b'b = this let ?g' = "?g(a' \ b')" let ?b' = "map_default b ?g'" let ?b = "map_default b ?g" from fing have fing': "finite (dom ?g')" by auto from bran b'b have bnrang': "b \ ran ?g'" by(auto simp add: ran_def) have ffmg': "map_default b ?g' = y(a' := b')" by(auto simp add: map_default_def restrict_map_def) with f y have f_Abs: "f(a' $:= b') = Abs_finfun (map_default b ?g')" by(auto simp add: finfun_update_def) have g': "The (?the (f(a' $:= b'))) = ?g'" proof (rule the_equality) from fing' bnrang' f_Abs show "?the (f(a' $:= b')) ?g'" by(auto simp add: finfun_update_def restrict_map_def) next fix g' assume "?the (f(a' $:= b')) g'" hence f': "f(a' $:= b') = Abs_finfun (map_default b g')" and fin': "finite (dom g')" and brang': "b \ ran g'" by auto from fing' fin' have "map_default b ?g' \ finfun" "map_default b g' \ finfun" by(auto intro: map_default_in_finfun) with f' f_Abs have "map_default b g' = map_default b ?g'" by simp with fin' brang' fing' bnrang' show "g' = ?g'" by(rule map_default_inject[OF disjI2[OF refl]]) qed have dom: "dom (((\a. Some (y a)) |` {a. y a \ b})(a' \ b')) = insert a' (dom ((\a. Some (y a)) |` {a. y a \ b}))" by auto show ?thesis proof(cases "y a' = b") case True hence a'ndomg: "a' \ dom ?g" by auto from f y b'b True have yff: "y = map_default b (?g' |` dom ?g)" by(auto simp add: restrict_map_def map_default_def intro!: ext) hence f': "f = Abs_finfun (map_default b (?g' |` dom ?g))" using f by simp interpret g'wf: comp_fun_commute "\a. upd a (?b' a)" by(rule upd_left_comm) from upd_left_comm upd_left_comm fing have "Finite_Set.fold (\a. upd a (?b a)) (cnst b) (dom ?g) = Finite_Set.fold (\a. upd a (?b' a)) (cnst b) (dom ?g)" by(rule finite_rec_cong1)(auto simp add: restrict_map_def b'b True map_default_def) thus ?thesis unfolding finfun_rec_def Let_def finfun_default_update_const b[symmetric] unfolding g' g[symmetric] gg g'wf.fold_insert[OF fing a'ndomg, of "cnst b", folded dom] by -(rule arg_cong2[where f="upd a'"], simp_all add: map_default_def) next case False hence "insert a' (dom ?g) = dom ?g" by auto moreover { let ?g'' = "?g(a' := None)" let ?b'' = "map_default b ?g''" from False have domg: "dom ?g = insert a' (dom ?g'')" by auto from False have a'ndomg'': "a' \ dom ?g''" by auto have fing'': "finite (dom ?g'')" by(rule finite_subset[OF _ fing]) auto have bnrang'': "b \ ran ?g''" by(auto simp add: ran_def restrict_map_def) interpret gwf: comp_fun_commute "\a. upd a (?b a)" by(rule upd_left_comm) interpret g'wf: comp_fun_commute "\a. upd a (?b' a)" by(rule upd_left_comm) have "upd a' b' (Finite_Set.fold (\a. upd a (?b a)) (cnst b) (insert a' (dom ?g''))) = upd a' b' (upd a' (?b a') (Finite_Set.fold (\a. upd a (?b a)) (cnst b) (dom ?g'')))" unfolding gwf.fold_insert[OF fing'' a'ndomg''] f .. also have g''leg: "?g |` dom ?g'' \\<^sub>m ?g" by(auto simp add: map_le_def) have "dom (?g |` dom ?g'') = dom ?g''" by auto note map_default_update_twice[where d=b and f = "?g |` dom ?g''" and a=a' and d'="?b a'" and d''=b' and g="?g", unfolded this, OF fing'' a'ndomg'' g''leg] also have b': "b' = ?b' a'" by(auto simp add: map_default_def) from upd_left_comm upd_left_comm fing'' have "Finite_Set.fold (\a. upd a (?b a)) (cnst b) (dom ?g'') = Finite_Set.fold (\a. upd a (?b' a)) (cnst b) (dom ?g'')" by(rule finite_rec_cong1)(auto simp add: restrict_map_def b'b map_default_def) with b' have "upd a' b' (Finite_Set.fold (\a. upd a (?b a)) (cnst b) (dom ?g'')) = upd a' (?b' a') (Finite_Set.fold (\a. upd a (?b' a)) (cnst b) (dom ?g''))" by simp also note g'wf.fold_insert[OF fing'' a'ndomg'', symmetric] finally have "upd a' b' (Finite_Set.fold (\a. upd a (?b a)) (cnst b) (dom ?g)) = Finite_Set.fold (\a. upd a (?b' a)) (cnst b) (dom ?g)" unfolding domg . } ultimately have "Finite_Set.fold (\a. upd a (?b' a)) (cnst b) (insert a' (dom ?g)) = upd a' b' (Finite_Set.fold (\a. upd a (?b a)) (cnst b) (dom ?g))" by simp thus ?thesis unfolding finfun_rec_def Let_def finfun_default_update_const b[symmetric] g[symmetric] g' dom[symmetric] using b'b gg by(simp add: map_default_insert) qed qed qed end locale finfun_rec_wf = finfun_rec_wf_aux + assumes const_update_all: "finite (UNIV :: 'a set) \ Finite_Set.fold (\a. upd a b') (cnst b) (UNIV :: 'a set) = cnst b'" begin lemma finfun_rec_const [simp]: "finfun_rec cnst upd (K$ c) = cnst c" including finfun proof(cases "finite (UNIV :: 'a set)") case False hence "finfun_default ((K$ c) :: 'a \f 'b) = c" by(simp add: finfun_default_const) moreover have "(THE g :: 'a \ 'b. (K$ c) = Abs_finfun (map_default c g) \ finite (dom g) \ c \ ran g) = Map.empty" proof (rule the_equality) show "(K$ c) = Abs_finfun (map_default c Map.empty) \ finite (dom Map.empty) \ c \ ran Map.empty" by(auto simp add: finfun_const_def) next fix g :: "'a \ 'b" assume "(K$ c) = Abs_finfun (map_default c g) \ finite (dom g) \ c \ ran g" hence g: "(K$ c) = Abs_finfun (map_default c g)" and fin: "finite (dom g)" and ran: "c \ ran g" by blast+ from g map_default_in_finfun[OF fin, of c] have "map_default c g = (\a. c)" by(simp add: finfun_const_def) moreover have "map_default c Map.empty = (\a. c)" by simp ultimately show "g = Map.empty" by-(rule map_default_inject[OF disjI2[OF refl] fin ran], auto) qed ultimately show ?thesis by(simp add: finfun_rec_def) next case True hence default: "finfun_default ((K$ c) :: 'a \f 'b) = undefined" by(simp add: finfun_default_const) let ?the = "\g :: 'a \ 'b. (K$ c) = Abs_finfun (map_default undefined g) \ finite (dom g) \ undefined \ ran g" show ?thesis proof(cases "c = undefined") case True have the: "The ?the = Map.empty" proof (rule the_equality) from True show "?the Map.empty" by(auto simp add: finfun_const_def) next fix g' assume "?the g'" hence fg: "(K$ c) = Abs_finfun (map_default undefined g')" and fin: "finite (dom g')" and g: "undefined \ ran g'" by simp_all from fin have "map_default undefined g' \ finfun" by(rule map_default_in_finfun) with fg have "map_default undefined g' = (\a. c)" by(auto simp add: finfun_const_def intro: Abs_finfun_inject[THEN iffD1, symmetric]) with True show "g' = Map.empty" by -(rule map_default_inject(2)[OF _ fin g], auto) qed show ?thesis unfolding finfun_rec_def using \finite UNIV\ True unfolding Let_def the default by(simp) next case False have the: "The ?the = (\a :: 'a. Some c)" proof (rule the_equality) from False True show "?the (\a :: 'a. Some c)" by(auto simp add: map_default_def [abs_def] finfun_const_def dom_def ran_def) next fix g' :: "'a \ 'b" assume "?the g'" hence fg: "(K$ c) = Abs_finfun (map_default undefined g')" and fin: "finite (dom g')" and g: "undefined \ ran g'" by simp_all from fin have "map_default undefined g' \ finfun" by(rule map_default_in_finfun) with fg have "map_default undefined g' = (\a. c)" by(auto simp add: finfun_const_def intro: Abs_finfun_inject[THEN iffD1]) with True False show "g' = (\a::'a. Some c)" by - (rule map_default_inject(2)[OF _ fin g], auto simp add: dom_def ran_def map_default_def [abs_def]) qed show ?thesis unfolding finfun_rec_def using True False unfolding Let_def the default by(simp add: dom_def map_default_def const_update_all) qed qed end subsection \Weak induction rule and case analysis for FinFuns\ lemma finfun_weak_induct [consumes 0, case_names const update]: assumes const: "\b. P (K$ b)" and update: "\f a b. P f \ P (f(a $:= b))" shows "P x" including finfun proof(induct x rule: Abs_finfun_induct) case (Abs_finfun y) then obtain b where "finite {a. y a \ b}" unfolding finfun_def by blast thus ?case using \y \ finfun\ proof(induct "{a. y a \ b}" arbitrary: y rule: finite_induct) case empty hence "\a. y a = b" by blast hence "y = (\a. b)" by(auto) hence "Abs_finfun y = finfun_const b" unfolding finfun_const_def by simp thus ?case by(simp add: const) next case (insert a A) note IH = \\y. \ A = {a. y a \ b}; y \ finfun \ \ P (Abs_finfun y)\ note y = \y \ finfun\ with \insert a A = {a. y a \ b}\ \a \ A\ have "A = {a'. (y(a := b)) a' \ b}" "y(a := b) \ finfun" by auto from IH[OF this] have "P (finfun_update (Abs_finfun (y(a := b))) a (y a))" by(rule update) thus ?case using y unfolding finfun_update_def by simp qed qed lemma finfun_exhaust_disj: "(\b. x = finfun_const b) \ (\f a b. x = finfun_update f a b)" by(induct x rule: finfun_weak_induct) blast+ lemma finfun_exhaust: obtains b where "x = (K$ b)" | f a b where "x = f(a $:= b)" by(atomize_elim)(rule finfun_exhaust_disj) lemma finfun_rec_unique: fixes f :: "'a \f 'b \ 'c" assumes c: "\c. f (K$ c) = cnst c" and u: "\g a b. f (g(a $:= b)) = upd g a b (f g)" and c': "\c. f' (K$ c) = cnst c" and u': "\g a b. f' (g(a $:= b)) = upd g a b (f' g)" shows "f = f'" proof fix g :: "'a \f 'b" show "f g = f' g" by(induct g rule: finfun_weak_induct)(auto simp add: c u c' u') qed subsection \Function application\ notation finfun_apply (infixl "$" 999) interpretation finfun_apply_aux: finfun_rec_wf_aux "\b. b" "\a' b c. if (a = a') then b else c" by(unfold_locales) auto interpretation finfun_apply: finfun_rec_wf "\b. b" "\a' b c. if (a = a') then b else c" proof(unfold_locales) fix b' b :: 'a assume fin: "finite (UNIV :: 'b set)" { fix A :: "'b set" interpret comp_fun_commute "\a'. If (a = a') b'" by(rule finfun_apply_aux.upd_left_comm) from fin have "finite A" by(auto intro: finite_subset) hence "Finite_Set.fold (\a'. If (a = a') b') b A = (if a \ A then b' else b)" by induct auto } from this[of UNIV] show "Finite_Set.fold (\a'. If (a = a') b') b UNIV = b'" by simp qed lemma finfun_apply_def: "($) = (\f a. finfun_rec (\b. b) (\a' b c. if (a = a') then b else c) f)" proof(rule finfun_rec_unique) fix c show "($) (K$ c) = (\a. c)" by(simp add: finfun_const.rep_eq) next fix g a b show "($) g(a $:= b) = (\c. if c = a then b else g $ c)" by(auto simp add: finfun_update_def fun_upd_finfun Abs_finfun_inverse finfun_apply) qed auto lemma finfun_upd_apply: "f(a $:= b) $ a' = (if a = a' then b else f $ a')" and finfun_upd_apply_code [code]: "(finfun_update_code f a b) $ a' = (if a = a' then b else f $ a')" by(simp_all add: finfun_apply_def) lemma finfun_const_apply [simp, code]: "(K$ b) $ a = b" by(simp add: finfun_apply_def) lemma finfun_upd_apply_same [simp]: "f(a $:= b) $ a = b" by(simp add: finfun_upd_apply) lemma finfun_upd_apply_other [simp]: "a \ a' \ f(a $:= b) $ a' = f $ a'" by(simp add: finfun_upd_apply) lemma finfun_ext: "(\a. f $ a = g $ a) \ f = g" by(auto simp add: finfun_apply_inject[symmetric]) lemma expand_finfun_eq: "(f = g) = (($) f = ($) g)" by(auto intro: finfun_ext) lemma finfun_upd_triv [simp]: "f(x $:= f $ x) = f" by(simp add: expand_finfun_eq fun_eq_iff finfun_upd_apply) lemma finfun_const_inject [simp]: "(K$ b) = (K$ b') \ b = b'" by(simp add: expand_finfun_eq fun_eq_iff) lemma finfun_const_eq_update: "((K$ b) = f(a $:= b')) = (b = b' \ (\a'. a \ a' \ f $ a' = b))" by(auto simp add: expand_finfun_eq fun_eq_iff finfun_upd_apply) subsection \Function composition\ definition finfun_comp :: "('a \ 'b) \ 'c \f 'a \ 'c \f 'b" (infixr "\$" 55) where [code del]: "g \$ f = finfun_rec (\b. (K$ g b)) (\a b c. c(a $:= g b)) f" notation (ASCII) finfun_comp (infixr "o$" 55) interpretation finfun_comp_aux: finfun_rec_wf_aux "(\b. (K$ g b))" "(\a b c. c(a $:= g b))" by(unfold_locales)(auto simp add: finfun_upd_apply intro: finfun_ext) interpretation finfun_comp: finfun_rec_wf "(\b. (K$ g b))" "(\a b c. c(a $:= g b))" proof fix b' b :: 'a assume fin: "finite (UNIV :: 'c set)" { fix A :: "'c set" from fin have "finite A" by(auto intro: finite_subset) hence "Finite_Set.fold (\(a :: 'c) c. c(a $:= g b')) (K$ g b) A = Abs_finfun (\a. if a \ A then g b' else g b)" by induct (simp_all add: finfun_const_def, auto simp add: finfun_update_def Abs_finfun_inverse_finite fun_upd_def Abs_finfun_inject_finite fun_eq_iff fin) } from this[of UNIV] show "Finite_Set.fold (\(a :: 'c) c. c(a $:= g b')) (K$ g b) UNIV = (K$ g b')" by(simp add: finfun_const_def) qed lemma finfun_comp_const [simp, code]: "g \$ (K$ c) = (K$ g c)" by(simp add: finfun_comp_def) lemma finfun_comp_update [simp]: "g \$ (f(a $:= b)) = (g \$ f)(a $:= g b)" and finfun_comp_update_code [code]: "g \$ (finfun_update_code f a b) = finfun_update_code (g \$ f) a (g b)" by(simp_all add: finfun_comp_def) lemma finfun_comp_apply [simp]: "($) (g \$ f) = g \ ($) f" by(induct f rule: finfun_weak_induct)(auto simp add: finfun_upd_apply) lemma finfun_comp_comp_collapse [simp]: "f \$ g \$ h = (f \ g) \$ h" by(induct h rule: finfun_weak_induct) simp_all lemma finfun_comp_const1 [simp]: "(\x. c) \$ f = (K$ c)" by(induct f rule: finfun_weak_induct)(auto intro: finfun_ext simp add: finfun_upd_apply) lemma finfun_comp_id1 [simp]: "(\x. x) \$ f = f" "id \$ f = f" by(induct f rule: finfun_weak_induct) auto lemma finfun_comp_conv_comp: "g \$ f = Abs_finfun (g \ ($) f)" including finfun proof - have "(\f. g \$ f) = (\f. Abs_finfun (g \ ($) f))" proof(rule finfun_rec_unique) { fix c show "Abs_finfun (g \ ($) (K$ c)) = (K$ g c)" by(simp add: finfun_comp_def o_def)(simp add: finfun_const_def) } { fix g' a b show "Abs_finfun (g \ ($) g'(a $:= b)) = (Abs_finfun (g \ ($) g'))(a $:= g b)" proof - obtain y where y: "y \ finfun" and g': "g' = Abs_finfun y" by(cases g') moreover from g' have "(g \ ($) g') \ finfun" by(simp add: finfun_left_compose) moreover have "g \ y(a := b) = (g \ y)(a := g b)" by(auto) ultimately show ?thesis by(simp add: finfun_comp_def finfun_update_def) qed } qed auto thus ?thesis by(auto simp add: fun_eq_iff) qed definition finfun_comp2 :: "'b \f 'c \ ('a \ 'b) \ 'a \f 'c" (infixr "$\" 55) where [code del]: "g $\ f = Abs_finfun (($) g \ f)" notation (ASCII) finfun_comp2 (infixr "$o" 55) lemma finfun_comp2_const [code, simp]: "finfun_comp2 (K$ c) f = (K$ c)" including finfun by(simp add: finfun_comp2_def finfun_const_def comp_def) lemma finfun_comp2_update: assumes inj: "inj f" shows "finfun_comp2 (g(b $:= c)) f = (if b \ range f then (finfun_comp2 g f)(inv f b $:= c) else finfun_comp2 g f)" including finfun proof(cases "b \ range f") case True from inj have "\x. (($) g)(f x := c) \ f = (($) g \ f)(x := c)" by(auto intro!: ext dest: injD) with inj True show ?thesis by(auto simp add: finfun_comp2_def finfun_update_def finfun_right_compose) next case False hence "(($) g)(b := c) \ f = ($) g \ f" by(auto simp add: fun_eq_iff) with False show ?thesis by(auto simp add: finfun_comp2_def finfun_update_def) qed subsection \Universal quantification\ definition finfun_All_except :: "'a list \ 'a \f bool \ bool" where [code del]: "finfun_All_except A P \ \a. a \ set A \ P $ a" lemma finfun_All_except_const: "finfun_All_except A (K$ b) \ b \ set A = UNIV" by(auto simp add: finfun_All_except_def) lemma finfun_All_except_const_finfun_UNIV_code [code]: "finfun_All_except A (K$ b) = (b \ is_list_UNIV A)" by(simp add: finfun_All_except_const is_list_UNIV_iff) lemma finfun_All_except_update: "finfun_All_except A f(a $:= b) = ((a \ set A \ b) \ finfun_All_except (a # A) f)" by(fastforce simp add: finfun_All_except_def finfun_upd_apply) lemma finfun_All_except_update_code [code]: fixes a :: "'a :: card_UNIV" shows "finfun_All_except A (finfun_update_code f a b) = ((a \ set A \ b) \ finfun_All_except (a # A) f)" by(simp add: finfun_All_except_update) definition finfun_All :: "'a \f bool \ bool" where "finfun_All = finfun_All_except []" lemma finfun_All_const [simp]: "finfun_All (K$ b) = b" by(simp add: finfun_All_def finfun_All_except_def) lemma finfun_All_update: "finfun_All f(a $:= b) = (b \ finfun_All_except [a] f)" by(simp add: finfun_All_def finfun_All_except_update) lemma finfun_All_All: "finfun_All P = All (($) P)" by(simp add: finfun_All_def finfun_All_except_def) definition finfun_Ex :: "'a \f bool \ bool" where "finfun_Ex P = Not (finfun_All (Not \$ P))" lemma finfun_Ex_Ex: "finfun_Ex P = Ex (($) P)" unfolding finfun_Ex_def finfun_All_All by simp lemma finfun_Ex_const [simp]: "finfun_Ex (K$ b) = b" by(simp add: finfun_Ex_def) subsection \A diagonal operator for FinFuns\ definition finfun_Diag :: "'a \f 'b \ 'a \f 'c \ 'a \f ('b \ 'c)" ("(1'($_,/ _$'))" [0, 0] 1000) where [code del]: "($f, g$) = finfun_rec (\b. Pair b \$ g) (\a b c. c(a $:= (b, g $ a))) f" interpretation finfun_Diag_aux: finfun_rec_wf_aux "\b. Pair b \$ g" "\a b c. c(a $:= (b, g $ a))" by(unfold_locales)(simp_all add: expand_finfun_eq fun_eq_iff finfun_upd_apply) interpretation finfun_Diag: finfun_rec_wf "\b. Pair b \$ g" "\a b c. c(a $:= (b, g $ a))" proof fix b' b :: 'a assume fin: "finite (UNIV :: 'c set)" { fix A :: "'c set" interpret comp_fun_commute "\a c. c(a $:= (b', g $ a))" by(rule finfun_Diag_aux.upd_left_comm) from fin have "finite A" by(auto intro: finite_subset) hence "Finite_Set.fold (\a c. c(a $:= (b', g $ a))) (Pair b \$ g) A = Abs_finfun (\a. (if a \ A then b' else b, g $ a))" by(induct)(simp_all add: finfun_const_def finfun_comp_conv_comp o_def, auto simp add: finfun_update_def Abs_finfun_inverse_finite fun_upd_def Abs_finfun_inject_finite fun_eq_iff fin) } from this[of UNIV] show "Finite_Set.fold (\a c. c(a $:= (b', g $ a))) (Pair b \$ g) UNIV = Pair b' \$ g" by(simp add: finfun_const_def finfun_comp_conv_comp o_def) qed lemma finfun_Diag_const1: "($K$ b, g$) = Pair b \$ g" by(simp add: finfun_Diag_def) text \ Do not use @{thm finfun_Diag_const1} for the code generator because @{term "Pair b"} is injective, i.e. if @{term g} is free of redundant updates, there is no need to check for redundant updates as is done for @{term "(\$)"}. \ lemma finfun_Diag_const_code [code]: "($K$ b, K$ c$) = (K$ (b, c))" "($K$ b, finfun_update_code g a c$) = finfun_update_code ($K$ b, g$) a (b, c)" by(simp_all add: finfun_Diag_const1) lemma finfun_Diag_update1: "($f(a $:= b), g$) = ($f, g$)(a $:= (b, g $ a))" and finfun_Diag_update1_code [code]: "($finfun_update_code f a b, g$) = ($f, g$)(a $:= (b, g $ a))" by(simp_all add: finfun_Diag_def) lemma finfun_Diag_const2: "($f, K$ c$) = (\b. (b, c)) \$ f" by(induct f rule: finfun_weak_induct)(auto intro!: finfun_ext simp add: finfun_upd_apply finfun_Diag_const1 finfun_Diag_update1) lemma finfun_Diag_update2: "($f, g(a $:= c)$) = ($f, g$)(a $:= (f $ a, c))" by(induct f rule: finfun_weak_induct)(auto intro!: finfun_ext simp add: finfun_upd_apply finfun_Diag_const1 finfun_Diag_update1) lemma finfun_Diag_const_const [simp]: "($K$ b, K$ c$) = (K$ (b, c))" by(simp add: finfun_Diag_const1) lemma finfun_Diag_const_update: "($K$ b, g(a $:= c)$) = ($K$ b, g$)(a $:= (b, c))" by(simp add: finfun_Diag_const1) lemma finfun_Diag_update_const: "($f(a $:= b), K$ c$) = ($f, K$ c$)(a $:= (b, c))" by(simp add: finfun_Diag_def) lemma finfun_Diag_update_update: "($f(a $:= b), g(a' $:= c)$) = (if a = a' then ($f, g$)(a $:= (b, c)) else ($f, g$)(a $:= (b, g $ a))(a' $:= (f $ a', c)))" by(auto simp add: finfun_Diag_update1 finfun_Diag_update2) lemma finfun_Diag_apply [simp]: "($) ($f, g$) = (\x. (f $ x, g $ x))" by(induct f rule: finfun_weak_induct)(auto simp add: finfun_Diag_const1 finfun_Diag_update1 finfun_upd_apply) lemma finfun_Diag_conv_Abs_finfun: "($f, g$) = Abs_finfun ((\x. (f $ x, g $ x)))" including finfun proof - have "(\f :: 'a \f 'b. ($f, g$)) = (\f. Abs_finfun ((\x. (f $ x, g $ x))))" proof(rule finfun_rec_unique) { fix c show "Abs_finfun (\x. ((K$ c) $ x, g $ x)) = Pair c \$ g" by(simp add: finfun_comp_conv_comp o_def finfun_const_def) } { fix g' a b show "Abs_finfun (\x. (g'(a $:= b) $ x, g $ x)) = (Abs_finfun (\x. (g' $ x, g $ x)))(a $:= (b, g $ a))" by(auto simp add: finfun_update_def fun_eq_iff simp del: fun_upd_apply) simp } qed(simp_all add: finfun_Diag_const1 finfun_Diag_update1) thus ?thesis by(auto simp add: fun_eq_iff) qed lemma finfun_Diag_eq: "($f, g$) = ($f', g'$) \ f = f' \ g = g'" by(auto simp add: expand_finfun_eq fun_eq_iff) definition finfun_fst :: "'a \f ('b \ 'c) \ 'a \f 'b" where [code]: "finfun_fst f = fst \$ f" lemma finfun_fst_const: "finfun_fst (K$ bc) = (K$ fst bc)" by(simp add: finfun_fst_def) lemma finfun_fst_update: "finfun_fst (f(a $:= bc)) = (finfun_fst f)(a $:= fst bc)" and finfun_fst_update_code: "finfun_fst (finfun_update_code f a bc) = (finfun_fst f)(a $:= fst bc)" by(simp_all add: finfun_fst_def) lemma finfun_fst_comp_conv: "finfun_fst (f \$ g) = (fst \ f) \$ g" by(simp add: finfun_fst_def) lemma finfun_fst_conv [simp]: "finfun_fst ($f, g$) = f" by(induct f rule: finfun_weak_induct)(simp_all add: finfun_Diag_const1 finfun_fst_comp_conv o_def finfun_Diag_update1 finfun_fst_update) lemma finfun_fst_conv_Abs_finfun: "finfun_fst = (\f. Abs_finfun (fst \ ($) f))" by(simp add: finfun_fst_def [abs_def] finfun_comp_conv_comp) definition finfun_snd :: "'a \f ('b \ 'c) \ 'a \f 'c" where [code]: "finfun_snd f = snd \$ f" lemma finfun_snd_const: "finfun_snd (K$ bc) = (K$ snd bc)" by(simp add: finfun_snd_def) lemma finfun_snd_update: "finfun_snd (f(a $:= bc)) = (finfun_snd f)(a $:= snd bc)" and finfun_snd_update_code [code]: "finfun_snd (finfun_update_code f a bc) = (finfun_snd f)(a $:= snd bc)" by(simp_all add: finfun_snd_def) lemma finfun_snd_comp_conv: "finfun_snd (f \$ g) = (snd \ f) \$ g" by(simp add: finfun_snd_def) lemma finfun_snd_conv [simp]: "finfun_snd ($f, g$) = g" apply(induct f rule: finfun_weak_induct) apply(auto simp add: finfun_Diag_const1 finfun_snd_comp_conv o_def finfun_Diag_update1 finfun_snd_update finfun_upd_apply intro: finfun_ext) done lemma finfun_snd_conv_Abs_finfun: "finfun_snd = (\f. Abs_finfun (snd \ ($) f))" by(simp add: finfun_snd_def [abs_def] finfun_comp_conv_comp) lemma finfun_Diag_collapse [simp]: "($finfun_fst f, finfun_snd f$) = f" by(induct f rule: finfun_weak_induct)(simp_all add: finfun_fst_const finfun_snd_const finfun_fst_update finfun_snd_update finfun_Diag_update_update) subsection \Currying for FinFuns\ definition finfun_curry :: "('a \ 'b) \f 'c \ 'a \f 'b \f 'c" where [code del]: "finfun_curry = finfun_rec (finfun_const \ finfun_const) (\(a, b) c f. f(a $:= (f $ a)(b $:= c)))" interpretation finfun_curry_aux: finfun_rec_wf_aux "finfun_const \ finfun_const" "\(a, b) c f. f(a $:= (f $ a)(b $:= c))" apply(unfold_locales) apply(auto simp add: split_def finfun_update_twist finfun_upd_apply split_paired_all finfun_update_const_same) done interpretation finfun_curry: finfun_rec_wf "finfun_const \ finfun_const" "\(a, b) c f. f(a $:= (f $ a)(b $:= c))" proof(unfold_locales) fix b' b :: 'b assume fin: "finite (UNIV :: ('c \ 'a) set)" hence fin1: "finite (UNIV :: 'c set)" and fin2: "finite (UNIV :: 'a set)" unfolding UNIV_Times_UNIV[symmetric] by(fastforce dest: finite_cartesian_productD1 finite_cartesian_productD2)+ note [simp] = Abs_finfun_inverse_finite[OF fin] Abs_finfun_inverse_finite[OF fin1] Abs_finfun_inverse_finite[OF fin2] { fix A :: "('c \ 'a) set" interpret comp_fun_commute "\a :: 'c \ 'a. (\(a, b) c f. f(a $:= (f $ a)(b $:= c))) a b'" by(rule finfun_curry_aux.upd_left_comm) from fin have "finite A" by(auto intro: finite_subset) hence "Finite_Set.fold (\a :: 'c \ 'a. (\(a, b) c f. f(a $:= (f $ a)(b $:= c))) a b') ((finfun_const \ finfun_const) b) A = Abs_finfun (\a. Abs_finfun (\b''. if (a, b'') \ A then b' else b))" by induct (simp_all, auto simp add: finfun_update_def finfun_const_def split_def intro!: arg_cong[where f="Abs_finfun"] ext) } from this[of UNIV] show "Finite_Set.fold (\a :: 'c \ 'a. (\(a, b) c f. f(a $:= (f $ a)(b $:= c))) a b') ((finfun_const \ finfun_const) b) UNIV = (finfun_const \ finfun_const) b'" by(simp add: finfun_const_def) qed lemma finfun_curry_const [simp, code]: "finfun_curry (K$ c) = (K$ K$ c)" by(simp add: finfun_curry_def) lemma finfun_curry_update [simp]: "finfun_curry (f((a, b) $:= c)) = (finfun_curry f)(a $:= (finfun_curry f $ a)(b $:= c))" and finfun_curry_update_code [code]: "finfun_curry (finfun_update_code f (a, b) c) = (finfun_curry f)(a $:= (finfun_curry f $ a)(b $:= c))" by(simp_all add: finfun_curry_def) lemma finfun_Abs_finfun_curry: assumes fin: "f \ finfun" shows "(\a. Abs_finfun (curry f a)) \ finfun" including finfun proof - from fin obtain c where c: "finite {ab. f ab \ c}" unfolding finfun_def by blast have "{a. \b. f (a, b) \ c} = fst ` {ab. f ab \ c}" by(force) hence "{a. curry f a \ (\x. c)} = fst ` {ab. f ab \ c}" by(auto simp add: curry_def fun_eq_iff) with fin c have "finite {a. Abs_finfun (curry f a) \ (K$ c)}" by(simp add: finfun_const_def finfun_curry) thus ?thesis unfolding finfun_def by auto qed lemma finfun_curry_conv_curry: fixes f :: "('a \ 'b) \f 'c" shows "finfun_curry f = Abs_finfun (\a. Abs_finfun (curry (finfun_apply f) a))" including finfun proof - have "finfun_curry = (\f :: ('a \ 'b) \f 'c. Abs_finfun (\a. Abs_finfun (curry (finfun_apply f) a)))" proof(rule finfun_rec_unique) fix c show "finfun_curry (K$ c) = (K$ K$ c)" by simp fix f a show "finfun_curry (f(a $:= c)) = (finfun_curry f)(fst a $:= (finfun_curry f $ (fst a))(snd a $:= c))" by(cases a) simp show "Abs_finfun (\a. Abs_finfun (curry (finfun_apply (K$ c)) a)) = (K$ K$ c)" by(simp add: finfun_curry_def finfun_const_def curry_def) fix g b show "Abs_finfun (\aa. Abs_finfun (curry (($) g(a $:= b)) aa)) = (Abs_finfun (\a. Abs_finfun (curry (($) g) a)))( fst a $:= ((Abs_finfun (\a. Abs_finfun (curry (($) g) a))) $ (fst a))(snd a $:= b))" by(cases a)(auto intro!: ext arg_cong[where f=Abs_finfun] simp add: finfun_curry_def finfun_update_def finfun_Abs_finfun_curry) qed thus ?thesis by(auto simp add: fun_eq_iff) qed subsection \Executable equality for FinFuns\ lemma eq_finfun_All_ext: "(f = g) \ finfun_All ((\(x, y). x = y) \$ ($f, g$))" by(simp add: expand_finfun_eq fun_eq_iff finfun_All_All o_def) instantiation finfun :: ("{card_UNIV,equal}",equal) equal begin definition eq_finfun_def [code]: "HOL.equal f g \ finfun_All ((\(x, y). x = y) \$ ($f, g$))" instance by(intro_classes)(simp add: eq_finfun_All_ext eq_finfun_def) end lemma [code nbe]: "HOL.equal (f :: _ \f _) f \ True" by (fact equal_refl) subsection \An operator that explicitly removes all redundant updates in the generated representations\ definition finfun_clearjunk :: "'a \f 'b \ 'a \f 'b" where [simp, code del]: "finfun_clearjunk = id" lemma finfun_clearjunk_const [code]: "finfun_clearjunk (K$ b) = (K$ b)" by simp lemma finfun_clearjunk_update [code]: "finfun_clearjunk (finfun_update_code f a b) = f(a $:= b)" by simp subsection \The domain of a FinFun as a FinFun\ definition finfun_dom :: "('a \f 'b) \ ('a \f bool)" where [code del]: "finfun_dom f = Abs_finfun (\a. f $ a \ finfun_default f)" lemma finfun_dom_const: "finfun_dom ((K$ c) :: 'a \f 'b) = (K$ finite (UNIV :: 'a set) \ c \ undefined)" unfolding finfun_dom_def finfun_default_const by(auto)(simp_all add: finfun_const_def) text \ @{term "finfun_dom" } raises an exception when called on a FinFun whose domain is a finite type. For such FinFuns, the default value (and as such the domain) is undefined. \ lemma finfun_dom_const_code [code]: "finfun_dom ((K$ c) :: ('a :: card_UNIV) \f 'b) = (if CARD('a) = 0 then (K$ False) else Code.abort (STR ''finfun_dom called on finite type'') (\_. finfun_dom (K$ c)))" by(simp add: finfun_dom_const card_UNIV card_eq_0_iff) lemma finfun_dom_finfunI: "(\a. f $ a \ finfun_default f) \ finfun" using finite_finfun_default[of f] by(simp add: finfun_def exI[where x=False]) lemma finfun_dom_update [simp]: "finfun_dom (f(a $:= b)) = (finfun_dom f)(a $:= (b \ finfun_default f))" including finfun unfolding finfun_dom_def finfun_update_def apply(simp add: finfun_default_update_const finfun_dom_finfunI) apply(fold finfun_update.rep_eq) apply(simp add: finfun_upd_apply fun_eq_iff fun_upd_def finfun_default_update_const) done lemma finfun_dom_update_code [code]: "finfun_dom (finfun_update_code f a b) = finfun_update_code (finfun_dom f) a (b \ finfun_default f)" by(simp) lemma finite_finfun_dom: "finite {x. finfun_dom f $ x}" proof(induct f rule: finfun_weak_induct) case (const b) thus ?case by (cases "finite (UNIV :: 'a set) \ b \ undefined") (auto simp add: finfun_dom_const UNIV_def [symmetric] Set.empty_def [symmetric]) next case (update f a b) have "{x. finfun_dom f(a $:= b) $ x} = (if b = finfun_default f then {x. finfun_dom f $ x} - {a} else insert a {x. finfun_dom f $ x})" by (auto simp add: finfun_upd_apply split: if_split_asm) thus ?case using update by simp qed subsection \The domain of a FinFun as a sorted list\ definition finfun_to_list :: "('a :: linorder) \f 'b \ 'a list" where "finfun_to_list f = (THE xs. set xs = {x. finfun_dom f $ x} \ sorted xs \ distinct xs)" lemma set_finfun_to_list [simp]: "set (finfun_to_list f) = {x. finfun_dom f $ x}" (is ?thesis1) and sorted_finfun_to_list: "sorted (finfun_to_list f)" (is ?thesis2) and distinct_finfun_to_list: "distinct (finfun_to_list f)" (is ?thesis3) proof (atomize (full)) show "?thesis1 \ ?thesis2 \ ?thesis3" unfolding finfun_to_list_def by(rule theI')(rule finite_sorted_distinct_unique finite_finfun_dom)+ qed lemma finfun_const_False_conv_bot: "($) (K$ False) = bot" by auto lemma finfun_const_True_conv_top: "($) (K$ True) = top" by auto lemma finfun_to_list_const: "finfun_to_list ((K$ c) :: ('a :: {linorder} \f 'b)) = (if \ finite (UNIV :: 'a set) \ c = undefined then [] else THE xs. set xs = UNIV \ sorted xs \ distinct xs)" by(auto simp add: finfun_to_list_def finfun_const_False_conv_bot finfun_const_True_conv_top finfun_dom_const) lemma finfun_to_list_const_code [code]: "finfun_to_list ((K$ c) :: ('a :: {linorder, card_UNIV} \f 'b)) = (if CARD('a) = 0 then [] else Code.abort (STR ''finfun_to_list called on finite type'') (\_. finfun_to_list ((K$ c) :: ('a \f 'b))))" by(auto simp add: finfun_to_list_const card_UNIV card_eq_0_iff) lemma remove1_insort_insert_same: "x \ set xs \ remove1 x (insort_insert x xs) = xs" -by (metis insort_insert_insort remove1_insort) +by (metis insort_insert_insort remove1_insort_key) lemma finfun_dom_conv: "finfun_dom f $ x \ f $ x \ finfun_default f" by(induct f rule: finfun_weak_induct)(auto simp add: finfun_dom_const finfun_default_const finfun_default_update_const finfun_upd_apply) lemma finfun_to_list_update: "finfun_to_list (f(a $:= b)) = (if b = finfun_default f then List.remove1 a (finfun_to_list f) else List.insort_insert a (finfun_to_list f))" proof(subst finfun_to_list_def, rule the_equality) fix xs assume "set xs = {x. finfun_dom f(a $:= b) $ x} \ sorted xs \ distinct xs" hence eq: "set xs = {x. finfun_dom f(a $:= b) $ x}" and [simp]: "sorted xs" "distinct xs" by simp_all show "xs = (if b = finfun_default f then remove1 a (finfun_to_list f) else insort_insert a (finfun_to_list f))" proof(cases "b = finfun_default f") case [simp]: True show ?thesis proof(cases "finfun_dom f $ a") case True have "finfun_to_list f = insort_insert a xs" unfolding finfun_to_list_def proof(rule the_equality) have "set (insort_insert a xs) = insert a (set xs)" by(simp add: set_insort_insert) also note eq also have "insert a {x. finfun_dom f(a $:= b) $ x} = {x. finfun_dom f $ x}" using True by(auto simp add: finfun_upd_apply split: if_split_asm) finally show 1: "set (insort_insert a xs) = {x. finfun_dom f $ x} \ sorted (insort_insert a xs) \ distinct (insort_insert a xs)" by(simp add: sorted_insort_insert distinct_insort_insert) fix xs' assume "set xs' = {x. finfun_dom f $ x} \ sorted xs' \ distinct xs'" thus "xs' = insort_insert a xs" using 1 by(auto dest: sorted_distinct_set_unique) qed with eq True show ?thesis by(simp add: remove1_insort_insert_same) next case False hence "f $ a = b" by(auto simp add: finfun_dom_conv) hence f: "f(a $:= b) = f" by(simp add: expand_finfun_eq fun_eq_iff finfun_upd_apply) from eq have "finfun_to_list f = xs" unfolding f finfun_to_list_def by(auto elim: sorted_distinct_set_unique intro!: the_equality) with eq False show ?thesis unfolding f by(simp add: remove1_idem) qed next case False show ?thesis proof(cases "finfun_dom f $ a") case True have "finfun_to_list f = xs" unfolding finfun_to_list_def proof(rule the_equality) have "finfun_dom f = finfun_dom f(a $:= b)" using False True by(simp add: expand_finfun_eq fun_eq_iff finfun_upd_apply) with eq show 1: "set xs = {x. finfun_dom f $ x} \ sorted xs \ distinct xs" by(simp del: finfun_dom_update) fix xs' assume "set xs' = {x. finfun_dom f $ x} \ sorted xs' \ distinct xs'" thus "xs' = xs" using 1 by(auto elim: sorted_distinct_set_unique) qed thus ?thesis using False True eq by(simp add: insort_insert_triv) next case False have "finfun_to_list f = remove1 a xs" unfolding finfun_to_list_def proof(rule the_equality) have "set (remove1 a xs) = set xs - {a}" by simp also note eq also have "{x. finfun_dom f(a $:= b) $ x} - {a} = {x. finfun_dom f $ x}" using False by(auto simp add: finfun_upd_apply split: if_split_asm) finally show 1: "set (remove1 a xs) = {x. finfun_dom f $ x} \ sorted (remove1 a xs) \ distinct (remove1 a xs)" by(simp add: sorted_remove1) fix xs' assume "set xs' = {x. finfun_dom f $ x} \ sorted xs' \ distinct xs'" thus "xs' = remove1 a xs" using 1 by(blast intro: sorted_distinct_set_unique) qed thus ?thesis using False eq \b \ finfun_default f\ by (simp add: insort_insert_insort insort_remove1) qed qed qed (auto simp add: distinct_finfun_to_list sorted_finfun_to_list sorted_remove1 set_insort_insert sorted_insort_insert distinct_insort_insert finfun_upd_apply split: if_split_asm) lemma finfun_to_list_update_code [code]: "finfun_to_list (finfun_update_code f a b) = (if b = finfun_default f then List.remove1 a (finfun_to_list f) else List.insort_insert a (finfun_to_list f))" by(simp add: finfun_to_list_update) text \More type class instantiations\ lemma card_eq_1_iff: "card A = 1 \ A \ {} \ (\x\A. \y\A. x = y)" (is "?lhs \ ?rhs") proof assume ?lhs moreover { fix x y assume A: "x \ A" "y \ A" and neq: "x \ y" have "finite A" using \?lhs\ by(simp add: card_ge_0_finite) from neq have "2 = card {x, y}" by simp also have "\ \ card A" using A \finite A\ by(auto intro: card_mono) finally have False using \?lhs\ by simp } ultimately show ?rhs by auto next assume ?rhs hence "A = {THE x. x \ A}" by safe (auto intro: theI the_equality[symmetric]) also have "card \ = 1" by simp finally show ?lhs . qed lemma card_UNIV_finfun: defines "F == finfun :: ('a \ 'b) set" shows "CARD('a \f 'b) = (if CARD('a) \ 0 \ CARD('b) \ 0 \ CARD('b) = 1 then CARD('b) ^ CARD('a) else 0)" proof(cases "0 < CARD('a) \ 0 < CARD('b) \ CARD('b) = 1") case True from True have "F = UNIV" proof assume b: "CARD('b) = 1" hence "\x :: 'b. x = undefined" by(auto simp add: card_eq_1_iff simp del: One_nat_def) thus ?thesis by(auto simp add: finfun_def F_def intro: exI[where x=undefined]) qed(auto simp add: finfun_def card_gt_0_iff F_def intro: finite_subset[where B=UNIV]) moreover have "CARD('a \f 'b) = card F" unfolding type_definition.Abs_image[OF type_definition_finfun, symmetric] by(auto intro!: card_image inj_onI simp add: Abs_finfun_inject F_def) ultimately show ?thesis by(simp add: card_fun) next case False hence infinite: "\ (finite (UNIV :: 'a set) \ finite (UNIV :: 'b set))" and b: "CARD('b) \ 1" by(simp_all add: card_eq_0_iff) from b obtain b1 b2 :: 'b where "b1 \ b2" by(auto simp add: card_eq_1_iff simp del: One_nat_def) let ?f = "\a a' :: 'a. if a = a' then b1 else b2" from infinite have "\ finite (UNIV :: ('a \f 'b) set)" proof(rule contrapos_nn[OF _ conjI]) assume finite: "finite (UNIV :: ('a \f 'b) set)" hence "finite F" unfolding type_definition.Abs_image[OF type_definition_finfun, symmetric] F_def by(rule finite_imageD)(auto intro: inj_onI simp add: Abs_finfun_inject) hence "finite (range ?f)" by(rule finite_subset[rotated 1])(auto simp add: F_def finfun_def \b1 \ b2\ intro!: exI[where x=b2]) thus "finite (UNIV :: 'a set)" by(rule finite_imageD)(auto intro: inj_onI simp add: fun_eq_iff \b1 \ b2\ split: if_split_asm) from finite have "finite (range (\b. ((K$ b) :: 'a \f 'b)))" by(rule finite_subset[rotated 1]) simp thus "finite (UNIV :: 'b set)" by(rule finite_imageD)(auto intro!: inj_onI) qed with False show ?thesis by auto qed lemma finite_UNIV_finfun: "finite (UNIV :: ('a \f 'b) set) \ (finite (UNIV :: 'a set) \ finite (UNIV :: 'b set) \ CARD('b) = 1)" (is "?lhs \ ?rhs") proof - have "?lhs \ CARD('a \f 'b) > 0" by(simp add: card_gt_0_iff) also have "\ \ CARD('a) > 0 \ CARD('b) > 0 \ CARD('b) = 1" by(simp add: card_UNIV_finfun) also have "\ = ?rhs" by(simp add: card_gt_0_iff) finally show ?thesis . qed instantiation finfun :: (finite_UNIV, card_UNIV) finite_UNIV begin definition "finite_UNIV = Phantom('a \f 'b) (let cb = of_phantom (card_UNIV :: 'b card_UNIV) in cb = 1 \ of_phantom (finite_UNIV :: 'a finite_UNIV) \ cb \ 0)" instance by intro_classes (auto simp add: finite_UNIV_finfun_def Let_def card_UNIV finite_UNIV finite_UNIV_finfun card_gt_0_iff) end instantiation finfun :: (card_UNIV, card_UNIV) card_UNIV begin definition "card_UNIV = Phantom('a \f 'b) (let ca = of_phantom (card_UNIV :: 'a card_UNIV); cb = of_phantom (card_UNIV :: 'b card_UNIV) in if ca \ 0 \ cb \ 0 \ cb = 1 then cb ^ ca else 0)" instance by intro_classes (simp add: card_UNIV_finfun_def card_UNIV Let_def card_UNIV_finfun) end subsubsection \Bundles for concrete syntax\ bundle finfun_syntax begin type_notation finfun ("(_ \f /_)" [22, 21] 21) notation finfun_const ("K$/ _" [0] 1) and finfun_update ("_'(_ $:= _')" [1000, 0, 0] 1000) and finfun_apply (infixl "$" 999) and finfun_comp (infixr "\$" 55) and finfun_comp2 (infixr "$\" 55) and finfun_Diag ("(1'($_,/ _$'))" [0, 0] 1000) notation (ASCII) finfun_comp (infixr "o$" 55) and finfun_comp2 (infixr "$o" 55) end bundle no_finfun_syntax begin no_type_notation finfun ("(_ \f /_)" [22, 21] 21) no_notation finfun_const ("K$/ _" [0] 1) and finfun_update ("_'(_ $:= _')" [1000, 0, 0] 1000) and finfun_apply (infixl "$" 999) and finfun_comp (infixr "\$" 55) and finfun_comp2 (infixr "$\" 55) and finfun_Diag ("(1'($_,/ _$'))" [0, 0] 1000) no_notation (ASCII) finfun_comp (infixr "o$" 55) and finfun_comp2 (infixr "$o" 55) end unbundle no_finfun_syntax end diff --git a/thys/IFC_Tracking/IFC.thy b/thys/IFC_Tracking/IFC.thy --- a/thys/IFC_Tracking/IFC.thy +++ b/thys/IFC_Tracking/IFC.thy @@ -1,4580 +1,4580 @@ section \Definitions\ text \ This section contains all necessary definitions of this development. Section~\ref{sec:pm} contains the structural definition of our program model which includes the security specification as well as abstractions of control flow and data. Executions of our program model are defined in section~\ref{sec:ex}. Additional well-formedness properties are defined in section~\ref{sec:wf}. Our security property is defined in section~\ref{sec:sec}. Our characterisation of how information is propagated by executions of our program model is defined in section~\ref{sec:char-cp}, for which the correctness result can be found in section~\ref{sec:cor-cp}. Section~\ref{sec:char-scp} contains an additional approximation of this characterisation whose correctness result can be found in section~\ref{sec:cor-scp}. \ theory IFC imports Main begin subsection \Program Model\ text_raw \\label{sec:pm}\ text \Our program model contains all necessary components for the remaining development and consists of:\ record ('n, 'var, 'val, 'obs) ifc_problem = \ \A set of nodes representing program locations:\ nodes :: \'n set\ \ \An initial node where all executions start:\ entry :: \'n\ \ \A final node where executions can terminate:\ return :: \'n\ \ \An abstraction of control flow in the form of an edge relation:\ edges :: \('n \ 'n) set\ \ \An abstraction of variables written at program locations:\ writes :: \'n \ 'var set\ \ \An abstraction of variables read at program locations:\ reads :: \'n \ 'var set\ \ \A set of variables containing the confidential information in the initial state:\ hvars :: \'var set\ \ \A step function on location state pairs:\ step :: \('n \ ('var \ 'val)) \ ('n \ ('var \ 'val))\ \ \An attacker model producing observations based on the reached state at certain locations:\ att :: \'n \ (('var \ 'val) \ 'obs)\ text \We fix a program in the following in order to define the central concepts. The necessary well-formedness assumptions will be made in section~\ref{sec:wf}.\ locale IFC_def = fixes prob :: \('n, 'var, 'val, 'obs) ifc_problem\ begin text \Some short hands to the components of the program which we will utilise exclusively in the following.\ definition nodes where \nodes = ifc_problem.nodes prob\ definition entry where \entry = ifc_problem.entry prob\ definition return where \return = ifc_problem.return prob\ definition edges where \edges = ifc_problem.edges prob\ definition writes where \writes = ifc_problem.writes prob\ definition reads where \reads = ifc_problem.reads prob\ definition hvars where \hvars = ifc_problem.hvars prob\ definition step where \step = ifc_problem.step prob\ definition att where \att = ifc_problem.att prob\ text \The components of the step function for convenience.\ definition suc where \suc n \ = fst (step (n, \))\ definition sem where \sem n \ = snd (step (n, \))\ lemma step_suc_sem: \step (n,\) = (suc n \, sem n \)\ unfolding suc_def sem_def by auto subsubsection \Executions\ text \\label{sec:ex}\ text \In order to define what it means for a program to be well-formed, we first require concepts of executions and program paths.\ text \The sequence of nodes visited by the execution corresponding to an input state.\ definition path where \path \ k= fst ((step^^k) (entry,\))\ text \The sequence of states visited by the execution corresponding to an input state.\ definition kth_state ( \_\<^bsup>_\<^esup>\ [111,111] 110) where \\\<^bsup>k\<^esup> = snd ((step^^k) (entry,\))\ text \A predicate asserting that a sequence of nodes is a valid program path according to the control flow graph.\ definition is_path where \is_path \ = (\ n. (\ n, \ (Suc n)) \ edges)\ end subsubsection \Well-formed Programs\ text_raw \\label{sec:wf}\ text \The following assumptions define our notion of valid programs.\ locale IFC = IFC_def \prob\ for prob:: \('n, 'var, 'val, 'out) ifc_problem\ + assumes ret_is_node[simp,intro]: \return \ nodes\ and entry_is_node[simp,intro]: \entry \ nodes\ and writes: \\ v n. (\\. \ v \ sem n \ v) \ v \ writes n\ and writes_return: \writes return = {}\ and uses_writes: \\ n \ \'. (\ v \ reads n. \ v = \' v) \ \ v \ writes n. sem n \ v = sem n \' v\ and uses_suc: \\ n \ \'. (\ v \ reads n. \ v = \' v) \ suc n \ = suc n \'\ and uses_att: \\ n f \ \'. att n = Some f \ (\ v \ reads n. \ v = \' v) \ f \ = f \'\ and edges_complete[intro,simp]: \\m \. m \ nodes \ (m,suc m \) \ edges\ and edges_return : \\x. (return,x) \ edges \ x = return \ and edges_nodes: \edges \ nodes \ nodes\ and reaching_ret: \\ x. x \ nodes \ \ \ n. is_path \ \ \ 0 = x \ \ n = return\ subsection \Security\ text_raw \\label{sec:sec}\ text \We define our notion of security, which corresponds to what Bohannon et al.~\cite{Bohannon:2009:RN:1653662.1653673} refer to as indistinguishable security. In order to do so we require notions of observations made by the attacker, termination and equivalence of input states.\ context IFC_def begin subsubsection \Observations\ text_raw \\label{sec:obs}\ text \The observation made at a given index within an execution.\ definition obsp where \obsp \ k = (case att(path \ k) of Some f \ Some (f (\\<^bsup>k\<^esup>)) | None \ None)\ text \The indices within a path where an observation is made.\ definition obs_ids :: \(nat \ 'n) \ nat set\ where \obs_ids \ = {k. att (\ k) \ None}\ text \A predicate relating an observable index to the number of observations made before.\ definition is_kth_obs :: \(nat \ 'n) \ nat \ nat \ bool\where \is_kth_obs \ k i = (card (obs_ids \ \ {.. att (\ i) \ None)\ text \The final sequence of observations made for an execution.\ definition obs where \obs \ k = (if (\i. is_kth_obs (path \) k i) then obsp \ (THE i. is_kth_obs (path \) k i) else None)\ text \Comparability of observations.\ definition obs_prefix :: \(nat \ 'obs option) \ (nat \ 'obs option) \ bool\ (infix \\\ 50) where \a \ b \ \ i. a i \ None \ a i = b i\ definition obs_comp (infix \\\ 50) where \a \ b \ a \ b \ b \ a\ subsubsection \Low equivalence of input states\ definition restrict (infix \\\ 100 ) where \f\U = (\ n. if n \ U then f n else undefined)\ text \Two input states are low equivalent if they coincide on the non high variables.\ definition loweq (infix \=\<^sub>L\ 50) where \\ =\<^sub>L \' = (\\(-hvars) = \'\(-hvars))\ subsubsection \Termination\ text \An execution terminates iff it reaches the terminal node at any point.\ definition terminates where \terminates \ \ \ i. path \ i = return\ subsubsection \Security Property\ text \The fixed program is secure if and only if for all pairs of low equivalent inputs the observation sequences are comparable and if the execution for an input state terminates then the observation sequence is not missing any observations.\ definition secure where \secure \ \ \ \'. \ =\<^sub>L \' \ (obs \ \ obs \' \ (terminates \ \ obs \' \ obs \))\ subsection \Characterisation of Information Flows\ text \We now define our characterisation of information flows which tracks data and control dependencies within executions. To do so we first require some additional concepts.\ subsubsection \Post Dominance\ text \We utilise the post dominance relation in order to define control dependence.\ text \The basic post dominance relation.\ definition is_pd (infix \pd\\ 50) where \y pd\ x \ x \ nodes \ (\ \ n. is_path \ \ \ (0::nat) = x \ \ n = return \ (\k\n. \ k = y))\ text \The immediate post dominance relation.\ definition is_ipd (infix \ipd\\ 50)where \y ipd\ x \ x \ y \ y pd\ x \ (\ z. z\x \ z pd\ x \ z pd\ y)\ definition ipd where \ipd x = (THE y. y ipd\ x)\ text \The post dominance tree.\ definition pdt where \pdt = {(x,y). x\y \ y pd\ x}\ subsubsection \Control Dependence\ text \An index on an execution path is control dependent upon another if the path does not visit the immediate post domiator of the node reached by the smaller index.\ definition is_cdi (\_ cd\<^bsup>_\<^esup>\ _\ [51,51,51]50) where \i cd\<^bsup>\\<^esup>\ k \ is_path \ \ k < i \ \ i \ return \ (\ j \ {k..i}. \ j \ ipd (\ k))\ text \The largest control dependency of an index is the immediate control dependency.\ definition is_icdi (\_ icd\<^bsup>_\<^esup>\ _\ [51,51,51]50) where \n icd\<^bsup>\\<^esup>\ n' \ is_path \ \ n cd\<^bsup>\\<^esup>\ n' \ (\ m \ {n'<.. n cd\<^bsup>\\<^esup>\ m)\ text \For the definition of the control slice, which we will define next, we require the uniqueness of the immediate control dependency.\ lemma icd_uniq: assumes \m icd\<^bsup>\\<^esup>\ n\ \ m icd\<^bsup>\\<^esup>\ n'\ shows \n = n'\ proof - { fix n n' assume *: \m icd\<^bsup>\\<^esup>\ n\ \ m icd\<^bsup>\\<^esup>\ n'\ \n < n'\ have \n' using * unfolding is_icdi_def is_cdi_def by auto hence \\ m cd\<^bsup>\\<^esup>\ n'\ using * unfolding is_icdi_def by auto with *(2) have \False\ unfolding is_icdi_def by auto } thus ?thesis using assms by (metis linorder_neqE_nat) qed subsubsection \Control Slice\ text \We utilise the control slice, that is the sequence of nodes visited by the control dependencies of an index, to match indices between executions.\ function cs:: \(nat \ 'n) \ nat \ 'n list\ (\cs\<^bsup>_\<^esup> _\ [51,70] 71) where \cs\<^bsup>\\<^esup> n = (if (\ m. n icd\<^bsup>\\<^esup>\ m) then (cs \ (THE m. n icd\<^bsup>\\<^esup>\ m))@[\ n] else [\ n])\ by pat_completeness auto termination \cs\ proof show \wf (measure snd)\ by simp fix \ n define m where \m == (The (is_icdi n \))\ assume \Ex (is_icdi n \)\ hence \n icd\<^bsup>\\<^esup>\ m\ unfolding m_def by (metis (full_types) icd_uniq theI') hence \m < n\ unfolding is_icdi_def is_cdi_def by simp thus \((\, The (is_icdi n \)), \, n) \ measure snd\ by (metis in_measure m_def snd_conv) qed inductive cs_less (infix \\\ 50) where \length xs < length ys \ take (length xs) ys = xs \ xs \ ys\ definition cs_select (infix \\\ 50) where \\\xs = (THE k. cs\<^bsup>\\<^esup> k = xs)\ subsubsection \Data Dependence\ text \Data dependence is defined straight forward. An index is data dependent upon another, if the index reads a variable written by the earlier index and the variable in question has not been written by any index in between.\ definition is_ddi (\_ dd\<^bsup>_,_\<^esup>\ _\ [51,51,51,51] 50) where \n dd\<^bsup>\,v\<^esup>\ m \ is_path \ \ m < n \ v \ reads (\ n) \ (writes (\ m)) \ (\ l \ {m<.. writes (\ l))\ subsubsection \Characterisation via Critical Paths\ text_raw \\label{sec:char-cp}\ text \With the above we define the set of critical paths which as we will prove characterise the matching points in executions where diverging data is read.\ inductive_set cp where \ \Any pair of low equivalent input states and indices where a diverging high variable is first read is critical.\ \\\ =\<^sub>L \'; cs\<^bsup>path \\<^esup> n = cs\<^bsup>path \'\<^esup> n'; h \ reads(path \ n); (\\<^bsup>n\<^esup>) h \ (\'\<^bsup>n'\<^esup>) h; \ kwrites(path \ k); \ k'writes(path \' k') \ \ ((\,n),(\',n')) \ cp\ | \ \If from a pair of critical indices in two executions there exist data dependencies from both indices to a pair of matching indices where the variable diverges, the later pair of indices is critical.\ \\((\,k),(\',k')) \ cp; n dd\<^bsup>path \,v\<^esup>\ k; n' dd\<^bsup>path \',v\<^esup>\ k'; cs\<^bsup>path \\<^esup> n = cs\<^bsup>path \'\<^esup> n'; (\\<^bsup>n\<^esup>) v \ (\'\<^bsup>n'\<^esup>) v \ \ ((\,n),(\',n')) \ cp\ | \ \If from a pair of critical indices the executions take different branches and one of the critical indices is a control dependency of an index that is data dependency of a matched index where diverging data is read and the variable in question is not written by the other execution after the executions first reached matching indices again, then the later matching pair of indices is critical.\ \\((\,k),(\',k')) \ cp; n dd\<^bsup>path \,v\<^esup>\ l; l cd\<^bsup>path \\<^esup>\ k; cs\<^bsup>path \\<^esup> n = cs\<^bsup>path \'\<^esup> n'; path \ (Suc k) \ path \' (Suc k'); (\\<^bsup>n\<^esup>) v \ (\'\<^bsup>n'\<^esup>) v; \j'\{(LEAST i'. k' < i' \ (\i. cs\<^bsup>path \\<^esup> i = cs\<^bsup>path \'\<^esup> i'))..writes (path \' j') \ \ ((\,n),(\',n')) \ cp\ | \ \The relation is symmetric.\ \\((\,k),(\',k')) \ cp\ \ ((\',k'),(\,k)) \ cp\ text \Based on the set of critical paths, the critical observable paths are those that either directly reach observable nodes or are diverging control dependencies of an observable index.\ inductive_set cop where \\((\,n),(\',n')) \ cp; path \ n \ dom att \ \ ((\,n),(\',n')) \ cop\ | \\((\,k),(\',k')) \ cp; n cd\<^bsup>path \\<^esup>\ k; path \ (Suc k) \ path \' (Suc k'); path \ n \ dom att \ \ ((\,n),(\',k')) \ cop\ subsubsection \Approximation via Single Critical Paths\ text_raw \\label{sec:char-scp}\ text \For applications we also define a single execution approximation.\ definition is_dcdi_via (\_ dcd\<^bsup>_,_\<^esup>\ _ via _ _\ [51,51,51,51,51,51] 50) where \n dcd\<^bsup>\,v\<^esup>\ m via \' m' = (is_path \ \ m < n \ (\ l' n'. cs\<^bsup>\\<^esup> m = cs\<^bsup>\'\<^esup> m' \ cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n' \ n' dd\<^bsup>\',v\<^esup>\ l' \ l' cd\<^bsup>\'\<^esup>\ m') \ (\ l \ {m.. writes(\ l)))\ inductive_set scp where \\h \ hvars; h \ reads (path \ n); (\ k writes(path \ k))\ \ (path \,n) \ scp\ | \\(\,m) \ scp; n cd\<^bsup>\\<^esup>\ m\ \ (\,n) \ scp\| \\(\,m) \ scp; n dd\<^bsup>\,v\<^esup>\ m\ \ (\,n) \ scp\| \\(\,m) \ scp; (\',m') \ scp; n dcd\<^bsup>\,v\<^esup>\ m via \' m'\ \ (\,n) \ scp\ inductive_set scop where \\(\,n) \ scp; \ n \ dom att\ \ (\,n) \ scop\ subsubsection \Further Definitions\ text \The following concepts are utilised by the proofs.\ inductive contradicts (infix \\\ 50) where \\cs\<^bsup>\'\<^esup> k' \ cs\<^bsup>\\<^esup> k ; \ = path \; \' = path \' ; \ (Suc (\\cs\<^bsup>\'\<^esup> k')) \ \' (Suc k')\ \ (\', k') \ (\, k)\| \\cs\<^bsup>\'\<^esup> k' = cs\<^bsup>\\<^esup> k ; \ = path \; \' = path \' ; \\<^bsup>k\<^esup> \ (reads (\ k)) \ \'\<^bsup>k'\<^esup> \ (reads (\ k))\ \ (\',k') \ (\,k)\ definition path_shift (infixl \\\ 51) where [simp]: \\\m = (\ n. \ (m+n))\ definition path_append :: \(nat \ 'n) \ nat \ (nat \ 'n) \ (nat \ 'n)\ (\_ @\<^bsup>_\<^esup> _\ [0,0,999] 51) where [simp]: \\ @\<^bsup>m\<^esup> \' = (\n.(if n \ m then \ n else \' (n-m)))\ definition eq_up_to :: \(nat \ 'n) \ nat \ (nat \ 'n) \ bool\ (\_ =\<^bsub>_\<^esub> _\ [55,55,55] 50) where \\ =\<^bsub>k\<^esub> \' = (\ i \ k. \ i = \' i)\ end (* End of locale IFC_def *) section \Proofs\ text_raw \\label{sec:proofs}\ subsection \Miscellaneous Facts\ lemma option_neq_cases: assumes \x \ y\ obtains (none1) a where \x = None\ \y = Some a\ | (none2) a where \x = Some a\ \y = None\ | (some) a b where \x = Some a\ \y = Some b\ \a \ b\ using assms by fastforce lemmas nat_sym_cases[case_names less sym eq] = linorder_less_wlog lemma mod_bound_instance: assumes \j < (i::nat)\ obtains j' where \k < j'\ and \j' mod i = j\ proof - have \k < Suc k * i + j\ using assms less_imp_Suc_add by fastforce moreover have \(Suc k * i + j) mod i = j\ by (metis assms mod_less mod_mult_self3) ultimately show \thesis\ using that by auto qed lemma list_neq_prefix_cases: assumes \ls \ ls'\ and \ls \ Nil\ and \ls' \ Nil\ obtains (diverge) xs x x' ys ys' where \ls = xs@[x]@ys\ \ls' = xs@[x']@ys'\ \x \ x'\ | (prefix1) xs where \ls = ls'@xs\ and \xs \ Nil\ | (prefix2) xs where \ls@xs = ls'\ and \xs \ Nil\ using assms proof (induct \length ls\ arbitrary: \ls\ \ls'\ rule: less_induct) case (less ls ls') obtain z zs z' zs' where lz: \ls = z#zs\ \ls' = z'#zs'\ by (metis list.exhaust less(6,7)) show \?case\ proof cases assume zz: \z = z'\ hence zsz: \zs \ zs'\ using less(5) lz by auto have lenz: \length zs < length ls\ using lz by auto show \?case\ proof(cases \zs = Nil\) assume zs: \zs = Nil\ hence \zs' \ Nil\ using zsz by auto moreover have \ls@zs' = ls'\ using zs lz zz by auto ultimately show \thesis\ using less(4) by blast next assume zs: \zs \ Nil\ show \thesis\ proof (cases \zs' = Nil\) assume \zs' = Nil\ hence \ls = ls'@zs\ using lz zz by auto thus \thesis\ using zs less(3) by blast next assume zs': \zs' \ Nil\ { fix xs x ys x' ys' assume \zs = xs @ [x] @ ys\ \zs' = xs @ [x'] @ ys'\ and xx: \x \ x'\ hence \ls = (z#xs) @ [x] @ ys\ \ls' = (z#xs) @ [x'] @ ys'\ using lz zz by auto hence \thesis\ using less(2) xx by blast } note * = this { fix xs assume \zs = zs' @ xs\ and xs: \xs \ []\ hence \ls = ls' @ xs\ using lz zz by auto hence \thesis\ using xs less(3) by blast } note ** = this { fix xs assume \zs@xs = zs'\ and xs: \xs \ []\ hence \ls@xs = ls'\ using lz zz by auto hence \thesis\ using xs less(4) by blast } note *** = this have \(\xs x ys x' ys'. zs = xs @ [x] @ ys \ zs' = xs @ [x'] @ ys' \ x \ x' \ thesis) \ (\xs. zs = zs' @ xs \ xs \ [] \ thesis) \ (\xs. zs @ xs = zs' \ xs \ [] \ thesis) \ thesis\ using less(1)[OF lenz _ _ _ zsz zs zs' ] . thus \thesis\ using * ** *** by blast qed qed next assume \z \ z'\ moreover have \ls = []@[z]@zs\ \ls' = []@[z']@zs'\ using lz by auto ultimately show \thesis\ using less(2) by blast qed qed lemma three_cases: assumes \A \ B \ C\ obtains \A\ | \B\ | \C\ using assms by auto lemma insort_greater: \\ x \ set ls. x < y \ insort y ls = ls@[y]\ by (induction \ls\,auto) lemma insort_append_first: assumes \\ y \ set ys. x \ y\ shows \insort x (xs@ys) = insort x xs @ ys\ using assms by (induction \xs\,auto,metis insort_is_Cons) lemma sorted_list_of_set_append: assumes \finite xs\ \finite ys\ \\ x \ xs. \ y \ ys. x < y\ shows \sorted_list_of_set (xs \ ys) = sorted_list_of_set xs @ (sorted_list_of_set ys)\ using assms(1,3) proof (induction \xs\) case empty thus \?case\ by simp next case (insert x xs) hence iv: \sorted_list_of_set (xs \ ys) = sorted_list_of_set xs @ sorted_list_of_set ys\ by blast have le: \\ y \ set (sorted_list_of_set ys). x < y\ using insert(4) assms(2) sorted_list_of_set by auto have \sorted_list_of_set (insert x xs \ ys) = sorted_list_of_set (insert x (xs \ ys))\ by auto also - have \\ = insort x (sorted_list_of_set (xs \ ys))\ by (metis Un_iff assms(2) finite_Un insert.hyps(1) insert.hyps(2) insert.prems insertI1 less_irrefl sorted_list_of_set.insert) + have \\ = insort x (sorted_list_of_set (xs \ ys))\ by (metis Un_iff assms(2) finite_Un insert.hyps(1) insert.hyps(2) insert.prems insertI1 less_irrefl sorted_list_of_set_insert) also have \\ = insort x (sorted_list_of_set xs @ sorted_list_of_set ys)\ using iv by simp also have \\ = insort x (sorted_list_of_set xs) @ sorted_list_of_set ys\ by (metis le insort_append_first less_le_not_le) also have \\ = sorted_list_of_set (insert x xs) @ sorted_list_of_set ys\ using sorted_list_of_set_insert[OF insert(1),of \x\] insert(2) by auto finally show \?case\ . qed lemma filter_insort: \sorted xs \ filter P (insort x xs) = (if P x then insort x (filter P xs) else filter P xs)\ by (induction \xs\, simp) (metis filter_insort filter_insort_triv map_ident) lemma filter_sorted_list_of_set: assumes \finite xs\ shows \filter P (sorted_list_of_set xs) = sorted_list_of_set {x \ xs. P x}\ using assms proof(induction \xs\) case empty thus \?case\ by simp next case (insert x xs) have *: \set (sorted_list_of_set xs) = xs\ \sorted (sorted_list_of_set xs)\ \distinct (sorted_list_of_set xs)\ by (auto simp add: insert.hyps(1)) have **: \P x \ {y \ insert x xs. P y} = insert x {y \ xs. P y}\ by auto have ***: \\ P x \ {y \ insert x xs. P y} = {y \ xs. P y}\ by auto note filter_insort[OF *(2),of \P\ \x\] sorted_list_of_set_insert[OF insert(1), of \x\] insert(2,3) ** *** - thus \?case\ by (metis (mono_tags) "*"(1) List.finite_set distinct_filter distinct_insort distinct_sorted_list_of_set set_filter sorted_list_of_set.insert) + thus \?case\ by (metis (mono_tags) "*"(1) List.finite_set distinct_filter distinct_insort distinct_sorted_list_of_set set_filter sorted_list_of_set_insert) qed lemma unbounded_nat_set_infinite: assumes \\ (i::nat). \ j\i. j \ A\ shows \\ finite A\ using assms by (metis finite_nat_set_iff_bounded_le not_less_eq_eq) lemma infinite_ascending: assumes nf: \\ finite (A::nat set)\ obtains f where \range f = A\ \\ i. f i < f (Suc i)\ proof let \?f\ = \\ i. (LEAST a. a \ A \ card (A \ {.. { fix i obtain a where \a \ A\ \card (A \ {.. proof (induction \i\ arbitrary: \thesis\) case 0 let \?a0\ = \(LEAST a. a \ A)\ have \?a0 \ A\ by (metis LeastI empty_iff finite.emptyI nf set_eq_iff) moreover have \\b. b \ A \ ?a0 \ b\ by (metis Least_le) hence \card (A \ {.. by force ultimately show \?case\ using 0 by blast next case (Suc i) obtain a where aa: \a \ A\ and card: \card (A \ {.. using Suc.IH by metis have nf': \~ finite (A - {..a})\ using nf by auto let \?b\ = \LEAST b. b \ A - {..a}\ have bin: \?b \ A-{..a}\ by (metis LeastI empty_iff finite.emptyI nf' set_eq_iff) have le: \\c. c \ A-{..a} \ ?b \ c\ by (metis Least_le) have ab: \a < ?b\ using bin by auto have \\ c. c \ A \ c < ?b \ c \ a\ using le by force hence \A \ {.. {.. using bin ab aa by force hence \card (A \{.. using card by auto thus \?case\ using Suc.prems bin by auto qed note \\ thesis. ((\a. a \ A \ card (A \ {.. thesis) \ thesis)\ } note ex = this { fix i obtain a where a: \a \ A \ card (A \{.. using ex by blast have ina: \?f i \ A\ and card: \card (A \{.. using LeastI[of \\ a. a \ A \ card (A \{.. \a\, OF a] by auto obtain b where b: \b \ A \ card (A \{.. using ex by blast have inab: \?f (Suc i) \ A\ and cardb: \card (A \{.. using LeastI[of \\ a. a \ A \ card (A \{.. \b\, OF b] by auto have \?f i < ?f (Suc i)\ proof (rule ccontr) assume \\ ?f i < ?f (Suc i)\ hence \A \{.. A \{.. by auto moreover have \finite (A \{.. by auto ultimately have \card(A \{.. card (A \{.. by (metis (erased, lifting) card_mono) thus \False\ using card cardb by auto qed note this ina } note b = this thus \\ i. ?f i < ?f (Suc i)\ by auto have *: \range ?f \ A\ using b by auto moreover { fix a assume ina: \a \ A\ let \?i\ = \card (A \ {.. obtain b where b: \b \ A \ card (A \{.. using ex by blast have inab: \?f ?i \ A\ and cardb: \card (A \{.. using LeastI[of \\ a. a \ A \ card (A \{.. \b\, OF b] by auto have le: \?f ?i \ a\ using Least_le[of \\ a. a \ A \ card (A \{.. \a\] ina by auto have \a = ?f ?i\ proof (rule ccontr) have fin: \finite (A \ {.. by auto assume \a \ ?f ?i\ hence \?f ?i < a\ using le by simp hence \?f ?i \ A \ {.. using inab by auto moreover have \A \ {.. A \ {.. using le by auto hence \A \ {.. {.. using cardb card_subset_eq[OF fin] by auto ultimately show \False\ by auto qed hence \a \ range ?f\ by auto } hence \A \ range ?f\ by auto ultimately show \range ?f = A\ by auto qed lemma mono_ge_id: \\ i. f i < f (Suc i) \ i \ f i\ apply (induction \i\,auto) by (metis not_le not_less_eq_eq order_trans) lemma insort_map_mono: assumes mono: \\ n m. n < m \ f n < f m\ shows \map f (insort n ns) = insort (f n) (map f ns)\ apply (induction \ns\) apply auto apply (metis not_less not_less_iff_gr_or_eq mono) apply (metis antisym_conv1 less_imp_le mono) apply (metis mono not_less) by (metis mono not_less) lemma sorted_list_of_set_map_mono: assumes mono: \\ n m. n < m \ f n < f m\ and fin: \finite A\ shows \map f (sorted_list_of_set A) = sorted_list_of_set (f`A)\ using fin proof (induction) case empty thus \?case\ by simp next case (insert x A) - have [simp]:\sorted_list_of_set (insert x A) = insort x (sorted_list_of_set A)\ using insert sorted_list_of_set.insert by simp + have [simp]:\sorted_list_of_set (insert x A) = insort x (sorted_list_of_set A)\ using insert sorted_list_of_set_insert by simp have \f ` insert x A = insert (f x) (f ` A)\ by auto moreover have \f x \ f`A\ apply (rule ccontr) using insert(2) mono apply auto by (metis insert.hyps(2) mono neq_iff) ultimately - have \sorted_list_of_set (f ` insert x A) = insort (f x) (sorted_list_of_set (f`A))\ using insert(1) sorted_list_of_set.insert by simp + have \sorted_list_of_set (f ` insert x A) = insort (f x) (sorted_list_of_set (f`A))\ using insert(1) sorted_list_of_set_insert by simp also have \\ = insort (f x) (map f (sorted_list_of_set A))\ using insert.IH by auto also have \\ = map f (insort x (sorted_list_of_set A))\ using insort_map_mono[OF mono] by auto finally show \map f (sorted_list_of_set (insert x A)) = sorted_list_of_set (f ` insert x A)\ by simp qed lemma GreatestIB: fixes n :: \nat\ and P assumes a:\\k\n. P k\ shows GreatestBI: \P (GREATEST k. k\n \ P k)\ and GreatestB: \(GREATEST k. k\n \ P k) \ n\ proof - show \P (GREATEST k. k\n \ P k)\ using GreatestI_ex_nat[OF assms] by auto show \(GREATEST k. k\n \ P k) \ n\ using GreatestI_ex_nat[OF assms] by auto qed lemma GreatestB_le: fixes n :: \nat\ assumes \x\n\ and \P x\ shows \x \ (GREATEST k. k\n \ P k)\ proof - have *: \\ y. y\n \ P y \ y by auto then show \x \ (GREATEST k. k\n \ P k)\ using assms by (blast intro: Greatest_le_nat) qed lemma LeastBI_ex: assumes \\k \ n. P k\ shows \P (LEAST k::'c::wellorder. P k)\ and \(LEAST k. P k) \ n\ proof - from assms guess k .. hence k: \k \ n\ \P k\ by auto thus \P (LEAST k. P k)\ using LeastI[of \P\ \k\] by simp show \(LEAST k. P k) \ n\ using Least_le[of \P\ \k\] k by auto qed lemma allB_atLeastLessThan_lower: assumes \(i::nat) \ j\ \\ x\{i.. shows \\ x\{j.. proof fix x assume \x\{j.. hence \x\{i.. using assms(1) by simp thus \P x\ using assms(2) by auto qed subsection \Facts about Paths\ context IFC begin lemma path0: \path \ 0 = entry\ unfolding path_def by auto lemma path_in_nodes[intro]: \path \ k \ nodes\ proof (induction \k\) case (Suc k) hence \\ \'. (path \ k, suc (path \ k) \') \ edges\ by auto hence \(path \ k, path \ (Suc k)) \ edges\ unfolding path_def by (metis suc_def comp_apply funpow.simps(2) prod.collapse) thus \?case\ using edges_nodes by force qed (auto simp add: path_def) lemma path_is_path[simp]: \is_path (path \)\ unfolding is_path_def path_def using step_suc_sem apply auto by (metis path_def suc_def edges_complete path_in_nodes prod.collapse) lemma term_path_stable: assumes \is_path \\ \\ i = return\ and le: \i \ j\ shows \\ j = return\ using le proof (induction \j\) case (Suc j) show \?case\ proof cases assume \i\j\ hence \\ j = return\ using Suc by simp hence \(return, \ (Suc j)) \ edges\ using assms(1) unfolding is_path_def by metis thus \\ (Suc j) = return\ using edges_return by auto next assume \\ i \ j\ hence \Suc j = i\ using Suc by auto thus \?thesis\ using assms(2) by auto qed next case 0 thus \?case\ using assms by simp qed lemma path_path_shift: assumes \is_path \\ shows \is_path (\\m)\ using assms unfolding is_path_def by simp lemma path_cons: assumes \is_path \\ \is_path \'\ \\ m = \' 0\ shows \is_path (\ @\<^bsup>m\<^esup> \')\ unfolding is_path_def proof(rule,cases) fix n assume \m < n\ thus \((\ @\<^bsup>m\<^esup> \') n, (\ @\<^bsup>m\<^esup> \') (Suc n)) \ edges\ using assms(2) unfolding is_path_def path_append_def by (auto,metis Suc_diff_Suc diff_Suc_Suc less_SucI) next fix n assume *: \\ m < n\ thus \((\ @\<^bsup>m\<^esup> \') n, (\ @\<^bsup>m\<^esup> \') (Suc n)) \ edges\ proof cases assume [simp]: \n = m\ thus \?thesis\ using assms unfolding is_path_def path_append_def by force next assume \n \ m\ hence \Suc n \ m\ \n\ m\ using * by auto with assms(1) show \?thesis\ unfolding is_path_def by auto qed qed lemma is_path_loop: assumes \is_path \\ \0 < i\ \\ i = \ 0\ shows \is_path (\ n. \ (n mod i))\ unfolding is_path_def proof (rule,cases) fix n assume \0 < Suc n mod i\ hence \Suc n mod i = Suc (n mod i)\ by (metis mod_Suc neq0_conv) moreover have \(\ (n mod i), \ (Suc (n mod i))) \ edges\ using assms(1) unfolding is_path_def by auto ultimately show \(\ (n mod i), \ (Suc n mod i)) \ edges\ by simp next fix n assume \\ 0 < Suc n mod i\ hence \Suc n mod i = 0\ by auto moreover hence \n mod i = i - 1\ using assms(2) by (metis Zero_neq_Suc diff_Suc_1 mod_Suc) ultimately show \(\(n mod i), \ (Suc n mod i)) \ edges\ using assms(1) unfolding is_path_def by (metis assms(3) mod_Suc) qed lemma path_nodes: \is_path \ \ \ k \ nodes\ unfolding is_path_def using edges_nodes by force lemma direct_path_return': assumes \is_path \ \ \\ 0 = x\ \x \ return\ \\ n = return\ obtains \' n' where \is_path \'\ \\' 0 = x\ \\' n' = return\ \\ i> 0. \' i \ x\ using assms proof (induction \n\ arbitrary: \\\ rule: less_induct) case (less n \) hence ih: \\ n' \'. n' < n \ is_path \' \ \' 0 = x \ \' n' = return \ thesis\ using assms by auto show \thesis\ proof cases assume \\ i>0. \ i \ x\ thus \thesis\ using less by auto next assume \\ (\ i>0. \ i \ x)\ then obtain i where \0 \\ i = x\ by auto hence \(\\i) 0 = x\ by auto moreover have \i < n\ using less(3,5,6) \\ i = x\ by (metis linorder_neqE_nat term_path_stable less_imp_le) hence \(\\i) (n-i) = return\ using less(6) by auto moreover have \is_path (\\i)\ using less(3) by (metis path_path_shift) moreover have \n - i < n\ using \0 \i < n\ by auto ultimately show \thesis\ using ih by auto qed qed lemma direct_path_return: assumes \x \ nodes\ \x \ return\ obtains \ n where \is_path \\ \\ 0 = x\ \\ n = return\ \\ i> 0. \ i \ x\ using direct_path_return'[of _ \x\] reaching_ret[OF assms(1)] assms(2) by blast lemma path_append_eq_up_to: \(\ @\<^bsup>k\<^esup> \') =\<^bsub>k\<^esub> \\ unfolding eq_up_to_def by auto lemma eq_up_to_le: assumes \k \ n\ \\ =\<^bsub>n\<^esub> \'\ shows \\ =\<^bsub>k\<^esub> \'\ using assms unfolding eq_up_to_def by auto lemma eq_up_to_refl: shows \\ =\<^bsub>k\<^esub> \\ unfolding eq_up_to_def by auto lemma eq_up_to_sym: assumes \\ =\<^bsub>k\<^esub> \'\ shows \\' =\<^bsub>k\<^esub> \\ using assms unfolding eq_up_to_def by auto lemma eq_up_to_apply: assumes \\ =\<^bsub>k\<^esub> \'\ \j \ k\ shows \\ j = \' j\ using assms unfolding eq_up_to_def by auto lemma path_swap_ret: assumes \is_path \\ obtains \' n where \is_path \'\ \\ =\<^bsub>k\<^esub> \'\ \\' n = return\ proof - have nd: \\ k \ nodes\ using assms path_nodes by simp obtain \' n where *: \is_path \'\ \\' 0 = \ k\ \\' n = return\ using reaching_ret[OF nd] by blast have \\ =\<^bsub>k\<^esub> (\@\<^bsup>k\<^esup> \')\ by (metis eq_up_to_sym path_append_eq_up_to) moreover have \is_path (\@\<^bsup>k\<^esup> \')\ using assms * path_cons by metis moreover have \(\@\<^bsup>k\<^esup> \') (k + n) = return\ using * by auto ultimately show \thesis\ using that by blast qed lemma path_suc: \path \ (Suc k) = fst (step (path \ k, \\<^bsup>k\<^esup>))\ by (induction \k\, auto simp: path_def kth_state_def) lemma kth_state_suc: \\\<^bsup>Suc k\<^esup> = snd (step (path \ k, \\<^bsup>k\<^esup>))\ by (induction \k\, auto simp: path_def kth_state_def) subsection \Facts about Post Dominators\ lemma pd_trans: assumes 1: \y pd\ x\ and 2: \z pd\y\ shows \z pd\x\ proof - { fix \ n assume 3[simp]: \is_path \\ \\ 0 = x\ \\ n = return\ then obtain k where \\ k = y\ and 7: \k \ n\ using 1 unfolding is_pd_def by blast then have \(\\k) 0 = y\ and \(\\k) (n-k) = return\ by auto moreover have \is_path (\\k)\ by(metis 3(1) path_path_shift) ultimately obtain k' where 8: \(\\k) k' = z\ and \k' \ n-k\ using 2 unfolding is_pd_def by blast hence \k+k'\n\ and \\ (k+ k') = z\ using 7 by auto hence \\k\n. \ k = z\ using path_nodes by auto } thus \?thesis\ using 1 unfolding is_pd_def by blast qed lemma pd_path: assumes \y pd\ x\ obtains \ n k where \is_path \\ and \\ 0 = x\ and \\ n = return\ and \\ k = y\ and \k \ n\ using assms unfolding is_pd_def using reaching_ret[of \x\] by blast lemma pd_antisym: assumes xpdy: \x pd\ y\ and ypdx: \y pd\ x\ shows \x = y\ proof - obtain \ n where path: \is_path \\ and \0: \\ 0 = x\ and \n: \\ n = return\ using pd_path[OF ypdx] by metis hence kex: \\k\n. \ k = y\ using ypdx unfolding is_pd_def by auto obtain k where k: \k = (GREATEST k. k\n \ \ k = y)\ by simp have \k: \\ k = y\ and kn: \k \ n\ using k kex by (auto intro: GreatestIB) have kpath: \is_path (\\k)\ by (metis path_path_shift path) moreover have k0: \(\\k) 0 = y\ using \k by simp moreover have kreturn: \(\\k) (n-k) = return\ using kn \n by simp ultimately have ky': \\k'\(n-k).(\\k) k' = x\ using xpdy unfolding is_pd_def by simp obtain k' where k': \k' = (GREATEST k'. k'\(n-k) \ (\\k) k' = x)\ by simp with ky' have \k': \(\\k) k' = x\ and kn': \k' \ (n-k)\ by (auto intro: GreatestIB) have k'path: \is_path (\\k\k')\ using kpath by(metis path_path_shift) moreover have k'0: \(\\k\k') 0 = x\ using \k' by simp moreover have k'return: \(\\k\k') (n-k-k') = return\ using kn' kreturn by (metis path_shift_def le_add_diff_inverse) ultimately have ky'': \\k''\(n-k-k').(\\k\k') k'' = y\ using ypdx unfolding is_pd_def by blast obtain k'' where k'': \k''= (GREATEST k''. k''\(n-k-k') \ (\\k\k') k'' = y)\ by simp with ky'' have \k'': \(\\k\k') k'' = y\ and kn'': \k'' \ (n-k-k')\ by (auto intro: GreatestIB) from this(1) have \\ (k + k' + k'') = y\ by (metis path_shift_def add.commute add.left_commute) moreover have \k + k' +k'' \ n\ using kn'' kn' kn by simp ultimately have \k + k' + k''\ k\ using k by(auto simp: GreatestB_le) hence \k' = 0\ by simp with k0 \k' show \x = y\ by simp qed lemma pd_refl[simp]: \x \ nodes \ x pd\ x\ unfolding is_pd_def by blast lemma pdt_trans_in_pdt: \(x,y) \ pdt\<^sup>+ \ (x,y) \ pdt\ proof (induction rule: trancl_induct) case base thus \?case\ by simp next case (step y z) show \?case\ unfolding pdt_def proof (simp) have *: \y pd\ x\ \z pd\ y\ using step unfolding pdt_def by auto hence [simp]: \z pd\ x\ using pd_trans[where x=\x\ and y=\y\ and z=\z\] by simp have \x\z\ proof assume \x = z\ hence \z pd\ y\ \y pd\ z\ using * by auto hence \z = y\ using pd_antisym by auto thus \False\ using step(2) unfolding pdt_def by simp qed thus \x \ z \ z pd\ x\ by auto qed qed lemma pdt_trancl_pdt: \pdt\<^sup>+ = pdt\ using pdt_trans_in_pdt by fast lemma trans_pdt: \trans pdt\ by (metis pdt_trancl_pdt trans_trancl) definition [simp]: \pdt_inv = pdt\\ lemma wf_pdt_inv: \wf (pdt_inv)\ proof (rule ccontr) assume \\ wf (pdt_inv)\ then obtain f where \\i. (f (Suc i), f i) \ pdt\\ using wf_iff_no_infinite_down_chain by force hence *: \\ i. (f i, f (Suc i)) \ pdt\ by simp have **:\\ i. \ j>i. (f i, f j) \ pdt\ proof(rule,rule,rule) fix i j assume \i < (j::nat)\ thus \(f i, f j) \ pdt\ proof (induction \j\ rule: less_induct) case (less k) show \?case\ proof (cases \Suc i < k\) case True hence k:\k-1 < k\ \i < k-1\ and sk: \Suc (k-1) = k\ by auto show \?thesis\ using less(1)[OF k] *[rule_format,of \k-1\,unfolded sk] trans_pdt[unfolded trans_def] by blast next case False hence \Suc i = k\ using less(2) by auto then show \?thesis\ using * by auto qed qed qed hence ***:\\ i. \ j > i. f j pd\ f i\ \\ i. \ j > i. f i \ f j\ unfolding pdt_def by auto hence ****:\\ i>0. f i pd\ f 0\ by simp hence \f 0 \ nodes\ using * is_pd_def by fastforce then obtain \ n where \:\is_path \\ \\ 0 = f 0\ \\ n = return\ using reaching_ret by blast hence \\ i>0. \ k\n. \ k = f i\ using ***(1) \f 0 \ nodes\ unfolding is_pd_def by blast hence \f:\\ i. \ k\n. \ k = f i\ using \(2) by (metis le0 not_gr_zero) have \range f \ \ ` {..n}\ proof(rule subsetI) fix x assume \x \ range f\ then obtain i where \x = f i\ by auto then obtain k where \x = \ k\ \k \ n\ using \f by metis thus \x \ \ ` {..n}\ by simp qed hence f:\finite (range f)\ using finite_surj by auto hence fi:\\ i. infinite {j. f j = f i}\ using pigeonhole_infinite[OF _ f] by auto obtain i where \infinite {j. f j = f i}\ using fi .. thus \False\ by (metis (mono_tags, lifting) "***"(2) bounded_nat_set_is_finite gt_ex mem_Collect_eq nat_neq_iff) qed lemma return_pd: assumes \x \ nodes\ shows \return pd\ x\ unfolding is_pd_def using assms by blast lemma pd_total: assumes xz: \x pd\ z\ and yz: \y pd\ z\ shows \x pd\ y \ y pd\x\ proof - obtain \ n where path: \is_path \\ and \0: \\ 0 = z\ and \n: \\ n = return\ using xz reaching_ret unfolding is_pd_def by force have *: \\ k\n. (\ k = x \ \ k = y)\ (is \\ k\n. ?P k\) using path \0 \n xz yz unfolding is_pd_def by auto obtain k where k: \k = (LEAST k. \ k = x \ \ k = y)\ by simp hence kn: \k\n\ and \k: \\ k = x \ \ k = y\ using LeastBI_ex[OF *] by auto note k_le = Least_le[where P = \?P\] show \?thesis\ proof (cases) assume kx: \\ k = x\ have k_min: \\ k'. \ k' = y \ k \ k'\ using k_le unfolding k by auto { fix \' and n' :: \nat\ assume path': \is_path \'\ and \'0: \\' 0 = x\ and \'n': \\' n' = return\ have path'': \is_path (\ @\<^bsup>k\<^esup> \')\ using path_cons[OF path path'] kx \'0 by auto have \''0: \(\ @\<^bsup>k\<^esup> \') 0 = z\ using \0 by simp have \''n: \(\ @\<^bsup>k\<^esup> \') (k+n') = return\ using \'n' kx \'0 by auto obtain k' where k': \k' \ k + n'\ \(\ @\<^bsup>k\<^esup> \') k' = y\ using yz path'' \''0 \''n unfolding is_pd_def by blast have **: \k \ k'\ proof (rule ccontr) assume \\ k \ k'\ hence \k' < k\ by simp moreover hence \\ k' = y\ using k' by auto ultimately show \False\ using k_min by force qed hence \\' (k' - k) = y\ using k' \'0 kx by auto moreover have \(k' - k) \ n'\ using k' by auto ultimately have \\ k\ n'. \' k = y\ by auto } hence \y pd\ x\ using kx path_nodes path unfolding is_pd_def by auto thus \?thesis\ .. next \ \This is analogous argument\ assume kx: \\ k \ x\ hence ky: \\ k = y\ using \k by auto have k_min: \\ k'. \ k' = x \ k \ k'\ using k_le unfolding k by auto { fix \' and n' :: \nat\ assume path': \is_path \'\ and \'0: \\' 0 = y\ and \'n': \\' n' = return\ have path'': \is_path (\ @\<^bsup>k\<^esup> \')\ using path_cons[OF path path'] ky \'0 by auto have \''0: \(\ @\<^bsup>k\<^esup> \') 0 = z\ using \0 by simp have \''n: \(\ @\<^bsup>k\<^esup> \') (k+n') = return\ using \'n' ky \'0 by auto obtain k' where k': \k' \ k + n'\ \(\ @\<^bsup>k\<^esup> \') k' = x\ using xz path'' \''0 \''n unfolding is_pd_def by blast have **: \k \ k'\ proof (rule ccontr) assume \\ k \ k'\ hence \k' < k\ by simp moreover hence \\ k' = x\ using k' by auto ultimately show \False\ using k_min by force qed hence \\' (k' - k) = x\ using k' \'0 ky by auto moreover have \(k' - k) \ n'\ using k' by auto ultimately have \\ k\ n'. \' k = x\ by auto } hence \x pd\ y\ using ky path_nodes path unfolding is_pd_def by auto thus \?thesis\ .. qed qed lemma pds_finite: \finite {y . (x,y) \ pdt}\ proof cases assume \x \ nodes\ then obtain \ n where \:\is_path \\ \\ 0 = x\ \\ n = return\ using reaching_ret by blast have *: \\ y \ {y. (x,y)\ pdt}. y pd\ x\ using pdt_def by auto have \\ y \ {y. (x,y)\ pdt}. \ k \ n. \ k = y\ using * \ is_pd_def by blast hence \{y. (x,y)\ pdt} \ \ ` {..n}\ by auto then show \?thesis\ using finite_surj by blast next assume \\ x\ nodes\ hence \{y. (x,y)\pdt} = {}\ unfolding pdt_def is_pd_def using path_nodes reaching_ret by fastforce then show \?thesis\ by simp qed lemma ipd_exists: assumes node: \x \ nodes\ and not_ret: \x\return\ shows \\y. y ipd\ x\ proof - let \?Q\ = \{y. x\y \ y pd\ x}\ have *: \return \ ?Q\ using assms return_pd by simp hence **: \\ x. x\ ?Q\ by auto have fin: \finite ?Q\ using pds_finite unfolding pdt_def by auto have tot: \\ y z. y\?Q \ z \ ?Q \ z pd\ y \ y pd\ z\ using pd_total by auto obtain y where ymax: \y\ ?Q\ \\ z\?Q. z = y \ z pd\ y\ using fin ** tot proof (induct) case empty then show \?case\ by auto next case (insert x F) show \thesis\ proof (cases \F = {}\) assume \F = {}\ thus \thesis\ using insert(4)[of \x\] by auto next assume \F \ {}\ hence \\ x. x\ F\ by auto have \\y. y \ F \ \z\F. z = y \ z pd\ y \ thesis\ proof - fix y assume a: \y \ F\ \\z\F. z = y \ z pd\ y\ have \x \ y\ using insert a by auto have \x pd\ y \ y pd\ x\ using insert(6) a(1) by auto thus \thesis\ proof assume \x pd\ y\ hence \\z\insert x F. z = y \ z pd\ y\ using a(2) by blast thus \thesis\ using a(1) insert(4) by blast next assume \y pd\ x\ have \\z\insert x F. z = x \ z pd\ x\ proof fix z assume \z\ insert x F\ thus \z = x \ z pd\ x\ proof(rule,simp) assume \z\F\ hence \z = y \ z pd\ y\ using a(2) by auto thus \z = x \ z pd\ x\ proof(rule,simp add: \y pd\ x\) assume \z pd\ y\ show \z = x \ z pd\ x\ using \y pd\ x\ \z pd\ y\ pd_trans by blast qed qed qed then show \thesis\ using insert by blast qed qed then show \thesis\ using insert by blast qed qed hence ***: \y pd\ x\ \x\y\ by auto have \\ z. z \ x \ z pd\ x \ z pd\ y\ proof (rule,rule) fix z assume a: \ z \ x \ z pd\ x\ hence b: \z \ ?Q\ by auto have \y pd\ z \ z pd\ y\ using pd_total ***(1) a by auto thus \z pd\ y\ proof assume c: \y pd\ z\ hence \y = z\ using b ymax pdt_def pd_antisym by auto thus \z pd\ y\ using c by simp qed simp qed with *** have \y ipd\ x\ unfolding is_ipd_def by simp thus \?thesis\ by blast qed lemma ipd_unique: assumes yipd: \y ipd\ x\ and y'ipd: \y' ipd\ x\ shows \y = y'\ proof - have 1: \y pd\ y'\ and 2: \y' pd\ y\ using yipd y'ipd unfolding is_ipd_def by auto show \?thesis\ using pd_antisym[OF 1 2] . qed lemma ipd_is_ipd: assumes \x \ nodes\ and \x\return\ shows \ipd x ipd\ x\ proof - from assms obtain y where \y ipd\ x\ using ipd_exists by auto moreover hence \\ z. z ipd\x \ z = y\ using ipd_unique by simp ultimately show \?thesis\ unfolding ipd_def by (auto intro: theI2) qed lemma is_ipd_in_pdt: \y ipd\ x \ (x,y) \ pdt\ unfolding is_ipd_def pdt_def by auto lemma ipd_in_pdt: \x \ nodes \ x\return \ (x,ipd x) \ pdt\ by (metis ipd_is_ipd is_ipd_in_pdt) lemma no_pd_path: assumes \x \ nodes\ and \\ y pd\ x\ obtains \ n where \is_path \\ and \\ 0 = x\ and \\ n = return\ and \\ k \ n. \ k \ y\ proof (rule ccontr) assume \\ thesis\ hence \\ \ n. is_path \ \ \ 0 = x \ \ n = return \ (\ k\n . \ k = y)\ using that by force thus \False\ using assms unfolding is_pd_def by auto qed lemma pd_pd_ipd: assumes \x \ nodes\ \x\return\ \y\x\ \y pd\ x\ shows \y pd\ ipd x\ proof - have \ipd x pd\ x\ by (metis assms(1,2) ipd_is_ipd is_ipd_def) hence \y pd\ ipd x \ ipd x pd\ y\ by (metis assms(4) pd_total) thus \?thesis\ proof have 1: \ipd x ipd\ x\ by (metis assms(1,2) ipd_is_ipd) moreover assume \ipd x pd\ y\ ultimately show \y pd\ ipd x\ unfolding is_ipd_def using assms(3,4) by auto qed auto qed lemma pd_nodes: assumes \y pd\ x\ shows pd_node1: \y \ nodes\ and pd_node2: \x \ nodes\ proof - obtain \ k where \is_path \\ \\ k = y\ using assms unfolding is_pd_def using reaching_ret by force thus \y \ nodes\ using path_nodes by auto show \x \ nodes\ using assms unfolding is_pd_def by simp qed lemma pd_ret_is_ret: \x pd\ return \ x = return\ by (metis pd_antisym pd_node1 return_pd) lemma ret_path_none_pd: assumes \x \ nodes\ \x\return\ obtains \ n where \is_path \\ \\ 0 = x\ \\ n = return\ \\ i>0. \ x pd\ \ i\ proof(rule ccontr) assume \\thesis\ hence *: \\ \ n. \is_path \; \ 0 = x; \ n = return\ \ \i>0. x pd\ \ i\ using that by blast obtain \ n where **: \is_path \\ \\ 0 = x\ \\ n = return\ \\ i>0. \ i \ x\ using direct_path_return[OF assms] by metis then obtain i where ***: \i>0\ \x pd\ \ i\ using * by blast hence \\ i \ return\ using pd_ret_is_ret assms(2) by auto hence \i < n\ using assms(2) term_path_stable ** by (metis linorder_neqE_nat less_imp_le) hence \(\\i)(n-i) = return\ using **(3) by auto moreover have \(\\i) (0) = \ i\ by simp moreover have \is_path (\\i)\ using **(1) path_path_shift by metis ultimately obtain k where \(\\i) k = x\ using ***(2) unfolding is_pd_def by metis hence \\ (i + k) = x\ by auto thus \False\ using **(4) \i>0\ by auto qed lemma path_pd_ipd0': assumes \is_path \\ and \\ n \ return\ \\ n \ \ 0\ and \\ n pd\ \ 0\ obtains k where \k \ n\ and \\ k = ipd(\ 0)\ proof(rule ccontr) have *: \\ n pd\ ipd (\ 0)\ by (metis is_pd_def assms(3,4) pd_pd_ipd pd_ret_is_ret) obtain \' n' where **: \is_path \'\ \\' 0 = \ n\ \\' n' = return\ \\ i>0. \ \ n pd\ \' i\ by (metis assms(2) assms(4) pd_node1 ret_path_none_pd) hence \\ i>0. \' i \ ipd (\ 0)\ using * by metis moreover assume \\ thesis\ hence \\ k\n. \ k \ ipd (\ 0)\ using that by blast ultimately have \\ i. (\@\<^bsup>n\<^esup> \') i \ ipd (\ 0)\ by (metis diff_is_0_eq neq0_conv path_append_def) moreover have \(\@\<^bsup>n\<^esup> \') (n + n') = return\ by (metis \\' 0 = \ n\ \\' n' = return\ add_diff_cancel_left' assms(2) diff_is_0_eq path_append_def) moreover have \(\@\<^bsup>n\<^esup> \') 0 = \ 0\ by (metis le0 path_append_def) moreover have \is_path (\@\<^bsup>n\<^esup> \')\ by (metis \\' 0 = \ n\ \is_path \'\ assms(1) path_cons) moreover have \ipd (\ 0) pd\ \ 0\ by (metis **(2,3,4) assms(2) assms(4) ipd_is_ipd is_ipd_def neq0_conv pd_node2) moreover have \\ 0 \ nodes\ by (metis assms(1) path_nodes) ultimately show \False\ unfolding is_pd_def by blast qed lemma path_pd_ipd0: assumes \is_path \\ and \\ 0 \ return\ \\ n \ \ 0\ and \\ n pd\ \ 0\ obtains k where \k \ n\ and \\ k = ipd(\ 0)\ proof cases assume *: \\ n = return\ have \ipd (\ 0) pd\ (\ 0)\ by (metis is_ipd_def is_pd_def assms(2,4) ipd_is_ipd) with assms(1,2,3) * show \thesis\ unfolding is_pd_def by (metis that) next assume \\ n \ return\ from path_pd_ipd0' [OF assms(1) this assms(3,4)] that show \thesis\ by auto qed lemma path_pd_ipd: assumes \is_path \\ and \\ k \ return\ \\ n \ \ k\ and \\ n pd\ \ k\ and kn: \k < n\ obtains l where \k < l\ and \l \ n\ and \\ l = ipd(\ k)\ proof - have \is_path (\ \ k)\ \(\ \ k) 0 \ return\ \(\ \ k) (n - k) \ (\ \ k) 0\ \(\ \ k) (n - k) pd\ (\ \ k) 0\ using assms path_path_shift by auto with path_pd_ipd0[of \\\k\ \n-k\] obtain ka where \ka \ n - k\ \(\ \ k) ka = ipd ((\ \ k) 0)\ . hence \k + ka \ n\ \\ (k + ka) = ipd (\ k)\ using kn by auto moreover hence \\ (k + ka) ipd\ \ k\ by (metis assms(1) assms(2) ipd_is_ipd path_nodes) hence \k < k + ka\ unfolding is_ipd_def by (metis nat_neq_iff not_add_less1) ultimately show \thesis\ using that[of \k+ka\] by auto qed lemma path_ret_ipd: assumes \is_path \\ and \\ k \ return\ \\ n = return\ obtains l where \k < l\ and \l \ n\ and \\ l = ipd(\ k)\ proof - have \\ n \ \ k\ using assms by auto moreover have \k \ n\ apply (rule ccontr) using term_path_stable assms by auto hence \k < n\ by (metis assms(2,3) dual_order.order_iff_strict) moreover have \\ n pd\ \ k\ by (metis assms(1,3) path_nodes return_pd) ultimately obtain l where \k < l\ \l \ n\ \\ l = ipd (\ k)\ using assms path_pd_ipd by blast thus \thesis\ using that by auto qed lemma pd_intro: assumes \l pd\ k\ \is_path \\ \\ 0 = k\ \\ n = return\ obtains i where \i \ n\ \\ i = l\ using assms unfolding is_pd_def by metis lemma path_pd_pd0: assumes path: \is_path \\ and lpdn: \\ l pd\ n\ and npd0: \n pd\ \ 0\ obtains k where \k \ l\ \\ k = n\ proof (rule ccontr) assume \\ thesis\ hence notn: \\ k. k \ l \ \ k \ n\ using that by blast have nret: \\ l \ return\ by (metis is_pd_def assms(1,3) notn) obtain \' n' where path': \is_path \'\ and \0': \\' 0 = \ l\ and \n': \\' n' = return\ and nonepd: \\ i>0. \ \ l pd\ \' i\ using nret path path_nodes ret_path_none_pd by metis have \\ l \ n\ using notn by simp hence \\ i. \' i \ n\ using nonepd \0' lpdn by (metis neq0_conv) hence notn': \\ i. (\@\<^bsup>l\<^esup> \') i \ n\ using notn \0' by auto have \is_path (\@\<^bsup>l\<^esup> \')\ using path path' by (metis \0' path_cons) moreover have \(\@\<^bsup>l\<^esup> \') 0 = \ 0\ by simp moreover have \(\@\<^bsup>l\<^esup> \') (n' + l) = return\ using \0' \n' by auto ultimately show \False\ using notn' npd0 unfolding is_pd_def by blast qed subsection \Facts about Control Dependencies\ lemma icd_imp_cd: \n icd\<^bsup>\\<^esup>\ k \ n cd\<^bsup>\\<^esup>\ k\ by (metis is_icdi_def) lemma ipd_impl_not_cd: assumes \j \ {k..i}\ and \\ j = ipd (\ k)\ shows \\ i cd\<^bsup>\\<^esup>\ k\ by (metis assms(1) assms(2) is_cdi_def) lemma cd_not_ret: assumes \i cd\<^bsup>\\<^esup>\ k \ shows \\ k \ return\ by (metis is_cdi_def assms nat_less_le term_path_stable) lemma cd_path_shift: assumes \j \ k\ \is_path \ \ shows \(i cd\<^bsup>\\<^esup>\ k) = (i - j cd\<^bsup>\\j\<^esup>\ k-j)\ proof assume a: \i cd\<^bsup>\\<^esup>\ k\ hence b: \k < i\ by (metis is_cdi_def) hence \is_path (\ \ j)\ \k - j < i - j\ using assms apply (metis path_path_shift) by (metis assms(1) b diff_less_mono) moreover have c: \\ j \ {k..i}. \ j \ ipd (\ k)\ by (metis a ipd_impl_not_cd) hence \\ ja \ {k - j..i - j}. (\ \ j) ja \ ipd ((\ \ j) (k - j))\ using b assms by auto fastforce moreover have \j < i\ using assms(1) b by auto hence \(\\j) (i - j) \ return\ using a unfolding is_cdi_def by auto ultimately show \i - j cd\<^bsup>\\j\<^esup>\ k-j\ unfolding is_cdi_def by simp next assume a: \i - j cd\<^bsup>\\j\<^esup>\ k-j\ hence b: \k - j < i-j\ by (metis is_cdi_def) moreover have c: \\ ja \ {k - j..i - j}. (\ \ j) ja \ ipd ((\ \ j) (k - j))\ by (metis a ipd_impl_not_cd) have \\ j \ {k..i}. \ j \ ipd (\ k)\ proof (rule,goal_cases) case (1 n) hence \n-j \ {k-j..i-j}\ using assms by auto hence \\ (j + (n-j)) \ ipd(\ (j + (k-j)))\ by (metis c path_shift_def) thus \?case\ using 1 assms(1) by auto qed moreover have \j < i\ using assms(1) b by auto hence \\ i \ return\ using a unfolding is_cdi_def by auto ultimately show \i cd\<^bsup>\\<^esup>\k\ unfolding is_cdi_def by (metis assms(1) assms(2) diff_is_0_eq' le_diff_iff nat_le_linear nat_less_le) qed lemma cd_path_shift0: assumes \is_path \\ shows \(i cd\<^bsup>\\<^esup>\ k) = (i-k cd\<^bsup>\\k\<^esup>\0)\ using cd_path_shift[OF _ assms] by (metis diff_self_eq_0 le_refl) lemma icd_path_shift: assumes \l \ k\ \is_path \\ shows \(i icd\<^bsup>\\<^esup>\ k) = (i - l icd\<^bsup>\\l\<^esup>\ k - l)\ proof - have \is_path (\\l)\ using path_path_shift assms(2) by auto moreover have \(i cd\<^bsup>\\<^esup>\ k) = (i - l cd\<^bsup>\\l\<^esup>\ k - l)\ using assms cd_path_shift by auto moreover have \(\ m \ {k<.. i cd\<^bsup>\\<^esup>\ m) = (\ m \ {k - l<.. i - l cd\<^bsup>\ \ l\<^esup>\ m)\ proof - {fix m assume *: \\ m \ {k - l<.. i - l cd\<^bsup>\ \ l\<^esup>\ m\ \m \ {k<.. hence \m-l \ {k-l<.. using assms(1) by auto hence \\ i - l cd\<^bsup>\\l\<^esup>\(m-l)\ using * by blast moreover have \l \ m\ using * assms by auto ultimately have \\ i cd\<^bsup>\\<^esup>\m\ using assms(2) cd_path_shift by blast } moreover {fix m assume *: \\ m \ {k<.. i cd\<^bsup>\\<^esup>\ m\ \m-l \ {k-l<.. hence \m \ {k<.. using assms(1) by auto hence \\ i cd\<^bsup>\\<^esup>\m\ using * by blast moreover have \l \ m\ using * assms by auto ultimately have \\ i - l cd\<^bsup>\\l\<^esup>\(m-l)\ using assms(2) cd_path_shift by blast } ultimately show \?thesis\ by auto (metis diff_add_inverse) qed ultimately show \?thesis\ unfolding is_icdi_def using assms by blast qed lemma icd_path_shift0: assumes \is_path \\ shows \(i icd\<^bsup>\\<^esup>\ k) = (i-k icd\<^bsup>\\k\<^esup>\0)\ using icd_path_shift[OF _ assms] by (metis diff_self_eq_0 le_refl) lemma cdi_path_swap: assumes \is_path \'\ \j cd\<^bsup>\\<^esup>\k\ \\ =\<^bsub>j\<^esub> \'\ shows \j cd\<^bsup>\'\<^esup>\k\ using assms unfolding eq_up_to_def is_cdi_def by auto lemma cdi_path_swap_le: assumes \is_path \'\ \j cd\<^bsup>\\<^esup>\k\ \\ =\<^bsub>n\<^esub> \'\ \j \ n\ shows \j cd\<^bsup>\'\<^esup>\k\ by (metis assms cdi_path_swap eq_up_to_le) lemma not_cd_impl_ipd: assumes \is_path \\ and \k < i\ and \\ i cd\<^bsup>\\<^esup>\ k\ and \\ i \ return\ obtains j where \j \ {k..i}\ and \\ j = ipd (\ k)\ by (metis assms(1) assms(2) assms(3) assms(4) is_cdi_def) lemma icd_is_the_icd: assumes \i icd\<^bsup>\\<^esup>\ k\ shows \k = (THE k. i icd\<^bsup>\\<^esup>\ k)\ using assms icd_uniq by (metis the1_equality) lemma all_ipd_imp_ret: assumes \is_path \\ and \\ i. \ i \ return \ (\ j>i. \ j = ipd (\ i))\ shows \\j. \ j = return\ proof - { fix x assume *: \\ 0 = x\ have \?thesis\ using wf_pdt_inv * assms proof(induction \x\ arbitrary: \\\ rule: wf_induct_rule ) case (less x \) show \?case\ proof (cases \x = return\) case True thus \?thesis\ using less(2) by auto next assume not_ret: \x \ return\ moreover then obtain k where k_ipd: \\ k = ipd x\ using less(2,4) by auto moreover have \x \ nodes\ using less(2,3) by (metis path_nodes) ultimately have \(x, \ k) \ pdt\ by (metis ipd_in_pdt) hence a: \(\ k, x) \ pdt_inv\ unfolding pdt_inv_def by simp have b: \is_path (\ \ k)\ by (metis less.prems(2) path_path_shift) have c: \\ i. (\\k) i \ return \ (\j>i. (\\k) j = ipd ((\\k) i))\ using less(4) apply auto by (metis (full_types) ab_semigroup_add_class.add_ac(1) less_add_same_cancel1 less_imp_add_positive) from less(1)[OF a _ b c] have \\j. (\\k) j = return\ by auto thus \\j. \ j = return\ by auto qed qed } thus \?thesis\ by simp qed lemma loop_has_cd: assumes \is_path \\ \0 < i\ \\ i = \ 0\ \\ 0 \ return\ shows \\ k < i. i cd\<^bsup>\\<^esup>\ k\ proof (rule ccontr) let \?\\ = \(\ n. \ (n mod i))\ assume \\ (\k\\<^esup>\ k)\ hence \\ k i cd\<^bsup>\\<^esup>\ k\ by blast hence *: \\ kj \ {k..i}. \ j = ipd (\ k))\ using assms(1,3,4) not_cd_impl_ipd by metis have \\ k. (\ j > k. ?\ j = ipd (?\ k))\ proof fix k have \k mod i < i\ using assms(2) by auto with * obtain j where \j \ {(k mod i)..i}\ \\ j = ipd (\ (k mod i))\ by auto then obtain j' where 1: \j' < i\ \\ j' = ipd (\ (k mod i))\ by (cases \j = i\, auto ,metis assms(2) assms(3),metis le_neq_implies_less) then obtain j'' where 2: \j'' > k\ \j'' mod i = j'\ by (metis mod_bound_instance) hence \?\ j'' = ipd (?\ k)\ using 1 by auto with 2(1) show \\ j > k. ?\ j = ipd (?\ k)\ by auto qed moreover have \is_path ?\\ by (metis assms(1) assms(2) assms(3) is_path_loop) ultimately obtain k where \?\ k = return\ by (metis (lifting) all_ipd_imp_ret) moreover have \k mod i < i\ by (simp add: assms(2)) ultimately have \\ i = return\ by (metis assms(1) term_path_stable less_imp_le) thus \False\ by (metis assms(3) assms(4)) qed lemma loop_has_cd': assumes \is_path \\ \j < i\ \\ i = \ j\ \\ j \ return\ shows \\ k \ {j..\\<^esup>\ k\ proof - have \\ k'< i-j. i-j cd\<^bsup>\\j\<^esup>\k'\ apply(rule loop_has_cd) apply (metis assms(1) path_path_shift) apply (auto simp add: assms less_imp_le) done then obtain k where k: \k \i-j cd\<^bsup>\\j\<^esup>\k\ by auto hence k': \(k+j) < i\ \i-j cd\<^bsup>\\j\<^esup>\ (k+j)-j\ by auto note cd_path_shift[OF _ assms(1)] hence \i cd\<^bsup>\\<^esup>\ k+j\ using k'(2) by (metis le_add1 add.commute) with k'(1) show \?thesis\ by force qed lemma claim'': assumes path\: \is_path \\ and path\': \is_path \'\ and \i: \\ i = \' i'\ and \j: \\ j = \' j'\ and not_cd: \\ k. \ j cd\<^bsup>\\<^esup>\ k\ \\ k. \ i' cd\<^bsup>\'\<^esup>\ k\ and nret: \\ i \ return\ and ilj: \i < j\ shows \i' < j'\ proof (rule ccontr) assume \\ i' < j'\ hence jlei: \j' \ i'\ by auto show \False\ proof (cases) assume j'li': \j' < i'\ define \'' where \\'' \ (\@\<^bsup>j\<^esup>(\'\j'))\i\ note \''_def[simp] have \\ j = (\' \ j') 0\ by (metis path_shift_def Nat.add_0_right \j) hence \is_path \''\ using path\ path\' \''_def path_path_shift path_cons by presburger moreover have \\'' (j-i+(i'-j')) = \'' 0\ using ilj jlei \i \j by (auto, metis add_diff_cancel_left' le_antisym le_diff_conv le_eq_less_or_eq) moreover have \\'' 0 \ return\ by (simp add: ilj less_or_eq_imp_le nret) moreover have \0 < j-i+(i'-j')\ by (metis add_is_0 ilj neq0_conv zero_less_diff) ultimately obtain k where k: \k < j-i+(i'-j')\ \j-i+(i'-j') cd\<^bsup>\''\<^esup>\ k\ by (metis loop_has_cd) hence *: \\ l \ {k..j-i+(i'-j')}. \'' l \ ipd (\'' k)\ by (metis is_cdi_def) show \False\ proof (cases \k < j-i\) assume a: \k < j - i\ hence b: \\'' k = \ (i + k)\ by auto have \\ l \ {i+k..j}. \ l \ ipd (\ (i+k))\ proof fix l assume l: \l \ {i + k..j}\ hence \\ l = \'' (l - i)\ by auto moreover from a l have \l-i \ {k .. j-i + (i'-j')}\ by force ultimately show \\ l \ ipd (\ (i + k))\ using * b by auto qed moreover have \i + k < j\ using a by simp moreover have \\ j \ return\ by (metis \i \j j'li' nret path\' term_path_stable less_imp_le) ultimately have \j cd\<^bsup>\\<^esup>\ i+k\ by (metis not_cd_impl_ipd path\) thus \False\ by (metis not_cd(1)) next assume \\ k < j - i\ hence a: \j - i \ k\ by simp hence b: \\'' k = \' (j' + (i + k) - j)\ unfolding \''_def path_shift_def path_append_def using ilj by(auto,metis \j add_diff_cancel_left' le_antisym le_diff_conv add.commute) have \\ l \ {j' + (i+k) - j..i'}. \' l \ ipd (\' (j' + (i+k) - j))\ proof fix l assume l: \l \ {j' + (i+k) - j..i'}\ hence \\' l = \'' (j + l - i - j')\ unfolding \''_def path_shift_def path_append_def using ilj by (auto, metis Nat.diff_add_assoc \j a add.commute add_diff_cancel_left' add_leD1 le_antisym le_diff_conv) moreover from a l have \j + l - i - j' \ {k .. j-i + (i'-j')}\ by force ultimately show \\' l \ ipd (\' (j' + (i + k) - j))\ using * b by auto qed moreover have \j' + (i+k) - j < i'\ using a j'li' ilj k(1) by linarith moreover have \\' i' \ return\ by (metis \i nret) ultimately have \i' cd\<^bsup>\'\<^esup>\ j' + (i+k) - j\ by (metis not_cd_impl_ipd path\') thus \False\ by (metis not_cd(2)) qed next assume \\ j' < i'\ hence \j' = i'\ by (metis \\ i' < j'\ linorder_cases) hence \\ i = \ j\ by (metis \i \j) thus \False\ by (metis ilj loop_has_cd' not_cd(1) nret path\) qed qed lemma other_claim': assumes path: \is_path \\ and eq: \\ i = \ j\ and \\ i \ return\ and icd: \\ k. \ i cd\<^bsup>\\<^esup>\ k\ and \\ k. \ j cd\<^bsup>\\<^esup>\ k\ shows \i = j\ proof (rule ccontr,cases) assume \i < j\ thus \False\ using assms claim'' by blast next assume \\ i < j\ \i \ j\ hence \j < i\ by auto thus \False\ using assms claim'' by (metis loop_has_cd') qed lemma icd_no_cd_path_shift: assumes \i icd\<^bsup>\\<^esup>\ 0\ shows \(\ k. \ i - 1 cd\<^bsup>\\1\<^esup>\ k)\ proof (rule,rule ccontr,goal_cases) case (1 k) hence *: \i - 1 cd\<^bsup>\ \ 1\<^esup>\ k\ by simp have **: \1 \ k + 1\ by simp have ***: \is_path \\ by (metis assms is_icdi_def) hence \i cd\<^bsup>\\<^esup>\ k+1\ using cd_path_shift[OF ** ***] * by auto moreover hence \k+1 < i\ unfolding is_cdi_def by simp moreover have \0 < k + 1\ by simp ultimately show \False\ using assms[unfolded is_icdi_def] by auto qed lemma claim': assumes path\: \is_path \\ and path\': \is_path \'\ and \i: \\ i = \' i'\ and \j: \\ j = \' j'\ and not_cd: \i icd\<^bsup>\\<^esup>\ 0\ \j icd\<^bsup>\\<^esup>\ 0\ \i' icd\<^bsup>\'\<^esup>\ 0\ \j' icd\<^bsup>\'\<^esup>\ 0\ and ilj: \i < j\ and nret: \\ i \ return\ shows \i' < j'\ proof - have g0: \0 < i\ \0 < j\ \0 < i'\ \0 < j'\using not_cd[unfolded is_icdi_def is_cdi_def] by auto have \(\ \ 1) (i - 1) = (\' \ 1) (i' - 1)\ \(\ \ 1) (j - 1) = (\' \ 1) (j' - 1)\ using \i \j g0 by auto moreover have \\ k. \ (j - 1) cd\<^bsup>\\1\<^esup>\ k\ \\ k. \ (i' - 1) cd\<^bsup>\'\1\<^esup>\ k\ by (metis icd_no_cd_path_shift not_cd(2)) (metis icd_no_cd_path_shift not_cd(3)) moreover have \is_path (\\1)\ \is_path (\'\1)\ using path\ path\' path_path_shift by blast+ moreover have \(\\1) (i - 1) \ return\ using g0 nret by auto moreover have \i - 1 < j - 1\ using g0 ilj by auto ultimately have \i' - 1 < j' - 1\ using claim'' by blast thus \i' by auto qed lemma other_claim: assumes path: \is_path \\ and eq: \\ i = \ j\ and \\ i \ return\ and icd: \i icd\<^bsup>\\<^esup>\ 0\ and \j icd\<^bsup>\\<^esup>\ 0\ shows \i = j\ proof (rule ccontr,cases) assume \i < j\ thus \False\ using assms claim' by blast next assume \\ i < j\ \i \ j\ hence \j < i\ by auto thus \False\ using assms claim' by (metis less_not_refl) qed lemma cd_trans0: assumes \j cd\<^bsup>\\<^esup>\ 0\ and \k cd\<^bsup>\\<^esup>\j\ shows \k cd\<^bsup>\\<^esup>\ 0\ proof (rule ccontr) have path: \is_path \\ and ij: \0 < j\ and jk: \j < k\ and nret: \\ j \ return\ \\ k \ return\ and noipdi: \\ l \ {0..j}. \ l \ ipd (\ 0)\ and noipdj: \\ l \ {j..k}. \ l \ ipd (\ j)\ using assms unfolding is_cdi_def by auto assume \\ k cd\<^bsup>\\<^esup>\ 0\ hence \\l \ {0..k}. \ l = ipd (\ 0)\ unfolding is_cdi_def using path ij jk nret by force then obtain l where \l \ {0..k}\ and l: \\ l = ipd (\ 0)\ by auto hence jl: \j and lk: \l\k\ using noipdi ij by auto have pdj: \ipd (\ 0) pd\ \ j\ proof (rule ccontr) have \\ j \ nodes\ using path by (metis path_nodes) moreover assume \\ ipd (\ 0) pd\ \ j\ ultimately obtain \' n where *: \is_path \'\ \\' 0 = \ j\ \\' n = return\ \\ k\n. \' k \ ipd(\ 0)\ using no_pd_path by metis hence path': \is_path (\ @\<^bsup>j\<^esup> \')\ by (metis path path_cons) moreover have \\ k \ j + n. (\@\<^bsup>j\<^esup> \') k \ ipd (\ 0)\ using noipdi *(4) by auto moreover have \(\@\<^bsup>j\<^esup> \') 0 = \ 0\ by auto moreover have \(\@\<^bsup>j\<^esup> \') (j + n) = return\ using *(2,3) by auto ultimately have \\ ipd (\ 0) pd\ \ 0\ unfolding is_pd_def by metis thus \False\ by (metis is_ipd_def ij ipd_is_ipd nret(1) path path_nodes term_path_stable less_imp_le) qed hence \(\\j) (l-j) pd\ (\\j) 0\ using jl l by auto moreover have \is_path (\\j)\ by (metis path path_path_shift) moreover have \\ l \ return\ by (metis lk nret(2) path term_path_stable) hence \(\\j) (l-j) \ return\ using jl by auto moreover have \\ j \ ipd (\ 0)\ using noipdi by force hence \(\\j) (l-j) \ (\\j) 0\ using jl l by auto ultimately obtain k' where \k' \ l-j\ and \(\\j) k' = ipd ((\\j) 0)\ using path_pd_ipd0' by blast hence \j + k' \ {j..k}\ \\ (j+k') = ipd (\ j)\ using jl lk by auto thus \False\ using noipdj by auto qed lemma cd_trans: assumes \j cd\<^bsup>\\<^esup>\ i\ and \k cd\<^bsup>\\<^esup>\j\ shows \k cd\<^bsup>\\<^esup>\ i\ proof - have path: \is_path \\ using assms is_cdi_def by auto have ij: \i using assms is_cdi_def by auto let \?\\ = \\\i\ have \j-i cd\<^bsup>?\\<^esup>\ 0\ using assms(1) cd_path_shift0 path by auto moreover have \k-i cd\<^bsup>?\\<^esup>\j-i\ by (metis assms(2) cd_path_shift is_cdi_def ij less_imp_le_nat) ultimately have \k-i cd\<^bsup>?\\<^esup>\ 0\ using cd_trans0 by auto thus \k cd\<^bsup>\\<^esup>\ i\ using path cd_path_shift0 by auto qed lemma excd_impl_exicd: assumes \\ k. i cd\<^bsup>\\<^esup>\k\ shows \\ k. i icd\<^bsup>\\<^esup>\k\ using assms proof(induction \i\ arbitrary: \\\ rule: less_induct) case (less i) then obtain k where k: \i cd\<^bsup>\\<^esup>\k\ by auto hence ip: \is_path \\ unfolding is_cdi_def by auto show \?case\ proof (cases) assume *: \\ m \ {k<.. i cd\<^bsup>\\<^esup>\ m\ hence \i icd\<^bsup>\\<^esup>\k\ using k ip unfolding is_icdi_def by auto thus \?case\ by auto next assume \\ (\ m \ {k<.. i cd\<^bsup>\\<^esup>\ m)\ then obtain m where m: \m \ {k<.. \i cd\<^bsup>\\<^esup>\ m\ by blast hence \i - m cd\<^bsup>\\m\<^esup>\ 0\ by (metis cd_path_shift0 is_cdi_def) moreover have \i - m < i\ using m by auto ultimately obtain k' where k': \i - m icd\<^bsup>\\m\<^esup>\ k'\ using less(1) by blast hence \i icd\<^bsup>\\<^esup>\ k' + m\ using ip by (metis add.commute add_diff_cancel_right' icd_path_shift le_add1) thus \?case\ by auto qed qed lemma cd_split: assumes \i cd\<^bsup>\\<^esup>\ k\ and \\ i icd\<^bsup>\\<^esup>\ k\ obtains m where \i icd\<^bsup>\\<^esup>\ m\ and \m cd\<^bsup>\\<^esup>\ k\ proof - have ki: \k < i\ using assms is_cdi_def by auto obtain m where m: \i icd\<^bsup>\\<^esup>\ m\ using assms(1) by (metis excd_impl_exicd) hence \k \ m\ unfolding is_icdi_def using ki assms(1) by force hence km: \k < m\using m assms(2) by (metis le_eq_less_or_eq) moreover have \\ m \ return\ using m unfolding is_icdi_def is_cdi_def by (simp, metis term_path_stable less_imp_le) moreover have \m using m unfolding is_cdi_def is_icdi_def by auto ultimately have \m cd\<^bsup>\\<^esup>\ k\ using assms(1) unfolding is_cdi_def by auto with m that show \thesis\ by auto qed lemma cd_induct[consumes 1, case_names base IS]: assumes prem: \i cd\<^bsup>\\<^esup>\ k\ and base: \\ i. i icd\<^bsup>\\<^esup>\k \ P i\ and IH: \\ k' i'. k' cd\<^bsup>\\<^esup>\ k \ P k' \ i' icd\<^bsup>\\<^esup>\ k' \ P i'\ shows \P i\ using prem IH proof (induction \i\ rule: less_induct,cases) case (less i) assume \i icd\<^bsup>\\<^esup>\ k\ thus \P i\ using base by simp next case (less i') assume \\ i' icd\<^bsup>\\<^esup>\ k\ then obtain k' where k': \ i' icd\<^bsup>\\<^esup>\ k'\ \k' cd\<^bsup>\\<^esup>\ k\ using less cd_split by blast hence icdk: \i' cd\<^bsup>\\<^esup>\ k'\ using is_icdi_def by auto note ih=less(3)[OF k'(2) _ k'(1)] have ki: \k' < i'\ using k' is_icdi_def is_cdi_def by auto have \P k'\ using less(1)[OF ki k'(2) ] less(3) by auto thus \P i'\ using ih by simp qed lemma cdi_prefix: \n cd\<^bsup>\\<^esup>\ m \ m < n' \ n' \ n \ n' cd\<^bsup>\\<^esup>\ m\ unfolding is_cdi_def by (simp, metis term_path_stable) lemma cr_wn': assumes 1: \n cd\<^bsup>\\<^esup>\ m\ and nc: \\ m' cd\<^bsup>\\<^esup>\ m\ and 3: \m < m'\ shows \n < m'\ proof (rule ccontr) assume \\ n < m'\ hence \m' \ n\ by simp hence \m' cd\<^bsup>\\<^esup>\ m\ by (metis 1 3 cdi_prefix) thus \False\ using nc by simp qed lemma cr_wn'': assumes \i cd\<^bsup>\\<^esup>\ m\ and \j cd\<^bsup>\\<^esup>\ n\ and \\ m cd\<^bsup>\\<^esup>\ n\ and \i \ j\ shows \m \ n\ proof (rule ccontr) assume \\m\n\ hence nm: \n < m\ by auto moreover have \m using assms(1) assms(4) unfolding is_cdi_def by auto ultimately have \m cd\<^bsup>\\<^esup>\ n\ using assms(2) cdi_prefix by auto thus \False\ using assms(3) by auto qed lemma ret_no_cd: assumes \\ n = return\ shows \\ n cd\<^bsup>\\<^esup>\ k\ by (metis assms is_cdi_def) lemma ipd_not_self: assumes \x \ nodes\ \x\ return\ shows \x \ ipd x\ by (metis is_ipd_def assms ipd_is_ipd) lemma icd_cs: assumes \l icd\<^bsup>\\<^esup>\k\ shows \cs\<^bsup>\\<^esup> l = cs\<^bsup>\\<^esup> k @ [\ l]\ proof - from assms have \k = (THE k. l icd\<^bsup>\\<^esup>\ k)\ by (metis icd_is_the_icd) with assms show \?thesis\ by auto qed lemma cd_not_pd: assumes \l cd\<^bsup>\\<^esup>\ k\ \\ l \ \ k\ shows \\ \ l pd\ \ k\ proof assume pd: \\ l pd\ \ k\ have nret: \\ k \ return\ by (metis assms(1) pd pd_ret_is_ret ret_no_cd) have kl: \k < l\ by (metis is_cdi_def assms(1)) have path: \is_path \\ by (metis is_cdi_def assms(1)) from path_pd_ipd[OF path nret assms(2) pd kl] obtain n where \k < n\ \n \ l\ \\ n = ipd (\ k)\ . thus \False\ using assms(1) unfolding is_cdi_def by auto qed lemma cd_ipd_is_cd: assumes \k \\ m = ipd (\ k)\ \\ n \ {k.. n \ ipd (\ k)\ and mcdj: \m cd\<^bsup>\\<^esup>\ j\ shows \k cd\<^bsup>\\<^esup>\ j\ proof cases assume \j < k\ thus \k cd\<^bsup>\\<^esup>\ j\ by (metis mcdj assms(1) cdi_prefix less_imp_le_nat) next assume \\ j < k\ hence kj: \k \ j\ by simp have \k < j\ apply (rule ccontr) using kj assms mcdj by (auto, metis is_cdi_def is_ipd_def cd_not_pd ipd_is_ipd path_nodes term_path_stable less_imp_le) moreover have \j < m\ using mcdj is_cdi_def by auto hence \\ n \ {k..j}. \ n \ ipd(\ k)\ using assms(3) by force ultimately have \j cd\<^bsup>\\<^esup>\ k\ by (metis mcdj is_cdi_def term_path_stable less_imp_le) hence \m cd\<^bsup>\\<^esup>\ k\ by (metis mcdj cd_trans) hence \False\ by (metis is_cdi_def is_ipd_def assms(2) cd_not_pd ipd_is_ipd path_nodes term_path_stable less_imp_le) thus \?thesis\ by simp qed lemma ipd_pd_cd0: assumes lcd: \n cd\<^bsup>\\<^esup>\ 0\ shows \ipd (\ 0) pd\ (\ n)\ proof - obtain k l where \0: \\ 0 = k\ and \n: \\ n = l\ and cdi: \n cd\<^bsup>\\<^esup>\ 0\ using lcd unfolding is_cdi_def by blast have nret: \k \ return\ by (metis is_cdi_def \0 cdi term_path_stable less_imp_le) have path: \is_path \\ and ipd: \\ i\n. \ i \ ipd k\ using cdi unfolding is_cdi_def \0 by auto { fix \' n' assume path': \is_path \'\ and \'0: \\' 0 = l\ and ret: \\' n' = return\ have \is_path (\ @\<^bsup>n\<^esup> \')\ using path path' \n \'0 by (metis path_cons) moreover have \(\ @\<^bsup>n\<^esup> \') (n+n') = return\ using ret \n \'0 by auto moreover have \(\ @\<^bsup>n\<^esup> \') 0 = k\ using \0 by auto moreover have \ipd k pd\ k\ by (metis is_ipd_def path \0 ipd_is_ipd nret path_nodes) ultimately obtain k' where k': \k' \ n+n'\ \(\ @\<^bsup>n\<^esup> \') k' = ipd k\ by (metis pd_intro) have \\ k'\ n\ proof assume \k' \ n\ hence \(\ @\<^bsup>n\<^esup> \') k' = \ k'\ by auto thus \False\ using k'(2) ipd by (metis \k' \ n\) qed hence \(\ @\<^bsup>n\<^esup> \') k' = \' (k' - n)\ by auto moreover have \(k' - n) \ n'\ using k' by simp ultimately have \\ k'\n'. \' k' = ipd k\ unfolding k' by auto } moreover have \l \ nodes\ by (metis \n path path_nodes) ultimately show \ipd (\ 0) pd\ (\ n)\ unfolding is_pd_def by (simp add: \0 \n) qed lemma ipd_pd_cd: assumes lcd: \l cd\<^bsup>\\<^esup>\ k\ shows \ipd (\ k) pd\ (\ l)\ proof - have \l-k cd\<^bsup>\\k\<^esup>\0\ using lcd cd_path_shift0 is_cdi_def by blast moreover note ipd_pd_cd0[OF this] moreover have \(\ \ k) 0 = \ k\ by auto moreover have \k < l\ using lcd unfolding is_cdi_def by simp then have \(\ \ k) (l - k) = \ l\ by simp ultimately show \?thesis\ by simp qed lemma cd_is_cd_ipd: assumes km: \k and ipd: \\ m = ipd (\ k)\ \\ n \ {k.. n \ ipd (\ k)\ and cdj: \k cd\<^bsup>\\<^esup>\ j\ and nipdj: \ipd (\ j) \ \ m\ shows \m cd\<^bsup>\\<^esup>\ j\ proof - have path: \is_path \\ and jk: \j < k\ and nretj: \\ k \ return\ and nipd: \\ l \ {j..k}. \ l \ ipd (\ j)\ using cdj is_cdi_def by auto have pd: \ipd (\ j) pd\ \ m\ by (metis atLeastAtMost_iff cdj ipd(1) ipd_pd_cd jk le_refl less_imp_le nipd nretj path path_nodes pd_pd_ipd) have nretm: \\ m \ return\ by (metis nipdj pd pd_ret_is_ret) have jm: \j < m\ using jk km by simp show \m cd\<^bsup>\\<^esup>\ j\ proof (rule ccontr) assume ncdj: \\ m cd\<^bsup>\\<^esup>\ j\ hence \\ l \ {j..m}. \ l = ipd (\ j)\ unfolding is_cdi_def by (metis jm nretm path) then obtain l where jl: \j \ l\ and \l \ m\ and lipd: \\ l = ipd (\ j)\ by force hence lm: \l < m\ using nipdj by (metis le_eq_less_or_eq) have npd: \\ ipd (\ k) pd\ \ l\ by (metis ipd(1) lipd nipdj pd pd_antisym) have nd: \\ l \ nodes\ using path path_nodes by simp from no_pd_path[OF nd npd] obtain \' n where path': \is_path \'\ and \'0: \\' 0 = \ l\ and \'n: \\' n = return\ and nipd: \\ ka\n. \' ka \ ipd (\ k)\ . let \?\\ = \(\@\<^bsup>l\<^esup> \') \ k\ have path'': \is_path ?\\ by (metis \'0 path path' path_cons path_path_shift) moreover have kl: \k < l\ using lipd cdj jl unfolding is_cdi_def by fastforce have \?\ 0 = \ k\ using kl by auto moreover have \?\ (l + n - k) = return\ using \'n \'0 kl by auto moreover have \ipd (\ k) pd\ \ k\ by (metis is_ipd_def ipd_is_ipd nretj path path_nodes) ultimately obtain l' where l': \l' \ (l + n - k)\ \?\ l' = ipd (\ k)\ unfolding is_pd_def by blast show \False\ proof (cases ) assume *: \k + l' \ l\ hence \\ (k + l') = ipd (\ k)\ using l' by auto moreover have \k + l' < m\ by (metis "*" dual_order.strict_trans2 lm) ultimately show \False\ using ipd(2) by simp next assume \\ k + l' \ l\ hence \\' (k + l' - l) = ipd (\ k)\ using l' by auto moreover have \k + l' - l \ n\ using l' kl by linarith ultimately show \False\ using nipd by auto qed qed qed lemma ipd_icd_greatest_cd_not_ipd: assumes ipd: \\ m = ipd (\ k)\ \\ n \ {k.. n \ ipd (\ k)\ and km: \k < m\ and icdj: \m icd\<^bsup>\\<^esup>\ j\ shows \j = (GREATEST j. k cd\<^bsup>\\<^esup>\ j \ ipd (\ j) \ \ m)\ proof - let \?j\ = \GREATEST j. k cd\<^bsup>\\<^esup>\ j \ ipd (\ j) \ \ m\ have kcdj: \k cd\<^bsup>\\<^esup>\ j\ using assms cd_ipd_is_cd is_icdi_def by blast have nipd: \ipd (\ j) \ \ m\ using icdj unfolding is_icdi_def is_cdi_def by auto have bound: \\ j. k cd\<^bsup>\\<^esup>\ j \ ipd (\ j) \ \ m \ j \ k\ unfolding is_cdi_def by simp have exists: \k cd\<^bsup>\\<^esup>\ j \ ipd (\ j) \ \ m\ (is \?P j\) using kcdj nipd by auto note GreatestI_nat[of \?P\ _ \k\, OF exists] Greatest_le_nat[of \?P\ \j\ \k\, OF exists] hence kcdj': \k cd\<^bsup>\\<^esup>\ ?j\ and ipd': \ipd (\ ?j) \ \ m\ and jj: \j \ ?j\ using bound by auto hence mcdj': \m cd\<^bsup>\\<^esup>\ ?j\ using ipd km cd_is_cd_ipd by auto show \j = ?j\ proof (rule ccontr) assume \j \ ?j\ hence jlj: \j < ?j\ using jj by simp moreover have \?j < m\ using kcdj' km unfolding is_cdi_def by auto ultimately show \False\ using icdj mcdj' unfolding is_icdi_def by auto qed qed lemma cd_impl_icd_cd: assumes \i cd\<^bsup>\\<^esup>\ l\ and \i icd\<^bsup>\\<^esup>\ k\ and \\ i icd\<^bsup>\\<^esup>\ l\ shows \k cd\<^bsup>\\<^esup>\ l\ using assms cd_split icd_uniq by metis lemma cdi_is_cd_icdi: assumes \k icd\<^bsup>\\<^esup>\ j\ shows \k cd\<^bsup>\\<^esup>\ i \ j cd\<^bsup>\\<^esup>\ i \ i = j\ by (metis assms cd_impl_icd_cd cd_trans icd_imp_cd icd_uniq) lemma same_ipd_stable: assumes \k cd\<^bsup>\\<^esup>\ i\ \k cd\<^bsup>\\<^esup>\ j\ \i \ipd (\ i) = ipd (\ k)\ shows \ipd (\ j) = ipd (\ k)\ proof - have jcdi: \j cd\<^bsup>\\<^esup>\ i\ by (metis is_cdi_def assms(1,2,3) cr_wn' le_antisym less_imp_le_nat) have 1: \ipd (\ j) pd\ \ k \ by (metis assms(2) ipd_pd_cd) have 2: \ipd (\ k) pd\ \ j \ by (metis assms(4) ipd_pd_cd jcdi) have 3: \ipd (\ k) pd\ (ipd (\ j))\ by (metis 2 IFC_def.is_cdi_def assms(1,2,4) atLeastAtMost_iff jcdi less_imp_le pd_node2 pd_pd_ipd) have 4: \ipd (\ j) pd\ (ipd (\ k))\ by (metis 1 2 IFC_def.is_ipd_def assms(2) cd_not_pd ipd_is_ipd jcdi pd_node2 ret_no_cd) show \?thesis\ using 3 4 pd_antisym by simp qed lemma icd_pd_intermediate': assumes icd: \i icd\<^bsup>\\<^esup>\ k\ and j: \k < j\ \j < i\ shows \\ i pd\ (\ j)\ using j proof (induction \i - j\ arbitrary: \j\ rule: less_induct) case (less j) have \\ i cd\<^bsup>\\<^esup>\ j\ using less.prems icd unfolding is_icdi_def by force moreover have \is_path \\ using icd by (metis is_icdi_def) moreover have \\ i \ return\ using icd by (metis is_icdi_def ret_no_cd) ultimately have \\ l. j \ l \ l \ i \ \ l = ipd (\ j)\ unfolding is_cdi_def using less.prems by auto then obtain l where l: \j \ l\ \l \ i\ \\ l = ipd (\ j)\ by blast hence lpd: \\ l pd\ (\ j)\ by (metis is_ipd_def \\ i \ return\ \is_path \\ ipd_is_ipd path_nodes term_path_stable) show \?case\ proof (cases) assume \l = i\ thus \?case\ using lpd by auto next assume \l \ i\ hence \l < i\ using l by simp moreover have \j \ l\ using l by (metis is_ipd_def \\ i \ return\ \is_path \\ ipd_is_ipd path_nodes term_path_stable) hence \j < l\ using l by simp moreover hence \i - l < i - j\ by (metis diff_less_mono2 less.prems(2)) moreover have \k < l\ by (metis l(1) less.prems(1) linorder_neqE_nat not_le order.strict_trans) ultimately have \\ i pd\ (\ l)\ using less.hyps by auto thus \?case\ using lpd by (metis pd_trans) qed qed lemma icd_pd_intermediate: assumes icd: \i icd\<^bsup>\\<^esup>\ k\ and j: \k < j\ \j \ i\ shows \\ i pd\ (\ j)\ using assms icd_pd_intermediate'[OF assms(1,2)] apply (cases \j < i\,metis) by (metis is_icdi_def le_neq_trans path_nodes pd_refl) lemma no_icd_pd: assumes path: \is_path \\ and noicd: \\ l\n. \ k icd\<^bsup>\\<^esup>\ l\ and nk: \n \ k\ shows \\ k pd\ \ n\ proof cases assume \\ k = return\ thus \?thesis\ by (metis path path_nodes return_pd) next assume nret: \\ k \ return\ have nocd: \\ l. n\l \ \ k cd\<^bsup>\\<^esup>\ l\ proof fix l assume kcd: \k cd\<^bsup>\\<^esup>\ l\ and nl: \n \ l\ hence \(k - n) cd\<^bsup>\\n\<^esup>\ (l - n)\ using cd_path_shift[OF nl path] by simp hence \\ l. (k - n) icd\<^bsup>\\n\<^esup>\ l\ using excd_impl_exicd by blast then guess l' .. hence \k icd\<^bsup>\\<^esup>\ (l' + n)\ using icd_path_shift[of \n\ \l' + n\ \\\ \k\] path by auto thus \False\ using noicd by auto qed hence \\l. n \ l \ l \ j \ {l..k}. \ j = ipd (\ l)\ using path nret unfolding is_cdi_def by auto thus \?thesis\ using nk proof (induction \k - n\ arbitrary: \n\ rule: less_induct,cases) case (less n) assume \n = k\ thus \?case\ using pd_refl path path_nodes by auto next case (less n) assume \n \ k\ hence nk: \n < k\ using less(3) by auto with less(2) obtain j where jnk: \j \ {n..k}\ and ipdj: \\ j = ipd (\ n)\ by blast have nretn: \\ n \ return\ using nk nret term_path_stable path by auto with ipd_is_ipd path path_nodes is_ipd_def ipdj have jpdn: \\ j pd\ \ n\ by auto show \?case\ proof cases assume \j = k\ thus \?case\ using jpdn by simp next assume \j \ k\ hence jk: \j < k\ using jnk by auto have \j \ n\ using ipdj by (metis ipd_not_self nretn path path_nodes) hence nj: \n < j\ using jnk by auto have *: \k - j < k - n\ using jk nj by auto with less(1)[OF *] less(2) jk nj have \\ k pd\ \ j\ by auto thus \?thesis\ using jpdn pd_trans by metis qed qed qed lemma first_pd_no_cd: assumes path: \is_path \\ and pd: \\ n pd\ \ 0\ and first: \\ l < n. \ l \ \ n\ shows \\ l. \ n cd\<^bsup>\\<^esup>\ l\ proof (rule ccontr, goal_cases) case 1 then obtain l where ncdl: \n cd\<^bsup>\\<^esup>\ l\ by blast hence ln: \l < n\ using is_cdi_def by auto have \\ \ n pd\ \ l\ using ncdl cd_not_pd by (metis ln first) then obtain \' n' where path': \is_path \'\ and \0: \\' 0 = \ l\ and \n: \\' n' = return\ and not\n: \\ j\ n'. \' j \ \ n\ unfolding is_pd_def using path path_nodes by auto let \?\\ = \\@\<^bsup>l\<^esup> \'\ have \is_path ?\\ by (metis \0 path path' path_cons) moreover have \?\ 0 = \ 0\ by auto moreover have \?\ (n' + l) = return\ using \0 \n by auto ultimately obtain j where j: \j \ n' + l\ and jn: \?\ j = \ n\ using pd unfolding is_pd_def by blast show \False\ proof cases assume \j \ l\ thus \False\ using jn first ln by auto next assume \\ j \ l\ thus \False\ using j jn not\n by auto qed qed lemma first_pd_no_icd: assumes path: \is_path \\ and pd: \\ n pd\ \ 0\ and first: \\ l < n. \ l \ \ n\ shows \\ l. \ n icd\<^bsup>\\<^esup>\ l\ by (metis first first_pd_no_cd icd_imp_cd path pd) lemma path_nret_ex_nipd: assumes \is_path \\ \\ i. \ i \ return\ shows \\ i. (\ j\i. (\ k>j. \ k \ ipd (\ j)))\ proof(rule, rule ccontr) fix i assume \\ (\j\i. \ k>j. \ k \ ipd (\ j))\ hence *: \\ j\i. (\k>j. \ k = ipd (\ j))\ by blast have \\ j. (\k>j. (\\i) k = ipd ((\\i) j))\ proof fix j have \i + j \ i\ by auto then obtain k where k: \k>i+j\ \\ k = ipd (\ (i+j))\ using * by blast hence \(\\i) (k - i) = ipd ((\\i) j)\ by auto moreover have \k - i > j\ using k by auto ultimately show \\k>j. (\\i) k = ipd ((\\i) j)\ by auto qed moreover have \is_path (\\i)\ using assms(1) path_path_shift by simp ultimately obtain k where \(\\i) k = return\ using all_ipd_imp_ret by blast thus \False\ using assms(2) by auto qed lemma path_nret_ex_all_cd: assumes \is_path \\ \\ i. \ i \ return\ shows \\ i. (\ j\i. (\ k>j. k cd\<^bsup>\\<^esup>\ j))\ unfolding is_cdi_def using assms path_nret_ex_nipd[OF assms] by (metis atLeastAtMost_iff ipd_not_self linorder_neqE_nat not_le path_nodes) lemma path_nret_inf_all_cd: assumes \is_path \\ \\ i. \ i \ return\ shows \\ finite {j. \ k>j. k cd\<^bsup>\\<^esup>\ j}\ using unbounded_nat_set_infinite path_nret_ex_all_cd[OF assms] by auto lemma path_nret_inf_icd_seq: assumes path: \is_path \\ and nret: \\ i. \ i \ return\ obtains f where \\ i. f (Suc i) icd\<^bsup>\\<^esup>\ f i\ \range f = {i. \ j>i. j cd\<^bsup>\\<^esup>\ i}\ \\ (\i. f 0 cd\<^bsup>\\<^esup>\ i)\ proof - note path_nret_inf_all_cd[OF assms] then obtain f where ran: \range f = {j. \ k>j. k cd\<^bsup>\\<^esup>\ j}\ and asc: \\ i. f i < f (Suc i)\ using infinite_ascending by blast have mono: \\ i j. i < j \ f i < f j\ using asc by (metis lift_Suc_mono_less) { fix i have cd: \f (Suc i) cd\<^bsup>\\<^esup>\ f i\ using ran asc by auto have \f (Suc i) icd\<^bsup>\\<^esup>\ f i\ proof (rule ccontr) assume \\ f (Suc i) icd\<^bsup>\\<^esup>\ f i\ then obtain m where im: \f i < m\ and mi: \ m < f (Suc i)\ and cdm: \f (Suc i) cd\<^bsup>\\<^esup>\ m\ unfolding is_icdi_def using assms(1) cd by auto have \\ k>m. k cd\<^bsup>\\<^esup>\m\ proof (rule,rule,cases) fix k assume \f (Suc i) < k\ hence \k cd\<^bsup>\\<^esup>\ f (Suc i)\ using ran by auto thus \k cd\<^bsup>\\<^esup>\ m\ using cdm cd_trans by metis next fix k assume mk: \m < k\ and \\ f (Suc i) < k\ hence ik: \k \ f (Suc i)\ by simp thus \k cd\<^bsup>\\<^esup>\ m\ using cdm by (metis cdi_prefix mk) qed hence \m \ range f\ using ran by blast then obtain j where m: \m = f j\ by blast show \False\ using im mi mono unfolding m by (metis Suc_lessI le_less not_le) qed } moreover { fix m assume cdm: \f 0 cd\<^bsup>\\<^esup>\ m\ have \\ k>m. k cd\<^bsup>\\<^esup>\m\ proof (rule,rule,cases) fix k assume \f 0 < k\ hence \k cd\<^bsup>\\<^esup>\ f 0\ using ran by auto thus \k cd\<^bsup>\\<^esup>\ m\ using cdm cd_trans by metis next fix k assume mk: \m < k\ and \\ f 0 < k\ hence ik: \k \ f 0\ by simp thus \k cd\<^bsup>\\<^esup>\ m\ using cdm by (metis cdi_prefix mk) qed hence \m \ range f\ using ran by blast then obtain j where m: \m = f j\ by blast hence fj0: \f j < f 0\ using cdm m is_cdi_def by auto hence \0 < j\ by (metis less_irrefl neq0_conv) hence \False\ using fj0 mono by fastforce } ultimately show \thesis\ using that ran by blast qed lemma cdi_iff_no_strict_pd: \i cd\<^bsup>\\<^esup>\ k \ is_path \ \ k < i \ \ i \ return \ (\ j \ {k..i}. \ (\ k, \ j) \ pdt)\ proof assume cd:\i cd\<^bsup>\\<^esup>\ k\ have 1: \is_path \ \ k < i \ \ i \ return\ using cd unfolding is_cdi_def by auto have 2: \\ j \ {k..i}. \ (\ k, \ j) \ pdt\ proof (rule ccontr) assume \ \ (\j\{k..i}. (\ k, \ j) \ pdt)\ then obtain j where \j \ {k..i}\ and \(\ k, \ j) \ pdt\ by auto hence \\ j \ \ k\ and \\ j pd\ \ k\ unfolding pdt_def by auto thus \False\ using path_pd_ipd by (metis \j \ {k..i}\ atLeastAtMost_iff cd cd_not_pd cdi_prefix le_eq_less_or_eq) qed show \is_path \ \ k < i \ \ i \ return \ (\ j \ {k..i}. \ (\ k, \ j) \ pdt)\ using 1 2 by simp next assume \is_path \ \ k < i \ \ i \ return \ (\ j \ {k..i}. \ (\ k, \ j) \ pdt)\ thus \i cd\<^bsup>\\<^esup>\ k\ by (metis ipd_in_pdt term_path_stable less_or_eq_imp_le not_cd_impl_ipd path_nodes) qed subsection \Facts about Control Slices\ lemma last_cs: \last (cs\<^bsup>\\<^esup> i) = \ i\ by auto lemma cs_not_nil: \cs\<^bsup>\\<^esup> n \ []\ by (auto) lemma cs_return: assumes \\ n = return\ shows \cs\<^bsup>\\<^esup> n = [\ n]\ by (metis assms cs.elims icd_imp_cd ret_no_cd) lemma cs_0[simp]: \cs\<^bsup>\\<^esup> 0 = [\ 0]\ using is_icdi_def is_cdi_def by auto lemma cs_inj: assumes \is_path \\ \\ n \ return\ \cs\<^bsup>\\<^esup> n = cs\<^bsup>\\<^esup> n'\ shows \n = n'\ using assms proof (induction \cs\<^bsup>\\<^esup> n\ arbitrary: \\\ \n\ \n'\ rule:rev_induct) case Nil hence \False\ using cs_not_nil by metis thus \?case\ by simp next case (snoc x xs \ n n') show \?case\ proof (cases \xs\) case Nil hence *: \\ (\ k. n icd\<^bsup>\\<^esup>\k)\ using snoc(2) cs_not_nil by (auto,metis append1_eq_conv append_Nil cs_not_nil) moreover have \[x] = cs\<^bsup>\\<^esup> n'\ using Nil snoc by auto hence **: \\ (\ k. n' icd\<^bsup>\\<^esup>\k)\ using cs_not_nil by (auto,metis append1_eq_conv append_Nil cs_not_nil) ultimately have \\ k. \ n cd\<^bsup>\\<^esup>\ k\ \\ k. \ n' cd\<^bsup>\\<^esup>\ k\ using excd_impl_exicd by auto blast+ moreover hence \\ n = \ n'\ using snoc(5,2) by auto (metis * ** list.inject) ultimately show \n = n'\ using other_claim' snoc by blast next case (Cons y ys) hence *: \\ k. n icd\<^bsup>\\<^esup>\k\ using snoc(2) by auto (metis append_is_Nil_conv list.distinct(1) list.inject) then obtain k where k: \n icd\<^bsup>\\<^esup>\k\ by auto have \k = (THE k . n icd\<^bsup>\\<^esup>\ k)\ using k by (metis icd_is_the_icd) hence xsk: \xs = cs\<^bsup>\\<^esup> k\ using * k snoc(2) unfolding cs.simps[of \\\ \n\] by auto have **: \\ k. n' icd\<^bsup>\\<^esup>\k\ using snoc(2)[unfolded snoc(5)] by auto (metis Cons append1_eq_conv append_Nil list.distinct(1)) then obtain k' where k': \n' icd\<^bsup>\\<^esup>\ k'\ by auto hence \k' = (THE k . n' icd\<^bsup>\\<^esup>\ k)\ using k' by (metis icd_is_the_icd) hence xsk': \xs = cs\<^bsup>\\<^esup> k'\ using ** k' snoc(5,2) unfolding cs.simps[of \\\ \n'\] by auto hence \cs\<^bsup>\\<^esup> k = cs\<^bsup>\\<^esup> k'\ using xsk by simp moreover have kn: \k < n\ using k by (metis is_icdi_def is_cdi_def) hence \\ k \ return\ using snoc by (metis term_path_stable less_imp_le) ultimately have kk'[simp]: \k' = k\ using snoc(1) xsk snoc(3) by metis have nk0: \n - k icd\<^bsup>\\k\<^esup>\ 0\ \n' - k icd\<^bsup>\\k\<^esup>\ 0\ using k k' icd_path_shift0 snoc(3) by auto moreover have nkr: \(\\k)(n-k) \ return\ using snoc(4) kn by auto moreover have \is_path (\\k)\ by (metis path_path_shift snoc.prems(1)) moreover have kn': \k < n'\ using k' kk' by (metis is_icdi_def is_cdi_def) have \\ n = \ n'\ using snoc(5) * ** by auto hence \(\\k)(n-k) = (\\k)(n'-k)\ using kn kn' by auto ultimately have \n - k = n' - k\ using other_claim by auto thus \n = n'\ using kn kn' by auto qed qed lemma cs_cases: fixes \ i obtains (base) \cs\<^bsup>\\<^esup> i = [\ i]\ and \\ k. \ i cd\<^bsup>\\<^esup>\ k\ | (depend) k where \cs\<^bsup>\\<^esup> i = (cs\<^bsup>\\<^esup> k)@[\ i]\ and \i icd\<^bsup>\\<^esup>\ k\ proof cases assume *: \\ k. i icd\<^bsup>\\<^esup>\ k\ then obtain k where k: \i icd\<^bsup>\\<^esup>\ k\ .. hence \k = (THE k. i icd\<^bsup>\\<^esup>\ k)\ by (metis icd_is_the_icd) hence \cs\<^bsup>\\<^esup> i = (cs\<^bsup>\\<^esup> k)@[\ i]\ using * by auto with k that show \thesis\ by simp next assume *: \\ (\ k. i icd\<^bsup>\\<^esup>\ k)\ hence \\ k. \ i cd\<^bsup>\\<^esup>\ k\ by (metis excd_impl_exicd) moreover have \cs\<^bsup>\\<^esup> i = [\ i]\ using * by auto ultimately show \thesis\ using that by simp qed lemma cs_length_one: assumes \length (cs\<^bsup>\\<^esup> i) = 1\ shows \cs\<^bsup>\\<^esup> i = [\ i]\ and \\ k. \ i cd\<^bsup>\\<^esup>\ k\ apply (cases \i\ \\\ rule: cs_cases) using assms cs_not_nil apply auto apply (cases \i\ \\\ rule: cs_cases) using assms cs_not_nil by auto lemma cs_length_g_one: assumes \length (cs\<^bsup>\\<^esup> i) \ 1\ obtains k where \cs\<^bsup>\\<^esup> i = (cs\<^bsup>\\<^esup> k)@[\ i]\ and \i icd\<^bsup>\\<^esup>\ k\ apply (cases \i\ \\\ rule: cs_cases) using assms cs_not_nil by auto lemma claim: assumes path: \is_path \\ \is_path \'\ and ii: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ and jj: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ and bl: \butlast (cs\<^bsup>\\<^esup> i) = butlast (cs\<^bsup>\\<^esup> j)\ and nret: \\ i \ return\ and ilj: \i < j\ shows \i' < j'\ proof (cases ) assume *: \length (cs\<^bsup>\\<^esup> i) = 1\ hence **: \length (cs\<^bsup>\\<^esup> i) = 1\ \length (cs\<^bsup>\\<^esup> j) = 1\ \length (cs\<^bsup>\'\<^esup> i') = 1\ \length (cs\<^bsup>\'\<^esup> j') = 1\ apply metis apply (metis "*" bl butlast.simps(2) butlast_snoc cs_length_g_one cs_length_one(1) cs_not_nil) apply (metis "*" ii) by (metis "*" bl butlast.simps(2) butlast_snoc cs_length_g_one cs_length_one(1) cs_not_nil jj) then obtain \cs\<^bsup>\\<^esup> i = [\ i]\ \cs\<^bsup>\\<^esup> j = [\ j]\ \cs\<^bsup>\'\<^esup> j' = [\' j']\ \cs\<^bsup>\'\<^esup> i'= [\' i']\ \\ k. \ j cd\<^bsup>\\<^esup>\ k\ \\ k. \ i' cd\<^bsup>\'\<^esup>\ k\ \\ k. \ j' cd\<^bsup>\'\<^esup>\ k\ by (metis cs_length_one ** ) moreover hence \\ i = \' i'\ \\ j = \' j'\ using assms by auto ultimately show \i' < j'\ using nret ilj path claim'' by blast next assume *: \length (cs\<^bsup>\\<^esup> i) \ 1\ hence **: \length (cs\<^bsup>\\<^esup> i) \ 1\ \length (cs\<^bsup>\\<^esup> j) \ 1\ \length (cs\<^bsup>\'\<^esup> i') \ 1\ \length (cs\<^bsup>\'\<^esup> j') \ 1\ apply metis apply (metis "*" bl butlast.simps(2) butlast_snoc cs_length_g_one cs_length_one(1) cs_not_nil) apply (metis "*" ii) by (metis "*" bl butlast.simps(2) butlast_snoc cs_length_g_one cs_length_one(1) cs_not_nil jj) obtain k l k' l' where ***: \cs\<^bsup>\\<^esup> i = (cs\<^bsup>\\<^esup> k)@[\ i]\ \cs\<^bsup>\\<^esup> j = (cs\<^bsup>\\<^esup> l)@[\ j]\ \cs\<^bsup>\'\<^esup> i' = (cs\<^bsup>\'\<^esup> k')@[\' i']\ \cs\<^bsup>\'\<^esup> j' = (cs\<^bsup>\'\<^esup> l')@[\' j']\ and icds: \i icd\<^bsup>\\<^esup>\ k\ \j icd\<^bsup>\\<^esup>\ l\ \i' icd\<^bsup>\'\<^esup>\ k'\ \j' icd\<^bsup>\'\<^esup>\ l'\ by (metis ** cs_length_g_one) hence \cs\<^bsup>\\<^esup> k = cs\<^bsup>\\<^esup> l\ \cs\<^bsup>\'\<^esup> k' = cs\<^bsup>\'\<^esup> l'\ using assms by auto moreover have \\ k \ return\ \\' k' \ return\ using nret apply (metis is_icdi_def icds(1) is_cdi_def term_path_stable less_imp_le) by (metis is_cdi_def is_icdi_def icds(3) term_path_stable less_imp_le) ultimately have lk[simp]: \l = k\ \l' = k'\ using path cs_inj by auto let \?\\ = \\ \ k\ let \?\'\ = \\'\k'\ have \i-k icd\<^bsup>?\\<^esup>\ 0\ \j-k icd\<^bsup>?\\<^esup>\ 0\ \i'-k' icd\<^bsup>?\'\<^esup>\ 0\ \j'-k' icd\<^bsup>?\'\<^esup>\ 0\ using icd_path_shift0 path icds by auto moreover have ki: \k < i\ using icds by (metis is_icdi_def is_cdi_def) hence \i-k < j-k\ by (metis diff_is_0_eq diff_less_mono ilj nat_le_linear order.strict_trans) moreover have \i: \\ i = \' i'\ \\ j = \' j'\ using assms *** by auto have \k' < i'\ \k' < j'\ using icds unfolding lk by (metis is_cdi_def is_icdi_def)+ hence \?\ (i-k) = ?\' (i'-k')\ \?\ (j-k) = ?\' (j'-k')\ using \i ki ilj by auto moreover have \?\ (i-k) \ return\ using nret ki by auto moreover have \is_path ?\\ \is_path ?\'\ using path path_path_shift by auto ultimately have \i'-k' < j' - k'\ using claim' by blast thus \i' < j'\ by (metis diff_is_0_eq diff_less_mono less_nat_zero_code linorder_neqE_nat nat_le_linear) qed lemma cs_split': assumes \cs\<^bsup>\\<^esup> i = xs@[x,x']@ys\ shows \\ m. cs\<^bsup>\\<^esup> m = xs@[x] \ i cd\<^bsup>\\<^esup>\ m\ using assms proof (induction \ys\ arbitrary: \i\ rule:rev_induct ) case (snoc y ys) hence \length (cs\<^bsup>\\<^esup> i) \ 1\ by auto then obtain i' where \cs\<^bsup>\\<^esup> i = (cs\<^bsup>\\<^esup> i') @ [\ i]\ and *: \i icd\<^bsup>\\<^esup>\ i'\ using cs_length_g_one[of \\\ \i\] by metis hence \cs\<^bsup>\\<^esup> i' = xs@[x,x']@ys\ using snoc(2) by (metis append1_eq_conv append_assoc) then obtain m where **: \cs\<^bsup>\\<^esup> m = xs @ [x]\ and \i' cd\<^bsup>\\<^esup>\ m\ using snoc(1) by blast hence \i cd\<^bsup>\\<^esup>\ m\ using * cd_trans by (metis is_icdi_def) with ** show \?case\ by blast next case Nil hence \length (cs\<^bsup>\\<^esup> i) \ 1\ by auto then obtain i' where a: \cs\<^bsup>\\<^esup> i = (cs\<^bsup>\\<^esup> i') @ [\ i]\ and *: \i icd\<^bsup>\\<^esup>\ i'\ using cs_length_g_one[of \\\ \i\] by metis have \cs\<^bsup>\\<^esup> i = (xs@[x])@[x']\ using Nil by auto hence \cs\<^bsup>\\<^esup> i' = xs@[x]\ using append1_eq_conv a by metis thus \?case\ using * unfolding is_icdi_def by blast qed lemma cs_split: assumes \cs\<^bsup>\\<^esup> i = xs@[x]@ys@[\ i]\ shows \\ m. cs\<^bsup>\\<^esup> m = xs@[x] \ i cd\<^bsup>\\<^esup>\ m\ proof - obtain x' ys' where \ys@[\ i] = [x']@ys'\ by (metis append_Cons append_Nil neq_Nil_conv) thus \?thesis\ using cs_split'[of \\\ \i\ \xs\ \x\ \x'\ \ys'\] assms by auto qed lemma cs_less_split: assumes \xs \ ys\ obtains a as where \ys = xs@a#as\ using assms unfolding cs_less.simps apply auto by (metis Cons_nth_drop_Suc append_take_drop_id) lemma cs_select_is_cs: assumes \is_path \\ \xs \ Nil\ \xs \ cs\<^bsup>\\<^esup> k\ shows \cs\<^bsup>\\<^esup> (\\xs) = xs\ \k cd\<^bsup>\\<^esup>\ (\\xs)\proof - obtain b bs where b: \cs\<^bsup>\\<^esup> k = xs@b#bs\ using assms cs_less_split by blast obtain a as where a: \xs = as@[a]\ using assms by (metis rev_exhaust) have \cs\<^bsup>\\<^esup> k = as@[a,b]@bs\ using a b by auto then obtain k' where csk: \cs\<^bsup>\\<^esup> k' = xs\ and is_cd: \k cd\<^bsup>\\<^esup>\ k'\ using cs_split' a by blast hence nret: \\ k' \ return\ by (metis is_cdi_def term_path_stable less_imp_le) show a: \cs\<^bsup>\\<^esup> (\\xs) = xs\ unfolding cs_select_def using cs_inj[OF assms(1) nret] csk the_equality[of _ \k'\] by (metis (mono_tags)) show \k cd\<^bsup>\\<^esup>\ (\\xs)\ unfolding cs_select_def by (metis a assms(1) cs_inj cs_select_def csk is_cd nret) qed lemma cd_in_cs: assumes \n cd\<^bsup>\\<^esup>\ m\ shows \\ ns. cs\<^bsup>\\<^esup> n = (cs\<^bsup>\\<^esup> m) @ ns @[\ n]\ using assms proof (induction rule: cd_induct) case (base n) thus \?case\ by (metis append_Nil cs.simps icd_is_the_icd) next case (IS k n) hence \cs\<^bsup>\\<^esup> n = cs\<^bsup>\\<^esup> k @ [\ n]\ by (metis cs.simps icd_is_the_icd) thus \?case\ using IS by force qed lemma butlast_cs_not_cd: assumes \butlast (cs\<^bsup>\\<^esup> m) = butlast (cs\<^bsup>\\<^esup> n)\ shows \\ m cd\<^bsup>\\<^esup>\n\ by (metis append_Cons append_Nil append_assoc assms cd_in_cs cs_not_nil list.distinct(1) self_append_conv snoc_eq_iff_butlast) lemma wn_cs_butlast: assumes \butlast (cs\<^bsup>\\<^esup> m) = butlast (cs\<^bsup>\\<^esup> n)\ \i cd\<^bsup>\\<^esup>\ m\ \j cd\<^bsup>\\<^esup>\ n\ \m shows \i proof (rule ccontr) assume \\ i < j\ moreover have \\ n cd\<^bsup>\\<^esup>\ m\ by (metis assms(1) butlast_cs_not_cd) ultimately have \n \ m\ using assms(2,3) cr_wn'' by auto thus \False\ using assms(4) by auto qed text \This is the central theorem making the control slice suitable for matching indices between executions.\ theorem cs_order: assumes path: \is_path \\ \is_path \'\ and csi: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ and csj: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ and nret: \\ i \ return\ and ilj: \i < j\ shows \i' proof - have \cs\<^bsup>\\<^esup> i \ cs\<^bsup>\\<^esup> j\ using cs_inj[OF path(1) nret] ilj by blast moreover have \cs\<^bsup>\\<^esup> i \ Nil\ \cs\<^bsup>\\<^esup> j \ Nil\ by (metis cs_not_nil)+ ultimately show \?thesis\ proof (cases rule: list_neq_prefix_cases) case (diverge xs x x' ys ys') note csx = \cs\<^bsup>\\<^esup> i = xs @ [x] @ ys\ note csx' = \cs\<^bsup>\\<^esup> j = xs @ [x'] @ ys'\ note xx = \x \ x'\ show \i' < j'\ proof (cases \ys\) assume ys: \ys = Nil\ show \?thesis\ proof (cases \ys'\) assume ys': \ys' = Nil\ have cs: \cs\<^bsup>\\<^esup> i = xs @ [x]\ \cs\<^bsup>\\<^esup> j = xs @ [x']\ by (metis append_Nil2 csx ys, metis append_Nil2 csx' ys') hence bl: \butlast (cs\<^bsup>\\<^esup> i) = butlast (cs\<^bsup>\\<^esup> j)\ by auto show \i' < j'\ using claim[OF path csi csj bl nret ilj] . next fix y' zs' assume ys': \ys' = y'#zs'\ have cs: \cs\<^bsup>\\<^esup> i = xs @ [x]\ \cs\<^bsup>\\<^esup> j = xs @ [x',y']@ zs'\ by (metis append_Nil2 csx ys, metis append_Cons append_Nil csx' ys') obtain n where n: \cs\<^bsup>\\<^esup> n = xs@[x']\ and jn: \j cd\<^bsup>\\<^esup>\ n\ using cs cs_split' by blast obtain n' where n': \cs\<^bsup>\'\<^esup> n' = xs@[x']\ and jn': \j' cd\<^bsup>\'\<^esup>\ n'\ using cs cs_split' unfolding csj by blast have csn : \cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and bl: \butlast (cs\<^bsup>\\<^esup> i) = butlast (cs\<^bsup>\\<^esup> n)\ using n n' cs by auto hence bl': \butlast (cs\<^bsup>\'\<^esup> i') = butlast (cs\<^bsup>\'\<^esup> n')\ using csi by auto have notcd: \\ i cd\<^bsup>\\<^esup>\ n\ by (metis butlast_cs_not_cd bl) have nin: \i \ n\ using cs n xx by auto have iln: \i < n\ apply (rule ccontr) using cr_wn'[OF jn notcd] nin ilj by auto note claim[OF path csi csn bl nret iln] hence \i' < n'\ . thus \i' < j'\ using jn' unfolding is_cdi_def by auto qed next fix y zs assume ys: \ys = y#zs\ show \?thesis\ proof (cases \ys'\) assume ys' : \ys' = Nil\ have cs: \cs\<^bsup>\\<^esup> i = xs @ [x,y]@zs\ \cs\<^bsup>\\<^esup> j = xs @ [x']\ by (metis append_Cons append_Nil csx ys, metis append_Nil2 csx' ys') obtain n where n: \cs\<^bsup>\\<^esup> n = xs@[x]\ and jn: \i cd\<^bsup>\\<^esup>\ n\ using cs cs_split' by blast obtain n' where n': \cs\<^bsup>\'\<^esup> n' = xs@[x]\ and jn': \i' cd\<^bsup>\'\<^esup>\ n'\ using cs cs_split' unfolding csi by blast have csn : \cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and bl: \butlast (cs\<^bsup>\\<^esup> n) = butlast (cs\<^bsup>\\<^esup> j)\ using n n' cs by auto hence bl': \butlast (cs\<^bsup>\'\<^esup> j') = butlast (cs\<^bsup>\'\<^esup> n')\ using csj by auto have notcd: \\ j' cd\<^bsup>\'\<^esup>\ n'\ by (metis butlast_cs_not_cd bl') have nin: \n < i\ using jn unfolding is_cdi_def by auto have nlj: \n < j\ using nin ilj by auto note claim[OF path csn csj bl _ nlj] hence nj': \n' < j'\ using term_path_stable[OF path(1) _] less_imp_le nin nret by auto show \i' < j'\ apply(rule ccontr) using cdi_prefix[OF jn' nj'] notcd by auto next fix y' zs' assume ys' : \ys' = y'#zs'\ have cs: \cs\<^bsup>\\<^esup> i = xs@[x,y]@zs\ \cs\<^bsup>\\<^esup> j = xs@[x',y']@zs'\ by (metis append_Cons append_Nil csx ys,metis append_Cons append_Nil csx' ys') have neq: \cs\<^bsup>\\<^esup> i \ cs\<^bsup>\\<^esup> j\ using cs_inj path nret ilj by blast obtain m where m: \cs\<^bsup>\\<^esup> m = xs@[x]\ and im: \i cd\<^bsup>\\<^esup>\ m\ using cs cs_split' by blast obtain n where n: \cs\<^bsup>\\<^esup> n = xs@[x']\ and jn: \j cd\<^bsup>\\<^esup>\ n\ using cs cs_split' by blast obtain m' where m': \cs\<^bsup>\'\<^esup> m' = xs@[x]\ and im': \i' cd\<^bsup>\'\<^esup>\ m'\ using cs cs_split' unfolding csi by blast obtain n' where n': \cs\<^bsup>\'\<^esup> n' = xs@[x']\ and jn': \j' cd\<^bsup>\'\<^esup>\ n'\ using cs cs_split' unfolding csj by blast have \m \ n\ using ilj m n wn_cs_butlast[OF _ jn im] by force moreover have \m \ n\ using m n xx by (metis last_snoc) ultimately have mn: \m < n\ by auto moreover have \\ m \ return\ by (metis last_cs last_snoc m mn n path(1) term_path_stable xx less_imp_le) moreover have \butlast (cs\<^bsup>\\<^esup> m) = butlast (cs\<^bsup>\\<^esup> n)\ \cs\<^bsup>\\<^esup> m = cs\<^bsup>\'\<^esup> m'\ \cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ using m n n' m' by auto ultimately have \m' < n'\ using claim path by blast thus \i' < j'\ using m' n' im' jn' wn_cs_butlast by (metis butlast_snoc) qed qed next case (prefix1 xs) note pfx = \cs\<^bsup>\\<^esup> i = cs\<^bsup>\\<^esup> j @ xs\ note xs = \xs \ []\ obtain a as where \xs = a#as\ using xs by (metis list.exhaust) moreover obtain bs b where bj: \cs\<^bsup>\\<^esup> j = bs@[b]\ using cs_not_nil by (metis rev_exhaust) ultimately have \cs\<^bsup>\\<^esup> i = bs@[b,a]@as\ using pfx by auto then obtain m where \cs\<^bsup>\\<^esup> m = bs@[b]\ and cdep: \i cd\<^bsup>\\<^esup>\ m\ using cs_split' by blast hence mi: \m = j\ using bj cs_inj by (metis is_cdi_def term_path_stable less_imp_le) hence \i cd\<^bsup>\\<^esup>\ j\ using cdep by auto hence \False\ using ilj unfolding is_cdi_def by auto thus \i' < j'\ .. next case (prefix2 xs) have pfx : \cs\<^bsup>\'\<^esup> i' @ xs = cs\<^bsup>\'\<^esup> j'\ using prefix2 csi csj by auto note xs = \xs \ []\ obtain a as where \xs = a#as\ using xs by (metis list.exhaust) moreover obtain bs b where bj: \cs\<^bsup>\'\<^esup> i' = bs@[b]\ using cs_not_nil by (metis rev_exhaust) ultimately have \cs\<^bsup>\'\<^esup> j' = bs@[b,a]@as\ using pfx by auto then obtain m where \cs\<^bsup>\'\<^esup> m = bs@[b]\ and cdep: \j' cd\<^bsup>\'\<^esup>\ m\ using cs_split' by blast hence mi: \m = i'\ using bj cs_inj by (metis is_cdi_def term_path_stable less_imp_le) hence \j' cd\<^bsup>\'\<^esup>\ i'\ using cdep by auto thus \i' < j'\ unfolding is_cdi_def by auto qed qed lemma cs_order_le: assumes path: \is_path \\ \is_path \'\ and csi: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ and csj: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ and nret: \\ i \ return\ and ilj: \i \ j\ shows \i'\j'\ proof cases assume \i < j\ with cs_order[OF assms(1,2,3,4,5)] show \?thesis\ by simp next assume \\ i < j\ hence \i = j\ using ilj by simp hence csij: \cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\'\<^esup> j'\ using csi csj by simp have nret': \\' i' \ return\ using nret last_cs csi by metis show \?thesis\ using cs_inj[OF path(2) nret' csij] by simp qed lemmas cs_induct[case_names cs] = cs.induct lemma icdi_path_swap: assumes \is_path \'\ \j icd\<^bsup>\\<^esup>\k\ \\ =\<^bsub>j\<^esub> \'\ shows \j icd\<^bsup>\'\<^esup>\k\ using assms unfolding eq_up_to_def is_icdi_def is_cdi_def by auto lemma icdi_path_swap_le: assumes \is_path \'\ \j icd\<^bsup>\\<^esup>\k\ \\ =\<^bsub>n\<^esub> \'\ \j \ n\ shows \j icd\<^bsup>\'\<^esup>\k\ by (metis assms icdi_path_swap eq_up_to_le) lemma cs_path_swap: assumes \is_path \\ \is_path \'\ \\ =\<^bsub>k\<^esub> \'\ shows \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k\ using assms(1,3) proof (induction \\\ \k\ rule:cs_induct,cases) case (cs \ k) let \?l\ = \(THE l. k icd\<^bsup>\\<^esup>\ l)\ assume *: \\l. k icd\<^bsup>\\<^esup>\ l\ have kicd: \k icd\<^bsup>\\<^esup>\ ?l\ by (metis "*" icd_is_the_icd) hence \?l < k\ unfolding is_cdi_def[of \k\ \\\ \?l\] is_icdi_def[of \k\ \\\ \?l\] by auto hence \\ i\?l. \ i = \' i\ using cs(2,3) unfolding eq_up_to_def by auto hence csl: \cs\<^bsup>\\<^esup> ?l = cs\<^bsup>\'\<^esup> ?l\ using cs(1,2) * unfolding eq_up_to_def by auto have kicd: \k icd\<^bsup>\\<^esup>\ ?l\ by (metis "*" icd_is_the_icd) hence csk: \cs\<^bsup>\\<^esup> k = cs\<^bsup>\\<^esup> ?l @ [\ k]\ using kicd by auto have kicd': \k icd\<^bsup>\'\<^esup>\ ?l\ using kicd icdi_path_swap[OF assms(2) _ cs(3)] by simp hence \?l = (THE l. k icd\<^bsup>\'\<^esup>\ l)\ by (metis icd_is_the_icd) hence csk': \cs\<^bsup>\'\<^esup> k = cs\<^bsup>\'\<^esup> ?l @ [\' k]\ using kicd' by auto have \\' k = \ k\ using cs(3) unfolding eq_up_to_def by auto with csl csk csk' show \?case\ by auto next case (cs \ k) assume *: \\ (\l. k icd\<^bsup>\\<^esup>\ l)\ hence csk: \cs\<^bsup>\\<^esup> k = [\ k]\ by auto have \\ (\l. k icd\<^bsup>\'\<^esup>\ l)\ apply (rule ccontr) using * icdi_path_swap_le[OF cs(2) _, of \k\ \\'\] cs(3) by (metis eq_up_to_sym le_refl) hence csk': \cs\<^bsup>\'\<^esup> k = [\' k]\ by auto with csk show \?case\ using cs(3) eq_up_to_apply by auto qed lemma cs_path_swap_le: assumes \is_path \\ \is_path \'\ \\ =\<^bsub>n\<^esub> \'\ \k \ n\ shows \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k\ by (metis assms cs_path_swap eq_up_to_le) lemma cs_path_swap_cd: assumes \is_path \\ and \is_path \'\ and \cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and \n cd\<^bsup>\\<^esup>\ k\ obtains k' where \n' cd\<^bsup>\'\<^esup>\ k'\ and \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ proof - from cd_in_cs[OF assms(4)] obtain ns where *: \cs\<^bsup>\\<^esup> n = cs\<^bsup>\\<^esup> k @ ns @ [\ n]\ by blast obtain xs x where csk: \cs\<^bsup>\\<^esup> k = xs @ [x]\ by (metis cs_not_nil rev_exhaust) have \\ n = \' n'\ using assms(3) last_cs by metis hence **: \cs\<^bsup>\'\<^esup> n' = xs@[x]@ns@[\' n']\ using * assms(3) csk by auto from cs_split[OF **] obtain k' where \cs\<^bsup>\'\<^esup> k' = xs @ [x]\ \n' cd\<^bsup>\'\<^esup>\ k'\ by blast thus \thesis\ using that csk by auto qed lemma path_ipd_swap: assumes \is_path \\ \\ k \ return\ \k < n\ obtains \' m where \is_path \'\ \\ =\<^bsub>n\<^esub> \'\ \k < m\ \\' m = ipd (\' k)\ \\ l \ {k..' l \ ipd (\' k)\ proof - obtain \' r where *: \\' 0 = \ n\ \is_path \'\ \\' r = return\ by (metis assms(1) path_nodes reaching_ret) let \?\\ = \\@\<^bsup>n\<^esup> \'\ have path: \is_path ?\\ and ret: \?\ (n + r) = return\ and equpto: \?\ =\<^bsub>n\<^esub> \\ using assms path_cons * path_append_eq_up_to by auto have \k: \?\ k = \ k\ by (metis assms(3) less_imp_le_nat path_append_def) obtain j where j: \k < j \ j \ (n + r) \ ?\ j = ipd (\ k)\ (is \?P j\ )by (metis \k assms(2) path path_ret_ipd ret) define m where m: \m \ LEAST m . ?P m\ have Pm: \?P m\ using LeastI[of \?P\ \j\] j m by auto hence km: \k < m\ \m \ (n + r)\ \?\ m = ipd (\ k)\ by auto have le: \\l. ?P l \ m \ l\ using Least_le[of \?P\] m by blast have \knipd: \?\ k \ ipd (\ k)\ by (metis \k assms(1) assms(2) ipd_not_self path_nodes) have nipd': \\l. k < l \ l < m \ ?\ l \ ipd (\ k)\ apply (rule ccontr) using le km(2) by force have \\ l \ {k.. l \ ipd(\ k)\ using \knipd nipd' by(auto, metis le_eq_less_or_eq,metis le_eq_less_or_eq) thus \thesis\ using that by (metis \k eq_up_to_sym km(1) km(3) path path_append_eq_up_to) qed lemma cs_sorted_list_of_cd': \cs\<^bsup>\\<^esup> k = map \ (sorted_list_of_set { i . k cd\<^bsup>\\<^esup>\ i}) @ [\ k]\ proof (induction \\\ \k\ rule: cs.induct, cases) case (1 \ k) assume \\ j. k icd\<^bsup>\\<^esup>\ j\ then guess j .. note j = this hence csj: \cs\<^bsup>\\<^esup> j = map \ (sorted_list_of_set {i. j cd\<^bsup>\\<^esup>\ i}) @ [\ j]\ by (metis "1.IH" icd_is_the_icd) have \{i. k cd\<^bsup>\\<^esup>\ i} = insert j {i. j cd\<^bsup>\\<^esup>\ i}\ using cdi_is_cd_icdi[OF j] by auto moreover have f: \finite {i. j cd\<^bsup>\\<^esup>\ i}\ unfolding is_cdi_def by auto moreover have \j \ {i. j cd\<^bsup>\\<^esup>\ i}\ unfolding is_cdi_def by auto ultimately have \sorted_list_of_set { i . k cd\<^bsup>\\<^esup>\ i} = insort j (sorted_list_of_set { i . j cd\<^bsup>\\<^esup>\ i})\ using sorted_list_of_set_insert by auto moreover have \\ x \ {i. j cd\<^bsup>\\<^esup>\ i}. x < j\ unfolding is_cdi_def by auto hence \\ x \ set (sorted_list_of_set {i. j cd\<^bsup>\\<^esup>\ i}). x < j\ by (simp add: f) ultimately have \sorted_list_of_set { i . k cd\<^bsup>\\<^esup>\ i} = (sorted_list_of_set { i . j cd\<^bsup>\\<^esup>\ i})@[j]\ using insort_greater by auto hence \cs\<^bsup>\\<^esup> j = map \ (sorted_list_of_set { i . k cd\<^bsup>\\<^esup>\ i})\ using csj by auto thus \?case\ by (metis icd_cs j) next case (1 \ k) assume *: \\ (\ j. k icd\<^bsup>\\<^esup>\ j)\ hence \cs\<^bsup>\\<^esup> k = [\ k]\ by (metis cs_cases) moreover have \{ i . k cd\<^bsup>\\<^esup>\ i} = {}\ by (auto, metis * excd_impl_exicd) ultimately show \?case\ by (metis append_Nil list.simps(8) sorted_list_of_set_empty) qed lemma cs_sorted_list_of_cd: \cs\<^bsup>\\<^esup> k = map \ (sorted_list_of_set ({ i . k cd\<^bsup>\\<^esup>\ i} \ {k}))\ proof - have le: \\ x \ {i. k cd\<^bsup>\\<^esup>\i}.\ y \ {k}. x < y\ unfolding is_cdi_def by auto have fin: \finite {i. k cd\<^bsup>\\<^esup>\i}\ \finite {k}\ unfolding is_cdi_def by auto show \?thesis\ unfolding cs_sorted_list_of_cd'[of \\\ \k\] sorted_list_of_set_append[OF fin le] by auto qed lemma cs_not_ipd: assumes \k cd\<^bsup>\\<^esup>\ j \ ipd (\ j) \ ipd (\ k)\ (is \?Q j\) shows \cs\<^bsup>\\<^esup> (GREATEST j. k cd\<^bsup>\\<^esup>\ j \ ipd (\ j) \ ipd (\ k)) = [n\cs\<^bsup>\\<^esup> k . ipd n \ ipd (\ k)]\ (is \cs\<^bsup>\\<^esup> ?j = filter ?P _\) proof - have csk: \cs\<^bsup>\\<^esup> k = map \ (sorted_list_of_set ({ i . k cd\<^bsup>\\<^esup>\ i } \ {k}))\ by (metis cs_sorted_list_of_cd) have csj: \cs\<^bsup>\\<^esup> ?j = map \ (sorted_list_of_set ({i. ?j cd\<^bsup>\\<^esup>\ i } \ {?j}))\ by (metis cs_sorted_list_of_cd) have bound: \\ j. k cd\<^bsup>\\<^esup>\ j \ ipd (\ j) \ ipd(\ k) \ j \ k\ unfolding is_cdi_def by simp have kcdj: \k cd\<^bsup>\\<^esup>\ ?j\ and ipd': \ipd (\ ?j) \ ipd(\ k)\ using GreatestI_nat[of \?Q\ \j\ \k\, OF assms] bound by auto have greatest: \\ j. k cd\<^bsup>\\<^esup>\ j \ ipd (\ j) \ ipd (\ k) \ j \ ?j\ using Greatest_le_nat[of \?Q\ _ \k\] bound by auto have less_not_ipdk: \\ j. k cd\<^bsup>\\<^esup>\ j \ j < ?j \ ipd (\ j) \ ipd (\ k)\ by (metis (lifting) ipd' kcdj same_ipd_stable) hence le_not_ipdk: \\ j. k cd\<^bsup>\\<^esup>\ j \ j \ ?j \ ipd (\ j) \ ipd (\ k)\ using kcdj ipd' by (case_tac \j = ?j\,auto) have *: \{j \ {i. k cd\<^bsup>\\<^esup>\i} \ {k}. ?P (\ j)} = insert ?j { i . ?j cd\<^bsup>\\<^esup>\ i} \ apply auto apply (metis (lifting, no_types) greatest cr_wn'' kcdj le_antisym le_refl) apply (metis kcdj) apply (metis ipd') apply (metis (full_types) cd_trans kcdj) apply (subgoal_tac \k cd\<^bsup>\\<^esup>\ x\) apply (metis (lifting, no_types) is_cdi_def less_not_ipdk) by (metis (full_types) cd_trans kcdj) have \finite ({i . k cd\<^bsup>\\<^esup>\ i} \ {k})\ unfolding is_cdi_def by auto note filter_sorted_list_of_set[OF this, of \?P o \\] hence \[n\cs\<^bsup>\\<^esup> k . ipd n \ ipd(\ k)] = map \ (sorted_list_of_set {j \ {i. k cd\<^bsup>\\<^esup>\i} \ {k}. ?P (\ j)})\ unfolding csk filter_map by auto also have \\ = map \ (sorted_list_of_set (insert ?j { i . ?j cd\<^bsup>\\<^esup>\ i}))\ unfolding * by auto also have \\ = cs\<^bsup>\\<^esup> ?j\ using csj by auto finally show \?thesis\ by metis qed lemma cs_ipd: assumes ipd: \\ m = ipd (\ k)\ \\ n \ {k.. n \ ipd (\ k)\ and km: \k < m\ shows \cs\<^bsup>\\<^esup> m = [n\cs\<^bsup>\\<^esup> k . ipd n \ \ m] @ [\ m]\ proof cases assume \\ j. m icd\<^bsup>\\<^esup>\ j\ then obtain j where jicd: \m icd\<^bsup>\\<^esup>\ j\ by blast hence *: \cs\<^bsup>\\<^esup> m = cs\<^bsup>\\<^esup> j @ [\ m]\ by (metis icd_cs) have j: \j = (GREATEST j. k cd\<^bsup>\\<^esup>\ j \ ipd (\ j) \ \ m)\ using jicd assms ipd_icd_greatest_cd_not_ipd by blast moreover have \ipd (\ j) \ ipd (\ k)\ by (metis is_cdi_def is_icdi_def is_ipd_def cd_not_pd ipd(1) ipd_is_ipd jicd path_nodes less_imp_le term_path_stable) moreover have \k cd\<^bsup>\\<^esup>\ j\ unfolding j by (metis (lifting, no_types) assms(3) cd_ipd_is_cd icd_imp_cd ipd(1) ipd(2) j jicd) ultimately have \cs\<^bsup>\\<^esup> j = [n\cs\<^bsup>\\<^esup> k . ipd n \ \ m]\ using cs_not_ipd[of \k\ \\\ \j\] ipd(1) by metis thus \?thesis\ using * by metis next assume noicd: \\ (\ j. m icd\<^bsup>\\<^esup>\ j)\ hence csm: \cs\<^bsup>\\<^esup> m = [\ m]\ by auto have \\j. k cd\<^bsup>\\<^esup>\j \ ipd(\ j) = \ m\ using cd_is_cd_ipd[OF km ipd] by (metis excd_impl_exicd noicd) hence *: \{j \ {i. k cd\<^bsup>\\<^esup>\ i} \ {k}. ipd (\ j) \ \ m} = {}\ using ipd(1) by auto have **: \((\n. ipd n \ \ m) o \) = (\n. ipd (\ n) \ \ m)\ by auto have fin: \finite ({i. k cd\<^bsup>\\<^esup>\ i} \ {k})\ unfolding is_cdi_def by auto note csk = cs_sorted_list_of_cd[of \\\ \k\] hence \[n\cs\<^bsup>\\<^esup> k . ipd n \ \ m] = [n\ (map \ (sorted_list_of_set ({i. k cd\<^bsup>\\<^esup>\ i} \ {k}))) . ipd n \ \ m]\ by simp also have \\ = map \ [n <- sorted_list_of_set ({i. k cd\<^bsup>\\<^esup>\ i} \ {k}). ipd (\ n) \ \ m]\ by (auto simp add: filter_map **) also have \\ = []\ unfolding * filter_sorted_list_of_set[OF fin, of \\n. ipd (\ n) \ \ m\] by auto finally show \?thesis\ using csm by (metis append_Nil) qed lemma converged_ipd_same_icd: assumes path: \is_path \\ \is_path \'\ and converge: \l < m\ \cs\<^bsup>\\<^esup> m = cs\<^bsup>\'\<^esup> m'\ and csk: \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ and icd: \l icd\<^bsup>\\<^esup>\ k\ and suc: \\ (Suc k) = \' (Suc k')\ and ipd: \\' m' = ipd (\ k)\ \\ n \ {k'..' n \ ipd (\ k)\ shows \\l'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ proof cases assume l: \l = Suc k\ hence \Suc k cd\<^bsup>\\<^esup>\ k\ using icd by (metis is_icdi_def) hence \\ (Suc k) \ ipd (\ k)\ unfolding is_cdi_def by auto hence \\' (Suc k') \ ipd (\' k')\ by (metis csk last_cs suc) moreover have \\' (Suc k') \ return\ by (metis \Suc k cd\<^bsup>\\<^esup>\ k\ ret_no_cd suc) ultimately have \Suc k' cd\<^bsup>\'\<^esup>\ k'\ unfolding is_cdi_def using path(2) apply auto by (metis ipd_not_self le_Suc_eq le_antisym path_nodes term_path_stable) hence \Suc k' icd\<^bsup>\'\<^esup>\ k'\ unfolding is_icdi_def using path(2) by fastforce hence \cs\<^bsup>\'\<^esup> (Suc k') = cs\<^bsup>\'\<^esup> k' @[\' (Suc k')]\ using icd_cs by auto moreover have \cs\<^bsup>\\<^esup> l = cs\<^bsup>\\<^esup> k @ [\ l]\ using icd icd_cs by auto ultimately have \cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> (Suc k')\ by (metis csk l suc) thus \?thesis\ by blast next assume nsuck: \l \ Suc k\ have kk'[simp]: \\' k' = \ k\ by (metis csk last_cs) have kl: \k < l\ using icd unfolding is_icdi_def is_cdi_def by auto hence skl: \Suc k < l\ by (metis Suc_lessI nsuck) hence lpd: \\ l pd\ \ (Suc k)\ using icd icd_pd_intermediate by auto have km: \k < m\ by (metis converge(1) kl order.strict_trans) have lcd: \l cd\<^bsup>\\<^esup>\ k\ using icd is_icdi_def by auto hence ipdk_pdl: \ipd (\ k) pd\ (\ l)\ by (metis ipd_pd_cd) have *: \ipd (\ k) \ nodes\ by (metis ipdk_pdl pd_node1) have nretk: \\ k \ return\ by (metis kl lcd path(1) ret_no_cd term_path_stable less_imp_le) have **: \\ (\ l) pd\ ipd (\ k)\ proof assume a: \\ l pd\ ipd (\ k)\ hence \\ l pd\ (\ k)\ by (metis is_ipd_def \k < l\ ipd_is_ipd ipdk_pdl path(1) path_nodes pd_antisym term_path_stable less_imp_le) moreover have \\ l \ (\ k)\ by (metis "*" a ipd_not_self ipdk_pdl lcd pd_antisym ret_no_cd) ultimately show \False\ using lcd cd_not_pd by auto qed have km': \k' < m'\ using cs_order[OF path csk converge(2) nretk km] . obtain \'' n'' where path'': \is_path \''\ and \''0: \\'' 0 = ipd (\ k)\ and \''n: \\'' n'' = return\ and not\l: \\ i\n''. \'' i \ \ l\ using no_pd_path[OF * **] . let \?\'\ = \(\' @\<^bsup>m'\<^esup> \'') \ Suc k'\ have \is_path ?\'\ by (metis \''0 ipd(1) path'' path(2) path_cons path_path_shift) moreover have \?\' 0 = \ (Suc k)\ using km' suc by auto moreover have \?\' (m' - Suc k' + n'') = return\ using \''n km' \''0 ipd(1) by auto ultimately obtain l'' where l'': \l'' \ m' - Suc k' + n''\ \?\' l'' = \ l\ using lpd unfolding is_pd_def by blast have l''m: \l'' \ m' - Suc k'\ apply (rule ccontr) using l'' not\l km' by (cases \Suc (k' + l'') \ m'\, auto) let \?l'\ = \Suc ( k' + l'')\ have lm': \?l' \ m'\ using l''m km' by auto \ \Now we have found our desired l'\ have 1: \\' ?l' = \ l\ using l'' l''m lm' by auto have 2: \k' < ?l'\ by simp have 3: \?l' < m'\ apply (rule ccontr) using lm' by (simp, metis "**" 1 ipd(1) ipdk_pdl) \ \Need the least such l'\ let \?P\ = \\ l'. \' l' = \ l \ k' < l' \ l' < m'\ have *: \?P ?l'\ using 1 2 3 by blast define l' where l': \l' == LEAST l'. ?P l'\ have \l': \\' l' = \ l\ using l' 1 2 3 LeastI[of \?P\] by blast have kl': \k' < l'\ using l' 1 2 3 LeastI[of \?P\] by blast have lm': \l' < m'\ using l' 1 2 3 LeastI[of \?P\] by blast have nretl': \\' l' \ return\ by (metis \''n \l' le_refl not\l) have nipd': \\ j \ {k'..l'}. \' j \ ipd (\' k')\ using lm' kk' ipd(2) kl' by force have lcd': \l' cd\<^bsup>\'\<^esup>\ k'\ by (metis is_cdi_def kl' nipd' nretl' path(2)) have licd: \l' icd\<^bsup>\'\<^esup>\ k'\ proof - have \\ m \ {k'<.. l' cd\<^bsup>\'\<^esup>\ m\ proof (rule ccontr) assume \\ (\ m \ {k'<.. l' cd\<^bsup>\'\<^esup>\ m)\ then obtain j' where kj': \k' < j'\ and jl': \j' < l'\ and lcdj': \l' cd\<^bsup>\'\<^esup>\ j'\ by force have jm': \j' by (metis jl' lm' order.strict_trans) have \\' j' \ \ l\ apply (rule ccontr) using l' kj' jm' jl' Least_le[of \?P\ \j'\] by auto hence \\ \' l' pd\ \' j'\ using cd_not_pd lcdj' \l' by metis moreover have \\' j' \ nodes\ using path(2) path_nodes by auto ultimately obtain \\<^sub>1 n\<^sub>1 where path\<^sub>1: \is_path \\<^sub>1\ and \0\<^sub>1: \\\<^sub>1 0 = \' j'\ and \n\<^sub>1: \\\<^sub>1 n\<^sub>1 = return\ and nl': \\ l \n\<^sub>1. \\<^sub>1 l \ \' l'\ unfolding is_pd_def by blast let \?\''\ = \(\'@\<^bsup>j'\<^esup> \\<^sub>1) \ Suc k'\ have \is_path ?\''\ by (metis \0\<^sub>1 path(2) path\<^sub>1 path_cons path_path_shift) moreover have \?\'' 0 = \ (Suc k)\ by (simp, metis kj' less_eq_Suc_le suc) moreover have kj': \Suc k' \ j'\ by (metis kj' less_eq_Suc_le) hence \?\'' (j' - Suc k' + n\<^sub>1) = return\ by (simp, metis \0\<^sub>1 \n\<^sub>1) ultimately obtain l'' where *: \?\'' l'' = \ l\ and **: \l'' \j' - Suc k' + n\<^sub>1\ using lpd is_pd_def by blast show \False\ proof (cases) assume a: \l'' \ j' - Suc k'\ hence \\' (l'' + Suc k') = \ l\ using * kj' by(simp, metis Nat.le_diff_conv2 add_Suc diff_add_inverse le_add1 le_add_diff_inverse2) moreover have \l'' + Suc k' < l'\ by (metis a jl' add_diff_cancel_right' kj' le_add_diff_inverse less_imp_diff_less ordered_cancel_comm_monoid_diff_class.le_diff_conv2) moreover have \l'' + Suc k' < m'\ by (metis Suc_lessD calculation(2) less_trans_Suc lm') moreover have \k' < l'' + Suc k'\ by simp ultimately show \False\ using Least_le[of \?P\ \l'' + Suc k'\] l' by auto next assume a: \\ l'' \ j' - Suc k'\ hence \\ Suc (k' + l'') \ j'\ by simp hence \\\<^sub>1 (Suc (k' + l'') - j') = \ l\ using * kj' by simp moreover have \Suc (k' + l'') - j' \ n\<^sub>1\ using ** kj' by simp ultimately show \False\ using nl' by (metis \l') qed qed thus \?thesis\ unfolding is_icdi_def using lcd' path(2) by simp qed hence \cs\<^bsup>\'\<^esup> l' = cs\<^bsup>\'\<^esup> k' @ [\' l']\ by (metis icd_cs) hence \cs\<^bsup>\'\<^esup> l' = cs\<^bsup>\\<^esup> l\ by (metis \l' csk icd icd_cs) thus \?thesis\ by metis qed lemma converged_same_icd: assumes path: \is_path \\ \is_path \'\ and converge: \l < n\ \cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and csk: \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ and icd: \l icd\<^bsup>\\<^esup>\ k\ and suc: \\ (Suc k) = \' (Suc k')\ shows \\l'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ proof - have nret: \\ k \ return\ using icd unfolding is_icdi_def is_cdi_def using term_path_stable less_imp_le by metis have kl: \k < l\ using icd unfolding is_icdi_def is_cdi_def by auto have kn: \k < n\ using converge kl by simp from path_ipd_swap[OF path(1) nret kn] obtain \ m where path\: \is_path \\ and \\: \\ =\<^bsub>n\<^esub> \\ and km: \k < m\ and ipd: \\ m = ipd (\ k)\ \\ l \ {k.. l \ ipd (\ k)\ . have csk1: \cs\<^bsup>\\<^esup> k = cs\<^bsup>\\<^esup> k\ using cs_path_swap_le path path\ \\ kn by auto have suc\: \\ (Suc k) = \ (Suc k)\ by (metis \\ eq_up_to_def kn less_eq_Suc_le) have nret': \\' k' \ return\ by (metis csk last_cs nret) have kn': \k' < n'\ using cs_order[OF path csk converge(2) nret kn] . from path_ipd_swap[OF path(2) nret' kn'] obtain \' m' where path\': \is_path \'\ and \\': \\' =\<^bsub>n'\<^esub> \'\ and km': \k' < m'\ and ipd': \\' m' = ipd (\' k')\ \\ l \ {k'..' l \ ipd (\' k')\ . have csk1': \cs\<^bsup>\'\<^esup> k' = cs\<^bsup>\'\<^esup> k'\ using cs_path_swap_le path path\' \\' kn' by auto have suc\': \\' (Suc k') = \' (Suc k')\ by (metis \\' eq_up_to_def kn' less_eq_Suc_le) have icd\: \l icd\<^bsup>\\<^esup>\ k\ using icdi_path_swap_le[OF path\ icd \\] converge by simp have lm: \l < m\ using ipd(1) icd\ km unfolding is_icdi_def is_cdi_def by auto have csk': \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ using csk1 csk1' csk by auto hence kk': \\' k' = \ k\ using last_cs by metis have suc': \\ (Suc k) = \' (Suc k')\ using suc suc\ suc\' by auto have mm': \\' m' = \ m\ using ipd(1) ipd'(1) kk' by auto from cs_ipd[OF ipd km] cs_ipd[OF ipd' km',unfolded mm', folded csk'] have csm: \cs\<^bsup>\\<^esup> m = cs\<^bsup>\'\<^esup> m'\ by metis from converged_ipd_same_icd[OF path\ path\' lm csm csk' icd\ suc' ipd'[unfolded kk']] obtain l' where csl: \cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ by blast have csl\: \cs\<^bsup>\\<^esup> l = cs\<^bsup>\\<^esup> l\ using \\ converge(1) cs_path_swap_le less_imp_le_nat path(1) path\ by blast have nretl: \\ l \ return\ by (metis icd\ icd_imp_cd ret_no_cd) have csn': \cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ using converge(2) cs_path_swap path path\ path\' \\ \\' by auto have ln': \l' < n'\ using cs_order[OF path\ path\' csl csn' nretl converge(1)] . have csl\': \cs\<^bsup>\'\<^esup> l' = cs\<^bsup>\'\<^esup> l'\ using cs_path_swap_le[OF path(2) path\' \\'] ln' by auto have csl': \cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ using csl\ csl\' csl by auto thus \?thesis\ by blast qed lemma cd_is_cs_less: assumes \l cd\<^bsup>\\<^esup>\ k\ shows \cs\<^bsup>\\<^esup> k \ cs\<^bsup>\\<^esup> l\ proof - obtain xs where csl: \cs\<^bsup>\\<^esup> l = cs\<^bsup>\\<^esup> k @ xs @[\ l]\ using cd_in_cs[OF assms] by blast hence len: \length(cs\<^bsup>\\<^esup> k) < length (cs\<^bsup>\\<^esup> l)\ by auto have take: \take (length (cs\<^bsup>\\<^esup> k)) (cs\<^bsup>\\<^esup> l) = cs\<^bsup>\\<^esup> k\ using csl by auto show \?thesis\ using cs_less.intros[OF len take] . qed lemma cs_select_id: assumes \is_path \\ \\ k \ return\ shows \\\cs\<^bsup>\\<^esup> k = k\ (is \?k = k\) proof - have *: \\ i . cs\<^bsup>\\<^esup> i = cs\<^bsup>\\<^esup> k \ i = k\ using cs_inj[OF assms] by metis hence \cs\<^bsup>\\<^esup> ?k = cs\<^bsup>\\<^esup> k\ unfolding cs_select_def using theI[of \\ i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\\<^esup> k\ \k\] by auto thus \?k = k\ using * by auto qed lemma cs_single_nocd: assumes \cs\<^bsup>\\<^esup> i = [x]\ shows \\ k. \ i cd\<^bsup>\\<^esup>\ k\ proof - have \\ (\ k. i icd\<^bsup>\\<^esup>\ k)\ apply (rule ccontr) using assms cs_not_nil by auto hence \\ (\ k. i cd\<^bsup>\\<^esup>\ k)\ by (metis excd_impl_exicd) thus \?thesis\ by blast qed lemma cs_single_pd_intermed: assumes \is_path \\ \cs\<^bsup>\\<^esup> n = [\ n]\ \k \ n\ shows \\ n pd\ \ k\ proof - have \\ l. \ n icd\<^bsup>\\<^esup>\ l\ by (metis assms(2) cs_single_nocd icd_imp_cd) thus \?thesis\ by (metis assms(1) assms(3) no_icd_pd) qed lemma cs_first_pd: assumes path: \is_path \\ and pd: \\ n pd\ \ 0\ and first: \\ l < n. \ l \ \ n\ shows \cs\<^bsup>\\<^esup> n = [\ n]\ by (metis cs_cases first first_pd_no_cd icd_imp_cd path pd) lemma converged_pd_cs_single: assumes path: \is_path \\ \is_path \'\ and converge: \l < m\ \cs\<^bsup>\\<^esup> m = cs\<^bsup>\'\<^esup> m'\ and \0: \\ 0 = \' 0\ and mpdl: \\ m pd\ \ l\ and csl: \cs\<^bsup>\\<^esup> l = [\ l]\ shows \\l'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ proof - have *: \\ l pd\ \' 0\ using cs_single_pd_intermed[OF path(1) csl] \0[symmetric] by auto have \m: \\ m = \' m'\ by (metis converge(2) last_cs) hence **: \\' m' pd\ \ l\ using mpdl by metis obtain l' where lm': \l' \ m'\ and \l: \\' l' = \ l\ (is \?P l'\) using path_pd_pd0[OF path(2) ** *] . let \?l\ = \(LEAST l'. \' l' = \ l)\ have \l': \\' ?l = \ l\ using LeastI[of \?P\,OF \l] . moreover have \\ i ' i \ \ l\ using Least_le[of \?P\] by (metis not_less) hence \\ i ' i \ \' ?l\ using \l' by metis moreover have \\' ?l pd\ \' 0\ using * \l' by metis ultimately have \cs\<^bsup>\'\<^esup> ?l = [\' ?l]\ using cs_first_pd[OF path(2)] by metis thus \?thesis\ using csl \l' by metis qed lemma converged_cs_single: assumes path: \is_path \\ \is_path \'\ and converge: \l < m\ \cs\<^bsup>\\<^esup> m = cs\<^bsup>\'\<^esup> m'\ and \0: \\ 0 = \' 0\ and csl: \cs\<^bsup>\\<^esup> l = [\ l]\ shows \\l'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ proof cases assume *: \\ l = return\ hence \\ m = return\ by (metis converge(1) path(1) term_path_stable less_imp_le) hence \cs\<^bsup>\\<^esup> m = [return]\ using cs_return by auto hence \cs\<^bsup>\'\<^esup> m' = [return]\ using converge by simp moreover have \cs\<^bsup>\\<^esup> l = [return]\ using * cs_return by auto ultimately show \?thesis\ by metis next assume nret: \\ l \ return\ have \m: \\ m = \' m'\ by (metis converge(2) last_cs) obtain \\<^sub>1 n where path1: \is_path \\<^sub>1\ and upto: \\ =\<^bsub>m\<^esub> \\<^sub>1\ and \n: \\\<^sub>1 n = return\ using path(1) path_swap_ret by blast obtain \\<^sub>1' n' where path1': \is_path \\<^sub>1'\ and upto': \\' =\<^bsub>m'\<^esub> \\<^sub>1'\ and \n': \\\<^sub>1' n' = return\ using path(2) path_swap_ret by blast have \1l: \\\<^sub>1 l = \ l\ using upto converge(1) by (metis eq_up_to_def nat_less_le) have cs1l: \cs\<^bsup>\\<^sub>1\<^esup> l = cs\<^bsup>\\<^esup> l\ using cs_path_swap_le upto path1 path(1) converge(1) by auto have csl1: \cs\<^bsup>\\<^sub>1\<^esup> l = [\\<^sub>1 l]\ by (metis \1l cs1l csl) have converge1: \cs\<^bsup>\\<^sub>1\<^esup> n = cs\<^bsup>\\<^sub>1'\<^esup> n'\ using \n \n' cs_return by auto have ln: \l < n\ using nret \n \1l term_path_stable[OF path1 \n] by (auto, metis linorder_neqE_nat less_imp_le) have \01: \\\<^sub>1 0 = \\<^sub>1' 0\ using \0 eq_up_to_apply[OF upto] eq_up_to_apply[OF upto'] by auto have pd: \\\<^sub>1 n pd\ \\<^sub>1 l\ using \n by (metis path1 path_nodes return_pd) obtain l' where csl: \cs\<^bsup>\\<^sub>1\<^esup> l = cs\<^bsup>\\<^sub>1'\<^esup> l'\ using converged_pd_cs_single[OF path1 path1' ln converge1 \01 pd csl1] by blast have cs1m: \cs\<^bsup>\\<^sub>1\<^esup> m = cs\<^bsup>\\<^esup> m\ using cs_path_swap upto path1 path(1) by auto have cs1m': \cs\<^bsup>\\<^sub>1'\<^esup> m' = cs\<^bsup>\'\<^esup> m'\ using cs_path_swap upto' path1' path(2) by auto hence converge1: \cs\<^bsup>\\<^sub>1\<^esup> m = cs\<^bsup>\\<^sub>1'\<^esup> m'\ using converge(2) cs1m by metis have nret1: \\\<^sub>1 l \ return\ using nret \1l by auto have lm': \l' < m'\ using cs_order[OF path1 path1' csl converge1 nret1 converge(1)] . have \cs\<^bsup>\'\<^esup> l' = cs\<^bsup>\\<^sub>1'\<^esup> l'\ using cs_path_swap_le[OF path(2) path1' upto'] lm' by auto moreover have \cs\<^bsup>\\<^esup> l = cs\<^bsup>\\<^sub>1\<^esup> l\ using cs_path_swap_le[OF path(1) path1 upto] converge(1) by auto ultimately have \cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ using csl by auto thus \?thesis\ by blast qed lemma converged_cd_same_suc: assumes path: \is_path \\ \is_path \'\ and init: \\ 0 = \' 0\ and cd_suc: \\ k k'. cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k' \ l cd\<^bsup>\\<^esup>\ k \ \ (Suc k) = \' (Suc k')\ and converge: \l < m\ \cs\<^bsup>\\<^esup> m = cs\<^bsup>\'\<^esup> m'\ shows \\l'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ using path init cd_suc converge proof (induction \\\ \l\ rule: cs_induct,cases) case (cs \ l) assume *: \\k. l icd\<^bsup>\\<^esup>\ k\ let \?k\ = \THE k. l icd\<^bsup>\\<^esup>\ k\ have icd: \l icd\<^bsup>\\<^esup>\ ?k\ by (metis "*" icd_is_the_icd) hence lcdk: \l cd\<^bsup>\\<^esup>\ ?k\ by (metis is_icdi_def) hence kl: \?k using is_cdi_def by metis have \\ j. ?k cd\<^bsup>\\<^esup>\ j \ l cd\<^bsup>\\<^esup>\ j\ using icd cd_trans is_icdi_def by fast hence suc': \\ j j'. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j' \ ?k cd\<^bsup>\\<^esup>\ j \ \ (Suc j) = \' (Suc j')\ using cs.prems(4) by blast from cs.IH[OF * cs(2) path(2) cs(4) suc'] cs.prems kl have \\k'. cs\<^bsup>\\<^esup> (THE k. l icd\<^bsup>\\<^esup>\ k) = cs\<^bsup>\'\<^esup> k'\ by (metis Suc_lessD less_trans_Suc) then obtain k' where csk: \cs\<^bsup>\\<^esup> ?k = cs\<^bsup>\'\<^esup> k'\ by blast have suc2: \\ (Suc ?k) = \' (Suc k')\ using cs.prems(4) lcdk csk by auto have km: \?k < m\ using kl cs.prems(5) by simp from converged_same_icd[OF cs(2) path(2) cs.prems(5) cs.prems(6) csk icd suc2] show \?case\ . next case (cs \ l) assume \\ (\k. l icd\<^bsup>\\<^esup>\ k)\ hence \cs\<^bsup>\\<^esup> l = [\ l]\ by auto with cs converged_cs_single show \?case\ by metis qed lemma converged_cd_diverge: assumes path: \is_path \\ \is_path \'\ and init: \\ 0 = \' 0\ and notin: \\ (\l'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l')\ and converge: \l < m\ \cs\<^bsup>\\<^esup> m = cs\<^bsup>\'\<^esup> m'\ obtains k k' where \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ \l cd\<^bsup>\\<^esup>\ k\ \\ (Suc k) \ \' (Suc k')\ using assms converged_cd_same_suc by blast lemma converged_cd_same_suc_return: assumes path: \is_path \\ \is_path \'\ and \0: \\ 0 = \' 0\ and cd_suc: \\ k k'. cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k' \ l cd\<^bsup>\\<^esup>\ k \ \ (Suc k) = \' (Suc k')\ and ret: \\' n' = return\ shows \\l'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\proof cases assume \\ l = return\ hence \cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> n'\ using ret cs_return by presburger thus \?thesis\ by blast next assume nretl: \\ l \ return\ have \\ l \ nodes\ using path path_nodes by auto then obtain \l n where ipl: \is_path \l\ and \l: \\ l = \l 0\ and retn: \\l n = return\ and notl: \\ i>0. \l i \ \ l\ by (metis direct_path_return nretl) hence ip: \is_path (\@\<^bsup>l\<^esup> \l)\ and l: \(\@\<^bsup>l\<^esup> \l) l = \ l\ and retl: \(\@\<^bsup>l\<^esup> \l) (l + n) = return\ and nl: \\ i>l. (\@\<^bsup>l\<^esup> \l) i \ \ l\ using path_cons[OF path(1) ipl \l] by auto have \0': \(\@\<^bsup>l\<^esup> \l) 0 = \' 0\ unfolding cs_0 using \l \0 by auto have csn: \cs\<^bsup>\@\<^bsup>l\<^esup> \l\<^esup> (l+n) = cs\<^bsup>\'\<^esup> n'\ using ret retl cs_return by metis have eql: \(\@\<^bsup>l\<^esup> \l) =\<^bsub>l\<^esub> \\ by (metis path_append_eq_up_to) have csl': \cs\<^bsup>\@\<^bsup>l\<^esup> \l\<^esup> l = cs\<^bsup>\\<^esup> l\ using eql cs_path_swap ip path(1) by metis have \0 < n\ using nretl[unfolded \l] retn by (metis neq0_conv) hence ln: \l < l + n\ by simp have *: \\ k k'. cs\<^bsup>\ @\<^bsup>l\<^esup> \l\<^esup> k = cs\<^bsup>\'\<^esup> k' \ l cd\<^bsup>\ @\<^bsup>l\<^esup> \l\<^esup>\ k \ (\ @\<^bsup>l\<^esup> \l) (Suc k) = \' (Suc k')\ proof (rule,rule,rule) fix k k' assume *: \cs\<^bsup>\ @\<^bsup>l\<^esup> \l\<^esup> k = cs\<^bsup>\'\<^esup> k' \ l cd\<^bsup>\ @\<^bsup>l\<^esup> \l\<^esup>\ k\ hence kl: \k < l\ using is_cdi_def by auto hence \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k' \ l cd\<^bsup>\\<^esup>\ k\ using eql * cs_path_swap_le[OF ip path(1) eql,of \k\] cdi_path_swap_le[OF path(1) _ eql,of \l\ \k\] by auto hence \\ (Suc k) = \' (Suc k')\ using cd_suc by blast then show \(\ @\<^bsup>l\<^esup> \l) (Suc k) = \' (Suc k')\ using cs_path_swap_le[OF ip path(1) eql,of \Suc k\] kl by auto qed obtain l' where \cs\<^bsup>\ @\<^bsup>l\<^esup> \l\<^esup> l = cs\<^bsup>\'\<^esup> l'\ using converged_cd_same_suc[OF ip path(2) \0' * ln csn] by blast moreover have \cs\<^bsup>\@\<^bsup>l\<^esup> \l\<^esup> l = cs\<^bsup>\\<^esup> l\ using eql by (metis cs_path_swap ip path(1)) ultimately show \?thesis\ by metis qed lemma converged_cd_diverge_return: assumes path: \is_path \\ \is_path \'\ and init: \\ 0 = \' 0\ and notin: \\ (\l'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l')\ and ret: \\' m' = return\ obtains k k' where \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ \l cd\<^bsup>\\<^esup>\ k\ \\ (Suc k) \ \' (Suc k')\ using converged_cd_same_suc_return[OF path init _ ret, of \l\] notin by blast lemma returned_missing_cd_or_loop: assumes path: \is_path \\ \is_path \'\ and \0: \\ 0 = \' 0\ and notin': \\(\ k'. cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k')\ and nret: \\ n'. \' n' \ return\ and ret: \\ n = return\ obtains i i' where \i \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ \\ (Suc i) \ \' (Suc i')\ \k cd\<^bsup>\\<^esup>\ i \ (\ j'> i'. j' cd\<^bsup>\'\<^esup>\ i')\ proof - obtain f where icdf: \\ i'. f (Suc i') icd\<^bsup>\'\<^esup>\ f i'\ and ran: \range f = {i'. \ j'>i'. j' cd\<^bsup>\'\<^esup>\ i'}\ and icdf0: \\ (\i'. f 0 cd\<^bsup>\'\<^esup>\ i')\ using path(2) path_nret_inf_icd_seq nret by blast show \thesis\ proof cases assume \\ j. \ (\ i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> (f j))\ then obtain j where ni\: \\ (\ i. cs\<^bsup>\'\<^esup> (f j) = cs\<^bsup>\\<^esup> i)\ by metis note converged_cd_diverge_return[OF path(2,1) \0[symmetric] ni\ ret] that then obtain i k' where csk: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> k'\ and cdj: \f j cd\<^bsup>\'\<^esup>\ k'\ and div: \\ (Suc i) \ \' (Suc k')\ by metis have \k' \ range f\ using cdj proof (induction \j\) case 0 thus \?case\ using icdf0 by blast next case (Suc j) have icdfj: \f (Suc j) icd\<^bsup>\'\<^esup>\ f j\ using icdf by auto show \?case\ proof cases assume \f (Suc j) icd\<^bsup>\'\<^esup>\ k'\ hence \k' = f j\ using icdfj by (metis icd_uniq) thus \?case\ by auto next assume \\ f (Suc j) icd\<^bsup>\'\<^esup>\ k'\ hence \f j cd\<^bsup>\'\<^esup>\ k'\ using cd_impl_icd_cd[OF Suc.prems icdfj] by auto thus \?case\ using Suc.IH by auto qed qed hence alldep: \\ i'>k'. i' cd\<^bsup>\'\<^esup>\ k'\ using ran by auto show \thesis\ proof cases assume \i < k\ with alldep that[OF _ csk div] show \thesis\ by blast next assume \\ i < k\ hence ki: \k\i\ by auto have \k \ i\ using notin' csk by auto hence ki': \k using ki by auto obtain ka k' where \cs\<^bsup>\\<^esup> ka = cs\<^bsup>\'\<^esup> k'\ \k cd\<^bsup>\\<^esup>\ ka\ \\ (Suc ka) \ \' (Suc k')\ using converged_cd_diverge[OF path \0 notin' ki' csk] by blast moreover hence \ka < k\ unfolding is_cdi_def by auto ultimately show \?thesis\ using that by blast qed next assume \\(\ j. \ (\ i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> (f j)))\ hence allin: \\ j. (\ i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> (f j))\ by blast define f' where f': \f' \ \ j. (SOME i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> (f j))\ have \\ i. f' i < f' (Suc i)\ proof fix i have csi: \cs\<^bsup>\'\<^esup> (f i) = cs\<^bsup>\\<^esup> (f' i)\ unfolding f' using allin by (metis (mono_tags) someI_ex) have cssuci: \cs\<^bsup>\'\<^esup> (f (Suc i)) = cs\<^bsup>\\<^esup> (f' (Suc i))\ unfolding f' using allin by (metis (mono_tags) someI_ex) have fi: \f i < f (Suc i)\ using icdf unfolding is_icdi_def is_cdi_def by auto have \f (Suc i) cd\<^bsup>\'\<^esup>\ f i\ using icdf unfolding is_icdi_def by blast hence nreti: \\' (f i) \ return\ by (metis cd_not_ret) show \f' i < f' (Suc i)\ using cs_order[OF path(2,1) csi cssuci nreti fi] . qed hence kle: \k < f' (Suc k)\ using mono_ge_id[of \f'\ \Suc k\] by auto have cssk: \cs\<^bsup>\\<^esup> (f' (Suc k)) = cs\<^bsup>\'\<^esup> (f (Suc k))\ unfolding f' using allin by (metis (mono_tags) someI_ex) obtain ka k' where \cs\<^bsup>\\<^esup> ka = cs\<^bsup>\'\<^esup> k'\ \k cd\<^bsup>\\<^esup>\ ka\ \\ (Suc ka) \ \' (Suc k')\ using converged_cd_diverge[OF path \0 notin' kle cssk] by blast moreover hence \ka < k\ unfolding is_cdi_def by auto ultimately show \?thesis\ using that by blast qed qed lemma missing_cd_or_loop: assumes path: \is_path \\ \is_path \'\ and \0: \\ 0 = \' 0\ and notin': \\(\ k'. cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k')\ obtains i i' where \i < k\ \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ \\ (Suc i) \ \' (Suc i')\ \k cd\<^bsup>\\<^esup>\ i \ (\ j'> i'. j' cd\<^bsup>\'\<^esup>\ i')\ proof cases assume \\ n'. \' n' = return\ then obtain n' where retn: \\' n' = return\ by blast note converged_cd_diverge_return[OF path \0 notin' retn] then obtain ka k' where \cs\<^bsup>\\<^esup> ka = cs\<^bsup>\'\<^esup> k'\ \k cd\<^bsup>\\<^esup>\ ka\ \\ (Suc ka) \ \' (Suc k')\ by blast moreover hence \ka < k\ unfolding is_cdi_def by auto ultimately show \thesis\ using that by simp next assume \\ (\ n'. \' n' = return)\ hence notret: \\ n'. \' n' \ return\ by auto then obtain \l n where ipl: \is_path \l\ and \l: \\ k = \l 0\ and retn: \\l n = return\ using reaching_ret path(1) path_nodes by metis hence ip: \is_path (\@\<^bsup>k\<^esup>\l)\ and l: \(\@\<^bsup>k\<^esup>\l) k = \ k\ and retl: \(\@\<^bsup>k\<^esup>\l) (k + n) = return\ using path_cons[OF path(1) ipl \l] by auto have \0': \(\@\<^bsup>k\<^esup>\l) 0 = \' 0\ unfolding cs_0 using \l \0 by auto have eql: \(\@\<^bsup>k\<^esup>\l) =\<^bsub>k\<^esub> \\ by (metis path_append_eq_up_to) have csl': \cs\<^bsup>\@\<^bsup>k\<^esup>\l\<^esup> k = cs\<^bsup>\\<^esup> k\ using eql cs_path_swap ip path(1) by metis hence notin: \\(\ k'. cs\<^bsup>\@\<^bsup>k\<^esup>\l\<^esup> k = cs\<^bsup>\'\<^esup> k')\ using notin' by auto obtain i i' where *: \i < k\ and csi: \cs\<^bsup>\@\<^bsup>k\<^esup>\l\<^esup> i = cs\<^bsup>\'\<^esup> i'\ and suci: \(\ @\<^bsup>k\<^esup> \l) (Suc i) \ \' (Suc i')\ and cdloop: \k cd\<^bsup>\@\<^bsup>k\<^esup>\l\<^esup>\ i \ (\ j'>i'. j' cd\<^bsup>\'\<^esup>\ i')\ using returned_missing_cd_or_loop[OF ip path(2) \0' notin notret retl] by blast have \i \ k\ using notin csi by auto hence ik: \i < k\ using * by auto hence \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ using csi cs_path_swap_le[OF ip path(1) eql] by auto moreover have \\ (Suc i) \ \' (Suc i')\ using ik eq_up_to_apply[OF eql, of \Suc i\] suci by auto moreover have \k cd\<^bsup>\\<^esup>\ i \ (\ j'>i'. j' cd\<^bsup>\'\<^esup>\ i')\ using cdloop cdi_path_swap_le[OF path(1) _ eql, of \k\ \i\] by auto ultimately show \thesis\ using that[OF *] by blast qed lemma path_shift_set_cd: assumes \is_path \\ shows \{k + j| j . n cd\<^bsup>\\k\<^esup>\ j } = {i. (k+n) cd\<^bsup>\\<^esup>\ i \ k \ i }\ proof - { fix i assume \i\{k+j | j . n cd\<^bsup>\\k\<^esup>\ j }\ then obtain j where \i = k+j\ \n cd\<^bsup>\\k\<^esup>\ j\ by auto hence \k+n cd\<^bsup>\\<^esup>\ i \ k \ i\ using cd_path_shift[OF _ assms, of \k\ \k+j\ \k+n\] by simp hence \i\{ i. k+n cd\<^bsup>\\<^esup>\ i \ k \ i }\ by blast } moreover { fix i assume \i\{ i. k+n cd\<^bsup>\\<^esup>\ i \ k \ i }\ hence *: \k+n cd\<^bsup>\\<^esup>\ i \ k \ i\ by blast then obtain j where i: \i = k+j\ by (metis le_Suc_ex) hence \k+n cd\<^bsup>\\<^esup>\ k+j\ using * by auto hence \n cd\<^bsup>\\k\<^esup>\ j\ using cd_path_shift[OF _ assms, of \k\ \k+j\ \k+n\] by simp hence \i\{k+j | j . n cd\<^bsup>\\k\<^esup>\ j }\ using i by simp } ultimately show \?thesis\ by blast qed lemma cs_path_shift_set_cd: assumes path: \is_path \\ shows \cs\<^bsup>\\k\<^esup> n = map \ (sorted_list_of_set {i. k+n cd\<^bsup>\\<^esup>\ i \ k \ i }) @ [\ (k+n)]\ proof - have mono:\\n m. n < m \ k + n < k + m\ by auto have fin: \finite {i. n cd\<^bsup>\ \ k\<^esup>\ i}\ unfolding is_cdi_def by auto have *: \(\ x. k+x)`{i. n cd\<^bsup>\\k\<^esup>\ i} = {k + i | i. n cd\<^bsup>\\k\<^esup>\ i}\ by auto have \cs\<^bsup>\\k\<^esup> n = map (\\k) (sorted_list_of_set {i. n cd\<^bsup>\\k\<^esup>\ i}) @ [(\\k) n]\ using cs_sorted_list_of_cd' by blast also have \\ = map \ (map (\ x. k+x) (sorted_list_of_set{i. n cd\<^bsup>\\k\<^esup>\ i})) @ [\ (k+n)]\ by auto also have \\ = map \ (sorted_list_of_set ((\ x. k+x)`{i. n cd\<^bsup>\\k\<^esup>\ i})) @ [\ (k+n)]\ using sorted_list_of_set_map_mono[OF mono fin] by auto also have \\ = map \ (sorted_list_of_set ({k + i | i. n cd\<^bsup>\\k\<^esup>\ i})) @ [\ (k+n)]\ using * by auto also have \\ = map \ (sorted_list_of_set ({i. k+n cd\<^bsup>\\<^esup>\ i \ k \ i})) @ [\ (k+n)]\ using path_shift_set_cd[OF path] by auto finally show \?thesis\ . qed lemma cs_split_shift_cd: assumes \n cd\<^bsup>\\<^esup>\ j\ and \j < k\ and \k < n\ and \\j'\\<^esup>\ j' \ j' \ j\ shows \cs\<^bsup>\\<^esup> n = cs\<^bsup>\\<^esup> j @ cs\<^bsup>\\k\<^esup> (n-k)\ proof - have path: \is_path \\ using assms unfolding is_cdi_def by auto have 1: \{i. n cd\<^bsup>\\<^esup>\ i} = {i. n cd\<^bsup>\\<^esup>\ i \ i < k} \ {i. n cd\<^bsup>\\<^esup>\ i \ k \ i}\ by auto have le: \\ i\ {i. n cd\<^bsup>\\<^esup>\ i \ i < k}. \ j\ {i. n cd\<^bsup>\\<^esup>\ i \ k \ i}. i < j\ by auto have 2: \{i. n cd\<^bsup>\\<^esup>\ i \ i < k} = {i . j cd\<^bsup>\\<^esup>\ i} \ {j}\ proof - { fix i assume \i\{i. n cd\<^bsup>\\<^esup>\ i \ i < k}\ hence cd: \n cd\<^bsup>\\<^esup>\ i\ and ik:\i < k\ by auto have \i\{i . j cd\<^bsup>\\<^esup>\ i} \ {j}\ proof cases assume \i < j\ hence \j cd\<^bsup>\\<^esup>\ i\ by (metis is_cdi_def assms(1) cd cdi_prefix nat_less_le) thus \?thesis\ by simp next assume \\ i < j\ moreover have \i \ j\ using assms(4) ik cd by auto ultimately show \?thesis\ by auto qed } moreover { fix i assume \i\{i . j cd\<^bsup>\\<^esup>\ i} \ {j}\ hence \j cd\<^bsup>\\<^esup>\ i \ i = j\ by auto hence \i\{i. n cd\<^bsup>\\<^esup>\ i \ i < k}\ using assms(1,2) cd_trans[OF _ assms(1)] apply auto unfolding is_cdi_def by (metis (poly_guards_query) diff_diff_cancel diff_is_0_eq le_refl le_trans nat_less_le) } ultimately show \?thesis\ by blast qed have \cs\<^bsup>\\<^esup> n = map \ (sorted_list_of_set {i. n cd\<^bsup>\\<^esup>\ i}) @ [\ n]\ using cs_sorted_list_of_cd' by simp also have \\ = map \ (sorted_list_of_set ({i. n cd\<^bsup>\\<^esup>\ i \ i < k} \ {i. n cd\<^bsup>\\<^esup>\ i \ k \ i})) @ [\ n]\ using 1 by metis also have \\ = map \ ((sorted_list_of_set {i. n cd\<^bsup>\\<^esup>\ i \ i < k}) @ (sorted_list_of_set {i. n cd\<^bsup>\\<^esup>\ i \ k \ i})) @ [\ n]\ using sorted_list_of_set_append[OF _ _ le] is_cdi_def by auto also have \\ = (map \ (sorted_list_of_set {i. n cd\<^bsup>\\<^esup>\ i \ i < k})) @ (map \ (sorted_list_of_set {i. n cd\<^bsup>\\<^esup>\ i \ k \ i})) @ [\ n]\ by auto also have \\ = cs\<^bsup>\\<^esup> j @ (map \ (sorted_list_of_set {i. n cd\<^bsup>\\<^esup>\ i \ k \ i})) @ [\ n]\ unfolding 2 using cs_sorted_list_of_cd by auto also have \\ = cs\<^bsup>\\<^esup> j @ cs\<^bsup>\\k\<^esup> (n-k)\ using cs_path_shift_set_cd[OF path, of \k\ \n-k\] assms(3) by auto finally show \?thesis\ . qed lemma cs_split_shift_nocd: assumes \is_path \\ and \k < n\ and \\j. n cd\<^bsup>\\<^esup>\ j \ k \ j\ shows \cs\<^bsup>\\<^esup> n = cs\<^bsup>\\k\<^esup> (n-k)\ proof - have path: \is_path \\ using assms by auto have 1: \{i. n cd\<^bsup>\\<^esup>\ i} = {i. n cd\<^bsup>\\<^esup>\ i \ i < k} \ {i. n cd\<^bsup>\\<^esup>\ i \ k \ i}\ by auto have le: \\ i\ {i. n cd\<^bsup>\\<^esup>\ i \ i < k}. \ j\ {i. n cd\<^bsup>\\<^esup>\ i \ k \ i}. i < j\ by auto have 2: \{i. n cd\<^bsup>\\<^esup>\ i \ i < k} = {}\ using assms by auto have \cs\<^bsup>\\<^esup> n = map \ (sorted_list_of_set {i. n cd\<^bsup>\\<^esup>\ i}) @ [\ n]\ using cs_sorted_list_of_cd' by simp also have \\ = map \ (sorted_list_of_set ({i. n cd\<^bsup>\\<^esup>\ i \ i < k} \ {i. n cd\<^bsup>\\<^esup>\ i \ k \ i})) @ [\ n]\ using 1 by metis also have \\ = map \ (sorted_list_of_set {i. n cd\<^bsup>\\<^esup>\ i \ k \ i}) @ [\ n]\ unfolding 2 by auto also have \\ = cs\<^bsup>\\k\<^esup> (n-k)\ using cs_path_shift_set_cd[OF path, of \k\ \n-k\] assms(2) by auto finally show \?thesis\ . qed lemma shifted_cs_eq_is_eq: assumes \is_path \\ and \is_path \'\ and \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ and \cs\<^bsup>\\k\<^esup> n = cs\<^bsup>\'\k'\<^esup> n'\ shows \cs\<^bsup>\\<^esup> (k+n) = cs\<^bsup>\'\<^esup> (k'+n')\ proof (rule ccontr) note path = assms(1,2) note csk = assms(3) note csn = assms(4) assume ne: \cs\<^bsup>\\<^esup> (k+n) \ cs\<^bsup>\'\<^esup> (k'+n')\ have nretkn:\\ (k+n) \ return\ proof assume 1:\\ (k+n) = return\ hence \(\\k) n = return\ by auto hence \(\'\k') n' = return\ using last_cs assms(4) by metis hence \\' (k' + n') = return\ by auto thus \False\ using ne 1 cs_return by auto qed hence nretk: \\ k \ return\ using term_path_stable[OF assms(1), of \k\ \k +n\] by auto have nretkn': \\' (k'+n') \ return\ proof assume 1:\\' (k'+n') = return\ hence \(\'\k') n' = return\ by auto hence \(\\k) n = return\ using last_cs assms(4) by metis hence \\ (k + n) = return\ by auto thus \False\ using ne 1 cs_return by auto qed hence nretk': \\' k' \ return\ using term_path_stable[OF assms(2), of \k'\ \k' +n'\] by auto have n0:\n > 0\ proof (rule ccontr) assume *: \\ 0 < n\ hence 1:\cs\<^bsup>\\k\<^esup> 0 = cs\<^bsup>\'\k'\<^esup> n'\ using assms(3,4) by auto have \(\\k) 0 = (\'\k') 0\ using assms(3) last_cs path_shift_def by (metis monoid_add_class.add.right_neutral) hence \cs\<^bsup>\'\k'\<^esup> 0 = cs\<^bsup>\'\k'\<^esup> n'\ using 1 cs_0 by metis hence n0': \n' = 0\ using cs_inj[of \\'\k'\ \0\ \n'\ ] * assms(2) by (metis path_shift_def assms(4) last_cs nretkn path_path_shift) thus \False\ using ne * assms(3) by fastforce qed have n0':\n' > 0\ proof (rule ccontr) assume *: \\ 0 < n'\ hence 1:\cs\<^bsup>\'\k'\<^esup> 0 = cs\<^bsup>\\k\<^esup> n\ using assms(3,4) by auto have \(\'\k') 0 = (\\k) 0\ using assms(3) last_cs path_shift_def by (metis monoid_add_class.add.right_neutral) hence \cs\<^bsup>\\k\<^esup> 0 = cs\<^bsup>\\k\<^esup> n\ using 1 cs_0 by metis hence n0: \n = 0\ using cs_inj[of \\\k\ \0\ \n\ ] * assms(1) by (metis path_shift_def assms(4) last_cs nretkn path_path_shift) thus \False\ using ne * assms(3) by fastforce qed have cdleswap': \\ j'\'\<^esup>\ j' \ (\j\\<^esup>\ j \ cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ proof (rule,rule,rule, rule ccontr) fix j' assume jk': \j' and ncdj': \(k'+n') cd\<^bsup>\'\<^esup>\ j'\ and ne: \\ (\j\\<^esup>\ j \ cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ hence kcdj': \k' cd\<^bsup>\'\<^esup>\ j'\ using cr_wn' by blast then obtain j where kcdj: \k cd\<^bsup>\\<^esup>\ j\ and csj: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ using csk cs_path_swap_cd path by metis hence jk: \j < k\ unfolding is_cdi_def by auto have ncdn: \\ (k+n) cd\<^bsup>\\<^esup>\ j\ using ne csj jk by blast obtain l' where lnocd': \l' = n' \ n' cd\<^bsup>\'\k'\<^esup>\ l'\ and cslsing': \cs\<^bsup>\'\k'\<^esup> l' = [(\'\k') l']\ proof cases assume \cs\<^bsup>\'\k'\<^esup> n' = [(\'\k') n']\ thus \thesis\ using that[of \n'\] by auto next assume *: \cs\<^bsup>\'\k'\<^esup> n' \ [(\'\k') n']\ then obtain x ys where \cs\<^bsup>\'\k'\<^esup> n' = [x]@ys@[(\'\k') n']\ by (metis append_Cons append_Nil cs_length_g_one cs_length_one(1) neq_Nil_conv) then obtain l' where \cs\<^bsup>\'\k'\<^esup> l' = [x]\ and cdl': \n' cd\<^bsup>\'\k'\<^esup>\ l'\ using cs_split[of \\'\k'\ \n'\ \Nil\ \x\ \ys\] by auto hence \cs\<^bsup>\'\k'\<^esup> l' = [(\'\k') l']\ using last_cs by (metis last.simps) thus \thesis\ using that cdl' by auto qed hence ln': \l'\n'\ unfolding is_cdi_def by auto hence lcdj': \k'+l' cd\<^bsup>\'\<^esup>\ j'\ using jk' ncdj' by (metis add_le_cancel_left cdi_prefix trans_less_add1) obtain l where lnocd: \l = n \ n cd\<^bsup>\\k\<^esup>\ l\ and csl: \cs\<^bsup>\\k\<^esup> l = cs\<^bsup>\'\k'\<^esup> l'\ using lnocd' proof assume \l' = n'\ thus \thesis\ using csn that[of \n\] by auto next assume \n' cd\<^bsup>\'\k'\<^esup>\ l'\ then obtain l where \n cd\<^bsup>\\k\<^esup>\ l\ \cs\<^bsup>\\k\<^esup> l = cs\<^bsup>\'\k'\<^esup> l'\ using cs_path_swap_cd path csn by (metis path_path_shift) thus \thesis\ using that by auto qed have cslsing: \cs\<^bsup>\\k\<^esup> l = [(\\k) l]\ using cslsing' last_cs csl last.simps by metis have ln: \l\n\ using lnocd unfolding is_cdi_def by auto hence nretkl: \\ (k + l) \ return\ using term_path_stable[of \\\ \k+l\ \k+n\] nretkn path(1) by auto have *: \n cd\<^bsup>\\k\<^esup>\ l \ k+n cd\<^bsup>\\<^esup>\ k+l\ using cd_path_shift[of \k\ \k+l\ \\\ \k+n\] path(1) by auto have ncdl: \\ (k+l) cd\<^bsup>\\<^esup>\ j\ apply rule using lnocd apply rule using ncdn apply blast using cd_trans ncdn * by blast hence \\ i\ {j..k+l}. \ i = ipd (\ j)\ unfolding is_cdi_def using path(1) jk nretkl by auto hence \\ i\ {k<..k+l}. \ i = ipd (\ j)\ using kcdj unfolding is_cdi_def by force then obtain i where ki: \k < i\ and il: \i \ k+l\ and ipdi: \\ i = ipd (\ j)\ by force hence \(\\k) (i-k) = ipd (\ j)\ \i-k \ l\ by auto hence pd: \(\\k) l pd\ ipd (\ j)\ using cs_single_pd_intermed[OF _ cslsing] path(1) path_path_shift by metis moreover have \(\\k) l = \' (k' + l')\ using csl last_cs by (metis path_shift_def) moreover have \\ j = \' j'\ using csj last_cs by metis ultimately have \\' (k'+l') pd\ ipd (\' j')\ by simp moreover have \ipd (\' j') pd\ \' (k'+l')\ using ipd_pd_cd[OF lcdj'] . ultimately have \\' (k'+l') = ipd (\' j')\ using pd_antisym by auto thus \False\ using lcdj' unfolding is_cdi_def by force qed \ \Symmetric version of the above statement\ have cdleswap: \\ j\\<^esup>\ j \ (\j'\'\<^esup>\ j' \ cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ proof (rule,rule,rule, rule ccontr) fix j assume jk: \j and ncdj: \(k+n) cd\<^bsup>\\<^esup>\ j\ and ne: \\ (\j'\'\<^esup>\ j' \ cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ hence kcdj: \k cd\<^bsup>\\<^esup>\ j\ using cr_wn' by blast then obtain j' where kcdj': \k' cd\<^bsup>\'\<^esup>\ j'\ and csj: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ using csk cs_path_swap_cd path by metis hence jk': \j' < k'\ unfolding is_cdi_def by auto have ncdn': \\ (k'+n') cd\<^bsup>\'\<^esup>\ j'\ using ne csj jk' by blast obtain l where lnocd: \l = n \ n cd\<^bsup>\\k\<^esup>\ l\ and cslsing: \cs\<^bsup>\\k\<^esup> l = [(\\k) l]\ proof cases assume \cs\<^bsup>\\k\<^esup> n = [(\\k) n]\ thus \thesis\ using that[of \n\] by auto next assume *: \cs\<^bsup>\\k\<^esup> n \ [(\\k) n]\ then obtain x ys where \cs\<^bsup>\\k\<^esup> n = [x]@ys@[(\\k) n]\ by (metis append_Cons append_Nil cs_length_g_one cs_length_one(1) neq_Nil_conv) then obtain l where \cs\<^bsup>\\k\<^esup> l = [x]\ and cdl: \n cd\<^bsup>\\k\<^esup>\ l\ using cs_split[of \\\k\ \n\ \Nil\ \x\ \ys\] by auto hence \cs\<^bsup>\\k\<^esup> l = [(\\k) l]\ using last_cs by (metis last.simps) thus \thesis\ using that cdl by auto qed hence ln: \l\n\ unfolding is_cdi_def by auto hence lcdj: \k+l cd\<^bsup>\\<^esup>\ j\ using jk ncdj by (metis add_le_cancel_left cdi_prefix trans_less_add1) obtain l' where lnocd': \l' = n' \ n' cd\<^bsup>\'\k'\<^esup>\ l'\ and csl: \cs\<^bsup>\\k\<^esup> l = cs\<^bsup>\'\k'\<^esup> l'\ using lnocd proof assume \l = n\ thus \thesis\ using csn that[of \n'\] by auto next assume \n cd\<^bsup>\\k\<^esup>\ l\ then obtain l' where \n' cd\<^bsup>\'\k'\<^esup>\ l'\ \cs\<^bsup>\\k\<^esup> l = cs\<^bsup>\'\k'\<^esup> l'\ using cs_path_swap_cd path csn by (metis path_path_shift) thus \thesis\ using that by auto qed have cslsing': \cs\<^bsup>\'\k'\<^esup> l' = [(\'\k') l']\ using cslsing last_cs csl last.simps by metis have ln': \l'\n'\ using lnocd' unfolding is_cdi_def by auto hence nretkl': \\' (k' + l') \ return\ using term_path_stable[of \\'\ \k'+l'\ \k'+n'\] nretkn' path(2) by auto have *: \n' cd\<^bsup>\'\k'\<^esup>\ l' \ k'+n' cd\<^bsup>\'\<^esup>\ k'+l'\ using cd_path_shift[of \k'\ \k'+l'\ \\'\ \k'+n'\] path(2) by auto have ncdl': \\ (k'+l') cd\<^bsup>\'\<^esup>\ j'\ apply rule using lnocd' apply rule using ncdn' apply blast using cd_trans ncdn' * by blast hence \\ i'\ {j'..k'+l'}. \' i' = ipd (\' j')\ unfolding is_cdi_def using path(2) jk' nretkl' by auto hence \\ i'\ {k'<..k'+l'}. \' i' = ipd (\' j')\ using kcdj' unfolding is_cdi_def by force then obtain i' where ki': \k' < i'\ and il': \i' \ k'+l'\ and ipdi: \\' i' = ipd (\' j')\ by force hence \(\'\k') (i'-k') = ipd (\' j')\ \i'-k' \ l'\ by auto hence pd: \(\'\k') l' pd\ ipd (\' j')\ using cs_single_pd_intermed[OF _ cslsing'] path(2) path_path_shift by metis moreover have \(\'\k') l' = \ (k + l)\ using csl last_cs by (metis path_shift_def) moreover have \\' j' = \ j\ using csj last_cs by metis ultimately have \\ (k+l) pd\ ipd (\ j)\ by simp moreover have \ipd (\ j) pd\ \ (k+l)\ using ipd_pd_cd[OF lcdj] . ultimately have \\ (k+l) = ipd (\ j)\ using pd_antisym by auto thus \False\ using lcdj unfolding is_cdi_def by force qed have cdle: \\j. (k+n) cd\<^bsup>\\<^esup>\ j \ j < k\ (is \\ j. ?P j\) proof (rule ccontr) assume \\ (\j. (k+n) cd\<^bsup>\\<^esup>\ j \ j < k)\ hence allge: \\j. (k+n) cd\<^bsup>\\<^esup>\ j \ k \ j\ by auto have allge': \\j'. (k'+n') cd\<^bsup>\'\<^esup>\ j' \ k' \ j'\ proof (rule, rule, rule ccontr) fix j' assume *: \k' + n' cd\<^bsup>\'\<^esup>\ j'\ and \\ k' \ j'\ then obtain j where \j \(k+n) cd\<^bsup>\\<^esup>\ j\ using cdleswap' by (metis le_neq_implies_less nat_le_linear) thus \False\ using allge by auto qed have \cs\<^bsup>\\<^esup> (k + n) = cs\<^bsup>\ \ k\<^esup> n\ using cs_split_shift_nocd[OF assms(1) _ allge] n0 by auto moreover have \cs\<^bsup>\'\<^esup> (k' + n') = cs\<^bsup>\' \ k'\<^esup> n'\ using cs_split_shift_nocd[OF assms(2) _ allge'] n0' by auto ultimately show \False\ using ne assms(4) by auto qed define j where \j == GREATEST j. (k+n) cd\<^bsup>\\<^esup>\ j \ j < k\ have cdj:\(k+n) cd\<^bsup>\\<^esup>\ j\ and jk: \j < k\ and jge:\\ j'< k. (k+n) cd\<^bsup>\\<^esup>\ j' \ j' \ j\ proof - have bound: \\ y. ?P y \ y \ k\ by auto show \(k+n) cd\<^bsup>\\<^esup>\ j\ using GreatestI_nat[of \?P\] j_def bound cdle by blast show \j < k\ using GreatestI_nat[of \?P\] bound j_def cdle by blast show \\ j'< k. (k+n) cd\<^bsup>\\<^esup>\ j' \ j' \ j\ using Greatest_le_nat[of \?P\] bound j_def by blast qed obtain j' where cdj':\(k'+n') cd\<^bsup>\'\<^esup>\ j'\ and csj: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ and jk': \j' < k'\ using cdleswap cdj jk by blast have jge':\\ i'< k'. (k'+n') cd\<^bsup>\'\<^esup>\ i' \ i' \ j'\ proof(rule,rule,rule) fix i' assume ik': \i' < k'\ and cdi': \k' + n' cd\<^bsup>\'\<^esup>\ i'\ then obtain i where cdi:\(k+n) cd\<^bsup>\\<^esup>\ i\ and csi: \ cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i\ and ik: \i using cdleswap' by force have ij: \i \ j\ using jge cdi ik by auto show \i' \ j'\ using cs_order_le[OF assms(1,2) csi[symmetric] csj _ ij] cd_not_ret[OF cdi] by simp qed have \cs\<^bsup>\\<^esup> (k + n) = cs\<^bsup>\\<^esup> j @ cs\<^bsup>\ \ k\<^esup> n\ using cs_split_shift_cd[OF cdj jk _ jge] n0 by auto moreover have \cs\<^bsup>\'\<^esup> (k' + n') = cs\<^bsup>\'\<^esup> j' @ cs\<^bsup>\' \ k'\<^esup> n'\ using cs_split_shift_cd[OF cdj' jk' _ jge'] n0' by auto ultimately have \cs\<^bsup>\\<^esup> (k+n) = cs\<^bsup>\'\<^esup> (k'+n')\ using csj assms(4) by auto thus \False\ using ne by simp qed lemma cs_eq_is_eq_shifted: assumes \is_path \\ and \is_path \'\ and \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ and \cs\<^bsup>\\<^esup> (k+n) = cs\<^bsup>\'\<^esup> (k'+n')\ shows \cs\<^bsup>\\k\<^esup> n = cs\<^bsup>\'\k'\<^esup> n'\ proof (rule ccontr) assume ne: \cs\<^bsup>\ \ k\<^esup> n \ cs\<^bsup>\' \ k'\<^esup> n'\ have nretkn:\\ (k+n) \ return\ proof assume 1:\\ (k+n) = return\ hence 2:\\' (k'+n') = return\ using assms(4) last_cs by metis hence \(\\k) n = return\ \(\'\k') n' = return\ using 1 by auto hence \cs\<^bsup>\ \ k\<^esup> n = cs\<^bsup>\' \ k'\<^esup> n'\ using cs_return by metis thus \False\ using ne by simp qed hence nretk: \\ k \ return\ using term_path_stable[OF assms(1), of \k\ \k +n\] by auto have nretkn': \\' (k'+n') \ return\ proof assume 1:\\' (k'+n') = return\ hence 2:\\ (k+n) = return\ using assms(4) last_cs by metis hence \(\\k) n = return\ \(\'\k') n' = return\ using 1 by auto hence \cs\<^bsup>\ \ k\<^esup> n = cs\<^bsup>\' \ k'\<^esup> n'\ using cs_return by metis thus \False\ using ne by simp qed hence nretk': \\' k' \ return\ using term_path_stable[OF assms(2), of \k'\ \k' +n'\] by auto have n0:\n > 0\ proof (rule ccontr) assume *: \\ 0 < n\ hence \cs\<^bsup>\'\<^esup> k' = cs\<^bsup>\'\<^esup> (k'+ n')\ using assms(3,4) by auto hence n0: \n = 0\ \n' = 0\ using cs_inj[OF assms(2) nretkn', of \k'\] * by auto have \cs\<^bsup>\ \ k\<^esup> n = cs\<^bsup>\' \ k'\<^esup> n'\ unfolding n0 cs_0 by (auto , metis last_cs assms(3)) thus \False\ using ne by simp qed have n0':\n' > 0\ proof (rule ccontr) assume *: \\ 0 < n'\ hence \cs\<^bsup>\\<^esup> k = cs\<^bsup>\\<^esup> (k+ n)\ using assms(3,4) by auto hence n0: \n = 0\ \n' = 0\ using cs_inj[OF assms(1) nretkn, of \k\] * by auto have \cs\<^bsup>\ \ k\<^esup> n = cs\<^bsup>\' \ k'\<^esup> n'\ unfolding n0 cs_0 by (auto , metis last_cs assms(3)) thus \False\ using ne by simp qed have cdle: \\j. (k+n) cd\<^bsup>\\<^esup>\ j \ j < k\ (is \\ j. ?P j\) proof (rule ccontr) assume \\ (\j. (k+n) cd\<^bsup>\\<^esup>\ j \ j < k)\ hence allge: \\j. (k+n) cd\<^bsup>\\<^esup>\ j \ k \ j\ by auto have allge': \\j'. (k'+n') cd\<^bsup>\'\<^esup>\ j' \ k' \ j'\ proof (rule, rule) fix j' assume *: \k' + n' cd\<^bsup>\'\<^esup>\ j'\ obtain j where cdj: \k+n cd\<^bsup>\\<^esup>\ j\ and csj: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ using cs_path_swap_cd[OF assms(2,1) assms(4)[symmetric] *] by metis hence kj:\k \ j\ using allge by auto thus kj': \k' \ j'\ using cs_order_le[OF assms(1,2,3) csj nretk] by simp qed have \cs\<^bsup>\\<^esup> (k + n) = cs\<^bsup>\ \ k\<^esup> n\ using cs_split_shift_nocd[OF assms(1) _ allge] n0 by auto moreover have \cs\<^bsup>\'\<^esup> (k' + n') = cs\<^bsup>\' \ k'\<^esup> n'\ using cs_split_shift_nocd[OF assms(2) _ allge'] n0' by auto ultimately show \False\ using ne assms(4) by auto qed define j where \j == GREATEST j. (k+n) cd\<^bsup>\\<^esup>\ j \ j < k\ have cdj:\(k+n) cd\<^bsup>\\<^esup>\ j\ and jk: \j < k\ and jge:\\ j'< k. (k+n) cd\<^bsup>\\<^esup>\ j' \ j' \ j\ proof - have bound: \\ y. ?P y \ y \ k\ by auto show \(k+n) cd\<^bsup>\\<^esup>\ j\ using GreatestI_nat[of \?P\] bound j_def cdle by blast show \j < k\ using GreatestI_nat[of \?P\] bound j_def cdle by blast show \\ j'< k. (k+n) cd\<^bsup>\\<^esup>\ j' \ j' \ j\ using Greatest_le_nat[of \?P\] bound j_def by blast qed obtain j' where cdj':\(k'+n') cd\<^bsup>\'\<^esup>\ j'\ and csj: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ using cs_path_swap_cd assms cdj by blast have jge':\\ i'< k'. (k'+n') cd\<^bsup>\'\<^esup>\ i' \ i' \ j'\ proof(rule,rule,rule) fix i' assume ik': \i' < k'\ and cdi': \k' + n' cd\<^bsup>\'\<^esup>\ i'\ then obtain i where cdi:\(k+n) cd\<^bsup>\\<^esup>\ i\ and csi: \ cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i\ using cs_path_swap_cd[OF assms(2,1) assms(4)[symmetric]] by blast have nreti': \\' i' \ return\ by (metis cd_not_ret cdi') have ik: \i < k\ using cs_order[OF assms(2,1) csi _ nreti' ik'] assms(3) by auto have ij: \i \ j\ using jge cdi ik by auto show \i' \ j'\ using cs_order_le[OF assms(1,2) csi[symmetric] csj _ ij] cd_not_ret[OF cdi] by simp qed have jk': \j' < k'\ using cs_order[OF assms(1,2) csj assms(3) cd_not_ret[OF cdj] jk] . have \cs\<^bsup>\\<^esup> (k + n) = cs\<^bsup>\\<^esup> j @ cs\<^bsup>\ \ k\<^esup> n\ using cs_split_shift_cd[OF cdj jk _ jge] n0 by auto moreover have \cs\<^bsup>\'\<^esup> (k' + n') = cs\<^bsup>\'\<^esup> j' @ cs\<^bsup>\' \ k'\<^esup> n'\ using cs_split_shift_cd[OF cdj' jk' _ jge'] n0' by auto ultimately have \cs\<^bsup>\\k\<^esup> n = cs\<^bsup>\'\k'\<^esup> n'\ using csj assms(4) by auto thus \False\ using ne by simp qed lemma converged_cd_diverge_cs: assumes \is_path \\ and \is_path \'\ and \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ and \j and \\ (\l'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l')\ and \l < m\ and \cs\<^bsup>\\<^esup> m = cs\<^bsup>\'\<^esup> m'\ obtains k k' where \j\k\ \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ and \l cd\<^bsup>\\<^esup>\ k\ and \\ (Suc k) \ \' (Suc k')\ proof - have \is_path (\\j)\ \is_path (\'\j')\ using assms(1,2) path_path_shift by auto moreover have \(\\j) 0 = (\'\j') 0\ using assms(3) last_cs by (metis path_shift_def add.right_neutral) moreover have \\(\l'. cs\<^bsup>\\j\<^esup> (l-j) = cs\<^bsup>\'\j'\<^esup> l')\ proof assume \\l'. cs\<^bsup>\ \ j\<^esup> (l - j) = cs\<^bsup>\' \ j'\<^esup> l'\ then obtain l' where csl: \cs\<^bsup>\\j\<^esup> (l - j) = cs\<^bsup>\'\j'\<^esup> l'\ by blast have \cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> (j' + l')\ using shifted_cs_eq_is_eq[OF assms(1,2,3) csl] assms(4) by auto thus \False\ using assms(5) by blast qed moreover have \l-j < m-j\ using assms by auto moreover have \\ j \ return\ using cs_return assms(1-5) term_path_stable by (metis nat_less_le) hence \j' using cs_order[OF assms(1,2,3,7)] assms by auto hence \cs\<^bsup>\\j\<^esup> (m-j) = cs\<^bsup>\'\j'\<^esup> (m'-j')\ using cs_eq_is_eq_shifted[OF assms(1,2,3),of \m-j\ \m'-j'\] assms(4,6,7) by auto ultimately obtain k k' where csk: \cs\<^bsup>\\j\<^esup> k = cs\<^bsup>\'\j'\<^esup> k'\ and lcdk: \l-j cd\<^bsup>\\j\<^esup>\ k\ and suc:\(\\j) (Suc k) \ (\'\j') (Suc k')\ using converged_cd_diverge by blast have \cs\<^bsup>\\<^esup> (j+k) = cs\<^bsup>\'\<^esup> (j'+k')\ using shifted_cs_eq_is_eq[OF assms(1-3) csk] . moreover have \l cd\<^bsup>\\<^esup>\ j+k\ using lcdk assms(1,2,4) by (metis add.commute add_diff_cancel_right' cd_path_shift le_add1) moreover have \\ (Suc (j+k)) \ \' (Suc (j'+ k'))\ using suc by auto moreover have \j \ j+k\ by auto ultimately show \thesis\ using that[of \j+k\ \j'+k'\] by auto qed lemma cs_ipd_conv: assumes csk: \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ and ipd: \\ l = ipd (\ k)\ \\' l' = ipd(\' k')\ and nipd: \\n\{k.. n \ ipd (\ k)\ \\n'\{k'..' n' \ ipd (\' k')\ and kl: \k < l\ \k' < l'\ shows \cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ using cs_ipd[OF ipd(1) nipd(1) kl(1)] cs_ipd[OF ipd(2) nipd(2) kl(2)] csk ipd by (metis (no_types) last_cs) lemma cp_eq_cs: assumes \((\,k),(\',k'))\cp\ shows \cs\<^bsup>path \\<^esup> k = cs\<^bsup>path \'\<^esup> k'\ using assms apply(induction rule: cp.induct) apply blast+ apply simp done lemma cd_cs_swap: assumes \l cd\<^bsup>\\<^esup>\ k\ \cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ shows \l' cd\<^bsup>\'\<^esup>\ k'\ proof - have \\ i. l icd\<^bsup>\\<^esup>\ i\ using assms(1) excd_impl_exicd by blast hence \cs\<^bsup>\\<^esup> l \ [\ l]\ by auto hence \cs\<^bsup>\'\<^esup> l' \ [\' l']\ using assms last_cs by metis hence \\ i'. l' icd\<^bsup>\'\<^esup>\ i'\ by (metis cs_cases) hence path': \is_path \'\ unfolding is_icdi_def is_cdi_def by auto from cd_in_cs[OF assms(1)] obtain ys where csl: \cs\<^bsup>\\<^esup> l = cs\<^bsup>\\<^esup> k @ ys @ [\ l]\ by blast obtain xs where csk: \cs\<^bsup>\\<^esup> k = xs@[\ k]\ by (metis append_butlast_last_id cs_not_nil last_cs) have \l: \\ l = \' l'\ using assms last_cs by metis have csl': \cs\<^bsup>\'\<^esup> l' = xs@[\ k]@ys@[\' l']\ by (metis \l append_eq_appendI assms(2) csk csl) from cs_split[of \\'\ \l'\ \xs\ \\ k\ \ys\] obtain m where csm: \cs\<^bsup>\'\<^esup> m = xs @ [\ k]\ and lcdm: \l' cd\<^bsup>\'\<^esup>\ m\ using csl' by metis have csm': \cs\<^bsup>\'\<^esup> m = cs\<^bsup>\'\<^esup> k'\ by (metis assms(3) csk csm) have \\' m \ return\ using lcdm unfolding is_cdi_def using term_path_stable by (metis nat_less_le) hence \m = k'\ using cs_inj path' csm' by auto thus \?thesis\ using lcdm by auto qed subsection \Facts about Observations\ lemma kth_obs_not_none: assumes \is_kth_obs (path \) k i\ obtains a where \obsp \ i = Some a\ using assms unfolding is_kth_obs_def obsp_def by auto lemma kth_obs_unique: \is_kth_obs \ k i \ is_kth_obs \ k j \ i = j\ proof (induction \i\ \j\ rule: nat_sym_cases) case sym thus \?case\ by simp next case eq thus \?case\ by simp next case (less i j) have \obs_ids \ \ {.. obs_ids \ \ {.. using less(1) by auto moreover have \i \ obs_ids \ \ {.. using less unfolding is_kth_obs_def obs_ids_def by auto moreover have \i \ obs_ids \ \ {.. by auto moreover have \card (obs_ids \ \ {.. \ {.. using less.prems unfolding is_kth_obs_def by auto moreover have \finite (obs_ids \ \ {.. \finite (obs_ids \ \ {.. by auto ultimately have \False\ by (metis card_subset_eq) thus \?case\ .. qed lemma obs_none_no_kth_obs: assumes \obs \ k = None\ shows \\ (\ i. is_kth_obs (path \) k i)\ apply rule using assms unfolding obs_def obsp_def apply (auto split: option.split_asm) by (metis assms kth_obs_not_none kth_obs_unique obs_def option.distinct(2) the_equality) lemma obs_some_kth_obs : assumes \obs \ k \ None\ obtains i where \is_kth_obs (path \) k i\ by (metis obs_def assms) lemma not_none_is_obs: assumes \att(\ i) \ None\ shows \is_kth_obs \ (card (obs_ids \ \ {.. unfolding is_kth_obs_def using assms by auto lemma in_obs_ids_is_kth_obs: assumes \i \ obs_ids \\ obtains k where \is_kth_obs \ k i\ proof have \att (\ i) \ None\ using assms obs_ids_def by auto thus \is_kth_obs \ (card (obs_ids \ \ {.. using not_none_is_obs by auto qed lemma kth_obs_stable: assumes \is_kth_obs \ l j\ \k < l\ shows \\ i. is_kth_obs \ k i\ using assms proof (induction \l\ arbitrary: \j\ rule: less_induct ) case (less l j) have cardl: \card (obs_ids \ \ {.. using less is_kth_obs_def by auto then obtain i where ex: \i \ obs_ids \ \ {.. (is \?P i\) using less(3) by (metis card.empty empty_iff less_irrefl subsetI subset_antisym zero_diff zero_less_diff) have bound: \\ i. i \ obs_ids \ \ {.. i \ j\ by auto let \?i\ = \GREATEST i. i \ obs_ids \ \ {.. have *: \?i < j\ \?i \ obs_ids \\ using GreatestI_nat[of \?P\ \i\ \j\] ex bound by auto have **: \\ i. i \ obs_ids \ \ i i \ ?i\ using Greatest_le_nat[of \?P\ _ \j\] ex bound by auto have \(obs_ids \ \ {.. {?i} = obs_ids \ \ {.. apply rule apply auto using *[simplified] apply simp+ using **[simplified] by auto moreover have \?i \ (obs_ids \ \ {.. by auto ultimately have \Suc (card (obs_ids \ \ {.. using cardl by (metis Un_empty_right Un_insert_right card_insert_disjoint finite_Int finite_lessThan) hence \card (obs_ids \ \ {.. by auto hence iko: \is_kth_obs \ (l - 1) ?i\ using *(2) unfolding is_kth_obs_def obs_ids_def by auto have ll: \l - 1 < l\ by (metis One_nat_def diff_Suc_less less.prems(2) not_gr0 not_less0) note IV=less(1)[OF ll iko] show \?thesis\ proof cases assume \k < l - 1\ thus \?thesis\ using IV by simp next assume \\ k < l - 1\ hence \k = l - 1\ using less by auto thus \?thesis\ using iko by blast qed qed lemma kth_obs_mono: assumes \is_kth_obs \ k i\ \is_kth_obs \ l j\ \k < l\ shows \i < j\ proof (rule ccontr) assume \\ i < j\ hence \{.. {.. by auto hence \obs_ids \ \ {.. obs_ids \ \ {.. by auto moreover have \finite (obs_ids \ \ {.. by auto ultimately have \card (obs_ids \ \ {.. card (obs_ids \ \ {.. by (metis card_mono) thus \False\ using assms unfolding is_kth_obs_def by auto qed lemma kth_obs_le_iff: assumes \is_kth_obs \ k i\ \is_kth_obs \ l j\ shows \k < l \ i < j\ by (metis assms kth_obs_unique kth_obs_mono not_less_iff_gr_or_eq) lemma ret_obs_all_obs: assumes path: \is_path \\ and iki: \is_kth_obs \ k i\ and ret: \\ i = return\ and kl: \k < l\ obtains j where \is_kth_obs \ l j\ proof- show \thesis\ using kl iki ret proof (induction \l - k\ arbitrary: \k\ \i\ rule: less_induct) case (less k i) note kl = \k < l\ note iki = \is_kth_obs \ k i\ note ret = \\ i = return\ have card: \card (obs_ids \ \ {.. and att_ret: \att return \ None\using iki ret unfolding is_kth_obs_def by auto have rets: \\ (Suc i) = return\ using path ret term_path_stable by auto hence attsuc: \att (\ (Suc i)) \ None\ using att_ret by auto hence *: \i \ obs_ids \\ using att_ret ret unfolding obs_ids_def by auto have \{..< Suc i} = insert i {.. by auto hence a: \obs_ids \ \ {..< Suc i} = insert i (obs_ids \ \ {.. using * by auto have b: \i \ obs_ids \ \ {.. by auto have \finite (obs_ids \ \ {.. by auto hence \card (obs_ids \ \ {.. by (metis card card_insert_disjoint a b) hence iksuc: \is_kth_obs \ (Suc k) (Suc i)\ using attsuc unfolding is_kth_obs_def by auto have suckl: \Suc k \ l\ using kl by auto note less thus \thesis\ proof (cases \Suc k < l\) assume skl: \Suc k < l\ from less(1)[OF _ skl iksuc rets] skl show \thesis\ by auto next assume \\ Suc k < l\ hence \Suc k = l\ using suckl by auto thus \thesis\ using iksuc that by auto qed qed qed lemma no_kth_obs_missing_cs: assumes path: \is_path \\ \is_path \'\ and iki: \is_kth_obs \ k i\ and not_in_\': \\(\i'. is_kth_obs \' k i')\ obtains l j where \is_kth_obs \ l j\ \\ (\ j'. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ proof (rule ccontr) assume \\ thesis\ hence all_in_\': \\ l j. is_kth_obs \ l j \ (\ j' . cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ using that by blast then obtain i' where csi: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ using assms by blast hence \att(\' i') \ None\ using iki by (metis is_kth_obs_def last_cs) then obtain k' where ik': \is_kth_obs \' k' i'\ by (metis not_none_is_obs) hence kk': \k' < k\ using not_in_\' kth_obs_stable by (auto, metis not_less_iff_gr_or_eq) show \False\ proof (cases \\ i = return\) assume \\ i \ return\ thus \False\ using kk' ik' csi iki proof (induction \k\ arbitrary: \i\ \i'\ \k'\ ) case 0 thus \?case\ by simp next case (Suc k i i' k') then obtain j where ikj: \is_kth_obs \ k j\ by (metis kth_obs_stable lessI) then obtain j' where csj: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ using all_in_\' by blast hence \att(\' j') \ None\ using ikj by (metis is_kth_obs_def last_cs) then obtain k2 where ik2: \is_kth_obs \' k2 j'\ by (metis not_none_is_obs) have ji: \j < i\ using kth_obs_mono [OF ikj \is_kth_obs \ (Suc k) i\] by auto hence nretj: \\ j \ return\ using Suc(2) term_path_stable less_imp_le path(1) by metis have ji': \j' < i'\ using cs_order[OF path _ _ nretj, of \j'\ \i\ \i'\] csj \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ ji by auto have \k2 \ k'\ using ik2 Suc(4) ji' kth_obs_unique[of \\'\ \k'\ \i'\ \j'\] by (metis less_irrefl) hence k2k': \k2 < k'\ using kth_obs_mono[OF \is_kth_obs \' k' i'\ ik2] ji' by (metis not_less_iff_gr_or_eq) hence k2k: \k2 < k\ using Suc by auto from Suc.IH[OF nretj k2k ik2 csj ikj] show \False\ . qed next assume \\ i = return\ hence reti': \\' i' = return\ by (metis csi last_cs) from ret_obs_all_obs[OF path(2) ik' reti' kk', of \False\] not_in_\' show \False\ by blast qed qed lemma kth_obs_cs_missing_cs: assumes path: \is_path \\ \is_path \'\ and iki: \is_kth_obs \ k i\ and iki': \is_kth_obs \' k i'\ and csi: \cs\<^bsup>\\<^esup> i \ cs\<^bsup>\'\<^esup> i'\ obtains l j where \j \ i\ \is_kth_obs \ l j\ \\ (\ j'. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ | l' j' where \j' \ i'\ \is_kth_obs \' l' j'\ \\ (\ j. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ proof (rule ccontr) assume nt: \\ thesis\ show \False\ using iki iki' csi that proof (induction \k\ arbitrary: \i\ \i'\ rule: less_induct) case (less k i i') hence all_in_\': \\ l j. j\i \ is_kth_obs \ l j \ (\ j' . cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ and all_in_\: \\ l' j'. j' \ i' \ is_kth_obs \' l' j' \ (\ j . cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ by (metis nt) (metis nt less(6)) obtain j j' where csji: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> i'\ and csij: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> j'\ using all_in_\ all_in_\' less by blast then obtain l l' where ilj: \is_kth_obs \ l j\ and ilj': \is_kth_obs \' l' j'\ by (metis is_kth_obs_def last_cs less.prems(1,2)) have lnk: \l \ k\ using ilj csji less(2) less(4) kth_obs_unique by auto have lnk': \l' \ k\ using ilj' csij less(3) less(4) kth_obs_unique by auto have cseq: \\ l j j'. l < k \ is_kth_obs \ l j \ is_kth_obs \' l j' \ cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ proof - { fix t p p' assume tk: \t < k\ and ikp: \is_kth_obs \ t p\ and ikp': \is_kth_obs \' t p'\ hence pi: \p < i\ and pi': \p' < i'\ by (metis kth_obs_mono less.prems(1)) (metis kth_obs_mono less.prems(2) tk ikp') have *: \\j l. j \ p \ is_kth_obs \ l j \ \j'. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ using pi all_in_\' by auto have **: \\j' l'. j' \ p' \ is_kth_obs \' l' j' \ \j. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ using pi' all_in_\ by auto have \cs\<^bsup>\\<^esup> p = cs\<^bsup>\'\<^esup> p'\ apply(rule ccontr) using less(1)[OF tk ikp ikp'] * ** by blast } thus \?thesis\ by blast qed have ii'nret: \\ i \ return \ \' i' \ return\ using less cs_return by auto have a: \k < l \ k < l'\ proof (rule ccontr) assume \\(k < l \ k < l')\ hence *: \l < k\ \l' < k\ using lnk lnk' by auto hence ji: \j < i\ and ji': \j' < i'\ using ilj ilj' less(2,3) kth_obs_mono by auto show \False\ using ii'nret proof assume nreti: \\ i \ return\ hence nretj': \\' j' \ return\ using last_cs csij by metis show \False\ using cs_order[OF path(2,1) csij[symmetric] csji[symmetric] nretj' ji'] ji by simp next assume nreti': \\' i' \ return\ hence nretj': \\ j \ return\ using last_cs csji by metis show \False\ using cs_order[OF path csji csij nretj' ji] ji' by simp qed qed have \l < k \ l' < k\ proof (rule ccontr) assume \\ (l< k \ l' < k)\ hence \k < l\ \k < l'\ using lnk lnk' by auto hence ji: \i < j\ and ji': \i' < j'\ using ilj ilj' less(2,3) kth_obs_mono by auto show \False\ using ii'nret proof assume nreti: \\ i \ return\ show \False\ using cs_order[OF path csij csji nreti ji] ji' by simp next assume nreti': \\' i' \ return\ show \False\ using cs_order[OF path(2,1) csji[symmetric] csij[symmetric] nreti' ji'] ji by simp qed qed hence \k < l \ l' < k \ k < l' \ l < k\ using a by auto thus \False\ proof assume \k < l \ l' < k\ hence kl: \k < l\ and lk': \l' < k\ by auto hence ij: \i < j\ and ji': \j' < i'\ using less(2,3) ilj ilj' kth_obs_mono by auto have nreti: \\ i \ return\ by (metis csji ii'nret ij last_cs path(1) term_path_stable less_imp_le) obtain h where ilh: \is_kth_obs \ l' h\ using ji' all_in_\ ilj' no_kth_obs_missing_cs path(1) path(2) by (metis kl lk' ilj kth_obs_stable) hence \cs\<^bsup>\\<^esup> h = cs\<^bsup>\'\<^esup> j'\ using cseq lk' ilj' by blast hence \cs\<^bsup>\\<^esup> i = cs\<^bsup>\\<^esup> h\ using csij by auto hence hi: \h = i\ using cs_inj nreti path(1) by metis have \l' = k\ using less(2) ilh unfolding hi by (metis is_kth_obs_def) thus \False\ using lk' by simp next assume \k < l' \ l < k\ hence kl': \k < l'\ and lk: \l < k\ by auto hence ij': \i' < j'\ and ji: \j < i\ using less(2,3) ilj ilj' kth_obs_mono by auto have nreti': \\' i' \ return\ by (metis csij ii'nret ij' last_cs path(2) term_path_stable less_imp_le) obtain h' where ilh': \is_kth_obs \' l h'\ using all_in_\' ilj no_kth_obs_missing_cs path(1) path(2) kl' lk ilj' kth_obs_stable by metis hence \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> h'\ using cseq lk ilj by blast hence \cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\'\<^esup> h'\ using csji by auto hence hi: \h' = i'\ using cs_inj nreti' path(2) by metis have \l = k\ using less(3) ilh' unfolding hi by (metis is_kth_obs_def) thus \False\ using lk by simp qed qed qed subsection \Facts about Data\ lemma reads_restrict1: \\ \ (reads n) = \' \ (reads n) \ \ x \ reads n. \ x = \' x\ by (metis restrict_def) lemma reads_restrict2: \\ x \ reads n. \ x = \' x \ \ \ (reads n) = \' \ (reads n)\ unfolding restrict_def by auto lemma reads_restrict: \(\ \ (reads n) = \' \ (reads n)) = (\ x \ reads n. \ x = \' x)\ using reads_restrict1 reads_restrict2 by metis lemma reads_restr_suc: \\ \ (reads n) = \' \ (reads n) \ suc n \ = suc n \'\ by (metis reads_restrict uses_suc) lemma reads_restr_sem: \\ \ (reads n) = \' \ (reads n) \ \ v \ writes n. sem n \ v = sem n \' v\ by (metis reads_restrict1 uses_writes) lemma reads_obsp: assumes \path \ k = path \' k'\ \\\<^bsup>k\<^esup> \ (reads (path \ k)) = \'\<^bsup>k'\<^esup> \ (reads (path \ k))\ shows \obsp \ k = obsp \' k'\ using assms(2) uses_att unfolding obsp_def assms(1) reads_restrict apply (cases \att (path \' k')\) by auto lemma no_writes_unchanged0: assumes \\ l writes(path \ l)\ shows \(\\<^bsup>k\<^esup>) v = \ v\ using assms proof (induction \k\) case 0 thus \?case\ by(auto simp add: kth_state_def) next case (Suc k) hence \(\\<^bsup>k\<^esup>) v = \ v\ by auto moreover have \\\<^bsup>Suc k\<^esup> = snd ( step (path \ k,\\<^bsup>k\<^esup>))\ by (metis kth_state_suc) hence \\\<^bsup>Suc k\<^esup> = sem (path \ k) (\\<^bsup>k\<^esup>)\ by (metis step_suc_sem snd_conv) moreover have \v \ writes (path \ k)\ using Suc.prems by blast ultimately show \?case\ using writes by metis qed lemma written_read_dd: assumes \is_path \\ \v \ reads (\ k) \ \v \ writes (\ j)\ \j obtains l where \k dd\<^bsup>\,v\<^esup>\ l\ proof - let \?l\ = \GREATEST l. l < k \ v \ writes (\ l)\ have \?l < k\ by (metis (no_types, lifting) GreatestI_ex_nat assms(3) assms(4) less_or_eq_imp_le) moreover have \v \ writes (\ ?l)\ by (metis (no_types, lifting) GreatestI_nat assms(3) assms(4) less_or_eq_imp_le) hence \v \ reads (\ k) \ writes (\ ?l)\ using assms(2) by auto moreover note is_ddi_def have \\ l \ {?l<.. writes (\ l)\ by (auto, metis (lifting, no_types) Greatest_le_nat le_antisym nat_less_le) ultimately have \k dd\<^bsup>\,v\<^esup>\ ?l\ using assms(1) unfolding is_ddi_def by blast thus \thesis\ using that by simp qed lemma no_writes_unchanged: assumes \k \ l\ \\ j \ {k.. writes(path \ j)\ shows \(\\<^bsup>l\<^esup>) v = (\\<^bsup>k\<^esup>) v\ using assms proof (induction \l - k\ arbitrary: \l\) case 0 thus \?case\ by(auto) next case (Suc lk l) hence kl: \k < l\ by auto then obtain l' where lsuc: \l = Suc l'\ using lessE by blast hence \lk = l' - k\ using Suc by auto moreover have \\ j \ {k.. writes (path \ j)\ using Suc(4) lsuc by auto ultimately have \(\\<^bsup>l'\<^esup>) v = (\\<^bsup>k\<^esup>) v\ using Suc(1)[of \l'\] lsuc kl by fastforce moreover have \\\<^bsup>l\<^esup> = snd ( step (path \ l',\\<^bsup>l'\<^esup>))\ by (metis kth_state_suc lsuc) hence \\\<^bsup>l\<^esup> = sem (path \ l') (\\<^bsup>l'\<^esup>)\ by (metis step_suc_sem snd_conv) moreover have \l' < l\ \k \ l'\ using kl lsuc by auto hence \v \ writes (path \ l')\ using Suc.prems(2) by auto ultimately show \?case\ using writes by metis qed lemma ddi_value: assumes \l dd\<^bsup>(path \),v\<^esup>\ k\ shows \(\\<^bsup>l\<^esup>) v = (\\<^bsup>Suc k\<^esup> ) v\ using assms no_writes_unchanged[of \Suc k\ \l\ \v\ \\\] unfolding is_ddi_def by auto lemma written_value: assumes \path \ l = path \' l'\ \\\<^bsup>l\<^esup> \ reads (path \ l) = \'\<^bsup>l'\<^esup> \ reads (path \ l)\ \v \ writes (path \ l)\ shows \(\\<^bsup>Suc l\<^esup> ) v = (\'\<^bsup>Suc l'\<^esup> ) v\ by (metis assms reads_restr_sem snd_conv step_suc_sem kth_state_suc) subsection \Facts about Contradicting Paths\ lemma obsp_contradict: assumes csk: \cs\<^bsup>path \\<^esup> k = cs\<^bsup>path \'\<^esup> k'\ and obs: \obsp \ k \ obsp \' k'\ shows \(\', k') \ (\, k)\ proof - have pk: \path \ k = path \' k'\ using assms last_cs by metis hence \\\<^bsup>k\<^esup>\(reads (path \ k)) \ \'\<^bsup>k'\<^esup>\(reads (path \ k))\ using obs reads_obsp[OF pk] by auto thus \(\',k') \ (\,k)\ using contradicts.intros(2)[OF csk[symmetric]] by auto qed lemma missing_cs_contradicts: assumes notin: \\(\ k'. cs\<^bsup>path \\<^esup> k = cs\<^bsup>path \'\<^esup> k')\ and converge: \k \cs\<^bsup>path \\<^esup> n = cs\<^bsup>path \'\<^esup> n'\ shows \\ j'. (\', j') \ (\, k)\ proof - let \?\\ = \path \\ let \?\'\ = \path \'\ have init: \?\ 0 = ?\' 0\ unfolding path_def by auto have path: \is_path ?\\ \is_path ?\'\ using path_is_path by auto obtain j j' where csj: \cs\<^bsup>?\\<^esup> j = cs\<^bsup>?\'\<^esup> j'\ and cd: \k cd\<^bsup>?\\<^esup>\j\ and suc: \?\ (Suc j) \ ?\' (Suc j')\ using converged_cd_diverge[OF path init notin converge] . have less: \cs\<^bsup>?\\<^esup> j \ cs\<^bsup>?\\<^esup> k\ using cd cd_is_cs_less by auto have nretj: \?\ j \ return\ by (metis cd is_cdi_def term_path_stable less_imp_le) have cs: \?\ \ cs\<^bsup>?\'\<^esup> j' = j\ using csj cs_select_id nretj path_is_path by metis have \(\',j') \ (\,k)\ using contradicts.intros(1)[of \?\'\ \j'\ \?\\ \k\ \\\ \\'\,unfolded cs] less suc csj by metis thus \?thesis\ by blast qed theorem obs_neq_contradicts_term: fixes \ \' defines \: \\ \ path \\ and \': \\' \ path \'\ assumes ret: \\ n = return\ \\' n' = return\ and obsne: \obs \ \ obs \'\ shows \\ k k'. ((\', k') \ (\ ,k) \ \ k \ dom (att)) \ ((\, k) \ (\' ,k') \ \' k' \ dom (att))\ proof - have path: \is_path \\ \is_path \'\ using \ \' path_is_path by auto obtain k1 where neq: \obs \ k1 \ obs \' k1\ using obsne ext[of \obs \\ \obs \'\] by blast hence \(\k i i'. is_kth_obs \ k i \ is_kth_obs \' k i' \ obsp \ i \ obsp \' i' \ cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i') \ (\ k i. is_kth_obs \ k i \ \ (\ i'. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i')) \ (\ k i'. is_kth_obs \' k i' \ \ (\ i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'))\ proof(cases rule: option_neq_cases) case (none2 x) have notin\': \\ (\ l. is_kth_obs \' k1 l)\ using none2(2) \' obs_none_no_kth_obs by auto obtain i where in\: \is_kth_obs \ k1 i\ using obs_some_kth_obs[of \\\ \k1\] none2(1) \ by auto obtain l j where \is_kth_obs \ l j\ \\ (\ j'. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ using path in\ notin\' by (metis no_kth_obs_missing_cs) thus \?thesis\ by blast next case (none1 x) have notin\: \\ (\ l. is_kth_obs \ k1 l)\ using none1(1) \ obs_none_no_kth_obs by auto obtain i' where in\': \is_kth_obs \' k1 i'\ using obs_some_kth_obs[of \\'\ \k1\] none1(2) \' by auto obtain l j where \is_kth_obs \' l j\ \\ (\ j'. cs\<^bsup>\\<^esup> j' = cs\<^bsup>\'\<^esup> j)\ using path in\' notin\ by (metis no_kth_obs_missing_cs) thus \?thesis\ by blast next case (some x y) obtain i where in\: \is_kth_obs \ k1 i\ using obs_some_kth_obs[of \\\ \k1\] some \ by auto obtain i' where in\': \is_kth_obs \' k1 i'\ using obs_some_kth_obs[of \\'\ \k1\] some \' by auto show \?thesis\ proof (cases) assume *: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ have \obsp \ i = obs \ k1\ by (metis obs_def \ in\ kth_obs_unique the_equality) moreover have \obsp \' i' = obs \' k1\ by (metis obs_def \' in\' kth_obs_unique the_equality) ultimately have \obsp \ i \ obsp \' i'\ using neq by auto thus \?thesis\ using * in\ in\' by blast next assume *: \cs\<^bsup>\\<^esup> i \ cs\<^bsup>\'\<^esup> i'\ note kth_obs_cs_missing_cs[OF path in\ in\' *] thus \?thesis\ by metis qed qed thus \?thesis\ proof (cases rule: three_cases) case 1 then obtain k i i' where iki: \is_kth_obs \ k i\ \is_kth_obs \' k i'\ and obsne: \obsp \ i \ obsp \' i'\ and csi: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ by auto note obsp_contradict[OF csi[unfolded \ \'] obsne] moreover have \\ i \ dom att\ using iki unfolding is_kth_obs_def by auto ultimately show \?thesis\ by blast next case 2 then obtain k i where iki: \is_kth_obs \ k i\ and notin\': \\ (\i'. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i')\ by auto let \?n\ = \Suc (max n i)\ have nn: \n < ?n\ by auto have iln: \i < ?n\ by auto have retn: \\ ?n = return\ using ret term_path_stable path by auto hence \cs\<^bsup>\\<^esup> ?n = cs\<^bsup>\'\<^esup> n'\ using ret(2) cs_return by auto then obtain i' where \(\',i') \ (\,i)\ using missing_cs_contradicts[OF notin\'[unfolded \ \'] iln] \ \' by auto moreover have \\ i \ dom att\ using iki is_kth_obs_def by auto ultimately show \?thesis\ by blast next case 3 then obtain k i' where iki: \is_kth_obs \' k i'\ and notin\': \\ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i')\ by auto let \?n\ = \Suc (max n' i')\ have nn: \n' < ?n\ by auto have iln: \i' < ?n\ by auto have retn: \\' ?n = return\ using ret term_path_stable path by auto hence \cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> ?n\ using ret(1) cs_return by auto then obtain i where \(\,i) \ (\',i')\ using missing_cs_contradicts notin\' iln \ \' by metis moreover have \\' i' \ dom att\ using iki is_kth_obs_def by auto ultimately show \?thesis\ by blast qed qed lemma obs_neq_some_contradicts': fixes \ \' defines \: \\ \ path \\ and \': \\' \ path \'\ assumes obsnecs: \obsp \ i \ obsp \' i' \ cs\<^bsup>\\<^esup> i \ cs\<^bsup>\'\<^esup> i'\ and iki: \is_kth_obs \ k i\ and iki': \is_kth_obs \' k i'\ shows \\ k k'. ((\', k') \ (\ ,k) \ \ k \ dom att) \ ((\, k) \ (\' ,k') \ \' k' \ dom att)\ using obsnecs iki iki' proof (induction \k\ arbitrary: \i\ \i'\ rule: less_induct ) case (less k i i') note iki = \is_kth_obs \ k i\ and iki' = \is_kth_obs \' k i'\ have domi: \\ i \ dom att\ by (metis is_kth_obs_def domIff iki) have domi': \\' i' \ dom att\ by (metis is_kth_obs_def domIff iki') note obsnecs = \obsp \ i \ obsp \' i' \ cs\<^bsup>\\<^esup> i \ cs\<^bsup>\'\<^esup> i'\ show \?thesis\ proof cases assume csi: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ hence *: \obsp \ i \ obsp \' i'\ using obsnecs by auto note obsp_contradict[OF _ *] csi domi \ \' thus \?thesis\ by blast next assume ncsi: \cs\<^bsup>\\<^esup> i \ cs\<^bsup>\'\<^esup> i'\ have path: \is_path \\ \is_path \'\ using \ \' path_is_path by auto have \0: \\ 0 = \' 0\ unfolding \ \' path_def by auto note kth_obs_cs_missing_cs[of \\\ \\'\ \k\ \i\ \i'\] \ \' path_is_path iki iki' ncsi hence \(\ l j .j \ i \ is_kth_obs \ l j \ \ (\ j'. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')) \ (\ l' j'. j' \ i' \ is_kth_obs \' l' j' \ \ (\ j. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'))\ by metis thus \?thesis\ proof assume \\l j. j \ i \ is_kth_obs \ l j \ \ (\j'. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ then obtain l j where ji: \j\i\ and iobs: \is_kth_obs \ l j\ and notin: \\ (\j'. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ by blast have dom: \\ j \ dom att\ using iobs is_kth_obs_def by auto obtain n n' where nj: \n < j\ and csn: \cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and sucn: \\ (Suc n) \ \' (Suc n')\ and cdloop: \j cd\<^bsup>\\<^esup>\ n \ (\ j'> n'. j' cd\<^bsup>\'\<^esup>\ n')\ using missing_cd_or_loop[OF path \0 notin] by blast show \?thesis\ using cdloop proof assume cdjn: \j cd\<^bsup>\\<^esup>\ n\ hence csnj: \cs\<^bsup>\'\<^esup> n' \ cs\<^bsup>\\<^esup> j\ using csn by (metis cd_is_cs_less) have cssel: \\ (Suc (\ \ cs\<^bsup>\'\<^esup> n')) = \ (Suc n)\ using csn by (metis cdjn cd_not_ret cs_select_id path(1)) have \(\',n') \ (\,j)\ using csnj apply(rule contradicts.intros(1)) using cssel \ \' sucn by auto thus \?thesis\ using dom by auto next assume loop: \\ j'>n'. j' cd\<^bsup>\'\<^esup>\ n'\ show \?thesis\ proof cases assume in': \i' \ n'\ have nreti': \\' i' \ return\ by( metis le_eq_less_or_eq lessI loop not_le path(2) ret_no_cd term_path_stable) show \?thesis\ proof cases assume \\ \. cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> \\ then obtain \ where cs\: \cs\<^bsup>\\<^esup> \ = cs\<^bsup>\'\<^esup> i'\ by metis have \n: \\ \ n\ using cs_order_le[OF path(2,1) cs\[symmetric] csn[symmetric] nreti' in'] . hence \i: \\ < i\ using nj ji by auto have dom\: \\ \ \ dom att\ using domi' cs\ last_cs by metis obtain \ where i\\: \is_kth_obs \ \ \\ using dom\ by (metis is_kth_obs_def domIff) hence \k: \\ < k\ using \i iki by (metis kth_obs_le_iff) obtain \' where i\\': \is_kth_obs \' \ \'\ using \k iki' by (metis kth_obs_stable) have \\' < i'\ using \k iki' i\\' by (metis kth_obs_le_iff) hence cs\': \cs\<^bsup>\\<^esup> \ \ cs\<^bsup>\'\<^esup> \'\ unfolding cs\ using cs_inj[OF path(2) nreti', of \\'\] by blast thus \?thesis\ using less(1)[OF \k _ i\\ i\\'] by auto next assume notin'': \\(\ \. cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> \)\ obtain \ \' where \i': \\' < i'\ and cs\: \cs\<^bsup>\\<^esup> \ = cs\<^bsup>\'\<^esup> \'\ and suc\: \\ (Suc \) \ \' (Suc \')\ and cdloop': \i' cd\<^bsup>\'\<^esup>\ \' \ (\ j>\. j cd\<^bsup>\\<^esup>\ \)\ using missing_cd_or_loop[OF path(2,1) \0[symmetric] notin''] by metis show \?thesis\ using cdloop' proof assume cdjn: \i' cd\<^bsup>\'\<^esup>\ \'\ hence csnj: \cs\<^bsup>\\<^esup> \ \ cs\<^bsup>\'\<^esup> i'\ using cs\ by (metis cd_is_cs_less) have cssel: \\' (Suc (\' \ cs\<^bsup>\\<^esup> \)) = \' (Suc \')\ using cs\ by (metis cdjn cd_not_ret cs_select_id path(2)) have \(\,\) \ (\',i')\ using csnj apply(rule contradicts.intros(1)) using cssel \ \' suc\ by auto thus \?thesis\ using domi' by auto next assume loop': \\ j>\. j cd\<^bsup>\\<^esup>\ \\ have \n': \\' < n'\ using in' \i' by auto have nret\': \\' \' \ return\ by (metis cs\ last_cs le_eq_less_or_eq lessI path(1) path(2) suc\ term_path_stable) have \\ < n\ using cs_order[OF path(2,1) cs\[symmetric] csn[symmetric] nret\' \n'] . hence \\ < i\ using nj ji by auto hence cdi\: \i cd\<^bsup>\\<^esup>\ \\ using loop' by auto hence cs\i: \cs\<^bsup>\'\<^esup> \' \ cs\<^bsup>\\<^esup> i\ using cs\ by (metis cd_is_cs_less) have cssel: \\ (Suc (\ \ cs\<^bsup>\'\<^esup> \')) = \ (Suc \)\ using cs\ by (metis cdi\ cd_not_ret cs_select_id path(1)) have \(\',\') \ (\,i)\ using cs\i apply(rule contradicts.intros(1)) using cssel \ \' suc\ by auto thus \?thesis\ using domi by auto qed qed next assume \\ i' \ n'\ hence ni': \n'< i'\ by simp hence cdin: \i' cd\<^bsup>\'\<^esup>\ n'\ using loop by auto hence csni: \cs\<^bsup>\\<^esup> n \ cs\<^bsup>\'\<^esup> i'\ using csn by (metis cd_is_cs_less) have cssel: \\' (Suc (\' \ cs\<^bsup>\\<^esup> n)) = \' (Suc n')\ using csn by (metis cdin cd_not_ret cs_select_id path(2)) have \(\,n) \ (\',i')\ using csni apply(rule contradicts.intros(1)) using cssel \ \' sucn by auto thus \?thesis\ using domi' by auto qed qed next \ \Symmetric case as above, indices might be messy.\ assume \\l j. j \ i' \ is_kth_obs \' l j \ \ (\j'. cs\<^bsup>\\<^esup> j' = cs\<^bsup>\'\<^esup> j)\ then obtain l j where ji': \j\i'\ and iobs: \is_kth_obs \' l j\ and notin: \\ (\j'. cs\<^bsup>\'\<^esup> j = cs\<^bsup>\\<^esup> j')\ by metis have dom: \\' j \ dom att\ using iobs is_kth_obs_def by auto obtain n n' where nj: \n < j\ and csn: \cs\<^bsup>\'\<^esup> n = cs\<^bsup>\\<^esup> n'\ and sucn: \\' (Suc n) \ \ (Suc n')\ and cdloop: \j cd\<^bsup>\'\<^esup>\ n \ (\ j'> n'. j' cd\<^bsup>\\<^esup>\ n')\ using missing_cd_or_loop[OF path(2,1) \0[symmetric] ] notin by metis show \?thesis\ using cdloop proof assume cdjn: \j cd\<^bsup>\'\<^esup>\ n\ hence csnj: \cs\<^bsup>\\<^esup> n' \ cs\<^bsup>\'\<^esup> j\ using csn by (metis cd_is_cs_less) have cssel: \\' (Suc (\' \ cs\<^bsup>\\<^esup> n')) = \' (Suc n)\ using csn by (metis cdjn cd_not_ret cs_select_id path(2)) have \(\,n') \ (\',j)\ using csnj apply(rule contradicts.intros(1)) using cssel \' \ sucn by auto thus \?thesis\ using dom by auto next assume loop: \\ j'>n'. j' cd\<^bsup>\\<^esup>\ n'\ show \?thesis\ proof cases assume in': \i \ n'\ have nreti: \\ i \ return\ by (metis le_eq_less_or_eq lessI loop not_le path(1) ret_no_cd term_path_stable) show \?thesis\ proof cases assume \\ \. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> \\ then obtain \ where cs\: \cs\<^bsup>\'\<^esup> \ = cs\<^bsup>\\<^esup> i\ by metis have \n: \\ \ n\ using cs_order_le[OF path cs\[symmetric] csn[symmetric] nreti in'] . hence \i': \\ < i'\ using nj ji' by auto have dom\: \\' \ \ dom att\ using domi cs\ last_cs by metis obtain \ where i\\: \is_kth_obs \' \ \\ using dom\ by (metis is_kth_obs_def domIff) hence \k: \\ < k\ using \i' iki' by (metis kth_obs_le_iff) obtain \' where i\\': \is_kth_obs \ \ \'\ using \k iki by (metis kth_obs_stable) have \\' < i\ using \k iki i\\' by (metis kth_obs_le_iff) hence cs\': \cs\<^bsup>\'\<^esup> \ \ cs\<^bsup>\\<^esup> \'\ unfolding cs\ using cs_inj[OF path(1) nreti, of \\'\] by blast thus \?thesis\ using less(1)[OF \k _ i\\' i\\] by auto next assume notin'': \\(\ \. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> \)\ obtain \ \' where \i: \\' < i\ and cs\: \cs\<^bsup>\'\<^esup> \ = cs\<^bsup>\\<^esup> \'\ and suc\: \\' (Suc \) \ \ (Suc \')\ and cdloop': \i cd\<^bsup>\\<^esup>\ \' \ (\ j>\. j cd\<^bsup>\'\<^esup>\ \)\ using missing_cd_or_loop[OF path \0 notin''] by metis show \?thesis\ using cdloop' proof assume cdjn: \i cd\<^bsup>\\<^esup>\ \'\ hence csnj: \cs\<^bsup>\'\<^esup> \ \ cs\<^bsup>\\<^esup> i\ using cs\ by (metis cd_is_cs_less) have cssel: \\ (Suc (\ \ cs\<^bsup>\'\<^esup> \)) = \ (Suc \')\ using cs\ by (metis cdjn cd_not_ret cs_select_id path(1)) have \(\',\) \ (\,i)\ using csnj apply(rule contradicts.intros(1)) using cssel \' \ suc\ by auto thus \?thesis\ using domi by auto next assume loop': \\ j>\. j cd\<^bsup>\'\<^esup>\ \\ have \n': \\' < n'\ using in' \i by auto have nret\': \\ \' \ return\ by (metis cs\ last_cs le_eq_less_or_eq lessI path(1) path(2) suc\ term_path_stable) have \\ < n\ using cs_order[OF path cs\[symmetric] csn[symmetric] nret\' \n'] . hence \\ < i'\ using nj ji' by auto hence cdi\: \i' cd\<^bsup>\'\<^esup>\ \\ using loop' by auto hence cs\i': \cs\<^bsup>\\<^esup> \' \ cs\<^bsup>\'\<^esup> i'\ using cs\ by (metis cd_is_cs_less) have cssel: \\' (Suc (\' \ cs\<^bsup>\\<^esup> \')) = \' (Suc \)\ using cs\ by (metis cdi\ cd_not_ret cs_select_id path(2)) have \(\,\') \ (\',i')\ using cs\i' apply(rule contradicts.intros(1)) using cssel \' \ suc\ by auto thus \?thesis\ using domi' by auto qed qed next assume \\ i \ n'\ hence ni: \n'< i\ by simp hence cdin: \i cd\<^bsup>\\<^esup>\ n'\ using loop by auto hence csni': \cs\<^bsup>\'\<^esup> n \ cs\<^bsup>\\<^esup> i\ using csn by (metis cd_is_cs_less) have cssel: \\ (Suc (\ \ cs\<^bsup>\'\<^esup> n)) = \ (Suc n')\ using csn by (metis cdin cd_not_ret cs_select_id path(1)) have \(\',n) \ (\,i)\ using csni' apply(rule contradicts.intros(1)) using cssel \' \ sucn by auto thus \?thesis\ using domi by auto qed qed qed qed qed theorem obs_neq_some_contradicts: fixes \ \' defines \: \\ \ path \\ and \': \\' \ path \'\ assumes obsne: \obs \ k \ obs \' k\ and not_none: \obs \ k \ None\ \obs \' k \ None\ shows \\ k k'. ((\', k') \ (\ ,k) \ \ k \ dom att) \ ((\, k) \ (\' ,k') \ \' k' \ dom att)\ proof - obtain i where iki: \is_kth_obs \ k i\ using not_none(1) by (metis \ obs_some_kth_obs) obtain i' where iki': \is_kth_obs \' k i'\ using not_none(2) by (metis \' obs_some_kth_obs) have \obsp \ i = obs \ k\ by (metis \ iki kth_obs_unique obs_def the_equality) moreover have \obsp \' i' = obs \' k\ by (metis \' iki' kth_obs_unique obs_def the_equality) ultimately have obspne: \obsp \ i \ obsp \' i'\ using obsne by auto show \?thesis\ using obs_neq_some_contradicts'[OF _ iki[unfolded \] iki'[unfolded \']] using obspne \ \' by metis qed theorem obs_neq_ret_contradicts: fixes \ \' defines \: \\ \ path \\ and \': \\' \ path \'\ assumes ret: \\ n = return\ and obsne: \obs \' i \ obs \ i\ and obs:\obs \' i \ None\ shows \\ k k'. ((\', k') \ (\ ,k) \ \ k \ dom (att)) \ ((\, k) \ (\' ,k') \ \' k' \ dom (att))\ proof (cases \\ j k'. is_kth_obs \' j k' \ (\ k. cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k')\) case True obtain l k' where jk': \is_kth_obs \' l k'\ and unmatched: \(\ k. cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k')\ using True by blast have \0: \\ 0 = \' 0\ using \ \' path0 by auto obtain j j' where csj: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ and cd: \k' cd\<^bsup>\'\<^esup>\j'\ and suc: \\ (Suc j) \ \' (Suc j')\ using converged_cd_diverge_return[of \\'\ \\\ \k'\ \n\] ret unmatched path_is_path \ \' \0 by metis hence *: \(\, j) \ (\' ,k')\ using contradicts.intros(1)[of \\\ \j\ \\'\ \k'\ \\'\ \\\, unfolded csj] \ \' using cd_is_cs_less cd_not_ret cs_select_id by auto have \\' k' \ dom(att)\ using jk' by (meson domIff is_kth_obs_def) thus \?thesis\ using * by blast next case False hence *: \\ j k'. is_kth_obs \' j k' \ \ k. cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ by auto obtain k' where k': \is_kth_obs \' i k'\ using obs \' obs_some_kth_obs by blast obtain l where \is_kth_obs \ i l\ using * \ \' k' no_kth_obs_missing_cs path_is_path by metis thus \?thesis\ using \ \' obs obs_neq_some_contradicts obs_none_no_kth_obs obsne by metis qed subsection \Facts about Critical Observable Paths\ lemma contradicting_in_cp: assumes leq:\\ =\<^sub>L \'\ and cseq: \cs\<^bsup>path \\<^esup> k = cs\<^bsup>path \'\<^esup> k'\ and readv: \v\reads(path \ k)\ and vneq: \(\\<^bsup>k\<^esup>) v \ (\'\<^bsup>k'\<^esup>) v\ shows \((\,k),(\',k')) \ cp\ using cseq readv vneq proof(induction \k+k'\ arbitrary: \k\ \k'\ \v\ rule: less_induct) fix k k' v assume csk: \cs\<^bsup>path \\<^esup> k = cs\<^bsup>path \'\<^esup> k'\ assume vread: \v \ reads (path \ k)\ assume vneq: \(\\<^bsup>k\<^esup>) v \ (\'\<^bsup>k'\<^esup>) v\ assume IH: \\ka k'a v. ka + k'a < k + k' \ cs\<^bsup>path \\<^esup> ka = cs\<^bsup>path \'\<^esup> k'a \ v \ reads (path \ ka) \ (\\<^bsup>ka\<^esup>) v \ (\'\<^bsup>k'a\<^esup>) v \ ((\, ka), \', k'a) \ cp\ define \ where \\ \ path \\ define \' where \\' \ path \'\ have path: \\ = path \\ \\' = path \'\ using \_def \'_def path_is_path by auto have ip: \is_path \\ \is_path \'\ using path path_is_path by auto have \0: \\' 0 = \ 0\ unfolding path path_def by auto have vread': \v \ reads (path \' k')\ using csk vread by (metis last_cs) have cseq: \cs\<^bsup>\'\<^esup> k' = cs\<^bsup>\\<^esup> k\ using csk path by simp show \((\, k), \', k') \ cp\ proof cases assume vnw: \\ l < k. v\writes (\ l)\ hence \v: \(\\<^bsup>k\<^esup>) v = \ v\ by (metis no_writes_unchanged0 path(1)) show \?thesis\ proof cases assume vnw': \\ l < k'. v\writes (\' l)\ hence \v': \(\'\<^bsup>k'\<^esup>) v = \' v\ by (metis no_writes_unchanged0 path(2)) with \v vneq have \\ v \ \' v\ by auto hence vhigh: \v \ hvars\ using leq unfolding loweq_def restrict_def by (auto,metis) thus \?thesis\ using cp.intros(1)[OF leq csk vread vneq] vnw vnw' path by simp next assume \\(\ l < k'. v\writes (\' l))\ then obtain l' where kddl': \k' dd\<^bsup>\',v\<^esup>\ l'\ using path(2) path_is_path written_read_dd vread' by blast hence lv': \v \ writes (\' l')\ unfolding is_ddi_def by auto have lk': \l' < k'\ by (metis is_ddi_def kddl') have nret: \\' l' \ return\ using lv' writes_return by auto have notin\: \\ (\l. cs\<^bsup>\'\<^esup> l' = cs\<^bsup>\\<^esup> l)\ proof assume \\l. cs\<^bsup>\'\<^esup> l' = cs\<^bsup>\\<^esup> l\ then guess l .. note csl = \cs\<^bsup>\'\<^esup> l' = cs\<^bsup>\\<^esup> l\ have lk: \l < k\ using lk' cseq ip cs_order[of \\'\ \\\ \l'\ \l\ \k'\ \k\] csl nret path by force have \v \ writes (\ l)\ using csl lv' last_cs by metis thus \False\ using lk vnw by blast qed from converged_cd_diverge[OF ip(2,1) \0 notin\ lk' cseq] obtain i i' where csi: \cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i\ and lcdi: \l' cd\<^bsup>\'\<^esup>\ i'\ and div: \\' (Suc i') \ \ (Suc i)\ . have 1: \\ (Suc i) = suc (\ i) (\\<^bsup>i\<^esup>)\ by (metis step_suc_sem fst_conv path(1) path_suc) have 2: \\' (Suc i') = suc (\' i') (\'\<^bsup>i'\<^esup>)\ by (metis step_suc_sem fst_conv path(2) path_suc) have 3: \\' i' = \ i\ using csi last_cs by metis have nreads: \\\<^bsup>i\<^esup> \ reads (\ i) \ \'\<^bsup>i'\<^esup> \ reads (\ i)\ by (metis 1 2 3 div reads_restr_suc) then obtain v' where v'read: \v'\ reads(path \ i)\ \(\\<^bsup>i\<^esup>) v' \ (\'\<^bsup>i'\<^esup>) v'\ unfolding path by (metis reads_restrict) have nreti: \\' i' \ return\ by (metis csi div ip(1) ip(2) last_cs lessI term_path_stable less_imp_le) have ik': \i' < k'\ using lcdi lk' is_cdi_def by auto have ik: \i < k\ using cs_order[OF ip(2,1) csi cseq nreti ik'] . have cpi: \((\, i), (\', i')) \ cp\ using IH[of \i\ \i'\] v'read csi ik ik' path by auto hence cpi': \((\', i'), (\, i)) \ cp\ using cp.intros(4) by blast have nwvi: \\j'\{LEAST i'. i < i' \ (\i. cs\<^bsup>path \'\<^esup> i = cs\<^bsup>path \\<^esup> i').. writes (path \ j')\ using vnw[unfolded path] by (metis (poly_guards_query) atLeastLessThan_iff) from cp.intros(3)[OF cpi' kddl'[unfolded path] lcdi[unfolded path] csk[symmetric] div[unfolded path] vneq[symmetric] nwvi] show \?thesis\ using cp.intros(4) by simp qed next assume wv: \\ (\ l writes (\ l))\ then obtain l where kddl: \k dd\<^bsup>\,v\<^esup>\ l\ using path(1) path_is_path written_read_dd vread by blast hence lv: \v \ writes (\ l)\ unfolding is_ddi_def by auto have lk: \l < k\ by (metis is_ddi_def kddl) have nret: \\ l \ return\ using lv writes_return by auto have nwb: \\ i \ {Suc l..< k}. v\writes(\ i)\ using kddl unfolding is_ddi_def by auto have \vk: \(\\<^bsup>k\<^esup>) v = (\\<^bsup>Suc l\<^esup> ) v\ using kddl ddi_value path(1) by auto show \?thesis\ proof cases assume vnw': \\ l < k'. v\writes (\' l)\ hence \v': \(\'\<^bsup>k'\<^esup>) v = \' v\ by (metis no_writes_unchanged0 path(2)) have notin\': \\ (\l'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l')\ proof assume \\l'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ then guess l' .. note csl = \cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ have lk: \l' < k'\ using lk cseq ip cs_order[of \\\ \\'\ \l\ \l'\ \k\ \k'\] csl nret by metis have \v \ writes (\' l')\ using csl lv last_cs by metis thus \False\ using lk vnw' by blast qed from converged_cd_diverge[OF ip(1,2) \0[symmetric] notin\' lk cseq[symmetric]] obtain i i' where csi: \cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i\ and lcdi: \l cd\<^bsup>\\<^esup>\ i\ and div: \\ (Suc i) \ \' (Suc i')\ by metis have 1: \\ (Suc i) = suc (\ i) (\\<^bsup>i\<^esup>)\ by (metis step_suc_sem fst_conv path(1) path_suc) have 2: \\' (Suc i') = suc (\' i') (\'\<^bsup>i'\<^esup>)\ by (metis step_suc_sem fst_conv path(2) path_suc) have 3: \\' i' = \ i\ using csi last_cs by metis have nreads: \\\<^bsup>i\<^esup> \ reads (\ i) \ \'\<^bsup>i'\<^esup> \ reads (\ i)\ by (metis 1 2 3 div reads_restr_suc) have contri: \(\',i') \ (\,i)\ using contradicts.intros(2)[OF csi path nreads] . have nreti: \\ i \ return\ by (metis csi div ip(1) ip(2) last_cs lessI term_path_stable less_imp_le) have ik: \i < k\ using lcdi lk is_cdi_def by auto have ik': \i' < k'\ using cs_order[OF ip(1,2) csi[symmetric] cseq[symmetric] nreti ik] . have nreads: \\\<^bsup>i\<^esup> \ reads (\ i) \ \'\<^bsup>i'\<^esup> \ reads (\ i)\ by (metis 1 2 3 div reads_restr_suc) then obtain v' where v'read: \v'\ reads(path \ i)\ \(\\<^bsup>i\<^esup>) v' \ (\'\<^bsup>i'\<^esup>) v'\ unfolding path by (metis reads_restrict) have cpi: \((\, i), (\', i')) \ cp\ using IH[of \i\ \i'\] v'read csi ik ik' path by auto hence cpi': \((\', i'), (\, i)) \ cp\ using cp.intros(4) by blast have vnwi: \\j'\{LEAST i'a. i' < i'a \ (\i. cs\<^bsup>path \\<^esup> i = cs\<^bsup>path \'\<^esup> i'a).. writes (path \' j')\ using vnw'[unfolded path] by (metis (poly_guards_query) atLeastLessThan_iff) from cp.intros(3)[OF cpi kddl[unfolded path] lcdi[unfolded path] csk div[unfolded path] vneq vnwi] show \?thesis\ using cp.intros(4) by simp next assume \\ (\ l writes (\' l))\ then obtain l' where kddl': \k' dd\<^bsup>\',v\<^esup>\ l'\ using path(2) path_is_path written_read_dd vread' by blast hence lv': \v \ writes (\' l')\ unfolding is_ddi_def by auto have lk': \l' < k'\ by (metis is_ddi_def kddl') have nretl': \\' l' \ return\ using lv' writes_return by auto have nwb': \\ i' \ {Suc l'..< k'}. v\writes(\' i')\ using kddl' unfolding is_ddi_def by auto have \vk': \(\'\<^bsup>k'\<^esup>) v = (\'\<^bsup>Suc l'\<^esup> ) v\ using kddl' ddi_value path(2) by auto show \?thesis\ proof cases assume csl: \cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ hence \l: \\ l = \' l'\ by (metis last_cs) have \vls: \(\\<^bsup>Suc l\<^esup> ) v \ (\'\<^bsup>Suc l'\<^esup> ) v\ by (metis \vk \vk' vneq) have r\: \\\<^bsup>l\<^esup> \ reads (\ l) \ \'\<^bsup>l'\<^esup> \ reads (\ l)\ using path \l \vls written_value lv by blast then obtain v' where v'read: \v'\ reads(path \ l)\ \(\\<^bsup>l\<^esup>) v' \ (\'\<^bsup>l'\<^esup>) v'\ unfolding path by (metis reads_restrict) have cpl: \((\, l), (\', l')) \ cp\ using IH[of \l\ \l'\] v'read csl lk lk' path by auto show \((\, k), (\', k')) \ cp\ using cp.intros(2)[OF cpl kddl[unfolded path] kddl'[unfolded path] csk vneq] . next assume csl: \cs\<^bsup>\\<^esup> l \ cs\<^bsup>\'\<^esup> l'\ show \?thesis\ proof cases assume \\ i'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> i'\ then obtain i' where csli': \cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> i'\ by blast have ilne': \i' \ l'\ using csl csli' by auto have ij': \i' < k'\ using cs_order[OF ip csli' cseq[symmetric] nret lk] . have iv': \v \ writes(\' i')\ using lv csli' last_cs by metis have il': \i' < l'\ using kddl' ilne' ij' iv' unfolding is_ddi_def by auto have nreti': \\' i' \ return\ using csli' nret last_cs by metis have l'notin\: \\(\i. cs\<^bsup>\'\<^esup> l' = cs\<^bsup>\\<^esup> i )\ proof assume \\i. cs\<^bsup>\'\<^esup> l' = cs\<^bsup>\\<^esup> i\ then obtain i where csil: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> l'\ by metis have ik: \i < k\ using cs_order[OF ip(2,1) csil[symmetric] cseq nretl' lk'] . have li: \l < i\ using cs_order[OF ip(2,1) csli'[symmetric] csil[symmetric] nreti' il'] . have iv: \v \ writes(\ i)\ using lv' csil last_cs by metis show \False\ using kddl ik li iv is_ddi_def by auto qed obtain n n' where csn: \cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and lcdn': \l' cd\<^bsup>\'\<^esup>\ n'\ and sucn: \\ (Suc n) \ \' (Suc n')\ and in': \i' \ n'\ using converged_cd_diverge_cs [OF ip(2,1) csli'[symmetric] il' l'notin\ lk' cseq] by metis \ \Can apply the IH to n and n'\ have 1: \\ (Suc n) = suc (\ n) (\\<^bsup>n\<^esup>)\ by (metis step_suc_sem fst_conv path(1) path_suc) have 2: \\' (Suc n') = suc (\' n') (\'\<^bsup>n'\<^esup>)\ by (metis step_suc_sem fst_conv path(2) path_suc) have 3: \\' n' = \ n\ using csn last_cs by metis have nreads: \\\<^bsup>n\<^esup> \ reads (\ n) \ \'\<^bsup>n'\<^esup> \ reads (\ n)\ by (metis 1 2 3 sucn reads_restr_suc) then obtain v' where v'read: \v'\reads (path \ n)\ \(\\<^bsup>n\<^esup>) v' \ (\'\<^bsup>n'\<^esup>) v'\ by (metis path(1) reads_restrict) moreover have nl': \n' < l'\ using lcdn' is_cdi_def by auto have nk': \n' < k'\ using nl' lk' by simp have nretn': \\' n' \ return\ by (metis ip(2) nl' nretl' term_path_stable less_imp_le) have nk: \n < k\ using cs_order[OF ip(2,1) csn[symmetric] cseq nretn' nk'] . hence lenn: \n+n' < k+k'\ using nk' by auto ultimately have \((\, n), (\', n')) \ cp\ using IH csn path by auto hence ncp: \((\', n'), (\, n)) \ cp\ using cp.intros(4) by auto have nles: \n < (LEAST i'. n < i' \ (\i. cs\<^bsup>\'\<^esup> i = cs\<^bsup>\\<^esup> i'))\ (is \_ < (LEAST i. ?P i)\) using nk cseq LeastI[of \?P\ \k\] by metis moreover have ln: \l \ n\ using cs_order_le[OF ip(2,1) csli'[symmetric] csn[symmetric] nreti' in'] . ultimately have lles: \Suc l \ (LEAST i'. n < i' \ (\i. cs\<^bsup>\'\<^esup> i = cs\<^bsup>\\<^esup> i'))\ by auto have nwcseq: \\j'\{LEAST i'. n < i' \ (\i. cs\<^bsup>\'\<^esup> i = cs\<^bsup>\\<^esup> i').. writes (\ j')\ proof fix j' assume *: \j' \ {LEAST i'. n < i' \ (\i. cs\<^bsup>\'\<^esup> i = cs\<^bsup>\\<^esup> i').. hence \(LEAST i'. n < i' \ (\i. cs\<^bsup>\'\<^esup> i = cs\<^bsup>\\<^esup> i')) \ j'\ by (metis (poly_guards_query) atLeastLessThan_iff) hence \Suc l \ j'\ using lles by auto moreover have \j' < k\ using * by (metis (poly_guards_query) atLeastLessThan_iff) ultimately have \j'\ {Suc l.. by (metis (poly_guards_query) atLeastLessThan_iff) thus \v\writes (\ j')\ using nwb by auto qed from cp.intros(3)[OF ncp,folded path,OF kddl' lcdn' cseq sucn[symmetric] vneq[symmetric] nwcseq] have \((\', k'), \, k) \ cp\ . thus \((\, k), (\', k')) \ cp\ using cp.intros(4) by auto next assume lnotin\': \\ (\i'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> i')\ show \?thesis\ proof cases assume \\ i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> l'\ then obtain i where csli: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> l'\ by blast have ilne: \i \ l\ using csl csli by auto have ij: \i < k\ using cs_order[OF ip(2,1) csli[symmetric] cseq nretl' lk'] . have iv: \v \ writes(\ i)\ using lv' csli last_cs by metis have il: \i < l\ using kddl ilne ij iv unfolding is_ddi_def by auto have nreti: \\ i \ return\ using csli nretl' last_cs by metis obtain n n' where csn: \cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and lcdn: \l cd\<^bsup>\\<^esup>\ n\ and sucn: \\ (Suc n) \ \' (Suc n')\ and ilen: \i \ n\ using converged_cd_diverge_cs [OF ip csli il lnotin\' lk cseq[symmetric]] by metis \ \Can apply the IH to n and n'\ have 1: \\ (Suc n) = suc (\ n) (\\<^bsup>n\<^esup>)\ by (metis step_suc_sem fst_conv path(1) path_suc) have 2: \\' (Suc n') = suc (\' n') (\'\<^bsup>n'\<^esup>)\ by (metis step_suc_sem fst_conv path(2) path_suc) have 3: \\' n' = \ n\ using csn last_cs by metis have nreads: \\\<^bsup>n\<^esup> \ reads (\ n) \ \'\<^bsup>n'\<^esup> \ reads (\ n)\ by (metis 1 2 3 sucn reads_restr_suc) then obtain v' where v'read: \v'\reads (path \ n)\ \(\\<^bsup>n\<^esup>) v' \ (\'\<^bsup>n'\<^esup>) v'\ by (metis path(1) reads_restrict) moreover have nl: \n < l\ using lcdn is_cdi_def by auto have nk: \n < k\ using nl lk by simp have nretn: \\ n \ return\ by (metis ip(1) nl nret term_path_stable less_imp_le) have nk': \n' < k'\ using cs_order[OF ip csn cseq[symmetric] nretn nk] . hence lenn: \n+n' < k+k'\ using nk by auto ultimately have ncp: \((\, n), (\', n')) \ cp\ using IH csn path by auto have nles': \n' < (LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'))\ (is \_ < (LEAST i. ?P i)\) using nk' cseq LeastI[of \?P\ \k'\] by metis moreover have ln': \l' \ n'\ using cs_order_le[OF ip csli csn nreti ilen] . ultimately have lles': \Suc l' \ (LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'))\ by auto have nwcseq': \\j'\{(LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i')).. writes (\' j')\ proof fix j' assume *: \j' \ {(LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i')).. hence \(LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i')) \ j'\ by (metis (poly_guards_query) atLeastLessThan_iff) hence \Suc l' \ j'\ using lles' by auto moreover have \j' < k'\ using * by (metis (poly_guards_query) atLeastLessThan_iff) ultimately have \j'\ {Suc l'.. by (metis (poly_guards_query) atLeastLessThan_iff) thus \v\writes (\' j')\ using nwb' by auto qed from cp.intros(3)[OF ncp,folded path, OF kddl lcdn cseq[symmetric] sucn vneq nwcseq'] show \((\, k), (\', k')) \ cp\ . next assume l'notin\: \\ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> l')\ define m where \m \ 0::nat\ define m' where \m' \ 0::nat\ have csm: \cs\<^bsup>\\<^esup> m = cs\<^bsup>\'\<^esup> m'\ unfolding m_def m'_def cs_0 by (metis \0) have ml: \m m' using csm csl unfolding m_def m'_def by (metis neq0_conv) have \\ n n'. cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n' \ \ (Suc n) \ \' (Suc n') \ (l cd\<^bsup>\\<^esup>\ n \ (\j'\{(LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'))..writes (\' j')) \ l' cd\<^bsup>\'\<^esup>\ n' \ (\j\{(LEAST i. n < i \ (\i'. cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i))..writes (\ j)))\ using csm ml proof (induction \k+k'-(m+m')\ arbitrary: \m\ \m'\ rule: less_induct) case (less m m') note csm = \cs\<^bsup>\\<^esup> m = cs\<^bsup>\'\<^esup> m'\ note lm = \m < l \ m' < l'\ note IH = \\ n n'. k + k' - (n + n') < k + k' - (m + m') \ cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n' \ n < l \ n' < l' \ ?thesis\ show \?thesis\ using lm proof assume ml: \m < l\ obtain n n' where mn: \m \ n\ and csn: \ cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and lcdn: \l cd\<^bsup>\\<^esup>\ n\ and suc: \\ (Suc n) \ \' (Suc n')\ using converged_cd_diverge_cs[OF ip csm ml lnotin\' lk cseq[symmetric]] . have nl: \n < l\ using lcdn is_cdi_def by auto hence nk: \n using lk by auto have nretn: \\ n \ return\ using lcdn by (metis cd_not_ret) have nk': \n' using cs_order[OF ip csn cseq[symmetric] nretn nk] . show \?thesis\ proof cases assume \\j'\{(LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'))..writes (\' j')\ thus \?thesis\ using lcdn csn suc by blast next assume \\(\j'\{(LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'))..writes (\' j'))\ then obtain j' where jin': \j'\{(LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i')).. and vwrite: \v\writes (\' j')\ by blast define i' where \i' \ LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i')\ have Pk': \n' < k' \ (\ k. cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k')\ (is \?P k'\) using nk' cseq[symmetric] by blast have ni': \n' < i'\ using LeastI[of \?P\, OF Pk'] i'_def by auto obtain i where csi: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ using LeastI[of \?P\, OF Pk'] i'_def by blast have ij': \i'\j'\ using jin'[folded i'_def] by auto have jk': \j' using jin'[folded i'_def] by auto have jl': \j' \ l'\ using kddl' jk' vwrite unfolding is_ddi_def by auto have nretn': \\' n' \ return\ using nretn csn last_cs by metis have iln: \n using cs_order[OF ip(2,1) csn[symmetric] csi[symmetric] nretn' ni'] . hence mi: \m < i\ using mn by auto have nretm: \\ m \ return\ by (metis ip(1) mn nretn term_path_stable) have mi': \m' using cs_order[OF ip csm csi nretm mi] . have ik': \i' < k'\ using ij' jk' by auto have nreti': \\' i' \ return\ by (metis ij' jl' nretl' ip(2) term_path_stable) have ik: \i < k\ using cs_order[OF ip(2,1) csi[symmetric] cseq nreti' ik'] . show \?thesis\ proof cases assume il:\i < l\ have le: \k + k' - (i +i') < k+k' - (m+m')\ using mi mi' ik ik' by auto show \?thesis\ using IH[OF le] using csi il by blast next assume \\ i < l\ hence li: \l \ i\ by auto have \i' \ l'\ using ij' jl' by auto hence il': \i' < l'\ using csi l'notin\ by fastforce obtain n n' where in': \i' \ n'\ and csn: \ cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and lcdn': \l' cd\<^bsup>\'\<^esup>\ n'\ and suc: \\ (Suc n) \ \' (Suc n')\ using converged_cd_diverge_cs[OF ip(2,1) csi[symmetric] il' _ lk' cseq] l'notin\ by metis have nk': \n' < k'\ using lcdn' is_cdi_def lk' by auto have nretn': \\' n' \ return\ by (metis cd_not_ret lcdn') have nk: \n < k\ using cs_order[OF ip(2,1) csn[symmetric] cseq nretn' nk'] . define j where \j \ LEAST j. n < j \ (\j'. cs\<^bsup>\'\<^esup> j' = cs\<^bsup>\\<^esup> j)\ have Pk: \n < k \ (\j'. cs\<^bsup>\'\<^esup> j' = cs\<^bsup>\\<^esup> k)\ (is \?P k\) using nk cseq by blast have nj: \n using LeastI[of \?P\, OF Pk] j_def by auto have ilen: \i \ n\ using cs_order_le[OF ip(2,1) csi[symmetric] csn[symmetric] nreti' in'] . hence lj: \l using li nj by simp have \\l\{l<.. writes (\ l)\ using kddl unfolding is_ddi_def by simp hence nw: \\l\{j.. writes (\ l)\ using lj by auto show \?thesis\ using csn lcdn' suc nw[unfolded j_def] by blast qed qed next assume ml': \m' < l'\ obtain n n' where mn': \m' \ n'\ and csn: \ cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and lcdn': \l' cd\<^bsup>\'\<^esup>\ n'\ and suc: \\ (Suc n) \ \' (Suc n')\ using converged_cd_diverge_cs[OF ip(2,1) csm[symmetric] ml' _ lk' cseq] l'notin\ by metis have nl': \n' < l'\ using lcdn' is_cdi_def by auto hence nk': \n' using lk' by auto have nretn': \\' n' \ return\ using lcdn' by (metis cd_not_ret) have nk: \n using cs_order[OF ip(2,1) csn[symmetric] cseq nretn' nk'] . show \?thesis\ proof cases assume \\j\{(LEAST i. n < i \ (\i'. cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i))..writes (\ j)\ thus \?thesis\ using lcdn' csn suc by blast next assume \\(\j\{(LEAST i. n < i \ (\i'. cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i))..writes (\ j))\ then obtain j where jin: \j\{(LEAST i. n < i \ (\i'. cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i)).. and vwrite: \v\writes (\ j)\ by blast define i where \i \ LEAST i. n < i \ (\i'. cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i)\ have Pk: \n < k \ (\ k'. cs\<^bsup>\'\<^esup> k' = cs\<^bsup>\\<^esup> k)\ (is \?P k\) using nk cseq by blast have ni: \n < i\ using LeastI[of \?P\, OF Pk] i_def by auto obtain i' where csi: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ using LeastI[of \?P\, OF Pk] i_def by metis have ij: \i\j\ using jin[folded i_def] by auto have jk: \j using jin[folded i_def] by auto have jl: \j \ l\ using kddl jk vwrite unfolding is_ddi_def by auto have nretn: \\ n \ return\ using nretn' csn last_cs by metis have iln': \n' using cs_order[OF ip csn csi nretn ni] . hence mi': \m' < i'\ using mn' by auto have nretm': \\' m' \ return\ by (metis ip(2) mn' nretn' term_path_stable) have mi: \m using cs_order[OF ip(2,1) csm[symmetric] csi[symmetric] nretm' mi'] . have ik: \i < k\ using ij jk by auto have nreti: \\ i \ return\ by (metis ij ip(1) jl nret term_path_stable) have ik': \i' < k'\ using cs_order[OF ip csi cseq[symmetric] nreti ik] . show \?thesis\ proof cases assume il':\i' < l'\ have le: \k + k' - (i +i') < k+k' - (m+m')\ using mi mi' ik ik' by auto show \?thesis\ using IH[OF le] using csi il' by blast next assume \\ i' < l'\ hence li': \l' \ i'\ by auto have \i \ l\ using ij jl by auto hence il: \i < l\ using csi lnotin\' by fastforce obtain n n' where ilen: \i \ n\ and csn: \ cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and lcdn: \l cd\<^bsup>\\<^esup>\ n\ and suc: \\ (Suc n) \ \' (Suc n')\ using converged_cd_diverge_cs[OF ip csi il _ lk cseq[symmetric]] lnotin\' by metis have nk: \n < k\ using lcdn is_cdi_def lk by auto have nretn: \\ n \ return\ by (metis cd_not_ret lcdn) have nk': \n' < k'\ using cs_order[OF ip csn cseq[symmetric] nretn nk] . define j' where \j' \ LEAST j'. n' < j' \ (\j. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ have Pk': \n' < k' \ (\j. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> k')\ (is \?P k'\) using nk' cseq[symmetric] by blast have nj': \n' using LeastI[of \?P\, OF Pk'] j'_def by auto have in': \i' \ n'\ using cs_order_le[OF ip csi csn nreti ilen] . hence lj': \l' using li' nj' by simp have \\l\{l'<.. writes (\' l)\ using kddl' unfolding is_ddi_def by simp hence nw': \\l\{j'.. writes (\' l)\ using lj' by auto show \?thesis\ using csn lcdn suc nw'[unfolded j'_def] by blast qed qed qed qed then obtain n n' where csn: \ cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and suc: \\ (Suc n) \ \' (Suc n')\ and cdor: \(l cd\<^bsup>\\<^esup>\ n \ (\j'\{(LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'))..writes (\' j')) \ l' cd\<^bsup>\'\<^esup>\ n' \ (\j\{(LEAST i. n < i \ (\i'. cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i))..writes (\ j)))\ by blast show \?thesis\ using cdor proof assume *: \l cd\<^bsup>\\<^esup>\ n \ (\j'\{LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i').. local.writes (\' j'))\ hence lcdn: \l cd\<^bsup>\\<^esup>\ n\ by blast have nowrite: \\j'\{LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i').. local.writes (\' j')\ using * by blast show \?thesis\ proof (rule cp.intros(3)[of \\\ \n\ \\'\ \n'\,folded path]) show \l cd\<^bsup>\\<^esup>\ n\ using lcdn . show \k dd\<^bsup>\,v\<^esup>\ l\ using kddl . show \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ using cseq by simp show \\ (Suc n) \ \' (Suc n')\ using suc by simp show \\j'\{LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i').. local.writes (\' j')\ using nowrite . show \(\\<^bsup>k\<^esup>) v \ (\'\<^bsup>k'\<^esup>) v\ using vneq . have nk: \n < k\ using lcdn lk is_cdi_def by auto have nretn: \\ n \ return\ using cd_not_ret lcdn by metis have nk': \n' < k'\ using cs_order[OF ip csn cseq[symmetric] nretn nk] . hence le: \n + n' < k + k'\ using nk by auto moreover have 1: \\ (Suc n) = suc (\ n) (\\<^bsup>n\<^esup>)\ by (metis step_suc_sem fst_conv path(1) path_suc) have 2: \\' (Suc n') = suc (\' n') (\'\<^bsup>n'\<^esup>)\ by (metis step_suc_sem fst_conv path(2) path_suc) have 3: \\' n' = \ n\ using csn last_cs by metis have nreads: \\\<^bsup>n\<^esup> \ reads (\ n) \ \'\<^bsup>n'\<^esup> \ reads (\ n)\ by (metis 1 2 3 suc reads_restr_suc) then obtain v' where v'read: \v'\reads (path \ n)\ \(\\<^bsup>n\<^esup>) v' \ (\'\<^bsup>n'\<^esup>) v'\ by (metis path(1) reads_restrict) ultimately show \((\, n), (\', n')) \ cp\ using IH csn path by auto qed next assume *: \l' cd\<^bsup>\'\<^esup>\ n' \ (\j\{(LEAST i. n < i \ (\i'. cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i))..writes (\ j))\ hence lcdn': \l' cd\<^bsup>\'\<^esup>\ n'\ by blast have nowrite: \\j\{(LEAST i. n < i \ (\i'. cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i))..writes (\ j)\ using * by blast show \?thesis\ proof (rule cp.intros(4), rule cp.intros(3)[of \\'\ \n'\ \\\ \n\,folded path]) show \l' cd\<^bsup>\'\<^esup>\ n'\ using lcdn' . show \k' dd\<^bsup>\',v\<^esup>\ l'\ using kddl' . show \cs\<^bsup>\'\<^esup> k' = cs\<^bsup>\\<^esup> k\ using cseq . show \\' (Suc n') \ \ (Suc n)\ using suc by simp show \\j\{(LEAST i. n < i \ (\i'. cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i))..writes (\ j)\ using nowrite . show \(\'\<^bsup>k'\<^esup>) v \ (\\<^bsup>k\<^esup>) v\ using vneq by simp have nk': \n' < k'\ using lcdn' lk' is_cdi_def by auto have nretn': \\' n' \ return\ using cd_not_ret lcdn' by metis have nk: \n < k\ using cs_order[OF ip(2,1) csn[symmetric] cseq nretn' nk'] . hence le: \n + n' < k + k'\ using nk' by auto moreover have 1: \\ (Suc n) = suc (\ n) (\\<^bsup>n\<^esup>)\ by (metis step_suc_sem fst_conv path(1) path_suc) have 2: \\' (Suc n') = suc (\' n') (\'\<^bsup>n'\<^esup>)\ by (metis step_suc_sem fst_conv path(2) path_suc) have 3: \\' n' = \ n\ using csn last_cs by metis have nreads: \\\<^bsup>n\<^esup> \ reads (\ n) \ \'\<^bsup>n'\<^esup> \ reads (\ n)\ by (metis 1 2 3 suc reads_restr_suc) then obtain v' where v'read: \v'\reads (path \ n)\ \(\\<^bsup>n\<^esup>) v' \ (\'\<^bsup>n'\<^esup>) v'\ by (metis path(1) reads_restrict) ultimately have \((\, n), (\', n')) \ cp\ using IH csn path by auto thus \((\', n'), \, n) \ cp\ using cp.intros(4) by simp qed qed qed qed qed qed qed qed theorem contradicting_in_cop: assumes \\ =\<^sub>L \'\ and \(\',k') \ (\,k)\ and \path \ k \ dom att\ shows \((\,k),\',k') \ cop\ using assms(2) proof(cases) case (1 \' \) define j where \j \ \ \ cs\<^bsup>\'\<^esup> k'\ have csj: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> k'\ unfolding j_def using 1 by (metis cs_not_nil cs_select_is_cs(1) path_is_path) have suc: \\ (Suc j) \ \' (Suc k')\ using 1 j_def by simp have kcdj: \k cd\<^bsup>\\<^esup>\ j\ by (metis cs_not_nil cs_select_is_cs(2) 1(1,2) j_def path_is_path) obtain v where readv: \v\reads(path \ j)\ and vneq: \(\\<^bsup>j\<^esup>) v \ (\'\<^bsup>k'\<^esup>) v\ using suc csj unfolding 1 by (metis IFC_def.suc_def 1(2) 1(3) last_cs path_suc reads_restr_suc reads_restrict) have \((\,j),\',k') \ cp\ apply (rule contradicting_in_cp[OF assms(1)]) using readv vneq csj 1 by auto thus \((\,k),\',k') \ cop\ using kcdj suc assms(3) cop.intros(2) unfolding 1 by auto next case (2 \' \) obtain v where readv: \v\reads(path \ k)\ and vneq: \(\\<^bsup>k\<^esup>) v \ (\'\<^bsup>k'\<^esup>) v\ using 2(2-4) by (metis reads_restrict) have \((\,k),\',k') \ cp\ apply (rule contradicting_in_cp[OF assms(1)]) using readv vneq 2 by auto thus \((\,k),\',k') \ cop\ using assms(3) cop.intros(1) unfolding 2 by auto qed theorem cop_correct_term: fixes \ \' defines \: \\ \ path \\ and \': \\' \ path \'\ assumes ret: \\ n = return\ \\' n' = return\ and obsne: \obs \ \ obs \'\ and leq: \\ =\<^sub>L \'\ shows \\ k k'. ((\,k),\',k')\ cop \ ((\',k'),\,k)\ cop\ proof - have *: \\ k k'. ((\', k') \ (\ ,k) \ \ k \ dom (att)) \ ((\, k) \ (\' ,k') \ \' k' \ dom (att))\ using obs_neq_contradicts_term ret obsne \ \' by auto have leq' :\\' =\<^sub>L \\ using leq unfolding loweq_def by auto from * contradicting_in_cop[OF leq] contradicting_in_cop[OF leq'] show \?thesis\ unfolding \ \' by metis qed theorem cop_correct_ret: fixes \ \' defines \: \\ \ path \\ and \': \\' \ path \'\ assumes ret: \\ n = return\ and obsne: \obs \ i \ obs \' i\ and obs: \obs \' i \ None\ and leq: \\ =\<^sub>L \'\ shows \\ k k'. ((\,k),\',k')\ cop \ ((\',k'),\,k)\ cop\ proof - have *: \\ k k'. ((\', k') \ (\ ,k) \ \ k \ dom (att)) \ ((\, k) \ (\' ,k') \ \' k' \ dom (att))\ by (metis (no_types, lifting) \ \' obs obs_neq_ret_contradicts obsne ret) have leq' :\\' =\<^sub>L \\ using leq unfolding loweq_def by auto from * contradicting_in_cop[OF leq] contradicting_in_cop[OF leq'] show \?thesis\ unfolding \ \' by metis qed theorem cop_correct_nterm: assumes obsne: \obs \ k \ obs \' k\ \obs \ k \ None\ \obs \' k \ None\ and leq: \\ =\<^sub>L \'\ shows \\ k k'. ((\,k),\',k')\ cop \ ((\',k'),\,k)\ cop\ proof - obtain k k' where \((\', k') \ (\ ,k) \ path \ k \ dom att) \ ((\, k) \ (\' ,k') \ path \' k' \ dom att)\ using obs_neq_some_contradicts[OF obsne] by metis thus \?thesis\ proof assume *: \(\', k') \ (\ ,k) \ path \ k \ dom att\ hence \((\,k),\',k') \ cop\ using leq by (metis contradicting_in_cop) thus \?thesis\ using * by blast next assume *: \(\, k) \ (\' ,k') \ path \' k' \ dom att\ hence \((\',k'),\,k) \ cop\ using leq by (metis contradicting_in_cop loweq_def) thus \?thesis\ using * by blast qed qed subsection \Correctness of the Characterisation\ text_raw \\label{sec:cor-cp}\ text \The following is our main correctness result. If there exist no critical observable paths, then the program is secure.\ theorem cop_correct: assumes \cop = empty\ shows \secure\ proof (rule ccontr) assume \\ secure\ then obtain \ \' where leq: \ \ =\<^sub>L \'\ and **: \\ obs \ \ obs \' \ (terminates \ \ \ obs \' \ obs \)\ unfolding secure_def by blast show \False\ using ** proof assume \\ obs \ \ obs \'\ then obtain k where \obs \ k \ obs \' k \ obs \ k \ None \ obs \' k \ None\ unfolding obs_comp_def obs_prefix_def by (metis kth_obs_stable linorder_neqE_nat obs_none_no_kth_obs obs_some_kth_obs) thus \False\ using cop_correct_nterm leq assms by auto next assume *: \terminates \ \ \ obs \' \ obs \\ then obtain n where ret: \path \ n = return\ unfolding terminates_def by auto obtain k where \obs \ k \ obs \' k \ obs \' k \ None\ using * unfolding obs_prefix_def by metis thus \False\ using cop_correct_ret ret leq assms by (metis empty_iff) qed qed text \Our characterisation is not only correct, it is also precise in the way that \cp\ characterises exactly the matching indices in executions for low equivalent input states where diverging data is read. This follows easily as the inverse implication to lemma \contradicting_in_cp\ can be shown by simple induction.\ theorem cp_iff_reads_contradict: \((\,k),(\',k')) \ cp \ \ =\<^sub>L \' \ cs\<^bsup>path \\<^esup> k = cs\<^bsup>path \'\<^esup> k' \ (\ v\reads(path \ k). (\\<^bsup>k\<^esup>) v \ (\'\<^bsup>k'\<^esup>) v)\ proof assume \\ =\<^sub>L \' \ cs\<^bsup>path \\<^esup> k = cs\<^bsup>path \'\<^esup> k' \ (\v\reads (path \ k). (\\<^bsup>k\<^esup>) v \ (\'\<^bsup>k'\<^esup>) v)\ thus \((\, k), \', k') \ cp\ using contradicting_in_cp by blast next assume \((\, k), \', k') \ cp\ thus \\ =\<^sub>L \' \ cs\<^bsup>path \\<^esup> k = cs\<^bsup>path \'\<^esup> k' \ (\v\reads (path \ k). (\\<^bsup>k\<^esup>) v \ (\'\<^bsup>k'\<^esup>) v)\ proof (induction) case (1 \ \' n n' h) then show \?case\ by blast next case (2 \ k \' k' n v n') have \v\reads (path \ n)\ using 2(2) unfolding is_ddi_def by auto then show \?case\ using 2 by auto next case (3 \ k \' k' n v l n') have \v\reads (path \ n)\ using 3(2) unfolding is_ddi_def by auto then show \?case\ using 3(4,6,8) by auto next case (4 \ k \' k') hence \cs\<^bsup>path \\<^esup> k = cs\<^bsup>path \'\<^esup> k'\ by simp hence \path \' k' = path \ k\ by (metis last_cs) moreover have \\' =\<^sub>L \\ using 4(2) unfolding loweq_def by simp ultimately show \?case\ using 4 by metis qed qed text \In the same way the inverse implication to \contradicting_in_cop\ follows easily such that we obtain the following characterisation of \cop\.\ theorem cop_iff_contradicting: \((\,k),(\',k')) \ cop \ \ =\<^sub>L \' \ (\',k') \ (\,k) \ path \ k \ dom att\ proof assume \\ =\<^sub>L \' \ (\', k') \ (\, k) \ path \ k \ dom att\ thus \((\,k),(\',k')) \ cop\ using contradicting_in_cop by simp next assume \((\,k),(\',k')) \ cop\ thus \ \ =\<^sub>L \' \ (\',k') \ (\,k) \ path \ k \ dom att\ proof (cases rule: cop.cases) case 1 then show \?thesis\ using cp_iff_reads_contradict contradicts.simps by (metis (full_types) reads_restrict1) next case (2 k) then show \?thesis\ using cp_iff_reads_contradict contradicts.simps by (metis cd_is_cs_less cd_not_ret contradicts.intros(1) cs_select_id path_is_path) qed qed subsection \Correctness of the Single Path Approximation\ text_raw \\label{sec:cor-scp}\ theorem cp_in_scp: assumes \((\,k),(\',k'))\cp\ shows \(path \,k)\scp \ (path \',k')\scp\ using assms proof(induction \\\ \k\ \\'\ \k'\ rule:cp.induct[case_names read_high dd dcd sym]) case (read_high \ \' k k' h) have \\ h = (\\<^bsup>k\<^esup>) h\ using read_high(5) by (simp add: no_writes_unchanged0) moreover have \\' h = (\'\<^bsup>k'\<^esup>) h\ using read_high(6) by (simp add: no_writes_unchanged0) ultimately have \\ h \ \' h\ using read_high(4) by simp hence *: \h\hvars\ using read_high(1) unfolding loweq_def by (metis Compl_iff IFC_def.restrict_def) have 1: \(path \,k)\scp\ using scp.intros(1) read_high(3,5) * by auto have \path \ k = path \' k'\ using read_high(2) by (metis last_cs) hence \(path \',k')\scp\ using scp.intros(1) read_high(3,6) * by auto thus \?case\ using 1 by auto next case dd show \?case\ using scp.intros(3) dd by auto next case sym thus \?case\ by blast next case (dcd \ k \' k' n v l n') note scp.intros(4) is_dcdi_via_def cd_cs_swap cs_ipd have 1: \(path \, n)\scp\ using dcd.IH dcd.hyps(2) dcd.hyps(3) scp.intros(2) scp.intros(3) by blast have csk: \cs\<^bsup>path \\<^esup> k = cs\<^bsup>path \'\<^esup> k'\ using cp_eq_cs[OF dcd(1)] . have kn: \k and kl: \k and ln: \l using dcd(2,3) unfolding is_ddi_def is_cdi_def by auto have nret: \path \ k \ return\ using cd_not_ret dcd.hyps(3) by auto have \k' < n'\ using kn csk dcd(4) cs_order nret path_is_path last_cs by blast have 2: \(path \', n')\scp\ proof cases assume j'ex: \\j'\{k'.. writes (path \' j')\ hence \\j'. j'\{k'.. v \ writes (path \' j')\ by auto note * = GreatestI_ex_nat[OF this] define j' where \j' == GREATEST j'. j'\{k'.. v \ writes (path \' j')\ note ** = *[of \j'\,folded j'_def] have \k' \ j'\ \j' and j'write: \v \ writes (path \' j')\ using "*" atLeastLessThan_iff j'_def nat_less_le by auto have nowrite: \\ i'\{j'<.. writes(path \' i')\ proof (rule, rule ccontr) fix i' assume \i' \ {j'<.. \\ v \ local.writes (path \' i')\ hence \i' \ {k'.. v \ local.writes (path \' i')\ using \k' \ j'\ by auto hence \i' \ j'\ using Greatest_le_nat by (metis (no_types, lifting) atLeastLessThan_iff j'_def nat_less_le) thus \False\ using \i' \ {j'<.. by auto qed have \path \' n' = path \ n\ using dcd(4) last_cs by metis hence \v\reads(path \' n')\ using dcd(2) unfolding is_ddi_def by auto hence nddj': \n' dd\<^bsup>path \',v\<^esup>\ j'\ using dcd(2) unfolding is_ddi_def using nowrite \j' j'write by auto show \?thesis\ proof cases assume \j' cd\<^bsup>path \'\<^esup>\ k'\ thus \(path \',n') \ scp\ using scp.intros(2) scp.intros(3) dcd.IH nddj' by fast next assume jcdk': \\ j' cd\<^bsup>path \'\<^esup>\ k'\ show \?thesis\ proof cases assume \j' = k'\ thus \?thesis\ using scp.intros(3) dcd.IH nddj' by fastforce next assume \j' \ k'\ hence \k' < j'\ using \k' \ j'\ by auto have \path \' j' \ return\ using j'write writes_return by auto hence ipdex':\\j. j \{k'..j'} \ path \' j = ipd (path \' k') \ using path_is_path \k' < j'\ jcdk' is_cdi_def by blast define i' where \i' == LEAST j. j\ {k'..j'} \ path \' j = ipd (path \' k')\ have iipd': \i'\ {k'..j'}\ \path \' i' = ipd (path \' k')\ unfolding i'_def using LeastI_ex[OF ipdex'] by simp_all have *:\\ i \ {k'..' i \ ipd (path \' k')\ proof (rule, rule ccontr) fix i assume *: \i \ {k'.. \\ path \' i \ ipd (path \' k')\ hence **: \i \{k'..j'} \ path \' i = ipd (path \' k')\ (is \?P i\) using iipd'(1) by auto thus \False\ using Least_le[of \?P\ \i\] i'_def * by auto qed have \i' \ k'\ using iipd'(2) by (metis csk last_cs nret path_in_nodes ipd_not_self) hence \k' using iipd'(1) by simp hence csi': \cs\<^bsup>path \'\<^esup> i' = [n\cs\<^bsup>path \'\<^esup> k' . ipd n \ path \' i'] @ [path \' i']\using cs_ipd[OF iipd'(2) *] by fast have ncdk': \\ n' cd\<^bsup>path \'\<^esup>\ k'\ using \j' < n'\ \k' < j'\ cdi_prefix jcdk' less_imp_le_nat by blast hence ncdk: \\ n cd\<^bsup>path \\<^esup>\ k\ using cd_cs_swap csk dcd(4) by blast have ipdex: \\i. i\{k..n} \ path \ i = ipd (path \ k)\ (is \\i. ?P i\) proof cases assume *:\path \ n = return\ from path_ret_ipd[of \path \\ \k\ \n\,OF path_is_path nret *] obtain i where \?P i\ by fastforce thus \?thesis\ by auto next assume *:\path \ n \ return\ show \?thesis\ using not_cd_impl_ipd [of \path \\ \k\ \n\, OF path_is_path \k ncdk *] by auto qed define i where \i == LEAST j. j\ {k..n} \ path \ j = ipd (path \ k)\ have iipd: \i\ {k..n}\ \path \ i = ipd (path \ k)\ unfolding i_def using LeastI_ex[OF ipdex] by simp_all have **:\\ i' \ {k.. i' \ ipd (path \ k)\ proof (rule, rule ccontr) fix i' assume *: \i' \ {k.. \\ path \ i' \ ipd (path \ k)\ hence **: \i' \{k..n} \ path \ i' = ipd (path \ k)\ (is \?P i'\) using iipd(1) by auto thus \False\ using Least_le[of \?P\ \i'\] i_def * by auto qed have \i \ k\ using iipd(2) by (metis nret path_in_nodes ipd_not_self) hence \k using iipd(1) by simp hence \cs\<^bsup>path \\<^esup> i = [n\cs\<^bsup>path \\<^esup> k . ipd n \ path \ i] @ [path \ i]\using cs_ipd[OF iipd(2) **] by fast hence csi: \cs\<^bsup>path \\<^esup> i = cs\<^bsup>path \'\<^esup> i'\ using csi' csk unfolding iipd'(2) iipd(2) by (metis last_cs) hence \(LEAST i'. k' < i' \ (\i. cs\<^bsup>path \\<^esup> i = cs\<^bsup>path \'\<^esup> i')) \ i'\ (is \(LEAST x. ?P x) \ _\) using \k' < i'\ Least_le[of \?P\ \i'\] by blast hence nw: \\j'\{i'.. writes (path \' j')\ using dcd(7) allB_atLeastLessThan_lower by blast moreover have \v \ writes (path \' j')\ using nddj' unfolding is_ddi_def by auto moreover have \i' \ j'\ using iipd'(1) by auto ultimately have \False\ using \j' < n'\ by auto thus \?thesis\ .. qed qed next assume \\ (\j'\{k'.. writes (path \' j'))\ hence \n' dcd\<^bsup>path \',v\<^esup>\ k' via (path \) k\ unfolding is_dcdi_via_def using dcd(2-4) csk \k' path_is_path by metis thus \?thesis\ using dcd.IH scp.intros(4) by blast qed with 1 show \?case\ .. qed theorem cop_in_scop: assumes \((\,k),(\',k'))\cop\ shows \(path \,k)\scop \ (path \',k')\scp\ using assms apply (induct rule: cop.induct) apply (simp add: cp_in_scp) using cp_in_scp scop.intros scp.intros(2) apply blast using cp_in_scp scop.intros scp.intros(2) apply blast done text \The main correctness result for out single execution approximation follows directly.\ theorem scop_correct: assumes \scop = empty\ shows \secure\ using cop_correct assms cop_in_scop by fast end end \ No newline at end of file diff --git a/thys/LTL/Disjunctive_Normal_Form.thy b/thys/LTL/Disjunctive_Normal_Form.thy --- a/thys/LTL/Disjunctive_Normal_Form.thy +++ b/thys/LTL/Disjunctive_Normal_Form.thy @@ -1,1092 +1,1092 @@ (* Author: Benedikt Seidl Author: Salomon Sickert License: BSD *) section \Disjunctive Normal Form of LTL formulas\ theory Disjunctive_Normal_Form imports LTL Equivalence_Relations "HOL-Library.FSet" begin text \ We use the propositional representation of LTL formulas to define the minimal disjunctive normal form of our formulas. For this purpose we define the minimal product \\\<^sub>m\ and union \\\<^sub>m\. In the end we show that for a set \\\ of literals, @{term "\ \\<^sub>P \"} if, and only if, there exists a subset of \\\ in the minimal DNF of \\\. \ subsection \Definition of Minimum Sets\ definition (in ord) min_set :: "'a set \ 'a set" where "min_set X = {y \ X. \x \ X. x \ y \ x = y}" lemma min_set_iff: "x \ min_set X \ x \ X \ (\y \ X. y \ x \ y = x)" unfolding min_set_def by blast lemma min_set_subset: "min_set X \ X" by (auto simp: min_set_def) lemma min_set_idem[simp]: "min_set (min_set X) = min_set X" by (auto simp: min_set_def) lemma min_set_empty[simp]: "min_set {} = {}" using min_set_subset by blast lemma min_set_singleton[simp]: "min_set {x} = {x}" by (auto simp: min_set_def) lemma min_set_finite: "finite X \ finite (min_set X)" by (simp add: min_set_def) lemma min_set_obtains_helper: "A \ B \ \C. C |\| A \ C \ min_set B" proof (induction "fcard A" arbitrary: A rule: less_induct) case less then have "(\A'. A' \ B \ \ A' |\| A \ A' = A) \ (\A'. A' |\| A \ A' \ min_set B)" by (metis (no_types) dual_order.trans order.not_eq_order_implies_strict pfsubset_fcard_mono) then show ?case using less.prems min_set_def by auto qed lemma min_set_obtains: assumes "A \ B" obtains C where "C |\| A" and "C \ min_set B" using min_set_obtains_helper assms by metis subsection \Minimal operators on sets\ definition product :: "'a fset set \ 'a fset set \ 'a fset set" (infixr "\" 65) where "A \ B = {a |\| b | a b. a \ A \ b \ B}" definition min_product :: "'a fset set \ 'a fset set \ 'a fset set" (infixr "\\<^sub>m" 65) where "A \\<^sub>m B = min_set (A \ B)" definition min_union :: "'a fset set \ 'a fset set \ 'a fset set" (infixr "\\<^sub>m" 65) where "A \\<^sub>m B = min_set (A \ B)" definition product_set :: "'a fset set set \ 'a fset set" ("\") where "\ X = Finite_Set.fold product {{||}} X" definition min_product_set :: "'a fset set set \ 'a fset set" ("\\<^sub>m") where "\\<^sub>m X = Finite_Set.fold min_product {{||}} X" lemma min_product_idem[simp]: "A \\<^sub>m A = min_set A" by (auto simp: min_product_def product_def min_set_def) fastforce lemma min_union_idem[simp]: "A \\<^sub>m A = min_set A" by (simp add: min_union_def) lemma product_empty[simp]: "A \ {} = {}" "{} \ A = {}" by (simp_all add: product_def) lemma min_product_empty[simp]: "A \\<^sub>m {} = {}" "{} \\<^sub>m A = {}" by (simp_all add: min_product_def) lemma min_union_empty[simp]: "A \\<^sub>m {} = min_set A" "{} \\<^sub>m A = min_set A" by (simp_all add: min_union_def) lemma product_empty_singleton[simp]: "A \ {{||}} = A" "{{||}} \ A = A" by (simp_all add: product_def) lemma min_product_empty_singleton[simp]: "A \\<^sub>m {{||}} = min_set A" "{{||}} \\<^sub>m A = min_set A" by (simp_all add: min_product_def) lemma product_singleton_singleton: "A \ {{|x|}} = finsert x ` A" "{{|x|}} \ A = finsert x ` A" unfolding product_def by blast+ lemma product_mono: "A \ B \ A \ C \ B \ C" "B \ C \ A \ B \ A \ C" unfolding product_def by auto lemma product_finite: "finite A \ finite B \ finite (A \ B)" by (simp add: product_def finite_image_set2) lemma min_product_finite: "finite A \ finite B \ finite (A \\<^sub>m B)" by (metis min_product_def product_finite min_set_finite) lemma min_union_finite: "finite A \ finite B \ finite (A \\<^sub>m B)" by (simp add: min_union_def min_set_finite) lemma product_set_infinite[simp]: "infinite X \ \ X = {{||}}" by (simp add: product_set_def) lemma min_product_set_infinite[simp]: "infinite X \ \\<^sub>m X = {{||}}" by (simp add: min_product_set_def) lemma product_comm: "A \ B = B \ A" unfolding product_def by blast lemma min_product_comm: "A \\<^sub>m B = B \\<^sub>m A" unfolding min_product_def by (simp add: product_comm) lemma min_union_comm: "A \\<^sub>m B = B \\<^sub>m A" unfolding min_union_def by (simp add: sup.commute) lemma product_iff: "x \ A \ B \ (\a \ A. \b \ B. x = a |\| b)" unfolding product_def by blast lemma min_product_iff: "x \ A \\<^sub>m B \ (\a \ A. \b \ B. x = a |\| b) \ (\a \ A. \b \ B. a |\| b |\| x \ a |\| b = x)" unfolding min_product_def min_set_iff product_iff product_def by blast lemma min_union_iff: "x \ A \\<^sub>m B \ x \ A \ B \ (\a \ A. a |\| x \ a = x) \ (\b \ B. b |\| x \ b = x)" unfolding min_union_def min_set_iff by blast lemma min_set_min_product_helper: "x \ (min_set A) \\<^sub>m B \ x \ A \\<^sub>m B" proof fix x assume "x \ (min_set A) \\<^sub>m B" then obtain a b where "a \ min_set A" and "b \ B" and "x = a |\| b" and 1: "\a \ min_set A. \b \ B. a |\| b |\| x \ a |\| b = x" unfolding min_product_iff by blast moreover { fix a' b' assume "a' \ A" and "b' \ B" and "a' |\| b' |\| x" then obtain a'' where "a'' |\| a'" and "a'' \ min_set A" using min_set_obtains by metis then have "a'' |\| b' = x" by (metis (full_types) 1 \b' \ B\ \a' |\| b' |\| x\ dual_order.trans le_sup_iff) then have "a' |\| b' = x" using \a' |\| b' |\| x\ \a'' |\| a'\ by blast } ultimately show "x \ A \\<^sub>m B" by (metis min_product_iff min_set_iff) next fix x assume "x \ A \\<^sub>m B" then have 1: "x \ A \ B" and "\y \ A \ B. y |\| x \ y = x" unfolding min_product_def min_set_iff by simp+ then have 2: "\y\min_set A \ B. y |\| x \ y = x" by (metis product_iff min_set_iff) then have "x \ min_set A \ B" by (metis 1 funion_mono min_set_obtains order_refl product_iff) then show "x \ min_set A \\<^sub>m B" by (simp add: 2 min_product_def min_set_iff) qed lemma min_set_min_product[simp]: "(min_set A) \\<^sub>m B = A \\<^sub>m B" "A \\<^sub>m (min_set B) = A \\<^sub>m B" using min_product_comm min_set_min_product_helper by blast+ lemma min_set_min_union[simp]: "(min_set A) \\<^sub>m B = A \\<^sub>m B" "A \\<^sub>m (min_set B) = A \\<^sub>m B" proof (unfold min_union_def min_set_def, safe) show "\x xa xb. \\xa\{y \ A. \x\A. x |\| y \ x = y} \ B. xa |\| x \ xa = x; x \ B; xa |\| x; xb |\| x; xa \ A\ \ xb |\| xa" by (metis (mono_tags) UnCI dual_order.trans fequalityI min_set_def min_set_obtains) next show "\x xa xb. \\xa\A \ {y \ B. \x\B. x |\| y \ x = y}. xa |\| x \ xa = x; x \ A; xa |\| x; xb |\| x; xa \ B\ \ xb |\| xa" by (metis (mono_tags) UnCI dual_order.trans fequalityI min_set_def min_set_obtains) qed blast+ lemma product_assoc[simp]: "(A \ B) \ C = A \ (B \ C)" proof (unfold product_def, safe) fix a b c assume "a \ A" and "c \ C" and "b \ B" then have "b |\| c \ {b |\| c |b c. b \ B \ c \ C}" by blast then show "\a' bc. a |\| b |\| c = a' |\| bc \ a' \ A \ bc \ {b |\| c |b c. b \ B \ c \ C}" using \a \ A\ by (metis (no_types) inf_sup_aci(5) sup_left_commute) qed (metis (mono_tags, lifting) mem_Collect_eq sup_assoc) lemma min_product_assoc[simp]: "(A \\<^sub>m B) \\<^sub>m C = A \\<^sub>m (B \\<^sub>m C)" unfolding min_product_def[of A B] min_product_def[of B C] by simp (simp add: min_product_def) lemma min_union_assoc[simp]: "(A \\<^sub>m B) \\<^sub>m C = A \\<^sub>m (B \\<^sub>m C)" unfolding min_union_def[of A B] min_union_def[of B C] by simp (simp add: min_union_def sup_assoc) lemma min_product_comp: "a \ A \ b \ B \ \c. c |\| (a |\| b) \ c \ A \\<^sub>m B" by (metis (mono_tags, lifting) mem_Collect_eq min_product_def product_def min_set_obtains) lemma min_union_comp: "a \ A \ \c. c |\| a \ c \ A \\<^sub>m B" by (metis Un_iff min_set_obtains min_union_def) interpretation product_set_thms: Finite_Set.comp_fun_commute product proof unfold_locales have "\x y z. x \ (y \ z) = y \ (x \ z)" by (simp only: product_assoc[symmetric]) (simp only: product_comm) then show "\x y. (\) y \ (\) x = (\) x \ (\) y" by fastforce qed interpretation min_product_set_thms: Finite_Set.comp_fun_idem min_product proof unfold_locales have "\x y z. x \\<^sub>m (y \\<^sub>m z) = y \\<^sub>m (x \\<^sub>m z)" by (simp only: min_product_assoc[symmetric]) (simp only: min_product_comm) then show "\x y. (\\<^sub>m) y \ (\\<^sub>m) x = (\\<^sub>m) x \ (\\<^sub>m) y" by fastforce next have "\x y. x \\<^sub>m (x \\<^sub>m y) = x \\<^sub>m y" by (simp add: min_product_assoc[symmetric]) then show "\x. (\\<^sub>m) x \ (\\<^sub>m) x = (\\<^sub>m) x" by fastforce qed interpretation min_union_set_thms: Finite_Set.comp_fun_idem min_union proof unfold_locales have "\x y z. x \\<^sub>m (y \\<^sub>m z) = y \\<^sub>m (x \\<^sub>m z)" by (simp only: min_union_assoc[symmetric]) (simp only: min_union_comm) then show "\x y. (\\<^sub>m) y \ (\\<^sub>m) x = (\\<^sub>m) x \ (\\<^sub>m) y" by fastforce next have "\x y. x \\<^sub>m (x \\<^sub>m y) = x \\<^sub>m y" by (simp add: min_union_assoc[symmetric]) then show "\x. (\\<^sub>m) x \ (\\<^sub>m) x = (\\<^sub>m) x" by fastforce qed lemma product_set_empty[simp]: "\ {} = {{||}}" "\ {{}} = {}" "\ {{{||}}} = {{||}}" by (simp_all add: product_set_def) lemma min_product_set_empty[simp]: "\\<^sub>m {} = {{||}}" "\\<^sub>m {{}} = {}" "\\<^sub>m {{{||}}} = {{||}}" by (simp_all add: min_product_set_def) lemma product_set_code[code]: "\ (set xs) = fold product (remdups xs) {{||}}" by (simp add: product_set_def product_set_thms.fold_set_fold_remdups) lemma min_product_set_code[code]: "\\<^sub>m (set xs) = fold min_product (remdups xs) {{||}}" by (simp add: min_product_set_def min_product_set_thms.fold_set_fold_remdups) lemma product_set_insert[simp]: "finite X \ \ (insert x X) = x \ (\ (X - {x}))" unfolding product_set_def product_set_thms.fold_insert_remove .. lemma min_product_set_insert[simp]: "finite X \ \\<^sub>m (insert x X) = x \\<^sub>m (\\<^sub>m X)" unfolding min_product_set_def min_product_set_thms.fold_insert_idem .. lemma min_product_subseteq: "x \ A \\<^sub>m B \ \a. a |\| x \ a \ A" by (metis funion_upper1 min_product_iff) lemma min_product_set_subseteq: "finite X \ x \ \\<^sub>m X \ A \ X \ \a \ A. a |\| x" by (induction X rule: finite_induct) (blast, metis finite_insert insert_absorb min_product_set_insert min_product_subseteq) lemma min_set_product_set: "\\<^sub>m A = min_set (\ A)" by (cases "finite A", induction A rule: finite_induct) (simp_all add: min_product_set_def product_set_def, metis min_product_def) lemma min_product_min_set[simp]: "min_set (A \\<^sub>m B) = A \\<^sub>m B" by (simp add: min_product_def) lemma min_union_min_set[simp]: "min_set (A \\<^sub>m B) = A \\<^sub>m B" by (simp add: min_union_def) lemma min_product_set_min_set[simp]: "finite X \ min_set (\\<^sub>m X) = \\<^sub>m X" by (induction X rule: finite_induct, auto simp add: min_product_set_def min_set_iff) lemma min_set_min_product_set[simp]: "finite X \ \\<^sub>m (min_set ` X) = \\<^sub>m X" by (induction X rule: finite_induct) simp_all lemma min_product_set_union[simp]: "finite X \ finite Y \ \\<^sub>m (X \ Y) = (\\<^sub>m X) \\<^sub>m (\\<^sub>m Y)" by (induction X rule: finite_induct) simp_all lemma product_set_finite: "(\x. x \ X \ finite x) \ finite (\ X)" by (cases "finite X", rotate_tac, induction X rule: finite_induct) (simp_all add: product_set_def, insert product_finite, blast) lemma min_product_set_finite: "(\x. x \ X \ finite x) \ finite (\\<^sub>m X)" by (cases "finite X", rotate_tac, induction X rule: finite_induct) (simp_all add: min_product_set_def, insert min_product_finite, blast) subsection \Disjunctive Normal Form\ fun dnf :: "'a ltln \ 'a ltln fset set" where "dnf true\<^sub>n = {{||}}" | "dnf false\<^sub>n = {}" | "dnf (\ and\<^sub>n \) = (dnf \) \ (dnf \)" | "dnf (\ or\<^sub>n \) = (dnf \) \ (dnf \)" | "dnf \ = {{|\|}}" fun min_dnf :: "'a ltln \ 'a ltln fset set" where "min_dnf true\<^sub>n = {{||}}" | "min_dnf false\<^sub>n = {}" | "min_dnf (\ and\<^sub>n \) = (min_dnf \) \\<^sub>m (min_dnf \)" | "min_dnf (\ or\<^sub>n \) = (min_dnf \) \\<^sub>m (min_dnf \)" | "min_dnf \ = {{|\|}}" lemma dnf_min_set: "min_dnf \ = min_set (dnf \)" by (induction \) (simp_all, simp_all only: min_product_def min_union_def) lemma dnf_finite: "finite (dnf \)" by (induction \) (auto simp: product_finite) lemma min_dnf_finite: "finite (min_dnf \)" by (induction \) (auto simp: min_product_finite min_union_finite) lemma dnf_Abs_fset[simp]: "fset (Abs_fset (dnf \)) = dnf \" by (simp add: dnf_finite Abs_fset_inverse) lemma min_dnf_Abs_fset[simp]: "fset (Abs_fset (min_dnf \)) = min_dnf \" by (simp add: min_dnf_finite Abs_fset_inverse) lemma dnf_prop_atoms: "\ \ dnf \ \ fset \ \ prop_atoms \" by (induction \ arbitrary: \) (auto simp: product_def, blast+) lemma min_dnf_prop_atoms: "\ \ min_dnf \ \ fset \ \ prop_atoms \" using dnf_min_set dnf_prop_atoms min_set_subset by blast lemma min_dnf_atoms_dnf: "\ \ min_dnf \ \ \ \ fset \ \ dnf \ = {{|\|}}" proof (induction \) case True_ltln then show ?case using min_dnf_prop_atoms prop_atoms_notin(1) by blast next case False_ltln then show ?case using min_dnf_prop_atoms prop_atoms_notin(2) by blast next case (And_ltln \1 \2) then show ?case using min_dnf_prop_atoms prop_atoms_notin(3) by force next case (Or_ltln \1 \2) then show ?case using min_dnf_prop_atoms prop_atoms_notin(4) by force qed auto lemma min_dnf_min_set[simp]: "min_set (min_dnf \) = min_dnf \" by (induction \) (simp_all add: min_set_def min_product_def min_union_def, blast+) lemma min_dnf_iff_prop_assignment_subset: "\ \\<^sub>P \ \ (\B. fset B \ \ \ B \ min_dnf \)" proof assume "\ \\<^sub>P \" then show "\B. fset B \ \ \ B \ min_dnf \" proof (induction \ arbitrary: \) case (And_ltln \\<^sub>1 \\<^sub>2) then obtain B\<^sub>1 B\<^sub>2 where 1: "fset B\<^sub>1 \ \ \ B\<^sub>1 \ min_dnf \\<^sub>1" and 2: "fset B\<^sub>2 \ \ \ B\<^sub>2 \ min_dnf \\<^sub>2" by fastforce then obtain C where "C |\| B\<^sub>1 |\| B\<^sub>2" and "C \ min_dnf \\<^sub>1 \\<^sub>m min_dnf \\<^sub>2" using min_product_comp by metis then show ?case by (metis 1 2 le_sup_iff min_dnf.simps(3) sup.absorb_iff1 sup_fset.rep_eq) next case (Or_ltln \\<^sub>1 \\<^sub>2) { assume "\ \\<^sub>P \\<^sub>1" then obtain B where 1: "fset B \ \ \ B \ min_dnf \\<^sub>1" using Or_ltln by fastforce then obtain C where "C |\| B" and "C \ min_dnf \\<^sub>1 \\<^sub>m min_dnf \\<^sub>2" using min_union_comp by metis then have ?case by (metis 1 dual_order.trans less_eq_fset.rep_eq min_dnf.simps(4)) } moreover { assume "\ \\<^sub>P \\<^sub>2" then obtain B where 2: "fset B \ \ \ B \ min_dnf \\<^sub>2" using Or_ltln by fastforce then obtain C where "C |\| B" and "C \ min_dnf \\<^sub>1 \\<^sub>m min_dnf \\<^sub>2" using min_union_comp min_union_comm by metis then have ?case by (metis 2 dual_order.trans less_eq_fset.rep_eq min_dnf.simps(4)) } ultimately show ?case using Or_ltln.prems by auto qed simp_all next assume "\B. fset B \ \ \ B \ min_dnf \" then obtain B where "fset B \ \" and "B \ min_dnf \" by auto then have "fset B \\<^sub>P \" by (induction \ arbitrary: B) (auto simp: min_set_def min_product_def product_def min_union_def, blast+) then show "\ \\<^sub>P \" using \fset B \ \\ by blast qed lemma ltl_prop_implies_min_dnf: "\ \\<^sub>P \ = (\A \ min_dnf \. \B \ min_dnf \. B |\| A)" by (meson less_eq_fset.rep_eq ltl_prop_implies_def min_dnf_iff_prop_assignment_subset order_refl dual_order.trans) lemma ltl_prop_equiv_min_dnf: "\ \\<^sub>P \ = (min_dnf \ = min_dnf \)" proof assume "\ \\<^sub>P \" then have "\x. x \ min_set (min_dnf \) \ x \ min_set (min_dnf \)" unfolding ltl_prop_implies_equiv ltl_prop_implies_min_dnf min_set_iff by fastforce then show "min_dnf \ = min_dnf \" by auto qed (simp add: ltl_prop_equiv_def min_dnf_iff_prop_assignment_subset) lemma min_dnf_rep_abs[simp]: "min_dnf (rep_ltln\<^sub>P (abs_ltln\<^sub>P \)) = min_dnf \" by (simp add: ltl_prop_equiv_min_dnf[symmetric] Quotient3_ltln\<^sub>P rep_abs_rsp_left) subsection \Folding of \and\<^sub>n\ and \or\<^sub>n\ over Finite Sets\ definition And\<^sub>n :: "'a ltln set \ 'a ltln" where "And\<^sub>n \ \ SOME \. fold_graph And_ltln True_ltln \ \" definition Or\<^sub>n :: "'a ltln set \ 'a ltln" where "Or\<^sub>n \ \ SOME \. fold_graph Or_ltln False_ltln \ \" lemma fold_graph_And\<^sub>n: "finite \ \ fold_graph And_ltln True_ltln \ (And\<^sub>n \)" unfolding And\<^sub>n_def by (rule someI2_ex[OF finite_imp_fold_graph]) lemma fold_graph_Or\<^sub>n: "finite \ \ fold_graph Or_ltln False_ltln \ (Or\<^sub>n \)" unfolding Or\<^sub>n_def by (rule someI2_ex[OF finite_imp_fold_graph]) lemma Or\<^sub>n_empty[simp]: "Or\<^sub>n {} = False_ltln" by (metis empty_fold_graphE finite.emptyI fold_graph_Or\<^sub>n) lemma And\<^sub>n_empty[simp]: "And\<^sub>n {} = True_ltln" by (metis empty_fold_graphE finite.emptyI fold_graph_And\<^sub>n) interpretation dnf_union_thms: Finite_Set.comp_fun_commute "\\. (\) (f \)" by unfold_locales fastforce interpretation dnf_product_thms: Finite_Set.comp_fun_commute "\\. (\) (f \)" by unfold_locales (simp add: product_set_thms.comp_fun_commute) \ \Copied from locale @{locale comp_fun_commute}\ lemma fold_graph_finite: assumes "fold_graph f z A y" shows "finite A" using assms by induct simp_all text \Taking the DNF of @{const And\<^sub>n} and @{const Or\<^sub>n} is the same as folding over the individual DNFs.\ lemma And\<^sub>n_dnf: "finite \ \ dnf (And\<^sub>n \) = Finite_Set.fold (\\. (\) (dnf \)) {{||}} \" proof (drule fold_graph_And\<^sub>n, induction rule: fold_graph.induct) case (insertI x A y) then have "finite A" using fold_graph_finite by fast then show ?case using insertI by auto qed simp lemma Or\<^sub>n_dnf: "finite \ \ dnf (Or\<^sub>n \) = Finite_Set.fold (\\. (\) (dnf \)) {} \" proof (drule fold_graph_Or\<^sub>n, induction rule: fold_graph.induct) case (insertI x A y) then have "finite A" using fold_graph_finite by fast then show ?case using insertI by auto qed simp text \@{const And\<^sub>n} and @{const Or\<^sub>n} are injective on finite sets.\ lemma And\<^sub>n_inj: "inj_on And\<^sub>n {s. finite s}" proof (standard, simp) fix x y :: "'a ltln set" assume "finite x" and "finite y" then have 1: "fold_graph And_ltln True_ltln x (And\<^sub>n x)" and 2: "fold_graph And_ltln True_ltln y (And\<^sub>n y)" using fold_graph_And\<^sub>n by blast+ assume "And\<^sub>n x = And\<^sub>n y" with 1 show "x = y" proof (induction rule: fold_graph.induct) case emptyI then show ?case using 2 fold_graph.cases by force next case (insertI x A y) with 2 show ?case proof (induction arbitrary: x A y rule: fold_graph.induct) case (insertI x A y) then show ?case by (metis fold_graph.cases insertI1 ltln.distinct(7) ltln.inject(3)) qed blast qed qed lemma Or\<^sub>n_inj: "inj_on Or\<^sub>n {s. finite s}" proof (standard, simp) fix x y :: "'a ltln set" assume "finite x" and "finite y" then have 1: "fold_graph Or_ltln False_ltln x (Or\<^sub>n x)" and 2: "fold_graph Or_ltln False_ltln y (Or\<^sub>n y)" using fold_graph_Or\<^sub>n by blast+ assume "Or\<^sub>n x = Or\<^sub>n y" with 1 show "x = y" proof (induction rule: fold_graph.induct) case emptyI then show ?case using 2 fold_graph.cases by force next case (insertI x A y) with 2 show ?case proof (induction arbitrary: x A y rule: fold_graph.induct) case (insertI x A y) then show ?case by (metis fold_graph.cases insertI1 ltln.distinct(27) ltln.inject(4)) qed blast qed qed text \The semantics of @{const And\<^sub>n} and @{const Or\<^sub>n} can be expressed using quantifiers.\ lemma And\<^sub>n_semantics: "finite \ \ w \\<^sub>n And\<^sub>n \ \ (\\ \ \. w \\<^sub>n \)" proof - assume "finite \" have "\\. fold_graph And_ltln True_ltln \ \ \ w \\<^sub>n \ \ (\\ \ \. w \\<^sub>n \)" by (rule fold_graph.induct) auto then show ?thesis using fold_graph_And\<^sub>n[OF \finite \\] by simp qed lemma Or\<^sub>n_semantics: "finite \ \ w \\<^sub>n Or\<^sub>n \ \ (\\ \ \. w \\<^sub>n \)" proof - assume "finite \" have "\\. fold_graph Or_ltln False_ltln \ \ \ w \\<^sub>n \ \ (\\ \ \. w \\<^sub>n \)" by (rule fold_graph.induct) auto then show ?thesis using fold_graph_Or\<^sub>n[OF \finite \\] by simp qed lemma And\<^sub>n_prop_semantics: "finite \ \ \ \\<^sub>P And\<^sub>n \ \ (\\ \ \. \ \\<^sub>P \)" proof - assume "finite \" have "\\. fold_graph And_ltln True_ltln \ \ \ \ \\<^sub>P \ \ (\\ \ \. \ \\<^sub>P \)" by (rule fold_graph.induct) auto then show ?thesis using fold_graph_And\<^sub>n[OF \finite \\] by simp qed lemma Or\<^sub>n_prop_semantics: "finite \ \ \ \\<^sub>P Or\<^sub>n \ \ (\\ \ \. \ \\<^sub>P \)" proof - assume "finite \" have "\\. fold_graph Or_ltln False_ltln \ \ \ \ \\<^sub>P \ \ (\\ \ \. \ \\<^sub>P \)" by (rule fold_graph.induct) auto then show ?thesis using fold_graph_Or\<^sub>n[OF \finite \\] by simp qed lemma Or\<^sub>n_And\<^sub>n_image_semantics: assumes "finite \" and "\\. \ \ \ \ finite \" shows "w \\<^sub>n Or\<^sub>n (And\<^sub>n ` \) \ (\\ \ \. \\ \ \. w \\<^sub>n \)" proof - have "w \\<^sub>n Or\<^sub>n (And\<^sub>n ` \) \ (\\ \ \. w \\<^sub>n And\<^sub>n \)" using Or\<^sub>n_semantics assms by auto then show ?thesis using And\<^sub>n_semantics assms by fast qed lemma Or\<^sub>n_And\<^sub>n_image_prop_semantics: assumes "finite \" and "\\. \ \ \ \ finite \" shows "\ \\<^sub>P Or\<^sub>n (And\<^sub>n ` \) \ (\\ \ \. \\ \ \. \ \\<^sub>P \)" proof - have "\ \\<^sub>P Or\<^sub>n (And\<^sub>n ` \) \ (\\ \ \. \ \\<^sub>P And\<^sub>n \)" using Or\<^sub>n_prop_semantics assms by blast then show ?thesis using And\<^sub>n_prop_semantics assms by metis qed subsection \DNF to LTL conversion\ definition ltln_of_dnf :: "'a ltln fset set \ 'a ltln" where "ltln_of_dnf \ = Or\<^sub>n (And\<^sub>n ` fset ` \)" lemma ltln_of_dnf_semantics: assumes "finite \" shows "w \\<^sub>n ltln_of_dnf \ \ (\\ \ \. \\. \ |\| \ \ w \\<^sub>n \)" proof - have "finite (fset ` \)" using assms by blast then have "w \\<^sub>n ltln_of_dnf \ \ (\\ \ fset ` \. \\ \ \. w \\<^sub>n \)" unfolding ltln_of_dnf_def using Or\<^sub>n_And\<^sub>n_image_semantics by fastforce then show ?thesis by (metis image_iff notin_fset) qed lemma ltln_of_dnf_prop_semantics: assumes "finite \" shows "\ \\<^sub>P ltln_of_dnf \ \ (\\ \ \. \\. \ |\| \ \ \ \\<^sub>P \)" proof - have "finite (fset ` \)" using assms by blast then have "\ \\<^sub>P ltln_of_dnf \ \ (\\ \ fset ` \. \\ \ \. \ \\<^sub>P \)" unfolding ltln_of_dnf_def using Or\<^sub>n_And\<^sub>n_image_prop_semantics by fastforce then show ?thesis by (metis image_iff notin_fset) qed lemma ltln_of_dnf_prop_equiv: "ltln_of_dnf (min_dnf \) \\<^sub>P \" unfolding ltl_prop_equiv_def proof fix \ have "\ \\<^sub>P ltln_of_dnf (min_dnf \) \ (\\ \ min_dnf \. \\. \ |\| \ \ \ \\<^sub>P \)" using ltln_of_dnf_prop_semantics min_dnf_finite by metis also have "\ \ (\\ \ min_dnf \. fset \ \ \)" by (metis min_dnf_prop_atoms prop_atoms_entailment_iff notin_fset subset_eq) also have "\ \ \ \\<^sub>P \" using min_dnf_iff_prop_assignment_subset by blast finally show "\ \\<^sub>P ltln_of_dnf (min_dnf \) = \ \\<^sub>P \" . qed lemma min_dnf_ltln_of_dnf[simp]: "min_dnf (ltln_of_dnf (min_dnf \)) = min_dnf \" using ltl_prop_equiv_min_dnf ltln_of_dnf_prop_equiv by blast subsection \Substitution in DNF formulas\ definition subst_clause :: "'a ltln fset \ ('a ltln \ 'a ltln) \ 'a ltln fset set" where "subst_clause \ m = \\<^sub>m {min_dnf (subst \ m) | \. \ \ fset \}" definition subst_dnf :: "'a ltln fset set \ ('a ltln \ 'a ltln) \ 'a ltln fset set" where "subst_dnf \ m = (\\ \ \. subst_clause \ m)" lemma subst_clause_empty[simp]: "subst_clause {||} m = {{||}}" by (simp add: subst_clause_def) lemma subst_dnf_empty[simp]: "subst_dnf {} m = {}" by (simp add: subst_dnf_def) lemma subst_clause_inner_finite: "finite {min_dnf (subst \ m) | \. \ \ \}" if "finite \" using that by simp lemma subst_clause_finite: "finite (subst_clause \ m)" unfolding subst_clause_def by (auto intro: min_dnf_finite min_product_set_finite) lemma subst_dnf_finite: "finite \ \ finite (subst_dnf \ m)" unfolding subst_dnf_def using subst_clause_finite by blast lemma subst_dnf_mono: "\ \ \ \ subst_dnf \ m \ subst_dnf \ m" unfolding subst_dnf_def by blast lemma subst_clause_min_set[simp]: "min_set (subst_clause \ m) = subst_clause \ m" unfolding subst_clause_def by simp lemma subst_clause_finsert[simp]: "subst_clause (finsert \ \) m = (min_dnf (subst \ m)) \\<^sub>m (subst_clause \ m)" proof - have "{min_dnf (subst \ m) | \. \ \ fset (finsert \ \)} = insert (min_dnf (subst \ m)) {min_dnf (subst \ m) | \. \ \ fset \}" by auto then show ?thesis by (simp add: subst_clause_def) qed lemma subst_clause_funion[simp]: "subst_clause (\ |\| \) m = (subst_clause \ m) \\<^sub>m (subst_clause \ m)" proof (induction \) case (insert x F) then show ?case using min_product_set_thms.fun_left_comm by fastforce qed simp text \For the proof of correctness, we redefine the @{const product} operator on lists.\ definition list_product :: "'a list set \ 'a list set \ 'a list set" (infixl "\\<^sub>l" 65) where "A \\<^sub>l B = {a @ b | a b. a \ A \ b \ B}" lemma list_product_fset_of_list[simp]: "fset_of_list ` (A \\<^sub>l B) = (fset_of_list ` A) \ (fset_of_list ` B)" unfolding list_product_def product_def image_def by fastforce lemma list_product_finite: "finite A \ finite B \ finite (A \\<^sub>l B)" unfolding list_product_def by (simp add: finite_image_set2) lemma list_product_iff: "x \ A \\<^sub>l B \ (\a b. a \ A \ b \ B \ x = a @ b)" unfolding list_product_def by blast lemma list_product_assoc[simp]: "A \\<^sub>l (B \\<^sub>l C) = A \\<^sub>l B \\<^sub>l C" unfolding set_eq_iff list_product_iff by fastforce text \Furthermore, we introduct DNFs where the clauses are represented as lists.\ fun list_dnf :: "'a ltln \ 'a ltln list set" where "list_dnf true\<^sub>n = {[]}" | "list_dnf false\<^sub>n = {}" | "list_dnf (\ and\<^sub>n \) = (list_dnf \) \\<^sub>l (list_dnf \)" | "list_dnf (\ or\<^sub>n \) = (list_dnf \) \ (list_dnf \)" | "list_dnf \ = {[\]}" definition list_dnf_to_dnf :: "'a list set \ 'a fset set" where "list_dnf_to_dnf X = fset_of_list ` X" lemma list_dnf_to_dnf_list_dnf[simp]: "list_dnf_to_dnf (list_dnf \) = dnf \" by (induction \) (simp_all add: list_dnf_to_dnf_def image_Un) lemma list_dnf_finite: "finite (list_dnf \)" by (induction \) (simp_all add: list_product_finite) text \We use this to redefine @{const subst_clause} and @{const subst_dnf} on list DNFs.\ definition subst_clause' :: "'a ltln list \ ('a ltln \ 'a ltln) \ 'a ltln list set" where "subst_clause' \ m = fold (\\ acc. acc \\<^sub>l list_dnf (subst \ m)) \ {[]}" definition subst_dnf' :: "'a ltln list set \ ('a ltln \ 'a ltln) \ 'a ltln list set" where "subst_dnf' \ m = (\\ \ \. subst_clause' \ m)" lemma subst_clause'_finite: "finite (subst_clause' \ m)" by (induction \ rule: rev_induct) (simp_all add: subst_clause'_def list_dnf_finite list_product_finite) lemma subst_clause'_nil[simp]: "subst_clause' [] m = {[]}" by (simp add: subst_clause'_def) lemma subst_clause'_cons[simp]: "subst_clause' (xs @ [x]) m = subst_clause' xs m \\<^sub>l list_dnf (subst x m)" by (simp add: subst_clause'_def) lemma subst_clause'_append[simp]: "subst_clause' (A @ B) m = subst_clause' A m \\<^sub>l subst_clause' B m" proof (induction B rule: rev_induct) case (snoc x xs) then show ?case by simp (metis append_assoc subst_clause'_cons) qed(simp add: list_product_def) lemma subst_dnf'_iff: "x \ subst_dnf' A m \ (\\ \ A. x \ subst_clause' \ m)" by (simp add: subst_dnf'_def) lemma subst_dnf'_product: "subst_dnf' (A \\<^sub>l B) m = (subst_dnf' A m) \\<^sub>l (subst_dnf' B m)" (is "?lhs = ?rhs") proof (unfold set_eq_iff, safe) fix x assume "x \ ?lhs" then obtain \ where "\ \ A \\<^sub>l B" and "x \ subst_clause' \ m" unfolding subst_dnf'_iff by blast then obtain a b where "a \ A" and "b \ B" and "\ = a @ b" unfolding list_product_def by blast then have "x \ (subst_clause' a m) \\<^sub>l (subst_clause' b m)" using \x \ subst_clause' \ m\ by simp then obtain a' b' where "a' \ subst_clause' a m" and "b' \ subst_clause' b m" and "x = a' @ b'" unfolding list_product_iff by blast then have "a' \ subst_dnf' A m" and "b' \ subst_dnf' B m" unfolding subst_dnf'_iff using \a \ A\ \b \ B\ by auto then have "\a\subst_dnf' A m. \b\subst_dnf' B m. x = a @ b" using \x = a' @ b'\ by blast then show "x \ ?rhs" unfolding list_product_iff by blast next fix x assume "x \ ?rhs" then obtain a b where "a \ subst_dnf' A m" and "b \ subst_dnf' B m" and "x = a @ b" unfolding list_product_iff by blast then obtain a' b' where "a' \ A" and "b' \ B" and a: "a \ subst_clause' a' m" and b: "b \ subst_clause' b' m" unfolding subst_dnf'_iff by blast then have "x \ (subst_clause' a' m) \\<^sub>l (subst_clause' b' m)" unfolding list_product_iff using \x = a @ b\ by blast moreover have "a' @ b' \ A \\<^sub>l B" unfolding list_product_iff using \a' \ A\ \b' \ B\ by blast ultimately show "x \ ?lhs" unfolding subst_dnf'_iff by force qed lemma subst_dnf'_list_dnf: "subst_dnf' (list_dnf \) m = list_dnf (subst \ m)" proof (induction \) case (And_ltln \1 \2) then show ?case by (simp add: subst_dnf'_product) qed (simp_all add: subst_dnf'_def subst_clause'_def list_product_def) lemma min_set_Union: "finite X \ min_set (\ (min_set ` X)) = min_set (\ X)" for X :: "'a fset set set" by (induction X rule: finite_induct) (force, metis Sup_insert image_insert min_set_min_union min_union_def) lemma min_set_Union_image: "finite X \ min_set (\x \ X. min_set (f x)) = min_set (\x \ X. f x)" for f :: "'b \ 'a fset set" proof - assume "finite X" then have *: "finite (f ` X)" by auto with min_set_Union show ?thesis unfolding image_image by fastforce qed lemma subst_clause_fset_of_list: "subst_clause (fset_of_list \) m = min_set (list_dnf_to_dnf (subst_clause' \ m))" unfolding list_dnf_to_dnf_def subst_clause'_def proof (induction \ rule: rev_induct) case (snoc x xs) then show ?case by simp (metis (no_types, lifting) dnf_min_set list_dnf_to_dnf_def list_dnf_to_dnf_list_dnf min_product_comm min_product_def min_set_min_product(1)) qed simp lemma min_set_list_dnf_to_dnf_subst_dnf': "finite X \ min_set (list_dnf_to_dnf (subst_dnf' X m)) = min_set (subst_dnf (list_dnf_to_dnf X) m)" by (simp add: subst_dnf'_def subst_dnf_def subst_clause_fset_of_list list_dnf_to_dnf_def min_set_Union_image image_Union) lemma subst_dnf_dnf: "min_set (subst_dnf (dnf \) m) = min_dnf (subst \ m)" unfolding dnf_min_set unfolding list_dnf_to_dnf_list_dnf[symmetric] unfolding subst_dnf'_list_dnf[symmetric] unfolding min_set_list_dnf_to_dnf_subst_dnf'[OF list_dnf_finite] by simp text \This is almost the lemma we need. However, we need to show that the same holds for @{term "min_dnf \"}, too.\ lemma fold_product: "Finite_Set.fold (\x. (\) {{|x|}}) {{||}} (fset x) = {x}" by (induction x) (simp_all add: notin_fset, simp add: product_singleton_singleton) lemma fold_union: "Finite_Set.fold (\x. (\) {x}) {} (fset x) = fset x" - by (induction x) (simp_all add: notin_fset comp_fun_idem.fold_insert_idem comp_fun_idem_insert) + by (induction x) (simp_all add: notin_fset comp_fun_idem_on.fold_insert_idem[OF comp_fun_idem_insert[unfolded comp_fun_idem_def']]) lemma fold_union_fold_product: assumes "finite X" and "\\ \. \ \ X \ \ \ fset \ \ dnf \ = {{|\|}}" shows "Finite_Set.fold (\x. (\) (Finite_Set.fold (\\. (\) (dnf \)) {{||}} (fset x))) {} X = X" (is "?lhs = X") proof - from assms have "?lhs = Finite_Set.fold (\x. (\) (Finite_Set.fold (\\. (\) {{|\|}}) {{||}} (fset x))) {} X" proof (induction X rule: finite_induct) case (insert \ X) from insert.prems have 1: "\\ \. \\ \ X; \ \ fset \\ \ dnf \ = {{|\|}}" by force from insert.prems have "Finite_Set.fold (\\. (\) (dnf \)) {{||}} (fset \) = Finite_Set.fold (\\. (\) {{|\|}}) {{||}} (fset \)" by (induction \) (force simp: notin_fset)+ with insert 1 show ?case by simp qed simp with \finite X\ show ?thesis unfolding fold_product by (metis fset_to_fset fold_union) qed lemma dnf_ltln_of_dnf_min_dnf: "dnf (ltln_of_dnf (min_dnf \)) = min_dnf \" proof - have 1: "finite (And\<^sub>n ` fset ` min_dnf \)" using min_dnf_finite by blast have 2: "inj_on And\<^sub>n (fset ` min_dnf \)" by (metis (mono_tags, lifting) And\<^sub>n_inj f_inv_into_f fset inj_onI inj_on_contraD) have 3: "inj_on fset (min_dnf \)" by (meson fset_inject inj_onI) show ?thesis unfolding ltln_of_dnf_def unfolding Or\<^sub>n_dnf[OF 1] unfolding fold_image[OF 2] unfolding fold_image[OF 3] unfolding comp_def unfolding And\<^sub>n_dnf[OF finite_fset] by (metis fold_union_fold_product min_dnf_finite min_dnf_atoms_dnf) qed lemma min_dnf_subst: "min_set (subst_dnf (min_dnf \) m) = min_dnf (subst \ m)" (is "?lhs = ?rhs") proof - let ?\' = "ltln_of_dnf (min_dnf \)" have "?lhs = min_set (subst_dnf (dnf ?\') m)" unfolding dnf_ltln_of_dnf_min_dnf .. also have "\ = min_dnf (subst ?\' m)" unfolding subst_dnf_dnf .. also have "\ = min_dnf (subst \ m)" using ltl_prop_equiv_min_dnf ltln_of_dnf_prop_equiv subst_respects_ltl_prop_entailment(2) by blast finally show ?thesis . qed end diff --git a/thys/MSO_Regex_Equivalence/Pi_Regular_Exp.thy b/thys/MSO_Regex_Equivalence/Pi_Regular_Exp.thy --- a/thys/MSO_Regex_Equivalence/Pi_Regular_Exp.thy +++ b/thys/MSO_Regex_Equivalence/Pi_Regular_Exp.thy @@ -1,423 +1,423 @@ (* Author: Dmitriy Traytel *) section \$\Pi$-Extended Regular Expressions\ (*<*) theory Pi_Regular_Exp imports Pi_Regular_Set "HOL-Library.List_Lexorder" "HOL-Library.Code_Target_Nat" Deriving.Compare_Instances begin (*>*) subsection \Syntax of regular expressions\ datatype 'a rexp = Zero | Full | One | Atom 'a | Plus "('a rexp)" "('a rexp)" | Times "('a rexp)" "('a rexp)" | Star "('a rexp)" | Not "('a rexp)" | Inter "('a rexp)" "('a rexp)" | Pr "('a rexp)" derive linorder rexp text \Lifting constructors to lists\ fun rexp_of_list where "rexp_of_list OPERATION N [] = N" | "rexp_of_list OPERATION N [x] = x" | "rexp_of_list OPERATION N (x # xs) = OPERATION x (rexp_of_list OPERATION N xs)" abbreviation "PLUS \ rexp_of_list Plus Zero" abbreviation "TIMES \ rexp_of_list Times One" abbreviation "INTERSECT \ rexp_of_list Inter Full" lemma list_singleton_induct [case_names nil single cons]: assumes nil: "P []" assumes single: "\x. P [x]" assumes cons: "\x y xs. P (y # xs) \ P (x # (y # xs))" shows "P xs" using assms proof (induct xs) case (Cons x xs) thus ?case by (cases xs) auto qed simp subsection \ACI normalization\ fun toplevel_summands :: "'a rexp \ 'a rexp set" where "toplevel_summands (Plus r s) = toplevel_summands r \ toplevel_summands s" | "toplevel_summands r = {r}" abbreviation (input) "flatten LISTOP X \ LISTOP (sorted_list_of_set X)" lemma toplevel_summands_nonempty[simp]: "toplevel_summands r \ {}" by (induct r) auto lemma toplevel_summands_finite[simp]: "finite (toplevel_summands r)" by (induct r) auto primrec ACI_norm :: "('a::linorder) rexp \ 'a rexp" ("\_\") where "\Zero\ = Zero" | "\Full\ = Full" | "\One\ = One" | "\Atom a\ = Atom a" | "\Plus r s\ = flatten PLUS (toplevel_summands (Plus \r\ \s\))" | "\Times r s\ = Times \r\ \s\" | "\Star r\ = Star \r\" | "\Not r\ = Not \r\" | "\Inter r s\ = Inter \r\ \s\" | "\Pr r\ = Pr \r\" lemma Plus_toplevel_summands: "Plus r s \ toplevel_summands t \ False" by (induct t) auto lemma toplevel_summands_not_Plus[simp]: "(\r s. x \ Plus r s) \ toplevel_summands x = {x}" by (induct x) auto lemma toplevel_summands_PLUS_strong: "\xs \ []; list_all (\x. \(\r s. x = Plus r s)) xs\ \ toplevel_summands (PLUS xs) = set xs" by (induct xs rule: list_singleton_induct) auto lemma toplevel_summands_flatten: "\X \ {}; finite X; \x \ X. \(\r s. x = Plus r s)\ \ toplevel_summands (flatten PLUS X) = X" using toplevel_summands_PLUS_strong[of "sorted_list_of_set X"] unfolding list_all_iff by fastforce lemma ACI_norm_Plus: "\r\ = Plus s t \ \s t. r = Plus s t" by (induct r) auto lemma toplevel_summands_flatten_ACI_norm_image: "toplevel_summands (flatten PLUS (ACI_norm ` toplevel_summands r)) = ACI_norm ` toplevel_summands r" by (intro toplevel_summands_flatten) (auto dest!: ACI_norm_Plus intro: Plus_toplevel_summands) lemma toplevel_summands_flatten_ACI_norm_image_Union: "toplevel_summands (flatten PLUS (ACI_norm ` toplevel_summands r \ ACI_norm ` toplevel_summands s)) = ACI_norm ` toplevel_summands r \ ACI_norm ` toplevel_summands s" by (intro toplevel_summands_flatten) (auto dest!: ACI_norm_Plus[OF sym] intro: Plus_toplevel_summands) lemma toplevel_summands_ACI_norm: "toplevel_summands \r\ = ACI_norm ` toplevel_summands r" by (induct r) (auto simp: toplevel_summands_flatten_ACI_norm_image_Union) lemma ACI_norm_flatten: "\r\ = flatten PLUS (ACI_norm ` toplevel_summands r)" by (induct r) (auto simp: image_Un toplevel_summands_flatten_ACI_norm_image) theorem ACI_norm_idem[simp]: "\\r\\ = \r\" proof (induct r) case (Plus r s) have "\\Plus r s\\ = \flatten PLUS (toplevel_summands \r\ \ toplevel_summands \s\)\" (is "_ = \flatten PLUS ?U\") by simp also have "\ = flatten PLUS (ACI_norm ` toplevel_summands (flatten PLUS ?U))" by (simp only: ACI_norm_flatten) also have "toplevel_summands (flatten PLUS ?U) = ?U" by (intro toplevel_summands_flatten) (auto intro: Plus_toplevel_summands) also have "flatten PLUS (ACI_norm ` ?U) = flatten PLUS (toplevel_summands \r\ \ toplevel_summands \s\)" by (simp only: image_Un toplevel_summands_ACI_norm[symmetric] Plus) finally show ?case by simp qed auto fun ACI_nPlus :: "'a::linorder rexp \ 'a rexp \ 'a rexp" where "ACI_nPlus (Plus r1 r2) s = ACI_nPlus r1 (ACI_nPlus r2 s)" | "ACI_nPlus r (Plus s1 s2) = (if r = s1 then Plus s1 s2 else if r < s1 then Plus r (Plus s1 s2) else Plus s1 (ACI_nPlus r s2))" | "ACI_nPlus r s = (if r = s then r else if r < s then Plus r s else Plus s r)" fun ACI_norm_alt where "ACI_norm_alt Zero = Zero" | "ACI_norm_alt Full = Full" | "ACI_norm_alt One = One" | "ACI_norm_alt (Atom a) = Atom a" | "ACI_norm_alt (Plus r s) = ACI_nPlus (ACI_norm_alt r) (ACI_norm_alt s)" | "ACI_norm_alt (Times r s) = Times (ACI_norm_alt r) (ACI_norm_alt s)" | "ACI_norm_alt (Star r) = Star (ACI_norm_alt r)" | "ACI_norm_alt (Not r) = Not (ACI_norm_alt r)" | "ACI_norm_alt (Inter r s) = Inter (ACI_norm_alt r) (ACI_norm_alt s)" | "ACI_norm_alt (Pr r) = Pr (ACI_norm_alt r)" lemma toplevel_summands_ACI_nPlus: "toplevel_summands (ACI_nPlus r s) = toplevel_summands (Plus r s)" by (induct r s rule: ACI_nPlus.induct) auto lemma toplevel_summands_ACI_norm_alt: "toplevel_summands (ACI_norm_alt r) = ACI_norm_alt ` toplevel_summands r" by (induct r) (auto simp: toplevel_summands_ACI_nPlus) lemma ACI_norm_alt_Plus: "ACI_norm_alt r = Plus s t \ \s t. r = Plus s t" by (induct r) auto lemma toplevel_summands_flatten_ACI_norm_alt_image: "toplevel_summands (flatten PLUS (ACI_norm_alt ` toplevel_summands r)) = ACI_norm_alt ` toplevel_summands r" by (intro toplevel_summands_flatten) (auto dest!: ACI_norm_alt_Plus intro: Plus_toplevel_summands) lemma ACI_norm_ACI_norm_alt: "\ACI_norm_alt r\ = \r\" proof (induction r) case (Plus r s) show ?case using ACI_norm_flatten [of "ACI_norm_alt (Plus r s)"] ACI_norm_flatten [of "Plus r s"] by (auto simp: image_Un toplevel_summands_ACI_nPlus) (metis Plus.IH toplevel_summands_ACI_norm) qed auto lemma ACI_nPlus_singleton_PLUS: "\xs \ []; sorted xs; distinct xs; \x \ {x} \ set xs. \(\r s. x = Plus r s)\ \ ACI_nPlus x (PLUS xs) = (if x \ set xs then PLUS xs else PLUS (insort x xs))" proof (induct xs rule: list_singleton_induct) case (single y) thus ?case by (cases x y rule: linorder_cases) (induct x y rule: ACI_nPlus.induct, auto)+ next case (cons y1 y2 ys) thus ?case by (cases x) (auto) qed simp lemma ACI_nPlus_PLUS: "\xs1 \ []; xs2 \ []; \x \ set (xs1 @ xs2). \(\r s. x = Plus r s); sorted xs2; distinct xs2\\ ACI_nPlus (PLUS xs1) (PLUS xs2) = flatten PLUS (set (xs1 @ xs2))" proof (induct xs1 arbitrary: xs2 rule: list_singleton_induct) case (single x1) thus ?case - apply (auto intro!: trans[OF ACI_nPlus_singleton_PLUS] simp: insert_absorb simp del: sorted_list_of_set_insert) + apply (auto intro!: trans[OF ACI_nPlus_singleton_PLUS] simp: insert_absorb simp del: sorted_list_of_set_insert_remove) apply (metis finite_set finite_sorted_distinct_unique sorted_list_of_set) apply (metis remdups_id_iff_distinct sorted_list_of_set_sort_remdups sorted_sort_id) done next case (cons x11 x12 xs1) thus ?case - apply (simp del: sorted_list_of_set_insert) + apply (simp del: sorted_list_of_set_insert_remove) apply (rule trans[OF ACI_nPlus_singleton_PLUS]) - apply (auto simp del: sorted_list_of_set_insert simp add: insert_commute[of x11]) + apply (auto simp del: sorted_list_of_set_insert_remove simp add: insert_commute[of x11]) apply (auto simp only: Un_insert_left[of x11, symmetric] insert_absorb) [] apply (auto simp only: Un_insert_right[of _ x11, symmetric] insert_absorb) [] apply (auto simp add: insert_commute[of x12]) done qed simp lemma ACI_nPlus_flatten_PLUS: "\X1 \ {}; X2 \ {}; finite X1; finite X2; \x \ X1 \ X2. \(\r s. x = Plus r s)\\ ACI_nPlus (flatten PLUS X1) (flatten PLUS X2) = flatten PLUS (X1 \ X2)" by (rule trans[OF ACI_nPlus_PLUS]) auto lemma ACI_nPlus_ACI_norm[simp]: "ACI_nPlus \r\ \s\ = \Plus r s\" using ACI_norm_flatten [of r] ACI_norm_flatten [of s] ACI_norm_flatten [of "Plus r s"] apply (auto intro!: trans [OF ACI_nPlus_flatten_PLUS]) apply (auto simp: image_Un Un_assoc intro!: trans [OF ACI_nPlus_flatten_PLUS]) apply (metis ACI_norm_Plus Plus_toplevel_summands)+ done lemma ACI_norm_alt: "ACI_norm_alt r = \r\" by (induct r) auto declare ACI_norm_alt[symmetric, code] subsection \Finality\ primrec final :: "'a rexp \ bool" where "final Zero = False" | "final Full = True" | "final One = True" | "final (Atom _) = False" | "final (Plus r s) = (final r \ final s)" | "final (Times r s) = (final r \ final s)" | "final (Star _) = True" | "final (Not r) = (~ final r)" | "final (Inter r1 r2) = (final r1 \ final r2)" | "final (Pr r) = final r" lemma toplevel_summands_final: "final s = (\r\toplevel_summands s. final r)" by (induct s) auto lemma final_PLUS: "final (PLUS xs) = (\r \ set xs. final r)" by (induct xs rule: list_singleton_induct) auto theorem ACI_norm_final[simp]: "final \r\ = final r" proof (induct r) case (Plus r1 r2) thus ?case using toplevel_summands_final by (auto simp: final_PLUS) qed auto subsection \Wellformedness w.r.t. an alphabet\ locale alphabet = fixes \ :: "nat \ 'a set" ("\ _") and wf_atom :: "nat \ 'b :: linorder \ bool" begin primrec wf :: "nat \ 'b rexp \ bool" where "wf n Zero = True" | "wf n Full = True" | "wf n One = True" | "wf n (Atom a) = (wf_atom n a)" | "wf n (Plus r s) = (wf n r \ wf n s)" | "wf n (Times r s) = (wf n r \ wf n s)" | "wf n (Star r) = wf n r" | "wf n (Not r) = wf n r" | "wf n (Inter r s) = (wf n r \ wf n s)" | "wf n (Pr r) = wf (n + 1) r" primrec wf_word where "wf_word n [] = True" | "wf_word n (w # ws) = ((w \ \ n) \ wf_word n ws)" lemma wf_word_snoc[simp]: "wf_word n (ws @ [w]) = ((w \ \ n) \ wf_word n ws)" by (induct ws) auto lemma wf_word_append[simp]: "wf_word n (ws @ vs) = (wf_word n ws \ wf_word n vs)" by (induct ws arbitrary: vs) auto lemma wf_word: "wf_word n w = (w \ lists (\ n))" by (induct w) auto lemma toplevel_summands_wf: "wf n s = (\r\toplevel_summands s. wf n r)" by (induct s) auto lemma wf_PLUS[simp]: "wf n (PLUS xs) = (\r \ set xs. wf n r)" by (induct xs rule: list_singleton_induct) auto lemma wf_TIMES[simp]: "wf n (TIMES xs) = (\r \ set xs. wf n r)" by (induct xs rule: list_singleton_induct) auto lemma wf_flatten_PLUS[simp]: "finite X \ wf n (flatten PLUS X) = (\r \ X. wf n r)" using wf_PLUS[of n "sorted_list_of_set X"] by fastforce theorem ACI_norm_wf[simp]: "wf n \r\ = wf n r" proof (induct r arbitrary: n) case (Plus r1 r2) thus ?case using toplevel_summands_wf[of n "\r1\"] toplevel_summands_wf[of n "\r2\"] by auto qed auto lemma wf_INTERSECT[simp]: "wf n (INTERSECT xs) = (\r \ set xs. wf n r)" by (induct xs rule: list_singleton_induct) auto lemma wf_flatten_INTERSECT[simp]: "finite X \ wf n (flatten INTERSECT X) = (\r \ X. wf n r)" using wf_INTERSECT[of n "sorted_list_of_set X"] by fastforce end subsection \Language\ locale project = alphabet \ wf_atom for \ :: "nat \ 'a set" and wf_atom :: "nat \ 'b :: linorder \ bool" + fixes project :: "'a \ 'a" and lookup :: "'b \ 'a \ bool" assumes project: "\a. a \ \ (Suc n) \ project a \ \ n" begin primrec lang :: "nat \ 'b rexp => 'a lang" where "lang n Zero = {}" | "lang n Full = lists (\ n)" | "lang n One = {[]}" | "lang n (Atom b) = {[x] | x. lookup b x \ x \ \ n}" | "lang n (Plus r s) = (lang n r) \ (lang n s)" | "lang n (Times r s) = conc (lang n r) (lang n s)" | "lang n (Star r) = star (lang n r)" | "lang n (Not r) = lists (\ n) - lang n r" | "lang n (Inter r s) = (lang n r \ lang n s)" | "lang n (Pr r) = map project ` lang (n + 1) r" lemma wf_word_map_project[simp]: "wf_word (Suc n) ws \ wf_word n (map project ws)" by (induct ws arbitrary: n) (auto intro: project) lemma wf_lang_wf_word: "wf n r \ \w \ lang n r. wf_word n w" by (induct r arbitrary: n) (auto elim: rev_subsetD[OF _ conc_mono] star_induct intro: iffD2[OF wf_word]) lemma lang_subset_lists: "wf n r \ lang n r \ lists (\ n)" proof (induct r arbitrary: n) case Pr thus ?case by (fastforce intro!: project) qed (auto simp: conc_subset_lists star_subset_lists) lemma toplevel_summands_lang: "r \ toplevel_summands s \ lang n r \ lang n s" by (induct s) auto lemma toplevel_summands_lang_UN: "lang n s = (\r\toplevel_summands s. lang n r)" by (induct s) auto lemma toplevel_summands_in_lang: "w \ lang n s = (\r\toplevel_summands s. w \ lang n r)" by (induct s) auto lemma lang_PLUS[simp]: "lang n (PLUS xs) = (\r \ set xs. lang n r)" by (induct xs rule: list_singleton_induct) auto lemma lang_TIMES[simp]: "lang n (TIMES xs) = foldr (@@) (map (lang n) xs) {[]}" by (induct xs rule: list_singleton_induct) auto lemma lang_flatten_PLUS: "finite X \ lang n (flatten PLUS X) = (\r \ X. lang n r)" using lang_PLUS[of n "sorted_list_of_set X"] by fastforce theorem ACI_norm_lang[simp]: "lang n \r\ = lang n r" proof (induct r arbitrary: n) case (Plus r1 r2) moreover from Plus[symmetric] have "lang n (Plus r1 r2) \ lang n \Plus r1 r2\" using toplevel_summands_in_lang[of _ n "\r1\"] toplevel_summands_in_lang[of _ n "\r2\"] by auto ultimately show ?case by (fastforce dest!: toplevel_summands_lang) qed auto lemma lang_final: "final r = ([] \ lang n r)" using concI[of "[]" _ "[]"] by (induct r arbitrary: n) auto lemma in_lang_INTERSECT: "wf_word n w \ w \ lang n (INTERSECT xs) = (\r \ set xs. w \ lang n r)" by (induct xs rule: list_singleton_induct) (auto simp: wf_word) lemma lang_INTERSECT: "lang n (INTERSECT xs) = (if xs = [] then lists (\ n) else \r \ set xs. lang n r)" by (induct xs rule: list_singleton_induct) auto lemma lang_flatten_INTERSECT[simp]: assumes "finite X" "X \ {}" "\r\X. wf n r" shows "w \ lang n (flatten INTERSECT X) = (\r \ X. w \ lang n r)" (is "?L = ?R") proof assume ?L thus ?R using in_lang_INTERSECT[OF bspec[OF wf_lang_wf_word[OF iffD2[OF wf_flatten_INTERSECT]]], OF assms(1,3) \?L\, of "sorted_list_of_set X"] assms(1) by auto next assume ?R with assms show ?L by (intro iffD2[OF in_lang_INTERSECT]) (auto dest: wf_lang_wf_word) qed end (*<*) end (*>*) diff --git a/thys/Ordinal_Partitions/Library_Additions.thy b/thys/Ordinal_Partitions/Library_Additions.thy --- a/thys/Ordinal_Partitions/Library_Additions.thy +++ b/thys/Ordinal_Partitions/Library_Additions.thy @@ -1,338 +1,338 @@ section \Library additions\ theory Library_Additions imports "ZFC_in_HOL.Ordinal_Exp" "HOL-Library.Ramsey" "Nash_Williams.Nash_Williams" begin lemma finite_enumerate_Diff_singleton: fixes S :: "'a::wellorder set" assumes "finite S" and i: "i < card S" "enumerate S i < x" shows "enumerate (S - {x}) i = enumerate S i" using i proof (induction i) case 0 have "(LEAST i. i \ S \ i\x) = (LEAST i. i \ S)" proof (rule Least_equality) have "\t. t \ S \ t\x" using 0 \finite S\ finite_enumerate_in_set by blast then show "(LEAST i. i \ S) \ S \ (LEAST i. i \ S) \ x" by (metis "0.prems"(2) LeastI enumerate_0 not_less_Least) qed (simp add: Least_le) then show ?case by (auto simp: enumerate_0) next case (Suc i) then have x: "enumerate S i < x" by (meson enumerate_step finite_enumerate_step less_trans) have cardSx: "Suc i < card (S - {x})" and "i < card S" using Suc \finite S\ card_Diff_singleton_if finite_enumerate_Ex by fastforce+ have "(LEAST s. s \ S \ s\x \ enumerate (S - {x}) i < s) = (LEAST s. s \ S \ enumerate S i < s)" (is "_ = ?r") proof (intro Least_equality conjI) show "?r \ S" by (metis (lifting) LeastI Suc.prems(1) assms(1) finite_enumerate_in_set finite_enumerate_step) show "?r \ x" using Suc.prems not_less_Least [of _ "\t. t \ S \ enumerate S i < t"] \finite S\ finite_enumerate_in_set finite_enumerate_step by blast show "enumerate (S - {x}) i < ?r" by (metis (full_types) Suc.IH Suc.prems(1) \i < card S\ enumerate_Suc'' enumerate_step finite_enumerate_Suc'' finite_enumerate_step x) show "\y. y \ S \ y \ x \ enumerate (S - {x}) i < y \ ?r \ y" by (simp add: Least_le Suc.IH \i < card S\ x) qed then show ?case using Suc assms by (simp add: finite_enumerate_Suc'' cardSx) qed lemma hd_lex: "\hd ms < hd ns; length ms = length ns; ns \ []\ \ (ms, ns) \ lex less_than" by (metis hd_Cons_tl length_0_conv less_than_iff lexord_cons_cons lexord_lex) lemma sorted_hd_le: assumes "sorted xs" "x \ list.set xs" shows "hd xs \ x" using assms by (induction xs) (auto simp: less_imp_le) lemma sorted_le_last: assumes "sorted xs" "x \ list.set xs" shows "x \ last xs" using assms by (induction xs) (auto simp: less_imp_le) lemma hd_list_of: assumes "finite A" "A \ {}" shows "hd (sorted_list_of_set A) = Min A" proof (rule antisym) have "Min A \ A" by (simp add: assms) then show "hd (sorted_list_of_set A) \ Min A" by (simp add: sorted_hd_le \finite A\) next show "Min A \ hd (sorted_list_of_set A)" by (metis Min_le assms hd_in_set set_sorted_list_of_set sorted_list_of_set_eq_Nil_iff) qed lemma sorted_hd_le_last: assumes "sorted xs" "xs \ []" shows "hd xs \ last xs" using assms by (simp add: sorted_hd_le) lemma sorted_list_of_set_set_of [simp]: "strict_sorted l \ sorted_list_of_set (list.set l) = l" by (simp add: strict_sorted_equal) lemma range_strict_mono_ext: fixes f::"nat \ 'a::linorder" assumes eq: "range f = range g" and sm: "strict_mono f" "strict_mono g" shows "f = g" proof fix n show "f n = g n" proof (induction n rule: less_induct) case (less n) obtain x y where xy: "f n = g y" "f x = g n" by (metis eq imageE rangeI) then have "n = y" by (metis (no_types) less.IH neq_iff sm strict_mono_less xy) then show ?case using xy by auto qed qed subsection \Other material\ definition strict_mono_sets :: "['a::order set, 'a::order \ 'b::order set] \ bool" where "strict_mono_sets A f \ \x\A. \y\A. x < y \ less_sets (f x) (f y)" lemma strict_mono_setsD: assumes "strict_mono_sets A f" "x < y" "x \ A" "y \ A" shows "less_sets (f x) (f y)" using assms by (auto simp: strict_mono_sets_def) lemma strict_mono_on_o: "\strict_mono_on r A; strict_mono_on s B; s ` B \ A\ \ strict_mono_on (r \ s) B" by (auto simp: image_subset_iff strict_mono_on_def) lemma strict_mono_sets_imp_disjoint: fixes A :: "'a::linorder set" assumes "strict_mono_sets A f" shows "pairwise (\x y. disjnt (f x) (f y)) A" using assms unfolding strict_mono_sets_def pairwise_def by (meson antisym_conv3 disjnt_sym less_sets_imp_disjnt) lemma strict_mono_sets_subset: assumes "strict_mono_sets B f" "A \ B" shows "strict_mono_sets A f" using assms by (auto simp: strict_mono_sets_def) lemma strict_mono_less_sets_Min: assumes "strict_mono_sets I f" "finite I" "I \ {}" shows "less_sets (f (Min I)) (\ (f ` (I - {Min I})))" using assms by (simp add: strict_mono_sets_def less_sets_UN2 dual_order.strict_iff_order) lemma pair_less_iff1 [simp]: "((x,y), (x,z)) \ pair_less \ y" "\\{}" "\A. A \ \ \ infinite A" and "\A B. \A \ \; B \ \\ \ A \ B \ \" shows "infinite (\\)" by (simp add: assms finite_Inf_in) lemma atLeast_less_sets: "\less_sets A {x}; B \ {x..}\ \ less_sets A B" by (force simp: less_sets_def subset_iff) subsection \The list-of function\ -lemma sorted_list_of_set_insert_cons: +lemma sorted_list_of_set_insert_remove_cons: assumes "finite A" "less_sets {a} A" shows "sorted_list_of_set (insert a A) = a # sorted_list_of_set A" proof - have "strict_sorted (a # sorted_list_of_set A)" using assms less_setsD by auto moreover have "list.set (a # sorted_list_of_set A) = insert a A" using assms by force moreover have "length (a # sorted_list_of_set A) = card (insert a A)" using assms card_insert_if less_setsD by fastforce ultimately show ?thesis by (metis \finite A\ finite_insert sorted_list_of_set_unique) qed lemma sorted_list_of_set_Un: assumes AB: "less_sets A B" and fin: "finite A" "finite B" shows "sorted_list_of_set (A \ B) = sorted_list_of_set A @ sorted_list_of_set B" proof - have "strict_sorted (sorted_list_of_set A @ sorted_list_of_set B)" using AB unfolding less_sets_def by (metis fin set_sorted_list_of_set sorted_wrt_append strict_sorted_list_of_set) moreover have "card A + card B = card (A \ B)" using less_sets_imp_disjnt [OF AB] by (simp add: assms card_Un_disjoint disjnt_def) ultimately show ?thesis by (simp add: assms strict_sorted_equal) qed lemma sorted_list_of_set_UN_lessThan: fixes k::nat assumes sm: "strict_mono_sets {..i. i < k \ finite (A i)" shows "sorted_list_of_set (\i A) (sorted_list_of_set {.. (A ` {.. (A ` {.. (A ` {.. A k)" by (simp add: Un_commute lessThan_Suc) also have "\ = sorted_list_of_set (\ (A ` {.. = concat (map (sorted_list_of_set \ A) (sorted_list_of_set {.. = concat (map (sorted_list_of_set \ A) (sorted_list_of_set {..i. i \ k \ finite (A i)" shows "sorted_list_of_set (\i\k. A i) = concat (map (sorted_list_of_set \ A) (sorted_list_of_set {..k}))" by (metis assms lessThan_Suc_atMost less_Suc_eq_le sorted_list_of_set_UN_lessThan) subsection \Monotonic enumeration of a countably infinite set\ abbreviation "enum \ enumerate" text \Could be generalised to infinite countable sets of any type\ lemma nat_infinite_iff: fixes N :: "nat set" shows "infinite N \ (\f::nat\nat. N = range f \ strict_mono f)" proof safe assume "infinite N" then show "\f. N = range (f::nat \ nat) \ strict_mono f" by (metis bij_betw_imp_surj_on bij_enumerate enumerate_mono strict_mono_def) next fix f :: "nat \ nat" assume "strict_mono f" and "N = range f" and "finite (range f)" then show False using range_inj_infinite strict_mono_imp_inj_on by blast qed lemma enum_works: fixes N :: "nat set" assumes "infinite N" shows "N = range (enum N) \ strict_mono (enum N)" by (metis assms bij_betw_imp_surj_on bij_enumerate enumerate_mono strict_monoI) lemma range_enum: "range (enum N) = N" and strict_mono_enum: "strict_mono (enum N)" if "infinite N" for N :: "nat set" using enum_works [OF that] by auto lemma enum_0_eq_Inf: fixes N :: "nat set" assumes "infinite N" shows "enum N 0 = Inf N" proof - have "enum N 0 \ N" using assms range_enum by auto moreover have "\x. x \ N \ enum N 0 \ x" by (metis (mono_tags, hide_lams) assms imageE le0 less_mono_imp_le_mono range_enum strict_monoD strict_mono_enum) ultimately show ?thesis by (metis cInf_eq_minimum) qed lemma enum_works_finite: fixes N :: "nat set" assumes "finite N" shows "N = enum N ` {.. strict_mono_on (enum N) {.. N" "finite N" obtains i where "i < card N" "x = enum N i" by (metis \x \ N\ \finite N\ enum_works_finite imageE lessThan_iff) lemma enum_0_eq_Inf_finite: fixes N :: "nat set" assumes "finite N" "N \ {}" shows "enum N 0 = Inf N" proof - have "enum N 0 \ N" by (metis Nat.neq0_conv assms empty_is_image enum_works_finite image_eqI lessThan_empty_iff lessThan_iff) moreover have "enum N 0 \ x" if "x \ N" for x proof - obtain i where "i < card N" "x = enum N i" by (metis \x \ N\ \finite N\ enum_obtain_index_finite) with assms show ?thesis by (metis Nat.neq0_conv finite_enumerate_mono less_or_eq_imp_le) qed ultimately show ?thesis by (metis cInf_eq_minimum) qed lemma greaterThan_less_enum: fixes N :: "nat set" assumes "N \ {x<..}" "infinite N" shows "x < enum N i" using assms range_enum by fastforce lemma atLeast_le_enum: fixes N :: "nat set" assumes "N \ {x..}" "infinite N" shows "x \ enum N i" using assms range_enum by fastforce lemma less_sets_empty1 [simp]: "less_sets {} A" and less_sets_empty2 [simp]: "less_sets A {}" by (simp_all add: less_sets_def) lemma less_sets_singleton1 [simp]: "less_sets {a} A \ (\x\A. a < x)" and less_sets_singleton2 [simp]: "less_sets A {a} \ (\x\A. x < a)" by (simp_all add: less_sets_def) lemma less_sets_atMost [simp]: "less_sets {..a} A \ (\x\A. a < x)" and less_sets_alLeast [simp]: "less_sets A {a..} \ (\x\A. x < a)" by (auto simp: less_sets_def) lemma less_sets_imp_strict_mono_sets: assumes "\i. less_sets (A i) (A (Suc i))" "\i. i>0 \ A i \ {}" shows "strict_mono_sets UNIV A" proof (clarsimp simp: strict_mono_sets_def) fix i j::nat assume "i < j" then show "less_sets (A i) (A j)" proof (induction "j-i" arbitrary: i j) case (Suc x) then show ?case by (metis Suc_diff_Suc Suc_inject Suc_mono assms less_Suc_eq less_sets_trans zero_less_Suc) qed auto qed lemma less_sets_Suc_Max: assumes "finite A" shows "less_sets A {Suc (Max A)..}" proof (cases "A = {}") case False then show ?thesis by (simp add: assms less_Suc_eq_le) qed auto lemma infinite_nat_greaterThan: fixes m::nat assumes "infinite N" shows "infinite (N \ {m<..})" proof - have "N \ -{m<..} \ (N \ {m<..})" by blast moreover have "finite (-{m<..})" by simp ultimately show ?thesis using assms finite_subset by blast qed end diff --git a/thys/Ordinal_Partitions/Omega_Omega.thy b/thys/Ordinal_Partitions/Omega_Omega.thy --- a/thys/Ordinal_Partitions/Omega_Omega.thy +++ b/thys/Ordinal_Partitions/Omega_Omega.thy @@ -1,4278 +1,4278 @@ section \An ordinal partition theorem by Jean A. Larson\ text \Jean A. Larson, A short proof of a partition theorem for the ordinal $\omega^\omega$. \emph{Annals of Mathematical Logic}, 6:129–145, 1973.\ theory Omega_Omega imports "HOL-Library.Product_Lexorder" Erdos_Milner begin abbreviation "list_of \ sorted_list_of_set" subsection \Cantor normal form for ordinals below @{term "\\\"}\ text \Unlike @{term Cantor_sum}, there is no list of ordinal exponents, which are instead taken as consecutive. We obtain an order-isomorphism between @{term "\\\"} and increasing lists of natural numbers (ordered lexicographically).\ fun omega_sum_aux where Nil: "omega_sum_aux 0 _ = 0" | Suc: "omega_sum_aux (Suc n) [] = 0" | Cons: "omega_sum_aux (Suc n) (m#ms) = (\\n) * (ord_of_nat m) + omega_sum_aux n ms" abbreviation omega_sum where "omega_sum ms \ omega_sum_aux (length ms) ms" text \A normal expansion has no leading zeroes\ inductive normal:: "nat list \ bool" where normal_Nil[iff]: "normal []" | normal_Cons: "m > 0 \ normal (m#ms)" inductive_simps normal_Cons_iff [simp]: "normal (m#ms)" lemma omega_sum_0_iff [simp]: "normal ns \ omega_sum ns = 0 \ ns = []" by (induction ns rule: normal.induct) auto lemma Ord_omega_sum_aux [simp]: "Ord (omega_sum_aux k ms)" by (induction rule: omega_sum_aux.induct) auto lemma Ord_omega_sum: "Ord (omega_sum ms)" by simp lemma omega_sum_less_\\ [intro]: "omega_sum ms < \\\" proof (induction ms) case (Cons m ms) have "\ \ (length ms) * ord_of_nat m \ elts (\ \ Suc (length ms))" using Ord_mem_iff_lt by auto then have "\\(length ms) * ord_of_nat m \ elts (\\\)" using Ord_ord_of_nat oexp_mono_le omega_nonzero ord_of_nat_le_omega by blast with Cons show ?case by (auto simp: mult_succ OrdmemD oexp_less indecomposableD indecomposable_\_power) qed (auto simp: zero_less_Limit) lemma omega_sum_aux_less: "omega_sum_aux k ms < \ \ k" proof (induction rule: omega_sum_aux.induct) case (3 n m ms) have " \\n * ord_of_nat m + \\n < \\n * \" by (metis Ord_ord_of_nat \_power_succ_gtr mult_succ oexp_succ ord_of_nat.simps(2)) with 3 show ?case using dual_order.strict_trans by force qed auto lemma omega_sum_less: "omega_sum ms < \ \ (length ms)" by (rule omega_sum_aux_less) lemma omega_sum_ge: "m \ 0 \ \ \ (length ms) \ omega_sum (m#ms)" apply clarsimp by (metis Ord_ord_of_nat add_le_cancel_left0 le_mult Nat.neq0_conv ord_of_eq_0_iff vsubsetD) lemma omega_sum_length_less: assumes "normal ns" "length ms < length ns" shows "omega_sum ms < omega_sum ns" using assms proof (induction rule: normal.induct) case (normal_Cons n ns') have "\ \ length ms \ \ \ length ns'" using normal_Cons oexp_mono_le by auto then show ?case by (metis gr_implies_not_zero less_le_trans normal_Cons.hyps omega_sum_aux_less omega_sum_ge) qed auto lemma omega_sum_length_leD: assumes "omega_sum ms \ omega_sum ns" "normal ms" shows "length ms \ length ns" by (meson assms leD leI omega_sum_length_less) lemma omega_sum_less_eqlen_iff_cases [simp]: assumes "length ms = length ns" shows "omega_sum (m#ms) < omega_sum (n#ns) \ m m=n \ omega_sum ms < omega_sum ns" (is "?lhs = ?rhs") proof assume L: ?lhs have "\ Suc n < Suc m" using omega_sum_less [of ms] omega_sum_less [of ns] L assms mult_nat_less_add_less by fastforce with L assms show ?rhs by auto qed (auto simp: mult_nat_less_add_less omega_sum_aux_less assms) lemma omega_sum_lex_less_iff_cases: "((length ms, omega_sum (m#ms)), (length ns, omega_sum (n#ns))) \ less_than <*lex*> VWF \ length ms < length ns \ length ms = length ns \ m m=n \ ((length ms, omega_sum ms), (length ns, omega_sum ns)) \ less_than <*lex*> VWF" using omega_sum_less_eqlen_iff_cases by force lemma omega_sum_less_iff_cases: assumes "m > 0" "n > 0" shows "omega_sum (m#ms) < omega_sum (n#ns) \ length ms < length ns \ length ms = length ns \ m length ms = length ns \ m=n \ omega_sum ms < omega_sum ns" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (metis Suc_less_eq \m>0\ length_Cons less_asym nat_neq_iff normal_Cons omega_sum_length_less omega_sum_less_eqlen_iff_cases) next assume ?rhs then show ?lhs by (metis (full_types) Suc_less_eq \n>0\ length_Cons normal_Cons omega_sum_length_less omega_sum_less_eqlen_iff_cases) qed lemma omega_sum_less_iff: "((length ms, omega_sum ms), (length ns, omega_sum ns)) \ less_than <*lex*> VWF \ (ms,ns) \ lenlex less_than" proof (induction ms arbitrary: ns) case (Cons m ms) then show ?case proof (induction ns) case (Cons n ns') show ?case using Cons.prems Cons_lenlex_iff omega_sum_less_eqlen_iff_cases by fastforce qed auto qed auto lemma eq_omega_sum_less_iff: assumes "length ms = length ns" shows "(omega_sum ms, omega_sum ns) \ VWF \ (ms,ns) \ lenlex less_than" by (metis assms in_lex_prod less_not_refl less_than_iff omega_sum_less_iff) lemma eq_omega_sum_eq_iff: assumes "length ms = length ns" shows "omega_sum ms = omega_sum ns \ ms=ns" proof assume "omega_sum ms = omega_sum ns" then have "(omega_sum ms, omega_sum ns) \ VWF" "(omega_sum ns, omega_sum ms) \ VWF" by auto then obtain "(ms,ns) \ lenlex less_than" "(ns,ms) \ lenlex less_than" using assms eq_omega_sum_less_iff by metis moreover have "total (lenlex less_than)" by (simp add: total_lenlex total_less_than) ultimately show "ms=ns" by (meson UNIV_I total_on_def) qed auto lemma inj_omega_sum: "inj_on omega_sum {l. length l = n}" unfolding inj_on_def using eq_omega_sum_eq_iff by fastforce lemma Ex_omega_sum: "\ \ elts (\\n) \ \ns. \ = omega_sum ns \ length ns = n" proof (induction n arbitrary: \) case 0 then show ?case by (rule_tac x="[]" in exI) auto next case (Suc n) then obtain k::nat where k: "\ \ elts (\ \ n * k)" and kmin: "\k'. k' \ \ elts (\ \ n * k')" by (metis Ord_ord_of_nat elts_mult_\E oexp_succ ord_of_nat.simps(2)) show ?case proof (cases k) case (Suc k') then obtain \ where \: "\ = (\ \ n * k') + \" by (metis lessI mult_succ ord_of_nat.simps(2) k kmin mem_plus_V_E) then have \in: "\ \ elts (\ \ n)" using Suc k mult_succ by auto then obtain ns where ns: "\ = omega_sum ns" and len: "length ns = n" using Suc.IH by auto moreover have "omega_sum ns < \\n" using OrdmemD ns \in by auto ultimately show ?thesis by (rule_tac x="k'#ns" in exI) (simp add: \) qed (use k in auto) qed lemma omega_sum_drop [simp]: "omega_sum (dropWhile (\n. n=0) ns) = omega_sum ns" by (induction ns) auto lemma normal_drop [simp]: "normal (dropWhile (\n. n=0) ns)" by (induction ns) auto lemma omega_sum_\\: assumes "\ \ elts (\\\)" obtains ns where "\ = omega_sum ns" "normal ns" proof - obtain ms where "\ = omega_sum ms" using assms Ex_omega_sum by (auto simp: oexp_Limit elts_\) show thesis proof show "\ = omega_sum (dropWhile (\n. n=0) ms)" by (simp add: \\ = omega_sum ms\) show "normal (dropWhile (\n. n=0) ms)" by auto qed qed definition Cantor_\\ :: "V \ nat list" where "Cantor_\\ \ \x. SOME ns. x = omega_sum ns \ normal ns" lemma assumes "\ \ elts (\\\)" shows Cantor_\\: "omega_sum (Cantor_\\ \) = \" and normal_Cantor_\\: "normal (Cantor_\\ \)" by (metis (mono_tags, lifting) Cantor_\\_def assms omega_sum_\\ someI)+ subsection \Larson's set $W(n)$\ definition WW :: "nat list set" where "WW \ {l. strict_sorted l}" fun into_WW :: "nat \ nat list \ nat list" where "into_WW k [] = []" | "into_WW k (n#ns) = (k+n) # into_WW (Suc (k+n)) ns" fun from_WW :: "nat \ nat list \ nat list" where "from_WW k [] = []" | "from_WW k (n#ns) = (n - k) # from_WW (Suc n) ns" lemma from_into_WW [simp]: "from_WW k (into_WW k ns) = ns" by (induction ns arbitrary: k) auto lemma inj_into_WW: "inj (into_WW k)" by (metis from_into_WW injI) lemma into_from_WW_aux: "\strict_sorted ns; \n\list.set ns. k \ n\ \ into_WW k (from_WW k ns) = ns" by (induction ns arbitrary: k) (auto simp: Suc_leI) lemma into_from_WW [simp]: "strict_sorted ns \ into_WW 0 (from_WW 0 ns) = ns" by (simp add: into_from_WW_aux) lemma into_WW_imp_ge: "y \ List.set (into_WW x ns) \ x \ y" by (induction ns arbitrary: x) fastforce+ lemma strict_sorted_into_WW: "strict_sorted (into_WW x ns)" by (induction ns arbitrary: x) (auto simp: dest: into_WW_imp_ge) lemma length_into_WW: "length (into_WW x ns) = length ns" by (induction ns arbitrary: x) auto lemma WW_eq_range_into: "WW = range (into_WW 0)" proof - have "\ns. strict_sorted ns \ ns \ range (into_WW 0)" by (metis into_from_WW rangeI) then show ?thesis by (auto simp: WW_def strict_sorted_into_WW) qed lemma into_WW_lenlex_iff: "(into_WW k ms, into_WW k ns) \ lenlex less_than \ (ms, ns) \ lenlex less_than" proof (induction ms arbitrary: ns k) case Nil then show ?case by simp (metis length_0_conv length_into_WW) next case (Cons m ms) then show ?case by (induction ns) (auto simp: Cons_lenlex_iff length_into_WW) qed lemma wf_llt [simp]: "wf (lenlex less_than)" and trans_llt [simp]: "trans (lenlex less_than)" by blast+ lemma total_llt [simp]: "total_on A (lenlex less_than)" by (meson UNIV_I total_lenlex total_less_than total_on_def) lemma omega_sum_1_less: assumes "(ms,ns) \ lenlex less_than" shows "omega_sum (1#ms) < omega_sum (1#ns)" proof - have "omega_sum (1#ms) < omega_sum (1#ns)" if "length ms < length ns" using omega_sum_less_iff_cases that zero_less_one by blast then show ?thesis using assms by (auto simp: mult_succ simp flip: omega_sum_less_iff) qed lemma ordertype_WW_1: "ordertype WW (lenlex less_than) \ ordertype UNIV (lenlex less_than)" by (rule ordertype_mono) auto lemma ordertype_WW_2: "ordertype UNIV (lenlex less_than) \ \\\" proof (rule ordertype_inc_le_Ord) show "range (\ms. omega_sum (1#ms)) \ elts (\\\)" by (meson Ord_\ Ord_mem_iff_lt Ord_oexp Ord_omega_sum image_subset_iff omega_sum_less_\\) qed (use omega_sum_1_less in auto) lemma ordertype_WW_3: "\\\ \ ordertype WW (lenlex less_than)" proof - define \ where "\ \ into_WW 0 \ Cantor_\\" have \\: "\\\ = tp (elts (\\\))" by simp also have "\ \ ordertype WW (lenlex less_than)" proof (rule ordertype_inc_le) fix \ \ assume \: "\ \ elts (\\\)" and \: "\ \ elts (\\\)" and "(\, \) \ VWF" then obtain *: "Ord \" "Ord \" "\<\" by (metis Ord_in_Ord Ord_ordertype VWF_iff_Ord_less \\) then have "length (Cantor_\\ \) \ length (Cantor_\\ \)" using \ \ by (simp add: Cantor_\\ normal_Cantor_\\ omega_sum_length_leD) with \ \ * have "(Cantor_\\ \, Cantor_\\ \) \ lenlex less_than" by (auto simp: Cantor_\\ simp flip: omega_sum_less_iff) then show "(\ \, \ \) \ lenlex less_than" by (simp add: \_def into_WW_lenlex_iff) qed (auto simp: \_def WW_def strict_sorted_into_WW) finally show "\\\ \ ordertype WW (lenlex less_than)" . qed lemma ordertype_WW: "ordertype WW (lenlex less_than) = \\\" and ordertype_UNIV_\\: "ordertype UNIV (lenlex less_than) = \\\" using ordertype_WW_1 ordertype_WW_2 ordertype_WW_3 by auto lemma ordertype_\\: fixes F :: "nat \ nat list set" assumes "\j::nat. ordertype (F j) (lenlex less_than) = \\j" shows "ordertype (\j. F j) (lenlex less_than) = \\\" proof (rule antisym) show "ordertype (\ (range F)) (lenlex less_than) \ \ \ \" by (metis ordertype_UNIV_\\ ordertype_mono small top_greatest trans_llt wf_llt) have "\n. \ \ ord_of_nat n \ ordertype (\ (range F)) (lenlex less_than)" by (metis TC_small Union_upper assms ordertype_mono rangeI trans_llt wf_llt) then show "\ \ \ \ ordertype (\ (range F)) (lenlex less_than)" by (auto simp: oexp_\_Limit ZFC_in_HOL.SUP_le_iff elts_\) qed definition WW_seg :: "nat \ nat list set" where "WW_seg n \ {l \ WW. length l = n}" lemma WW_seg_subset_WW: "WW_seg n \ WW" by (auto simp: WW_seg_def) lemma WW_eq_UN_WW_seg: "WW = (\ n. WW_seg n)" by (auto simp: WW_seg_def) lemma ordertype_list_seg: "ordertype {l. length l = n} (lenlex less_than) = \\n" proof - have "bij_betw omega_sum {l. length l = n} (elts (\\n))" unfolding WW_seg_def bij_betw_def by (auto simp: inj_omega_sum Ord_mem_iff_lt omega_sum_less dest: Ex_omega_sum) then show ?thesis by (force simp: ordertype_eq_iff simp flip: eq_omega_sum_less_iff) qed lemma ordertype_WW_seg: "ordertype (WW_seg n) (lenlex less_than) = \\n" (is "ordertype ?W ?R = \\n") proof - have "ordertype {l. length l = n} ?R = ordertype ?W ?R" proof (subst ordertype_eq_ordertype) show "\f. bij_betw f {l. length l = n} ?W \ (\x\{l. length l = n}. \y\{l. length l = n}. ((f x, f y) \ lenlex less_than) = ((x, y) \ lenlex less_than))" proof (intro exI conjI) have "inj_on (into_WW 0) {l. length l = n}" by (metis from_into_WW inj_onI) then show "bij_betw (into_WW 0) {l. length l = n} ?W" by (auto simp: bij_betw_def WW_seg_def WW_eq_range_into length_into_WW) qed (simp add: into_WW_lenlex_iff) qed auto then show ?thesis using ordertype_list_seg by auto qed subsection \Definitions required for the lemmas\ subsubsection \Larson's "$<$"-relation on ordered lists\ instantiation list :: (ord)ord begin definition "xs < ys \ xs \ [] \ ys \ [] \ last xs < hd ys" for xs ys :: "'a list" definition "xs \ ys \ xs < ys \ xs = ys" for xs ys :: "'a list" instance by standard end lemma less_Nil [simp]: "xs < []" "[] < xs" by (auto simp: less_list_def) lemma less_sets_imp_list_less: assumes "list.set xs \ list.set ys" shows "xs < ys" by (metis assms last_in_set less_list_def less_sets_def list.set_sel(1)) lemma less_sets_imp_sorted_list_of_set: assumes "A \ B" "finite A" "finite B" shows "list_of A < list_of B" by (simp add: assms less_sets_imp_list_less) lemma sorted_list_of_set_imp_less_sets: assumes "xs < ys" "sorted xs" "sorted ys" shows "list.set xs \ list.set ys" using assms sorted_hd_le sorted_le_last by (force simp: less_list_def less_sets_def intro: order.trans) lemma less_list_iff_less_sets: assumes "sorted xs" "sorted ys" shows "xs < ys \ list.set xs \ list.set ys" using assms sorted_hd_le sorted_le_last by (force simp: less_list_def less_sets_def intro: order.trans) lemma strict_sorted_append_iff: "strict_sorted (xs @ ys) \ xs < ys \ strict_sorted xs \ strict_sorted ys" (is "?lhs = ?rhs") by (metis less_list_iff_less_sets less_setsD sorted_wrt_append strict_sorted_imp_less_sets strict_sorted_imp_sorted) lemma singleton_less_list_iff: "sorted xs \ [n] < xs \ {..n} \ list.set xs = {}" apply (simp add: less_list_def disjoint_iff) by (metis empty_iff less_le_trans list.set(1) list.set_sel(1) not_le sorted_hd_le) lemma less_hd_imp_less: "xs < [hd ys] \ xs < ys" by (simp add: less_list_def) lemma strict_sorted_concat_I: assumes "\x. x \ list.set xs \ strict_sorted x" "\n. Suc n < length xs \ xs!n < xs!Suc n" "xs \ lists (- {[]})" shows "strict_sorted (concat xs)" using assms proof (induction xs) case (Cons x xs) then have "x < concat xs" apply (simp add: less_list_def) by (metis Compl_iff hd_concat insertI1 length_greater_0_conv length_pos_if_in_set list.sel(1) lists.cases nth_Cons_0) with Cons show ?case by (force simp: strict_sorted_append_iff) qed auto subsection \Nash Williams for lists\ subsubsection \Thin sets of lists\ inductive initial_segment :: "'a list \ 'a list \ bool" where "initial_segment xs (xs@ys)" definition thin :: "'a list set \ bool" where "thin A \ \ (\x y. x \ A \ y \ A \ x \ y \ initial_segment x y)" lemma initial_segment_ne: assumes "initial_segment xs ys" "xs \ []" shows "ys \ [] \ hd ys = hd xs" using assms by (auto elim!: initial_segment.cases) lemma take_initial_segment: assumes "initial_segment xs ys" "k \ length xs" shows "take k xs = take k ys" by (metis append_eq_conv_conj assms initial_segment.cases min_def take_take) lemma initial_segment_length_eq: assumes "initial_segment xs ys" "length xs = length ys" shows "xs = ys" using assms initial_segment.cases by fastforce lemma initial_segment_Nil [simp]: "initial_segment [] ys" by (simp add: initial_segment.simps) lemma initial_segment_Cons [simp]: "initial_segment (x#xs) (y#ys) \ x=y \ initial_segment xs ys" by (metis append_Cons initial_segment.simps list.inject) lemma init_segment_iff_initial_segment: assumes "strict_sorted xs" "strict_sorted ys" shows "init_segment (list.set xs) (list.set ys) \ initial_segment xs ys" (is "?lhs = ?rhs") proof assume ?lhs then obtain S' where S': "list.set ys = list.set xs \ S'" "list.set xs \ S'" by (auto simp: init_segment_def) then have "finite S'" by (metis List.finite_set finite_Un) have "ys = xs @ list_of S'" using S' \strict_sorted xs\ proof (induction xs) case Nil with \strict_sorted ys\ show ?case by auto next case (Cons a xs) with \finite S'\ have "ys = a # xs @ list_of S'" by (metis List.finite_set append_Cons assms(2) sorted_list_of_set_Un sorted_list_of_set_set_of) then show ?case by (auto simp: Cons) qed then show ?rhs using initial_segment.intros by blast next assume ?rhs then show ?lhs proof cases case (1 ys) with assms(2) show ?thesis by (metis init_segment_def set_append sorted_list_of_set_imp_less_sets strict_sorted_append_iff strict_sorted_imp_sorted) qed qed theorem Nash_Williams_WW: fixes h :: "nat list \ nat" assumes "infinite M" and h: "h ` {l \ A. List.set l \ M} \ {..<2}" and "thin A" "A \ WW" obtains i N where "i < 2" "infinite N" "N \ M" "h ` {l \ A. List.set l \ N} \ {i}" proof - define AM where "AM \ {l \ A. List.set l \ M}" have "thin_set (list.set ` A)" using \thin A\ \A \ WW\ unfolding thin_def thin_set_def WW_def by (auto simp: subset_iff init_segment_iff_initial_segment) then have "thin_set (list.set ` AM)" by (simp add: AM_def image_subset_iff thin_set_def) then have "Ramsey (list.set ` AM) 2" using Nash_Williams_2 by metis moreover have "(h \ list_of) \ list.set ` AM \ {..<2}" unfolding AM_def proof clarsimp fix l assume "l \ A" "list.set l \ M" then have "strict_sorted l" using WW_def \A \ WW\ by blast then show "h (list_of (list.set l)) < 2" using h \l \ A\ \list.set l \ M\ by auto qed ultimately obtain N i where N: "N \ M" "infinite N" "i<2" and "list.set ` AM \ Pow N \ (h \ list_of) -` {i}" unfolding Ramsey_eq by (metis \infinite M\) then have N_disjoint: "(h \ list_of) -` {1-i} \ (list.set ` AM) \ Pow N = {}" unfolding subset_vimage_iff less_2_cases_iff by force have "h ` {l \ A. list.set l \ N} \ {i}" proof clarify fix l assume "l \ A" and "list.set l \ N" then have "h l < 2" using h \N \ M\ by force with \i<2\ have "h l \ Suc 0 - i \ h l = i" by (auto simp: eval_nat_numeral less_Suc_eq) moreover have "strict_sorted l" using \A \ WW\ \l \ A\ unfolding WW_def by blast moreover have "h (list_of (list.set l)) = 1 - i \ \ (list.set l \ N)" using N_disjoint \N \ M\ \l \ A\ by (auto simp: AM_def) ultimately show "h l = i" using N \N \ M\ \l \ A\ \list.set l \ N\ by (auto simp: vimage_def set_eq_iff AM_def WW_def subset_iff) qed then show thesis using that \i<2\ N by auto qed subsection \Specialised functions on lists\ lemma mem_lists_non_Nil: "xss \ lists (- {[]}) \ (\x \ list.set xss. x \ [])" by auto fun acc_lengths :: "nat \ 'a list list \ nat list" where "acc_lengths acc [] = []" | "acc_lengths acc (l#ls) = (acc + length l) # acc_lengths (acc + length l) ls" lemma length_acc_lengths [simp]: "length (acc_lengths acc ls) = length ls" by (induction ls arbitrary: acc) auto lemma acc_lengths_eq_Nil_iff [simp]: "acc_lengths acc ls = [] \ ls = []" by (metis length_0_conv length_acc_lengths) lemma set_acc_lengths: assumes "ls \ lists (- {[]})" shows "list.set (acc_lengths acc ls) \ {acc<..}" using assms by (induction ls rule: acc_lengths.induct) fastforce+ text \Useful because @{text acc_lengths.simps} will sometimes be deleted from the simpset.\ lemma hd_acc_lengths [simp]: "hd (acc_lengths acc (l#ls)) = acc + length l" by simp lemma last_acc_lengths [simp]: "ls \ [] \ last (acc_lengths acc ls) = acc + sum_list (map length ls)" by (induction acc ls rule: acc_lengths.induct) auto lemma nth_acc_lengths [simp]: "\ls \ []; k < length ls\ \ acc_lengths acc ls ! k = acc + sum_list (map length (take (Suc k) ls))" by (induction acc ls arbitrary: k rule: acc_lengths.induct) (fastforce simp: less_Suc_eq nth_Cons')+ lemma acc_lengths_plus: "acc_lengths (m+n) as = map ((+)m) (acc_lengths n as)" by (induction n as arbitrary: m rule: acc_lengths.induct) (auto simp: add.assoc) lemma acc_lengths_shift: "NO_MATCH 0 acc \ acc_lengths acc as = map ((+)acc) (acc_lengths 0 as)" by (metis acc_lengths_plus add.comm_neutral) lemma length_concat_acc_lengths: "ls \ [] \ k + length (concat ls) \ list.set (acc_lengths k ls)" by (metis acc_lengths_eq_Nil_iff last_acc_lengths last_in_set length_concat) lemma strict_sorted_acc_lengths: assumes "ls \ lists (- {[]})" shows "strict_sorted (acc_lengths acc ls)" using assms proof (induction ls rule: acc_lengths.induct) case (2 acc l ls) then have "strict_sorted (acc_lengths (acc + length l) ls)" by auto then show ?case using set_acc_lengths "2.prems" by auto qed auto lemma acc_lengths_append: "acc_lengths acc (xs @ ys) = acc_lengths acc xs @ acc_lengths (acc + sum_list (map length xs)) ys" by (induction acc xs rule: acc_lengths.induct) (auto simp: add.assoc) lemma length_concat_ge: assumes "as \ lists (- {[]})" shows "length (concat as) \ length as" using assms proof (induction as) case (Cons a as) then have "length a \ Suc 0" "\l. l \ list.set as \ length l \ Suc 0" by (auto simp: Suc_leI) then show ?case using Cons.IH by force qed auto fun interact :: "'a list list \ 'a list list \ 'a list" where "interact [] ys = concat ys" | "interact xs [] = concat xs" | "interact (x#xs) (y#ys) = x @ y @ interact xs ys" lemma (in monoid_add) length_interact: "length (interact xs ys) = sum_list (map length xs) + sum_list (map length ys)" by (induction rule: interact.induct) (auto simp: length_concat) lemma length_interact_ge: assumes "xs \ lists (- {[]})" "ys \ lists (- {[]})" shows "length (interact xs ys) \ length xs + length ys" by (metis add_mono assms length_concat length_concat_ge length_interact) lemma set_interact [simp]: shows "list.set (interact xs ys) = list.set (concat xs) \ list.set (concat ys)" by (induction rule: interact.induct) auto lemma interact_eq_Nil_iff [simp]: assumes "xs \ lists (- {[]})" "ys \ lists (- {[]})" shows "interact xs ys = [] \ xs=[] \ ys=[]" using length_interact_ge [OF assms] by fastforce lemma interact_sing [simp]: "interact [x] ys = x @ concat ys" by (metis (no_types) concat.simps(2) interact.simps neq_Nil_conv) lemma hd_interact: "\xs \ []; hd xs \ []\ \ hd (interact xs ys) = hd (hd xs)" by (smt (verit, best) hd_append2 hd_concat interact.elims list.sel(1)) lemma acc_lengths_concat_injective: assumes "concat as' = concat as" "acc_lengths n as' = acc_lengths n as" shows "as' = as" using assms proof (induction as arbitrary: n as') case Nil then show ?case by (metis acc_lengths_eq_Nil_iff) next case (Cons a as) then obtain a' bs where "as' = a'#bs" by (metis Suc_length_conv length_acc_lengths) with Cons show ?case by simp qed lemma acc_lengths_interact_injective: assumes "interact as' bs' = interact as bs" "acc_lengths a as' = acc_lengths a as" "acc_lengths b bs' = acc_lengths b bs" shows "as' = as \ bs' = bs" using assms proof (induction as bs arbitrary: a b as' bs' rule: interact.induct) case (1 cs) then show ?case by (metis acc_lengths_concat_injective acc_lengths_eq_Nil_iff interact.simps(1)) next case (2 c cs) then show ?case by (metis acc_lengths_concat_injective acc_lengths_eq_Nil_iff interact.simps(2) list.exhaust) next case (3 x xs y ys) then obtain a' us b' vs where "as' = a'#us" "bs' = b'#vs" by (metis length_Suc_conv length_acc_lengths) with 3 show ?case by auto qed lemma strict_sorted_interact_I: assumes "length ys \ length xs" "length xs \ Suc (length ys)" "\x. x \ list.set xs \ strict_sorted x" "\y. y \ list.set ys \ strict_sorted y" "\n. n < length ys \ xs!n < ys!n" "\n. Suc n < length xs \ ys!n < xs!Suc n" assumes "xs \ lists (- {[]})" "ys \ lists (- {[]})" shows "strict_sorted (interact xs ys)" using assms proof (induction rule: interact.induct) case (3 x xs y ys) then have "x < y" by force moreover have "strict_sorted (interact xs ys)" using 3 by simp (metis Suc_less_eq nth_Cons_Suc) moreover have "y < interact xs ys" using 3 apply (simp add: less_list_def) by (metis hd_interact le_zero_eq length_greater_0_conv list.sel(1) list.set_sel(1) list.size(3) lists.simps mem_lists_non_Nil nth_Cons_0) ultimately show ?case using 3 by (simp add: strict_sorted_append_iff less_list_def) qed auto subsection \Forms and interactions\ subsubsection \Forms\ inductive Form_Body :: "[nat, nat, nat list, nat list, nat list] \ bool" where "Form_Body ka kb xs ys zs" if "length xs < length ys" "xs = concat (a#as)" "ys = concat (b#bs)" "a#as \ lists (- {[]})" "b#bs \ lists (- {[]})" "length (a#as) = ka" "length (b#bs) = kb" "c = acc_lengths 0 (a#as)" "d = acc_lengths 0 (b#bs)" "zs = concat [c, a, d, b] @ interact as bs" "strict_sorted zs" inductive Form :: "[nat, nat list set] \ bool" where "Form 0 {xs,ys}" if "length xs = length ys" "xs \ ys" | "Form (2*k-1) {xs,ys}" if "Form_Body k k xs ys zs" "k > 0" | "Form (2*k) {xs,ys}" if "Form_Body (Suc k) k xs ys zs" "k > 0" inductive_cases Form_0_cases_raw: "Form 0 u" lemma Form_elim_upair: assumes "Form l U" obtains xs ys where "xs \ ys" "U = {xs,ys}" "length xs \ length ys" using assms by (smt (verit, best) Form.simps Form_Body.cases less_or_eq_imp_le nat_neq_iff) lemma assumes "Form_Body ka kb xs ys zs" shows Form_Body_WW: "zs \ WW" and Form_Body_nonempty: "length zs > 0" and Form_Body_length: "length xs < length ys" using Form_Body.cases [OF assms] by (fastforce simp: WW_def)+ lemma form_cases: fixes l::nat obtains (zero) "l = 0" | (nz) ka kb where "l = ka+kb - 1" "0 < kb" "kb \ ka" "ka \ Suc kb" proof - have "l = 0 \ (\ka kb. l = ka+kb - 1 \ 0 < kb \ kb \ ka \ ka \ Suc kb)" by presburger then show thesis using nz zero by blast qed subsubsection \Interactions\ lemma interact: assumes "Form l U" "l>0" obtains ka kb xs ys zs where "l = ka+kb - 1" "U = {xs,ys}" "Form_Body ka kb xs ys zs" "0 < kb" "kb \ ka" "ka \ Suc kb" using assms unfolding Form.simps by (smt (verit, best) add_Suc diff_Suc_1 lessI mult_2 nat_less_le order_refl) definition inter_scheme :: "nat \ nat list set \ nat list" where "inter_scheme l U \ SOME zs. \k xs ys. U = {xs,ys} \ (l = 2*k-1 \ Form_Body k k xs ys zs \ l = 2*k \ Form_Body (Suc k) k xs ys zs)" lemma inter_scheme: assumes "Form l U" "l>0" obtains ka kb xs ys where "l = ka+kb - 1" "U = {xs,ys}" "Form_Body ka kb xs ys (inter_scheme l U)" "0 < kb" "kb \ ka" "ka \ Suc kb" using interact [OF \Form l U\] proof cases case (2 ka kb xs ys zs) then have \
: "\ka kb zs. \ Form_Body ka kb ys xs zs" using Form_Body_length less_asym' by blast have "Form_Body ka kb xs ys (inter_scheme l U)" proof (cases "ka = kb") case True with 2 have l: "\k. l \ k * 2" by presburger have [simp]: "\k. kb + kb - Suc 0 = k * 2 - Suc 0 \ k=kb" by auto show ?thesis unfolding inter_scheme_def using 2 l True by (auto simp: \
\l > 0\ Set.doubleton_eq_iff conj_disj_distribR ex_disj_distrib algebra_simps some_eq_ex) next case False with 2 have l: "\k. l \ k * 2 - Suc 0" and [simp]: "ka = Suc kb" by presburger+ have [simp]: "\k. kb + kb = k * 2 \ k=kb" by auto show ?thesis unfolding inter_scheme_def using 2 l False by (auto simp: \
\l > 0\ Set.doubleton_eq_iff conj_disj_distribR ex_disj_distrib algebra_simps some_eq_ex) qed then show ?thesis by (simp add: 2 that) qed (use \l > 0\ in auto) lemma inter_scheme_strict_sorted: assumes "Form l U" "l>0" shows "strict_sorted (inter_scheme l U)" using Form_Body.simps assms inter_scheme by fastforce lemma inter_scheme_simple: assumes "Form l U" "l>0" shows "inter_scheme l U \ WW \ length (inter_scheme l U) > 0" using inter_scheme [OF assms] by (meson Form_Body_WW Form_Body_nonempty) subsubsection \Injectivity of interactions\ proposition inter_scheme_injective: assumes "Form l U" "Form l U'" "l > 0" and eq: "inter_scheme l U' = inter_scheme l U" shows "U' = U" proof - obtain ka kb xs ys where l: "l = ka+kb - 1" and U: "U = {xs,ys}" and FB: "Form_Body ka kb xs ys (inter_scheme l U)" and kb: "0 < kb" "kb \ ka" "ka \ Suc kb" using assms inter_scheme by blast then obtain a as b bs c d where xs: "xs = concat (a#as)" and ys: "ys = concat (b#bs)" and len: "length (a#as) = ka" "length (b#bs) = kb" and c: "c = acc_lengths 0 (a#as)" and d: "d = acc_lengths 0 (b#bs)" and Ueq: "inter_scheme l U = concat [c, a, d, b] @ interact as bs" by (auto simp: Form_Body.simps) obtain ka' kb' xs' ys' where l': "l = ka'+kb' - 1" and U': "U' = {xs',ys'}" and FB': "Form_Body ka' kb' xs' ys' (inter_scheme l U')" and kb': "0 < kb'" "kb' \ ka'" "ka' \ Suc kb'" using assms inter_scheme by blast then obtain a' as' b' bs' c' d' where xs': "xs' = concat (a'#as')" and ys': "ys' = concat (b'#bs')" and len': "length (a'#as') = ka'" "length (b'#bs') = kb'" and c': "c' = acc_lengths 0 (a'#as')" and d': "d' = acc_lengths 0 (b'#bs')" and Ueq': "inter_scheme l U' = concat [c', a', d', b'] @ interact as' bs'" using Form_Body.simps by auto have [simp]: "ka' = ka \ kb' = kb" using \l > 0\ l l' kb kb' le_SucE le_antisym mult_2 by linarith have [simp]: "length c = length c'" "length d = length d'" using c c' d d' len' len by auto have c_off: "c' = c" "a' @ d' @ b' @ interact as' bs' = a @ d @ b @ interact as bs" using eq by (auto simp: Ueq Ueq') then have len_a: "length a' = length a" by (metis acc_lengths.simps(2) add.left_neutral c c' nth_Cons_0) with c_off have \
: "a' = a" "d' = d" "b' @ interact as' bs' = b @ interact as bs" by auto then have "length (interact as' bs') = length (interact as bs)" by (metis acc_lengths.simps(2) add_left_cancel append_eq_append_conv d d' list.inject) with \
have "b' = b" "interact as' bs' = interact as bs" by auto moreover have "acc_lengths 0 as' = acc_lengths 0 as" using \a' = a\ \c' = c\ by (simp add: c' c acc_lengths_shift) moreover have "acc_lengths 0 bs' = acc_lengths 0 bs" using \b' = b\ \d' = d\ by (simp add: d' d acc_lengths_shift) ultimately have "as' = as \ bs' = bs" using acc_lengths_interact_injective by blast then show ?thesis by (simp add: \a' = a\ U U' \b' = b\ xs xs' ys ys') qed lemma strict_sorted_interact_imp_concat: "strict_sorted (interact as bs) \ strict_sorted (concat as) \ strict_sorted (concat bs)" proof (induction as bs rule: interact.induct) case (3 x xs y ys) have "x < concat xs" using "3.prems" by (smt (verit, del_insts) Un_iff hd_in_set interact.simps(3) last_in_set less_list_def set_append set_interact sorted_wrt_append) moreover have "y < concat ys" using 3 sorted_wrt_append strict_sorted_append_iff by fastforce ultimately show ?case using 3 by (auto simp add: strict_sorted_append_iff) qed auto lemma strict_sorted_interact_hd: "\strict_sorted (interact cs ds); cs \ []; ds \ []; hd cs \ []; hd ds \ []\ \ hd (hd cs) < hd (hd ds)" by (metis append_is_Nil_conv hd_append2 hd_in_set interact.simps(3) list.exhaust_sel sorted_wrt_append) text \the lengths of the two lists can differ by one\ proposition interaction_scheme_unique_aux: assumes "concat as = concat as'" and ys': "concat bs = concat bs'" and "as \ lists (- {[]})" "bs \ lists (- {[]})" and "strict_sorted (interact as bs)" and "length bs \ length as" "length as \ Suc (length bs)" and "as' \ lists (- {[]})" "bs' \ lists (- {[]})" and "strict_sorted (interact as' bs')" and "length bs' \ length as'" "length as' \ Suc (length bs')" and "length as = length as'" "length bs = length bs'" shows "as = as' \ bs = bs'" using assms proof (induction "length as" arbitrary: as bs as' bs') case 0 then show ?case by auto next case (Suc k) show ?case proof (cases k) case 0 then obtain a a' where aa': "as = [a]" "as' = [a']" by (metis Suc.hyps(2) \length as = length as'\ Suc_length_conv length_0_conv) show ?thesis proof show "as = as'" using aa' \concat as = concat as'\ by force with Suc 0 show " bs = bs'" by (metis Suc_leI append_Nil2 concat.simps impossible_Cons le_antisym length_greater_0_conv list.exhaust) qed next case (Suc k') then obtain a cs b ds where eq: "as = a#cs" "bs = b#ds" using Suc.prems by (metis Suc.hyps(2) le0 list.exhaust list.size(3) not_less_eq_eq) have "length as' \ 0" using Suc.hyps(2) \length as = length as'\ by force then obtain a' cs' b' ds' where eq': "as' = a'#cs'" "bs' = b'#ds'" by (metis \length bs = length bs'\ eq(2) length_0_conv list.exhaust) obtain k: "k = length cs" "k \ Suc (length ds)" using eq \Suc k = length as\ \length bs \ length as\ \length as \ Suc (length bs)\ by auto case (Suc k') obtain [simp]: "b \ []" "b' \ []" "a \ []" "a' \ []" using Suc.prems by (simp add: eq eq') then have "hd b' = hd b" using Suc.prems(2) by (metis concat.simps(2) eq'(2) eq(2) hd_append2) have ss_ab: "strict_sorted (concat as)" "strict_sorted (concat bs)" using strict_sorted_interact_imp_concat Suc.prems(5) by blast+ have sw_ab: "strict_sorted (a @ b @ interact cs ds)" by (metis Suc.prems(5) eq interact.simps(3)) then obtain "a < b" "strict_sorted a" "strict_sorted b" by (metis append_assoc strict_sorted_append_iff) have b_cs: "strict_sorted (concat (b # cs))" by (metis append.simps(1) concat.simps(2) interact.simps(3) strict_sorted_interact_imp_concat sw_ab) then have "b < concat cs" using strict_sorted_append_iff by auto have "strict_sorted (a @ concat cs)" using eq(1) ss_ab(1) by force have "list.set a = list.set (concat as) \ {..< hd b}" proof - have "x \ list.set a" if "x < hd b" and "l \ list.set cs" and "x \ list.set l" for x l using b_cs sorted_hd_le strict_sorted_imp_sorted that by fastforce then show ?thesis using \b \ []\ sw_ab by (force simp: strict_sorted_append_iff sorted_wrt_append eq) qed moreover have ss_ab': "strict_sorted (concat as')" "strict_sorted (concat bs')" using strict_sorted_interact_imp_concat Suc.prems(10) by blast+ have sw_ab': "strict_sorted (a' @ b' @ interact cs' ds')" by (metis Suc.prems(10) eq' interact.simps(3)) then obtain "a' < b'" "strict_sorted a'" "strict_sorted b'" by (metis append_assoc strict_sorted_append_iff) have b_cs': "strict_sorted (concat (b' # cs'))" by (metis (no_types) Suc.prems(10) append_Nil eq' interact.simps(3) strict_sorted_append_iff strict_sorted_interact_imp_concat) then have "b' < concat cs'" by (simp add: strict_sorted_append_iff) then have "hd b' \ list.set (concat cs')" by (metis Un_iff \b' \ []\ list.set_sel(1) not_less_iff_gr_or_eq set_interact sorted_wrt_append sw_ab') have "strict_sorted (a' @ concat cs')" using eq'(1) ss_ab'(1) by force then have b_cs': "strict_sorted (b' @ concat cs')" using \b' < concat cs'\ eq'(2) ss_ab'(2) strict_sorted_append_iff by auto have "list.set a' = list.set (concat as') \ {..< hd b'}" proof - have "x \ list.set a'" if "x < hd b'" and "l \ list.set cs'" and "x \ list.set l" for x l using b_cs' sorted_hd_le strict_sorted_imp_sorted that by fastforce then show ?thesis using \b' \ []\ sw_ab' by (force simp: strict_sorted_append_iff sorted_wrt_append eq') qed ultimately have "a=a'" by (simp add: Suc.prems(1) \hd b' = hd b\ \strict_sorted a'\ \strict_sorted a\ strict_sorted_equal) moreover have ccat_cs_cs': "concat cs = concat cs'" using Suc.prems(1) \a = a'\ eq'(1) eq(1) by fastforce have "b=b'" proof (cases "ds = [] \ ds' = []") case True then show ?thesis using \length bs = length bs'\ Suc.prems(2) eq'(2) eq(2) by auto next case False then have "ds \ []" "ds' \ []" "sorted (concat ds)" "sorted (concat ds')" using eq(2) ss_ab(2) eq'(2) ss_ab'(2) strict_sorted_append_iff strict_sorted_imp_sorted by auto have "strict_sorted b" "strict_sorted b'" using b_cs b_cs' sorted_wrt_append by auto moreover have "cs \ []" using k local.Suc by auto then obtain "hd cs \ []" "hd ds \ []" using Suc.prems(3) Suc.prems(4) eq list.set_sel(1) by (simp add: \ds \ []\ mem_lists_non_Nil) then have "concat cs \ []" using \cs \ []\ hd_in_set by auto have "hd (concat cs) < hd (concat ds)" using strict_sorted_interact_hd by (metis \cs \ []\ \ds \ []\ \hd cs \ []\ \hd ds \ []\ hd_concat sorted_wrt_append sw_ab) have "list.set b = list.set (concat bs) \ {..< hd (concat cs)}" proof - have 1: "x \ list.set b" if "x < hd (concat cs)" and "l \ list.set ds" and "x \ list.set l" for x l using \hd (concat cs) < hd (concat ds)\ \sorted (concat ds)\ sorted_hd_le that by fastforce have 2: "l < hd (concat cs)" if "l \ list.set b" for l by (meson \b < concat cs\ \b \ []\ \concat cs \ []\ \strict_sorted b\ le_less_trans less_list_def sorted_le_last strict_sorted_imp_sorted that) show ?thesis using 1 2 by (auto simp: strict_sorted_append_iff sorted_wrt_append eq) qed moreover have "cs' \ []" using k local.Suc \concat cs \ []\ ccat_cs_cs' by auto then obtain "hd cs' \ []" "hd ds' \ []" using Suc.prems(8,9) \ds' \ []\ eq'(1) eq'(2) list.set_sel(1) by auto then have "concat cs' \ []" using \cs' \ []\ hd_in_set by auto have "hd (concat cs') < hd (concat ds')" using strict_sorted_interact_hd by (metis \cs' \ []\ \ds' \ []\ \hd cs' \ []\ \hd ds' \ []\ hd_concat sorted_wrt_append sw_ab') have "list.set b' = list.set (concat bs') \ {..< hd (concat cs')}" proof - have 1: "x \ list.set b'" if "x < hd (concat cs')" and "l \ list.set ds'" and "x \ list.set l" for x l using \hd (concat cs') < hd (concat ds')\ \sorted (concat ds')\ sorted_hd_le that by fastforce have 2: "l < hd (concat cs')" if "l \ list.set b'" for l by (metis \concat cs' \ []\ b_cs' list.set_sel(1) sorted_wrt_append that) show ?thesis using 1 2 by (auto simp: strict_sorted_append_iff sorted_wrt_append eq') qed ultimately show "b = b'" by (simp add: Suc.prems(2) ccat_cs_cs' strict_sorted_equal) qed moreover have "cs = cs' \ ds = ds'" proof (rule Suc.hyps) show "k = length cs" using eq Suc.hyps(2) by auto[1] show "concat ds = concat ds'" using Suc.prems(2) \b = b'\ eq'(2) eq(2) by auto show "strict_sorted (interact cs ds)" using eq Suc.prems(5) strict_sorted_append_iff by auto show "length ds \ length cs" "length cs \ Suc (length ds)" using eq Suc.hyps(2) Suc.prems(6) k by auto show "strict_sorted (interact cs' ds')" using eq' Suc.prems(10) strict_sorted_append_iff by auto show "length cs = length cs'" using Suc.hyps(2) Suc.prems(13) eq'(1) k(1) by force qed (use ccat_cs_cs' eq eq' Suc.prems in auto) ultimately show ?thesis by (simp add: \a = a'\ \b = b'\ eq eq') qed qed proposition Form_Body_unique: assumes "Form_Body ka kb xs ys zs" "Form_Body ka kb xs ys zs'" and "kb \ ka" "ka \ Suc kb" shows "zs' = zs" proof - obtain a as b bs c d where xs: "xs = concat (a#as)" and ys: "ys = concat (b#bs)" and ne: "a#as \ lists (- {[]})" "b#bs \ lists (- {[]})" and len: "length (a#as) = ka" "length (b#bs) = kb" and c: "c = acc_lengths 0 (a#as)" and d: "d = acc_lengths 0 (b#bs)" and Ueq: "zs = concat [c, a, d, b] @ interact as bs" and ss_zs: "strict_sorted zs" using Form_Body.cases [OF assms(1)] by (metis (no_types)) obtain a' as' b' bs' c' d' where xs': "xs = concat (a'#as')" and ys': "ys = concat (b'#bs')" and ne': "a'#as' \ lists (- {[]})" "b'#bs' \ lists (- {[]})" and len': "length (a'#as') = ka" "length (b'#bs') = kb" and c': "c' = acc_lengths 0 (a'#as')" and d': "d' = acc_lengths 0 (b'#bs')" and Ueq': "zs' = concat [c', a', d', b'] @ interact as' bs'" and ss_zs': "strict_sorted zs'" using Form_Body.cases [OF assms(2)] by (metis (no_types)) have [simp]: "length c = length c'" "length d = length d'" using c c' d d' len' len by auto note acc_lengths.simps [simp del] have "a < b" using ss_zs by (auto simp: Ueq strict_sorted_append_iff less_list_def c d) have "a' < b'" using ss_zs' by (auto simp: Ueq' strict_sorted_append_iff less_list_def c' d') have "a#as = a'#as' \ b#bs = b'#bs'" proof (rule interaction_scheme_unique_aux) show "strict_sorted (interact (a # as) (b # bs))" using ss_zs \a < b\ by (auto simp: Ueq strict_sorted_append_iff less_list_def d) show "strict_sorted (interact (a' # as') (b' # bs'))" using ss_zs' \a' < b'\ by (auto simp: Ueq' strict_sorted_append_iff less_list_def d') show "length (b # bs) \ length (a # as)" "length (b' # bs') \ length (a' # as')" using \kb \ ka\ len len' by auto show "length (a # as) \ Suc (length (b # bs))" using \ka \ Suc kb\ len by linarith then show "length (a' # as') \ Suc (length (b' # bs'))" using len len' by fastforce qed (use len len' xs xs' ys ys' ne ne' in fastforce)+ then show ?thesis using Ueq Ueq' c c' d d' by blast qed lemma Form_Body_imp_inter_scheme: assumes FB: "Form_Body ka kb xs ys zs" and "0 < kb" "kb \ ka" "ka \ Suc kb" shows "zs = inter_scheme ((ka+kb) - Suc 0) {xs,ys}" proof - have "length xs < length ys" by (meson Form_Body_length assms(1)) have [simp]: "a + a = b + b \ a=b" "a + a - Suc 0 = b + b - Suc 0 \ a=b" for a b::nat by auto show ?thesis proof (cases "ka = kb") case True show ?thesis unfolding inter_scheme_def apply (rule some_equality [symmetric], metis One_nat_def True FB mult_2) subgoal for zs' using assms \length xs < length ys\ by (auto simp: True mult_2 Set.doubleton_eq_iff Form_Body_unique dest: Form_Body_length, presburger) done next case False then have eq: "ka = Suc kb" using assms by linarith show ?thesis unfolding inter_scheme_def apply (rule some_equality [symmetric], use assms False mult_2 one_is_add eq in fastforce) subgoal for zs' using assms \length xs < length ys\ by (auto simp: eq mult_2 Set.doubleton_eq_iff Form_Body_unique dest: Form_Body_length, presburger) done qed qed subsection \For Lemma 3.8 AND PROBABLY 3.7\ definition grab :: "nat set \ nat \ nat set \ nat set" where "grab N n \ (N \ enumerate N ` {.. {enumerate N n..})" lemma grab_0 [simp]: "grab N 0 = ({}, N)" by (fastforce simp: grab_def enumerate_0 Least_le) lemma less_sets_grab: "infinite N \ fst (grab N n) \ snd (grab N n)" by (auto simp: grab_def less_sets_def intro: enumerate_mono less_le_trans) lemma finite_grab [iff]: "finite (fst (grab N n))" by (simp add: grab_def) lemma card_grab [simp]: assumes "infinite N" shows "card (fst (grab N n)) = n" proof - have "N \ enumerate N ` {.. N" using grab_def range_enum by fastforce lemma snd_grab_subset: "snd (grab N n) \ N" by (auto simp: grab_def) lemma grab_Un_eq: assumes "infinite N" shows "fst (grab N n) \ snd (grab N n) = N" proof show "N \ fst (grab N n) \ snd (grab N n)" unfolding grab_def using assms enumerate_Ex le_less_linear strict_mono_enum strict_mono_less by fastforce qed (simp add: grab_def) lemma finite_grab_iff [simp]: "finite (snd (grab N n)) \ finite N" by (metis finite_grab grab_Un_eq infinite_Un infinite_super snd_grab_subset) lemma grab_eqD: "\grab N n = (A,M); infinite N\ \ A \ M \ finite A \ card A = n \ infinite M \ A \ N \ M \ N" using card_grab grab_def less_sets_grab finite_grab_iff by auto lemma less_sets_fst_grab: "A \ N \ A \ fst (grab N n)" by (simp add: fst_grab_subset less_sets_weaken2) text\Possibly redundant, given @{term grab}\ definition nxt where "nxt \ \N. \n::nat. N \ {n<..}" lemma infinite_nxtN: "infinite N \ infinite (nxt N n)" by (simp add: infinite_nat_greaterThan nxt_def) lemma nxt_subset: "nxt N n \ N" unfolding nxt_def by blast lemma nxt_subset_greaterThan: "m \ n \ nxt N n \ {m<..}" by (auto simp: nxt_def) lemma nxt_subset_atLeast: "m \ n \ nxt N n \ {m..}" by (auto simp: nxt_def) lemma enum_nxt_ge: "infinite N \ a \ enum (nxt N a) n" by (simp add: atLeast_le_enum infinite_nxtN nxt_subset_atLeast) lemma inj_enum_nxt: "infinite N \ inj_on (enum (nxt N a)) A" by (simp add: infinite_nxtN strict_mono_enum strict_mono_imp_inj_on) subsection \Larson's Lemma 3.11\ text \Again from Jean A. Larson, A short proof of a partition theorem for the ordinal $\omega^\omega$. \emph{Annals of Mathematical Logic}, 6:129–145, 1973.\ lemma lemma_3_11: assumes "l > 0" shows "thin (inter_scheme l ` {U. Form l U})" using form_cases [of l] proof cases case zero then show ?thesis using assms by auto next case (nz ka kb) note acc_lengths.simps [simp del] show ?thesis unfolding thin_def proof clarify fix U U' assume ne: "inter_scheme l U \ inter_scheme l U'" and init: "initial_segment (inter_scheme l U) (inter_scheme l U')" assume "Form l U" then obtain kp kq xs ys where "l = kp+kq - 1" "U = {xs,ys}" and U: "Form_Body kp kq xs ys (inter_scheme l U)" and "0 < kq" "kq \ kp" "kp \ Suc kq" using assms inter_scheme by blast then have "kp = ka \ kq = kb" using nz by linarith then obtain a as b bs c d where len: "length (a#as) = ka" "length (b#bs) = kb" and c: "c = acc_lengths 0 (a#as)" and d: "d = acc_lengths 0 (b#bs)" and Ueq: "inter_scheme l U = concat [c, a, d, b] @ interact as bs" using U by (auto simp: Form_Body.simps) assume "Form l U'" then obtain kp' kq' xs' ys' where "l = kp'+kq' - 1" "U' = {xs',ys'}" and U': "Form_Body kp' kq' xs' ys' (inter_scheme l U')" and "0 < kq'" "kq' \ kp'" "kp' \ Suc kq'" using assms inter_scheme by blast then have "kp' = ka \ kq' = kb" using nz by linarith then obtain a' as' b' bs' c' d' where len': "length (a'#as') = ka" "length (b'#bs') = kb" and c': "c' = acc_lengths 0 (a'#as')" and d': "d' = acc_lengths 0 (b'#bs')" and Ueq': "inter_scheme l U' = concat [c', a', d', b'] @ interact as' bs'" using U' by (auto simp: Form_Body.simps) have [simp]: "length bs' = length bs" "length as' = length as" using len len' by auto have "inter_scheme l U \ []" "inter_scheme l U' \ []" using Form_Body_nonempty U U' by auto define u1 where "u1 \ hd (inter_scheme l U)" have u1_eq': "u1 = hd (inter_scheme l U')" using \inter_scheme l U \ []\ init u1_def initial_segment_ne by fastforce have au1: "u1 = length a" by (simp add: u1_def Ueq c) have au1': "u1 = length a'" by (simp add: u1_eq' Ueq' c') have len_eqk: "length c' = ka" "length d' = kb" "length c' = ka" "length d' = kb" using c d len c' d' len' by auto have take: "take (ka + u1 + kb) (c @ a @ d @ l) = c @ a @ d" "take (ka + u1 + kb) (c' @ a' @ d' @ l) = c' @ a' @ d'" for l using c d c' d' len by (simp_all flip: au1 au1') have leU: "ka + u1 + kb \ length (inter_scheme l U)" using c d len by (simp add: au1 Ueq) then have "take (ka + u1 + kb) (inter_scheme l U) = take (ka + u1 + kb) (inter_scheme l U')" using take_initial_segment init by blast then have \
: "c @ a @ d = c' @ a' @ d'" by (metis Ueq Ueq' append.assoc concat.simps(2) take) have "length (inter_scheme l U) = ka + (c @ a @ d)!(ka-1) + kb + last d" by (simp add: Ueq c d length_interact nth_append flip: len) moreover have "length (inter_scheme l U') = ka + (c' @ a' @ d')!(ka-1) + kb + last d'" by (simp add: Ueq' c' d' length_interact nth_append flip: len') moreover have "last d = last d'" using "\
" c d d' len'(1) len_eqk(1) by auto ultimately have "length (inter_scheme l U) = length (inter_scheme l U')" by (simp add: \
) then show False using init initial_segment_length_eq ne by blast qed qed subsection \Larson's Lemma 3.6\ proposition lemma_3_6: fixes g :: "nat list set \ nat" assumes g: "g \ [WW]\<^bsup>2\<^esup> \ {0,1}" obtains N j where "infinite N" and "\k u. \k > 0; u \ [WW]\<^bsup>2\<^esup>; Form k u; [enum N k] < inter_scheme k u; List.set (inter_scheme k u) \ N\ \ g u = j k" proof - define \ where "\ \ \m::nat. \M. infinite M \ m < Inf M" define \ where "\ \ \l m n::nat. \M N j. n > m \ N \ M \ n \ M \ (\U. Form l U \ U \ WW \ [n] < inter_scheme l U \ list.set (inter_scheme l U) \ N \ g U = j)" have *: "\n N j. \ n N \ \ l m n M N j" if "l > 0" "\ m M" for l m::nat and M :: "nat set" proof - define FF where "FF \ {U \ [WW]\<^bsup>2\<^esup>. Form l U}" define h where "h \ \zs. g (inv_into FF (inter_scheme l) zs)" have "thin (inter_scheme l ` FF)" using \l > 0\ lemma_3_11 by (simp add: thin_def FF_def) moreover have "inter_scheme l ` FF \ WW" using inter_scheme_simple \0 < l\ FF_def by blast moreover have "h ` {xs \ inter_scheme l ` FF. List.set xs \ M} \ {..<2}" using g inv_into_into[of concl: "FF" "inter_scheme l"] by (force simp: h_def FF_def Pi_iff) ultimately obtain j N where "j < 2" "infinite N" "N \ M" and hj: "h ` {xs \ inter_scheme l ` FF. List.set xs \ N} \ {j}" using \\ m M\ unfolding \_def by (blast intro: Nash_Williams_WW [of M]) let ?n = "Inf N" have "?n > m" using \\ m M\ \infinite N\ unfolding \_def Inf_nat_def infinite_nat_iff_unbounded by (metis LeastI_ex \N \ M\ le_less_trans not_less not_less_Least subsetD) have "g U = j" if "Form l U" "U \ WW" "[?n] < inter_scheme l U" "list.set (inter_scheme l U) \ N - {?n}" for U proof - obtain xs ys where xys: "xs \ ys" "U = {xs,ys}" using Form_elim_upair \Form l U\ by blast moreover have "inj_on (inter_scheme l) FF" using \0 < l\ inj_on_def inter_scheme_injective FF_def by blast moreover have "g (inv_into FF (inter_scheme l) (inter_scheme l U)) = j" using hj that xys subset_Diff_insert by (fastforce simp: h_def FF_def image_iff) ultimately show ?thesis using that FF_def by auto qed moreover have "?n < Inf (N - {?n})" by (metis Diff_iff Inf_nat_def Inf_nat_def1 \infinite N\ finite.emptyI infinite_remove linorder_neqE_nat not_less_Least singletonI) moreover have "?n \ M" by (metis Inf_nat_def1 \N \ M\ \infinite N\ finite.emptyI subsetD) ultimately have "\ ?n (N - {?n}) \ \ l m ?n M (N - {?n}) j" using \\ m M\ \infinite N\ \N \ M\ \?n > m\ by (auto simp: \_def \_def) then show ?thesis by blast qed have base: "\ 0 {0<..}" unfolding \_def by (metis infinite_Ioi Inf_nat_def1 greaterThan_iff greaterThan_non_empty) have step: "Ex (\(n,N,j). \ n N \ \ l m n M N j)" if "\ m M" "l > 0" for m M l using * [of l m M] that by (auto simp: \_def) define G where "G \ \l m M. @(n,N,j). \ n N \ \ (Suc l) m n M N j" have G\: "(\(n,N,j). \ n N) (G l m M)" and G\: "(\(n,N,j). \ (Suc l) m n M N j) (G l m M)" if "\ m M" for l m M using step [OF that, of "Suc l"] by (force simp: G_def dest: some_eq_imp)+ have G_increasing: "(\(n,N,j). n > m \ N \ M \ n \ M) (G l m M)" if "\ m M" for l m M using G\ [OF that, of l] that by (simp add: \_def split: prod.split_asm) define H where "H \ rec_nat (0,{0<..},0) (\l (m,M,j). G l m M)" have H_simps: "H 0 = (0,{0<..},0)" "\l. H (Suc l) = (case H l of (m,M,j) \ G l m M)" by (simp_all add: H_def) have H\: "(\(n,N,j). \ n N) (H l)" for l by (induction l) (use base G\ in \auto simp: H_simps split: prod.split_asm\) define \ where "\ \ (\l. case H l of (n,M,j) \ n)" have H_inc: "\ l \ l" for l proof (induction l) case (Suc l) then show ?case using H\ [of l] G_increasing [of "\ l"] apply (clarsimp simp: H_simps \_def split: prod.split) by (metis (no_types, lifting) case_prodD leD le_less_trans not_less_eq_eq) qed auto let ?N = "range \" define j where "j \ \l. case H l of (n,M,j) \ j" have H_increasing_Suc: "(case H k of (n, N, j') \ N) \ (case H (Suc k) of (n, N, j') \ insert n N)" for k using H\ [of k] by (force simp: H_simps split: prod.split dest: G_increasing [where l=k]) have H_increasing_superset: "(case H k of (n, N, j') \ N) \ (case H (n+k) of (n, N, j') \ N)" for k n proof (induction n) case (Suc n) then show ?case using H_increasing_Suc [of "n+k"] by (auto split: prod.split_asm) qed auto then have H_increasing_less: "(case H k of (n, N, j') \ N) \ (case H l of (n, N, j') \ insert n N)" if "k k < \ (Suc k)" for k using H\ [of k] unfolding \_def by (auto simp: H_simps split: prod.split dest: G_increasing [where l=k]) then have strict_mono_\: "strict_mono \" by (simp add: strict_mono_Suc_iff) then have enum_N: "enum ?N = \" by (metis enum_works nat_infinite_iff range_strict_mono_ext) have **: "?N \ {n<..} \ N'" if H: "H k = (n, N', j)" for n N' k j proof clarify fix l assume "n < \ l" then have False if "l \ k" using that strict_monoD [OF strict_mono_\, of l k ] H by (force simp: \_def) then have "k < l" using not_less by blast then obtain M j where Mj: "H l = (\ l,M,j)" unfolding \_def by (metis (mono_tags, lifting) case_prod_conv old.prod.exhaust) then show "\ l \ N'" using that H_increasing_less [OF \k] Mj by auto qed show thesis proof show "infinite (?N::nat set)" using H_inc infinite_nat_iff_unbounded_le by auto next fix l U assume "0 < l" and U: "U \ [WW]\<^bsup>2\<^esup>" and interU: "[enum ?N l] < inter_scheme l U" "Form l U" and sub: "list.set (inter_scheme l U) \ ?N" obtain k where k: "l = Suc k" using \0 < l\ gr0_conv_Suc by blast have "g U = v" if "H k = (m, M, j0)" and "G k m M = (n, N', v)" for m M j0 n N' v proof - have n: "\ (Suc k) = n" using that by (simp add: \_def H_simps) have "{..enum (range \) l} \ list.set (inter_scheme l U) = {}" using inter_scheme_strict_sorted \0 < l\ interU singleton_less_list_iff strict_sorted_iff by blast then have "list.set (inter_scheme (Suc k) U) \ N'" using that sub ** [of "Suc k" n N' v] Suc_le_eq not_less_eq_eq by (fastforce simp: k n enum_N H_simps) then show ?thesis using that interU U G\ [of m M k] H\ [of k] by (auto simp: \_def k enum_N H_simps n nsets_def) qed with U show "g U = j l" by (auto simp: k j_def H_simps split: prod.split) qed qed subsection \Larson's Lemma 3.7\ subsubsection \Preliminaries\ text \Analogous to @{thm [source] ordered_nsets_2_eq}, but without type classes\ lemma total_order_nsets_2_eq: assumes tot: "total_on A r" and irr: "irrefl r" shows "nsets A 2 = {{x,y} | x y. x \ A \ y \ A \ (x,y) \ r}" (is "_ = ?rhs") proof show "nsets A 2 \ ?rhs" using tot unfolding numeral_nat total_on_def nsets_def by (fastforce simp: card_Suc_eq Set.doubleton_eq_iff not_less) show "?rhs \ nsets A 2" using irr unfolding numeral_nat by (force simp: nsets_def card_Suc_eq irrefl_def) qed lemma lenlex_nsets_2_eq: "nsets A 2 = {{x,y} | x y. x \ A \ y \ A \ (x,y) \ lenlex less_than}" using total_order_nsets_2_eq by (simp add: total_order_nsets_2_eq irrefl_def) lemma sum_sorted_list_of_set_map: "finite I \ sum_list (map f (list_of I)) = sum f I" proof (induction "card I" arbitrary: I) case (Suc n I) then have [simp]: "I \ {}" by auto have "sum_list (map f (list_of (I - {Min I}))) = sum f (I - {Min I})" using Suc by auto then show ?case using Suc.prems sum.remove [of I "Min I" f] by (simp add: sorted_list_of_set_nonempty Suc) qed auto lemma sorted_list_of_set_UN_eq_concat: assumes I: "strict_mono_sets I f" "finite I" and fin: "\i. finite (f i)" shows "list_of (\i \ I. f i) = concat (map (list_of \ f) (list_of I))" using I proof (induction "card I" arbitrary: I) case (Suc n I) then have "I \ {}" and Iexp: "I = insert (Min I) (I - {Min I})" using Min_in Suc.hyps(2) Suc.prems(2) by fastforce+ have IH: "list_of (\ (f ` (I - {Min I}))) = concat (map (list_of \ f) (list_of (I - {Min I})))" using Suc unfolding strict_mono_sets_def by (metis DiffE Iexp card_Diff_singleton diff_Suc_1 finite_Diff insertI1) have "list_of (\ (f ` I)) = list_of (\ (f ` (insert (Min I) (I - {Min I}))))" using Iexp by auto also have "\ = list_of (f (Min I) \ \ (f ` (I - {Min I})))" by (metis Union_image_insert) also have "\ = list_of (f (Min I)) @ list_of (\ (f ` (I - {Min I})))" proof (rule sorted_list_of_set_Un) show "f (Min I) \ \ (f ` (I - {Min I}))" using Suc.prems \I \ {}\ strict_mono_less_sets_Min by blast show "finite (\ (f ` (I - {Min I})))" by (simp add: \finite I\ fin) qed (use fin in auto) also have "\ = list_of (f (Min I)) @ concat (map (list_of \ f) (list_of (I - {Min I})))" using IH by metis also have "\ = concat (map (list_of \ f) (list_of I))" by (simp add: Suc.prems(2) \I \ {}\ sorted_list_of_set_nonempty) finally show ?case . qed auto subsubsection \Lemma 3.7 of Jean A. Larson, ibid.\ proposition lemma_3_7: assumes "infinite N" "l > 0" obtains M where "M \ [WW]\<^bsup>m\<^esup>" "\U. U \ [M]\<^bsup>2\<^esup> \ Form l U \ List.set (inter_scheme l U) \ N" proof (cases "m < 2") case True obtain w where w: "w \ WW" using WW_def strict_sorted_into_WW by auto define M where "M \ if m=0 then {} else {w}" have M: "M \ [WW]\<^bsup>m\<^esup>" using True by (auto simp: M_def nsets_def w) have [simp]: "[M]\<^bsup>2\<^esup> = {}" using True by (auto simp: M_def nsets_def w dest: subset_singletonD) show ?thesis using M that by fastforce next case False then have "m \ 2" by auto have nonz: "(enum N \ Suc) i > 0" for i using assms(1) le_enumerate less_le_trans by fastforce note infinite_nxt_N = infinite_nxtN [OF \infinite N\, iff] note \infinite N\ [ iff] have [simp]: "{n<.. \k D. enum (nxt N (enum (nxt N (Max D)) (Inf D - 1))) ` {.. \k. rec_nat ((enum N \ Suc) ` {..r. DF_Suc k)" have DF_simps: "DF k 0 = (enum N \ Suc) ` {.. Suc) {..infinite N\ DF_simps DF_Suc_def card_image infinite_nxtN strict_mono_enum strict_mono_imp_inj_on) qed have DF_ne: "DF k i \ {}" for i k by (metis card_DF card_lessThan lessThan_empty_iff nat.simps(3)) have finite_DF: "finite (DF k i)" for i k by (induction i) (auto simp: DF_simps DF_Suc_def) have DF_Suc: "DF k i \ DF k (Suc i)" for i k unfolding less_sets_def by (force simp: finite_DF DF_simps DF_Suc_def intro!: greaterThan_less_enum nxt_subset_greaterThan atLeast_le_enum nxt_subset_atLeast infinite_nxtN [OF \infinite N\]) have DF_DF: "DF k i \ DF k j" if "i N" for i k proof (induction i) case 0 then show ?case using \infinite N\ range_enum by (auto simp: DF_simps) next case (Suc i) then show ?case unfolding DF_simps DF_Suc_def image_subset_iff by (metis IntE \infinite N\ enumerate_in_set infinite_nxtN nxt_def) qed have sm_enum_DF: "strict_mono_on (enum (DF k i)) {..k}" for k i by (metis card_DF enum_works_finite finite_DF lessThan_Suc_atMost) define AF where "AF \ \k i. enum (nxt N (Max (DF k i))) ` {.. {}" for i k by (auto simp: AF_def lessThan_empty_iff DF_gt0) have finite_AF [simp]: "finite (AF k i)" for i k by (simp add: AF_def) have card_AF: "card (AF k i) = \ (DF k i)" for k i by (simp add: AF_def card_image inj_enum_nxt) have DF_AF: "DF k i \ AF k i" for i k unfolding less_sets_def AF_def by (simp add: finite_DF greaterThan_less_enum nxt_subset_greaterThan) have E: "\x \ y; infinite M\ \ enum M x < enum (nxt N (enum M y)) z" for x y z M by (metis infinite_nxt_N dual_order.eq_iff enumerate_mono greaterThan_less_enum nat_less_le nxt_subset_greaterThan) have AF_DF_Suc: "AF k i \ DF k (Suc i)" for i k by (auto simp: DF_simps DF_Suc_def less_sets_def AF_def E) have AF_DF: "AF k p \ DF k q" if "p AF k (Suc i)" for i k using AF_DF_Suc DF_AF DF_ne less_sets_trans by blast then have sm_AF: "strict_mono_sets UNIV (AF k)" for k by (simp add: AF_ne less_sets_imp_strict_mono_sets) define del where "del \ \k i j. enum (DF k i) j - enum (DF k i) (j - 1)" define QF where "QF k \ wfrec pair_less (\f (j,i). if j=0 then AF k i else let r = (if i=0 then f (j-1,m-1) else f (j,i-1)) in enum (nxt N (Suc (Max r))) ` {..< del k (if j=k then m - Suc i else i) j})" for k note cut_apply [simp] have finite_QF [simp]: "finite (QF k p)" for p k using wf_pair_less proof (induction p rule: wf_induct_rule) case (less p) then show ?case by (simp add: def_wfrec [OF QF_def, of k p] split: prod.split) qed have del_gt_0: "\j < Suc k; 0 < j\ \ 0 < del k i j" for i j k by (simp add: card_DF del_def finite_DF) have QF_ne [simp]: "QF k (j,i) \ {}" if j: "j < Suc k" for j i k using wf_pair_less j proof (induction "(j,i)" rule: wf_induct_rule) case less then show ?case by (auto simp: def_wfrec [OF QF_def, of k "(j,i)"] AF_ne lessThan_empty_iff del_gt_0) qed have QF_0 [simp]: "QF k (0,i) = AF k i" for i k by (simp add: def_wfrec [OF QF_def]) have QF_Suc: "QF k (Suc j,0) = enum (nxt N (Suc (Max (QF k (j, m - 1))))) ` {..< del k (if Suc j = k then m - 1 else 0) (Suc j)}" for j k apply (simp add: def_wfrec [OF QF_def, of k "(Suc j,0)"] One_nat_def) apply (simp add: pair_less_def cut_def) done have QF_Suc_Suc: "QF k (Suc j, Suc i) = enum (nxt N (Suc (Max (QF k (Suc j, i))))) ` {..< del k (if Suc j = k then m - Suc(Suc i) else Suc i) (Suc j)}" for i j k by (simp add: def_wfrec [OF QF_def, of k "(Suc j,Suc i)"]) have less_QF1: "QF k (j, m - 1) \ QF k (Suc j,0)" for j k by (auto simp: def_wfrec [OF QF_def, of k "(Suc j,0)"] pair_lessI1 enum_nxt_ge intro!: less_sets_weaken2 [OF less_sets_Suc_Max]) have less_QF2: "QF k (j,i) \ QF k (j, Suc i)" for j i k by (auto simp: def_wfrec [OF QF_def, of k "(j, Suc i)"] pair_lessI2 enum_nxt_ge intro: less_sets_weaken2 [OF less_sets_Suc_Max] strict_mono_setsD [OF sm_AF]) have less_QF_same: "QF k (j,i') \ QF k (j,i)" if "i' < i" "j \ k" for i' i j k proof (rule strict_mono_setsD [OF less_sets_imp_strict_mono_sets]) show "QF k (j, i) \ QF k (j, Suc i)" for i by (simp add: less_QF2) show "QF k (j, i) \ {}" if "0 < i" for i using that by (simp add: \j \ k\ le_imp_less_Suc) qed (use that in auto) have less_QF_step: "QF k (j-1, i') \ QF k (j,i)" if "0 < j" "j \ k" "i' < m" for j i' i k proof - have less_QF1': "QF k (j - 1, m-1) \ QF k (j,0)" if "j > 0" for j by (metis less_QF1 that Suc_pred One_nat_def) have "QF k (j-1, i') \ QF k (j,0)" proof (cases "i' = m - 1") case True then show ?thesis using less_QF1' \0 < j\ by blast next case False show ?thesis using False that less_sets_trans [OF less_QF_same less_QF1' QF_ne] by auto qed then show ?thesis by (metis QF_ne less_QF_same less_Suc_eq_le less_sets_trans \j \ k\ zero_less_iff_neq_zero) qed have less_QF: "QF k (j',i') \ QF k (j,i)" if j: "j' < j" "j \ k" and i: "i' < m" "i < m" for j' j i' i k using j proof (induction "j-j'" arbitrary: j) case (Suc d) show ?case proof (cases "j' < j - 1") case True then have "QF k (j', i') \ QF k (j - 1, i)" using Suc.hyps Suc.prems(2) by force then show ?thesis by (rule less_sets_trans [OF _ less_QF_step QF_ne]) (use Suc i in auto) next case False then have "j' = j - 1" using \j' < j\ by linarith then show ?thesis using Suc.hyps \j \ k\ less_QF_step i by auto qed qed auto have sm_QF: "strict_mono_sets ({..k} \ {.. {..k} \ {.. {..k} \ {..: "p = (j',i')" "q = (j,i)" "i' < m" "i < m" "j' \ k" "j \ k" using surj_pair [of p] surj_pair [of q] by blast with \p < q\ have "j' < j \ j' = j \ i' < i" by auto then show "QF k p \ QF k q" using \
less_QF less_QF_same by presburger qed then have sm_QF1: "strict_mono_sets {..j. QF k (j,i))" if "i ka" "ka \ k" for ka k i proof - have "{.. {..k}" by (metis lessThan_Suc_atMost lessThan_subset_iff \Suc k \ ka\) then show ?thesis by (simp add: less_QF strict_mono_sets_def subset_iff that) qed have disjoint_QF: "i'=i \ j'=j" if "\ disjnt (QF k (j', i')) (QF k (j,i))" "j' \ k" "j \ k" "i' < m" "i < m" for i' i j' j k using that strict_mono_sets_imp_disjoint [OF sm_QF] by (force simp: pairwise_def) have card_QF: "card (QF k (j,i)) = (if j=0 then \ (DF k i) else del k (if j = k then m - Suc i else i) j)" for i k j proof (cases j) case 0 then show ?thesis by (simp add: AF_def card_image inj_enum_nxt) next case (Suc j') show ?thesis by (cases i; simp add: Suc One_nat_def QF_Suc QF_Suc_Suc card_image inj_enum_nxt) qed have AF_non_Nil: "list_of (AF k i) \ []" for k i by (simp add: AF_ne) have QF_non_Nil: "list_of (QF k (j,i)) \ []" if "j < Suc k" for i j k by (simp add: that) have AF_subset_N: "AF k i \ N" for i k unfolding AF_def image_subset_iff using nxt_subset enumerate_in_set infinite_nxtN \infinite N\ by blast have QF_subset_N: "QF k (j,i) \ N" for i j k proof (induction j) case (Suc j) show ?case by (cases i) (use nxt_subset enumerate_in_set in \(force simp: QF_Suc QF_Suc_Suc)+\) qed (use AF_subset_N in auto) obtain ka k where "k>0" and kka: "k \ ka" "ka \ Suc k" "l = ((ka+k) - 1)" by (metis One_nat_def assms(2) diff_add_inverse form_cases le0 le_refl) then have "ka > 0" using dual_order.strict_trans1 by blast have ka_k_or_Suc: "ka = k \ ka = Suc k" using kka by linarith have lessThan_k: "{..0" for k::nat using that by auto then have sorted_list_of_set_k: "list_of {..0" for k::nat - using sorted_list_of_set_insert_cons [of concl: 0 "{0<.. \j i. if j = k then QF k (j, m - Suc i) else QF k (j,i)" have RF_subset_N: "RF j i \ N" if "i0 < k\ by auto have disjoint_RF: "i'=i \ j'=j" if "\ disjnt (RF j' i') (RF j i)" "j' \ k" "j \ k" "i' < m" "i < m" for i' i j' j using disjoint_QF that by (auto simp: RF_def split: if_split_asm dest: disjoint_QF) have sum_card_RF [simp]: "(\j\n. card (RF j i)) = enum (DF k i) n" if "n \ k" "i < m" for i n using that proof (induction n) case 0 then show ?case using DF_ne [of k i] finite_DF [of k i] \k>0\ by (simp add: RF_def AF_def card_image inj_enum_nxt enum_0_eq_Inf_finite) next case (Suc n) then have "enum (DF k i) 0 \ enum (DF k i) n \ enum (DF k i) n \ enum (DF k i) (Suc n)" using sm_enum_DF [of k i] by (metis Suc_leD card_DF dual_order.eq_iff finite_DF finite_enumerate_mono le_imp_less_Suc less_imp_le_nat not_gr_zero) with Suc show ?case by (auto simp: RF_def card_QF del_def) qed have DF_in_N: "enum (DF k i) j \ N" if "j \ k" for i j by (metis DF_N card_DF finite_DF finite_enumerate_in_set le_imp_less_Suc subsetD that) have Inf_DF_N: "\(DF k p) \ N" for k p using DF_N DF_ne Inf_nat_def1 by blast have RF_in_N: "(\j\n. card (RF j i)) \ N" if "n \ k" "i < m" for i n by (auto simp: DF_in_N that) have "ka - 1 \ k" using kka(2) by linarith then have sum_card_RF' [simp]: "(\j0 < ka\ lessThan_Suc_atMost that) have enum_DF_le_iff [simp]: "enum (DF k i) j \ enum (DF k i') j \ i \ i'" (is "?lhs = _") if "j \ k" for i' i j k proof show "i \ i'" if ?lhs proof - have "enum (DF k i) j \ DF k i" by (simp add: card_DF finite_enumerate_in_set finite_DF le_imp_less_Suc \j \ k\) moreover have "enum (DF k i') j \ DF k i'" by (simp add: \j \ k\ card_DF finite_enumerate_in_set finite_DF le_imp_less_Suc that) ultimately have "enum (DF k i') j < enum (DF k i) j" if "i' < i" using sm_DF [of k] by (meson UNIV_I less_sets_def strict_mono_setsD that) then show ?thesis using not_less that by blast qed show ?lhs if "i \ i'" using sm_DF [of k] that \j \ k\ card_DF finite_enumerate_in_set finite_DF le_eq_less_or_eq by (force simp: strict_mono_sets_def less_sets_def finite_enumerate_in_set) qed then have enum_DF_eq_iff[simp]: "enum (DF k i) j = enum (DF k i') j \ i = i'" if "j \ k" for i' i j k by (metis le_antisym order_refl that) have enum_DF_less_iff [simp]: "enum (DF k i) j < enum (DF k i') j \ i < i'" if "j \ k" for i' i j k by (meson enum_DF_le_iff not_less that) have card_AF_sum: "card (AF k i) + (\j\{0<..k > 0\ \k \ ka\ \ka \ Suc k\ by (simp add: lessThan_k RF_0 flip: sum_card_RF') have sorted_list_of_set_iff [simp]: "list_of {0<.. k = 1" if "k>0" for k::nat proof - have "list_of {0<.. {0<.. \ k = 1" using \k > 0\ atLeastSucLessThan_greaterThanLessThan by fastforce finally show ?thesis . qed show thesis \\proof of main result\ proof have inj: "inj_on (\i. list_of (\jjjjj RF 0 x" using AF_ne QF_0 \0 < k\ Inf_nat_def1 \k \ ka\ by (force simp: RF_def) with eq \ka > 0\ obtain j' where "j' < ka" "n \ RF j' y" by blast then show ?thesis using disjoint_QF [of k 0 x j'] n \x < m\ \y < m\ \ka \ Suc k\ \0 < k\ by (force simp: RF_def disjnt_iff simp del: QF_0 split: if_split_asm) qed qed define M where "M \ (\i. list_of (\jk \ ka\ card_image inj) moreover have "M \ WW" by (force simp: M_def WW_def) ultimately show "M \ [WW]\<^bsup>m\<^esup>" by (simp add: nsets_def) have sm_RF: "strict_mono_sets {..j. RF j i)" if "i []" if "j < Suc k" for i j using that by (simp add: RF_def) have less_RF_same: "RF j i' \ RF j i" if "i' < i" "j < k" for i' i j using that by (simp add: less_QF_same RF_def) have less_RF_same_k: "RF k i' \ RF k i" \\reversed version for @{term k}\ if "i < i'" "i' < m" for i' i using that by (simp add: less_QF_same RF_def) show "Form l U \ list.set (inter_scheme l U) \ N" if "U \ [M]\<^bsup>2\<^esup>" for U proof - from that obtain x y where "U = {x,y}" "x \ M" "y \ M" and xy: "(x,y) \ lenlex less_than" by (auto simp: lenlex_nsets_2_eq) let ?R = "\p. list_of \ (\j. RF j p)" obtain p q where x: "x = list_of (\jjx \ M\ \y \ M\ by (auto simp: M_def) then have pq: "pk \ ka\ \ka \ Suc k\ lexl_not_refl [OF irrefl_less_than] by (auto simp: lenlex_def sm_RF sorted_list_of_set_UN_lessThan length_concat sum_sorted_list_of_set_map) moreover have xc: "x = concat (map (?R p) (list_of {..k \ ka\ \ka \ Suc k\ \p < m\ sm_RF) have yc: "y = concat (map (?R q) (list_of {..k \ ka\ \ka \ Suc k\ \q < m\ sm_RF) have enum_DF_AF: "enum (DF k p) (ka - 1) < hd (list_of (AF k p))" for p proof (rule less_setsD [OF DF_AF]) show "enum (DF k p) (ka - 1) \ DF k p" using \ka \ Suc k\ card_DF finite_DF by (auto simp: finite_enumerate_in_set) show "hd (list_of (AF k p)) \ AF k p" using AF_non_Nil finite_AF hd_in_set set_sorted_list_of_set by blast qed have less_RF_RF: "RF n p \ RF n q" if "n < k" for n using that \p by (simp add: less_RF_same) have less_RF_Suc: "RF n q \ RF (Suc n) q" if "n < k" for n using \q < m\ that by (auto simp: RF_def less_QF) have less_RF_k: "RF k q \ RF k p" using \q < m\ less_RF_same_k \p by blast have less_RF_k_ka: "RF (k-1) p \ RF (ka - 1) q" using ka_k_or_Suc less_RF_RF by (metis One_nat_def RF_def \0 < k\ \ka - 1 \ k\ \p < m\ diff_Suc_1 diff_Suc_less less_QF_step) have Inf_DF_eq_enum: "\ (DF k i) = enum (DF k i) 0" for k i by (simp add: Inf_nat_def enumerate_0) have Inf_DF_less: "\ (DF k i') < \ (DF k i)" if "i'x. x \ AF k i \ \ (DF k i') < x" if "i'\i" for i' i k using less_setsD [OF DF_AF] DF_ne that by (metis Inf_DF_less Inf_nat_def1 dual_order.order_iff_strict dual_order.strict_trans) show ?thesis \\The general case requires @{term\k>1\}, necessitating a painful special case\ proof (cases "k=1") case True with kka consider "ka=1" | "ka=2" by linarith then show ?thesis proof cases case 1 define zs where "zs = card (AF 1 p) # list_of (AF 1 p) @ card (AF 1 q) # list_of (AF 1 q)" have zs: "Form_Body ka k x y zs" proof (intro that exI conjI Form_Body.intros) show "x = concat ([list_of (AF k p)])" "y = concat ([list_of (AF k q)])" by (simp_all add: x y 1 lessThan_Suc RF_0) have "AF k p \ insert (\ (DF k q)) (AF k q)" by (metis AF_DF DF_ne Inf_nat_def1 RF_0 \0 < k\ insert_iff less_RF_RF less_sets_def pq(1)) then have "strict_sorted (list_of (AF k p) @ \ (DF k q) # list_of (AF k q))" by (auto simp: strict_sorted_append_iff intro: less_sets_imp_list_less AF_Inf_DF_less) moreover have "\x. x \ AF k q \ \ (DF k p) < x" by (meson AF_Inf_DF_less less_imp_le_nat \p < q\) moreover have "\x. x \ AF 1 p \ \ (DF 1 p) < x" by (meson DF_AF DF_ne Inf_nat_def1 less_setsD) ultimately show "strict_sorted zs" using \p < q\ True Inf_DF_less DF_AF DF_ne by (auto simp: zs_def less_sets_def card_AF AF_Inf_DF_less) qed (auto simp: \k=1\ \ka=1\ zs_def AF_ne \length x < length y\) have zs_N: "list.set zs \ N" using AF_subset_N by (auto simp: zs_def card_AF Inf_DF_N \k=1\) show ?thesis proof have "l = 1" using kka \k=1\ \ka=1\ by auto have "Form (2*1-1) {x,y}" using "1" Form.intros(2) True zs by fastforce then show "Form l U" by (simp add: \U = {x,y}\ \l = 1\ One_nat_def) show "list.set (inter_scheme l U) \ N" using kka zs zs_N \k=1\ Form_Body_imp_inter_scheme by (fastforce simp: \U = {x,y}\) qed next case 2 note True [simp] note 2 [simp] have [simp]: "{0<..<2} = {1::nat}" by auto have enum_DF1_eq: "enum (DF 1 i) 1 = card (AF 1 i) + card (RF 1 i)" if "i < m" for i using card_AF_sum that by (simp add: One_nat_def) have card_RF: "card (RF 1 i) = enum (DF 1 i) 1 - enum (DF 1 i) 0" if "i < m" for i using that by (auto simp: RF_def card_QF del_def) have list_of_AF_RF: "list_of (AF 1 q \ RF 1 q) = list_of (AF 1 q) @ list_of (RF 1 q)" by (metis One_nat_def RF_0 True \0 < k\ finite_RF less_RF_Suc sorted_list_of_set_Un) define zs where "zs = card (AF 1 p) # (card (AF 1 p) + card (RF 1 p)) # list_of (AF 1 p) @ (card (AF 1 q) + card (RF 1 q)) # list_of (AF 1 q) @ list_of (RF 1 q) @ list_of (RF 1 p)" have zs: "Form_Body ka k x y zs" proof (intro that exI conjI Form_Body.intros) have "x = list_of (RF 0 p \ RF 1 p)" by (simp add: x eval_nat_numeral lessThan_Suc RF_0 Un_commute One_nat_def) also have "\ = list_of (RF 0 p) @ list_of (RF 1 p)" using RF_def True \p < m\ less_QF_step by (metis QF_0 RF_0 diff_self_eq_0 finite_RF le_refl sorted_list_of_set_Un zero_less_one) finally show "x = concat ([list_of (AF 1 p),list_of (RF 1 p)])" by (simp add: RF_0) show "y = concat [list_of (RF 1 q \ AF 1 q)]" by (simp add: y eval_nat_numeral lessThan_Suc RF_0 One_nat_def) show zs: "zs = concat [[card (AF 1 p), card (AF 1 p) + card (RF 1 p)], list_of (AF 1 p), [card (AF 1 q) + card (RF 1 q)], list_of (RF 1 q \ AF 1 q)] @ interact [list_of (RF 1 p)] []" using list_of_AF_RF by (simp add: zs_def Un_commute) show "strict_sorted zs" proof (simp add: \p \q \p zs_def strict_sorted_append_iff, intro conjI strip) show "0 < card (RF 1 p)" using \p by (simp add: card_RF card_DF finite_DF) show "card (AF 1 p) < card (AF 1 q) + card (RF 1 q)" using \p \q by (simp add: Inf_DF_less card_AF trans_less_add1) show "card (AF 1 p) < x" if "x \ AF 1 p \ (AF 1 q \ (RF 1 q \ RF 1 p))" for x using that apply (simp add: card_AF) by (metis AF_ne DF_AF DF_ne less_RF_RF less_RF_Suc less_RF_k Inf_nat_def1 One_nat_def RF_0 RF_non_Nil True finite_RF lessI less_setsD less_sets_trans sorted_list_of_set_eq_Nil_iff) show "card (AF 1 p) + card (RF 1 p) < card (AF 1 q) + card (RF 1 q)" using \p < q\ \p < m\ \q < m\ by (metis enum_DF1_eq enum_DF_less_iff le_refl) show "card (AF 1 p) + card (RF 1 p) < x" if "x \ AF 1 p \ (AF 1 q \ (RF 1 q \ RF 1 p))" for x using that \p < m\ apply (simp flip: enum_DF1_eq) by (metis AF_ne DF_AF less_RF_RF less_RF_Suc less_RF_k One_nat_def RF_0 RF_non_Nil Suc_mono True \0 < k\ card_DF finite_enumerate_in_set finite_DF less_setsD less_sets_trans sorted_list_of_set_empty) have "list_of (AF 1 p) < list_of {enum (DF 1 q) 1}" proof (rule less_sets_imp_sorted_list_of_set) show "AF 1 p \ {enum (DF 1 q) 1}" by (metis AF_DF card_DF empty_subsetI finite_DF finite_enumerate_in_set insert_subset less_Suc_eq less_sets_weaken2 pq(1)) qed auto then show "list_of (AF 1 p) < (card (AF 1 q) + card (RF 1 q)) # list_of (AF 1 q) @ list_of (RF 1 q) @ list_of (RF 1 p)" using \q < m\ by (simp add: less_list_def enum_DF1_eq) show "card (AF 1 q) + card (RF 1 q) < x" if "x \ AF 1 q \ (RF 1 q \ RF 1 p)" for x using that \q < m\ apply (simp flip: enum_DF1_eq) by (metis AF_ne DF_AF less_RF_Suc less_RF_k One_nat_def RF_0 RF_non_Nil True card_DF finite_enumerate_in_set finite_DF finite_RF lessI less_setsD less_sets_trans sorted_list_of_set_eq_Nil_iff) have "list_of (AF 1 q) < list_of (RF 1 q)" proof (rule less_sets_imp_sorted_list_of_set) show "AF 1 q \ RF 1 q" by (metis less_RF_Suc One_nat_def RF_0 True \0 < k\) qed auto then show "list_of (AF 1 q) < list_of (RF 1 q) @ list_of (RF 1 p)" using RF_non_Nil by (auto simp: less_list_def) show "list_of (RF 1 q) < list_of (RF 1 p)" proof (rule less_sets_imp_sorted_list_of_set) show "RF 1 q \ RF 1 p" by (metis less_RF_k True) qed auto qed show "[list_of (AF 1 p), list_of (RF 1 p)] \ lists (- {[]})" using RF_non_Nil \0 < k\ by (auto simp: zs_def AF_ne) show "[card (AF 1 q) + card (RF 1 q)] = acc_lengths 0 [list_of (RF 1 q \ AF 1 q)]" using list_of_AF_RF by (auto simp: zs_def AF_ne sup_commute) qed (auto simp: zs_def AF_ne \length x < length y\) have zs_N: "list.set zs \ N" using \p < m\ \q < m\ DF_in_N enum_DF1_eq [symmetric] by (auto simp: zs_def card_AF AF_subset_N RF_subset_N Inf_DF_N) show ?thesis proof have "Form (2*1) {x,y}" by (metis "2" Form.simps Suc_1 True zero_less_one zs) with kka show "Form l U" by (simp add: \U = {x,y}\) show "list.set (inter_scheme l U) \ N" using kka zs zs_N \k=1\ Form_Body_imp_inter_scheme by (fastforce simp: \U = {x, y}\) qed qed next case False then have "k \ 2" "ka \ 2" using kka \k>0\ by auto then have k_minus_1 [simp]: "Suc (k - Suc (Suc 0)) = k - Suc 0" by auto have [simp]: "Suc (k - 2) = k-1" using \k \ 2\ by linarith define PP where "PP \ map (?R p) (list_of {0<.. map (?R q) (list_of {0<.. RF (ka-1) q)])" let ?INT = "interact PP QQ" \\No separate sets A and B as in the text, but instead we treat both cases as once\ have [simp]: "length PP = ka - 1" by (simp add: PP_def) have [simp]: "length QQ = k-1" using \k \ 2\ by (simp add: QQ_def) have PP_n: "PP ! n = list_of (RF (Suc n) p)" if "n < ka-1" for n using that kka by (auto simp: PP_def nth_sorted_list_of_set_greaterThanLessThan) have QQ_n: "QQ ! n = (if n < k-2 then list_of (RF (Suc n) q) else list_of (RF (k-1) q \ RF (ka - 1) q))" if "n < k-1" for n using that kka by (auto simp: QQ_def nth_append nth_sorted_list_of_set_greaterThanLessThan) have QQ_n_same: "QQ ! n = list_of (RF (Suc n) q)" if "n < k-1" "k=ka" for n using that kka Suc_diff_Suc by (fastforce simp: One_nat_def QQ_def nth_append nth_sorted_list_of_set_greaterThanLessThan) have split_nat_interval: "{0<.. 2" for n::nat using that by auto have split_list_interval: "list_of{0<.. 2" for n::nat proof (intro sorted_list_of_set_unique [THEN iffD1] conjI) have "list_of {0<..n \ 2\ in auto) have list_of_RF_Un: "list_of (RF (k-1) q \ RF k q) = list_of (RF (k-1) q) @ list_of (RF k q)" by (metis Suc_diff_1 \0 < k\ finite_RF lessI less_RF_Suc sorted_list_of_set_Un) have card_AF_sum_QQ: "card (AF k q) + sum_list (map length QQ) = (\j RF k q = {}" using less_RF_Suc [of "k-1"] \k > 0\ by (auto simp: less_sets_def) then have "card (RF (k-1) q \ RF k q) = card (RF (k-1) q) + card (RF k q)" by (simp add: card_Un_disjoint) then show ?thesis using \k\2\ \q < m\ apply (simp add: QQ_def True flip: RF_0) apply (simp add: lessThan_k split_nat_interval sum_sorted_list_of_set_map) done next case False with kka have "ka=k" by linarith with \k\2\ show ?thesis by (simp add: QQ_def lessThan_k split_nat_interval sum_sorted_list_of_set_map flip: RF_0) qed define LENS where "LENS \ \i. acc_lengths 0 (list_of (AF k i) # map (?R i) (list_of {0<.. N" if "i < m" for i proof - have eq: "(list_of (AF k i) # map (?R i) (list_of {0<..0 < ka\ sorted_list_of_set_k by auto let ?f = "rec_nat [card (AF k i)] (\n r. r @ [(\j\Suc n. card (RF j i))])" have f: "acc_lengths 0 (map (?R i) (list_of {..v})) = ?f v" for v by (induction v) (auto simp: RF_0 acc_lengths_append sum_sorted_list_of_set_map) have 3: "list.set (?f v) \ N" if "v \ k" for v using that proof (induction v) case 0 have "card (AF k i) \ N" by (metis DF_N DF_ne Inf_nat_def1 card_AF subsetD) with 0 show ?case by simp next case (Suc v) then have "enum (DF k i) (Suc v) \ N" by (metis DF_N card_DF finite_enumerate_in_set finite_DF in_mono le_imp_less_Suc) with Suc \i < m\ show ?case by (simp del: sum.atMost_Suc) qed show ?thesis unfolding LENS_def by (metis 3 Suc_pred' \0 < ka\ \ka - 1 \ k\ eq f lessThan_Suc_atMost) qed define LENS_QQ where "LENS_QQ \ acc_lengths 0 (list_of (AF k q) # QQ)" have LENS_QQ_subset: "list.set LENS_QQ \ list.set (LENS q)" proof (cases "ka = Suc k") case True with \k \ 2\ show ?thesis unfolding QQ_def LENS_QQ_def LENS_def by (auto simp: list_of_RF_Un split_list_interval acc_lengths_append) next case False then have "ka=k" using kka by linarith with \k \ 2\ show ?thesis by (simp add: QQ_def LENS_QQ_def LENS_def split_list_interval) qed have ss_INT: "strict_sorted ?INT" proof (rule strict_sorted_interact_I) fix n assume "n < length QQ" then have n: "n < k-1" by simp have "n = k - 2" if "\ n < k - 2" using n that by linarith moreover have "list_of (RF (Suc (k - 2)) p) < list_of (RF (k-1) q \ RF (ka - 1) q)" by (auto simp: less_sets_imp_sorted_list_of_set less_sets_Un2 less_RF_RF less_RF_k_ka \0 < k\) ultimately show "PP ! n < QQ ! n" using \k \ ka\ n by (auto simp: PP_n QQ_n less_sets_imp_sorted_list_of_set less_RF_RF) next fix n have V: "\Suc n < ka - 1\ \ list_of (RF (Suc n) q) < list_of (RF (Suc (Suc n)) p)" for n by (smt RF_def Suc_leI \ka - 1 \ k\ \q < m\ diff_Suc_1 finite_RF less_QF_step less_le_trans less_sets_imp_sorted_list_of_set nat_neq_iff zero_less_Suc) have "RF (k - 1) q \ RF k p" by (metis One_nat_def RF_non_Nil Suc_pred \0 < k\ finite_RF lessI less_RF_Suc less_RF_k less_sets_trans sorted_list_of_set_eq_Nil_iff) with kka have "RF (k-1) q \ RF (ka - 1) q \ RF k p" by (metis less_RF_k One_nat_def less_sets_Un1 antisym_conv2 diff_Suc_1 le_less_Suc_eq) then have VI: "list_of (RF (k-1) q \ RF (ka - 1) q) < list_of (RF k p)" by (rule less_sets_imp_sorted_list_of_set) auto assume "Suc n < length PP" with \ka \ Suc k\ VI show "QQ ! n < PP ! Suc n" apply (clarsimp simp: PP_n QQ_n V) by (metis One_nat_def Suc_1 Suc_lessI add.right_neutral add_Suc_right diff_Suc_Suc ka_k_or_Suc less_diff_conv) next show "PP \ lists (- {[]})" using RF_non_Nil kka by (clarsimp simp: PP_def) (metis RF_non_Nil less_le_trans) show "QQ \ lists (- {[]})" using RF_non_Nil kka by (clarsimp simp: QQ_def) (metis RF_non_Nil Suc_pred \0 < k\ less_SucI One_nat_def) qed (use kka PP_def QQ_def in auto) then have ss_QQ: "strict_sorted (concat QQ)" using strict_sorted_interact_imp_concat by blast obtain zs where zs: "Form_Body ka k x y zs" and zs_N: "list.set zs \ N" proof (intro that exI conjI Form_Body.intros [OF \length x < length y\]) show "x = concat (list_of (AF k p) # PP)" using \ka > 0\ by (simp add: PP_def RF_0 xc sorted_list_of_set_k) let ?YR = "(map (list_of \ (\j. RF j q)) (list_of {0<..ka - 1 \ k\ less_le_trans) then show "list_of (RF (list_of {0<.. lists (- {[]})" using RF_non_Nil \ka \ Suc k\ by (auto simp: mem_lists_non_Nil) qed auto show "list.set (concat ?YR) = list.set (concat QQ)" using ka_k_or_Suc proof assume "ka = k" then show "list.set (concat (map (list_of \ (\j. RF j q)) (list_of {0<..k\2\ by simp (simp add: split_nat_interval QQ_def) next assume "ka = Suc k" then show "list.set (concat (map (list_of \ (\j. RF j q)) (list_of {0<..k\2\ by simp (auto simp: QQ_def split_nat_interval) qed qed then show "y = concat (list_of (AF k q) # QQ)" using \ka > 0\ by (simp add: RF_0 yc sorted_list_of_set_k) show "list_of (AF k p) # PP \ lists (- {[]})" "list_of (AF k q) # QQ \ lists (- {[]})" using RF_non_Nil kka by (auto simp: AF_ne PP_def QQ_def eq_commute [of "[]"]) show "list.set ((LENS p @ list_of (AF k p) @ LENS_QQ @ list_of (AF k q) @ ?INT)) \ N" using AF_subset_N RF_subset_N LENS_subset_N \p < m\ \q < m\ LENS_QQ_subset by (auto simp: subset_iff PP_def QQ_def) show "length (list_of (AF k p) # PP) = ka" "length (list_of (AF k q) # QQ) = k" using \0 < ka\ \0 < k\ by auto show "LENS p = acc_lengths 0 (list_of (AF k p) # PP)" by (auto simp: LENS_def PP_def) show "strict_sorted (LENS p @ list_of (AF k p) @ LENS_QQ @ list_of (AF k q) @ ?INT)" unfolding strict_sorted_append_iff proof (intro conjI ss_INT) show "LENS p < list_of (AF k p) @ LENS_QQ @ list_of (AF k q) @ ?INT" using AF_non_Nil [of k p] \k \ ka\ \ka \ Suc k\ \p < m\ card_AF_sum enum_DF_AF by (simp add: enum_DF_AF less_list_def card_AF_sum LENS_def sum_sorted_list_of_set_map del: acc_lengths.simps) show "strict_sorted (LENS p)" unfolding LENS_def by (rule strict_sorted_acc_lengths) (use RF_non_Nil AF_non_Nil kka in \auto simp: in_lists_conv_set\) show "strict_sorted LENS_QQ" unfolding LENS_QQ_def QQ_def by (rule strict_sorted_acc_lengths) (use RF_non_Nil AF_non_Nil kka in \auto simp: in_lists_conv_set\) have last_AF_DF: "last (list_of (AF k p)) < \ (DF k q)" using AF_DF [OF \p < q\, of k] AF_non_Nil [of k p] DF_ne [of k q] by (metis Inf_nat_def1 finite_AF last_in_set less_sets_def set_sorted_list_of_set) then show "list_of (AF k p) < LENS_QQ @ list_of (AF k q) @ ?INT" by (simp add: less_list_def card_AF LENS_QQ_def) show "LENS_QQ < list_of (AF k q) @ ?INT" using AF_non_Nil [of k q] \q < m\ card_AF_sum enum_DF_AF card_AF_sum_QQ by (auto simp: less_list_def AF_ne hd_append card_AF_sum LENS_QQ_def) show "list_of (AF k q) < ?INT" proof - have "AF k q \ RF 1 p" using \0 < k\ \p < m\ \q < m\ by (simp add: RF_def less_QF flip: QF_0) then have "last (list_of (AF k q)) < hd (list_of (RF 1 p))" proof (rule less_setsD) show "last (list_of (AF k q)) \ AF k q" using AF_non_Nil finite_AF last_in_set set_sorted_list_of_set by blast show "hd (list_of (RF 1 p)) \ RF 1 p" by (metis One_nat_def RF_non_Nil \0 < k\ finite_RF hd_in_set not_less_eq set_sorted_list_of_set) qed with \k > 0\ \ka \ 2\ RF_non_Nil show ?thesis by (simp add: One_nat_def hd_interact less_list_def sorted_list_of_set_greaterThanLessThan PP_def QQ_def) qed qed auto qed (auto simp: LENS_QQ_def) show ?thesis proof (cases "ka = k") case True then have "l = 2*k-1" by (simp add: kka(3) mult_2) then show ?thesis by (metis One_nat_def Form.intros(2) Form_Body_imp_inter_scheme True \0 < k\ \U = {x, y}\ kka zs zs_N) next case False then have "l = 2*k" using kka by linarith then show ?thesis by (metis One_nat_def False Form.intros(3) Form_Body_imp_inter_scheme \0 < k\ \U = {x, y}\ antisym kka le_SucE zs zs_N) qed qed qed qed qed subsection \Larson's Lemma 3.8\ subsubsection \Primitives needed for the inductive construction of @{term b}\ definition IJ where "IJ \ \k. Sigma {..k} (\j::nat. {.. IJ k \ (\j i. u = (j,i) \ i j\k)" by (auto simp: IJ_def) lemma finite_IJ: "finite (IJ k)" by (auto simp: IJ_def) fun prev where "prev 0 0 = None" | "prev (Suc 0) 0 = None" | "prev (Suc j) 0 = Some (j, j - Suc 0)" | "prev j (Suc i) = Some (j,i)" lemma prev_eq_None_iff: "prev j i = None \ j \ Suc 0 \ i = 0" by (auto simp: le_Suc_eq elim: prev.elims) lemma prev_pair_less: "prev j i = Some ji' \ (ji', (j,i)) \ pair_less" by (auto simp: pair_lessI1 elim: prev.elims) lemma prev_Some_less: "\prev j i = Some (j',i'); i \ j\ \ i' < j'" by (auto elim: prev.elims) lemma prev_maximal: "\prev j i = Some (j',i'); (ji'', (j,i)) \ pair_less; ji'' \ IJ k\ \ (ji'', (j',i')) \ pair_less \ ji'' = (j',i')" by (force simp: IJ_def pair_less_def elim: prev.elims) lemma pair_less_prev: assumes "(u, (j,i)) \ pair_less" "u \ IJ k" shows "prev j i = Some u \ (\x. (u, x) \ pair_less \ prev j i = Some x)" proof (cases "prev j i") case None then show ?thesis using assms by (force simp: prev_eq_None_iff pair_less_def IJ_def split: prod.split) next case (Some a) then show ?thesis by (metis assms prev_maximal prod.exhaust_sel) qed subsubsection \Special primitives for the ordertype proof\ definition USigma :: "'a set set \ ('a set \ 'a set) \ 'a set set" where "USigma \ B \ \X\\. \y \ B X. {insert y X}" definition usplit where "usplit f A \ f (A - {Max A}) (Max A)" lemma USigma_empty [simp]: "USigma {} B = {}" by (auto simp: USigma_def) lemma USigma_iff: assumes "\I j. I \ \ \ I \ J I \ finite I" shows "x \ USigma \ J \ usplit (\I j. I \ \ \ j \ J I \ x = insert j I) x" proof - have [simp]: "\I j. \I \ \; j \ J I\ \ Max (insert j I) = j" by (meson Max_insert2 assms less_imp_le less_sets_def) show ?thesis proof - have \
: "j \ I" if "I \ \" "j \ J I" for I j using that by (metis assms less_irrefl less_sets_def) have "\I\\. \j\J I. x = insert j I" if "x - {Max x} \ \" and "Max x \ J (x - {Max x})" "x \ {}" using that by (metis Max_in assms infinite_remove insert_Diff) then show ?thesis by (auto simp: USigma_def usplit_def \
) qed qed proposition ordertype_append_image_IJ: assumes lenB [simp]: "\i j. i \ \ \ j \ J i \ length (B j) = c" and AB: "\i j. i \ \ \ j \ J i \ A i < B j" and IJ: "\i. i \ \ \ i \ J i \ finite i" and \: "\i. i \ \ \ ordertype (B ` J i) (lenlex less_than) = \" and A: "inj_on A \" shows "ordertype (usplit (\i j. A i @ B j) ` USigma \ J) (lenlex less_than) = \ * ordertype (A ` \) (lenlex less_than)" (is "ordertype ?AB ?R = _ * ?\") proof (cases "\ = {}") case False have "Ord \" using \ False wf_Ord_ordertype by fastforce show ?thesis proof (subst ordertype_eq_iff) define split where "split \ \l::nat list. (take (length l - c) l, (drop (length l - c) l))" have oB: "ordermap (B ` J i) ?R (B j) \ \" if \i \ \\ \j \ J i\ for i j using \ less_TC_iff that by fastforce then show "Ord (\ * ?\)" by (intro \Ord \\ wf_Ord_ordertype Ord_mult; simp) define f where "f \ \u. let (x,y) = split u in let i = inv_into \ A x in \ * ordermap (A`\) ?R x + ordermap (B`J i) ?R y" have inv_into_IA [simp]: "inv_into \ A (A i) = i" if "i \ \" for i by (simp add: A that) show "\f. bij_betw f ?AB (elts (\ * ?\)) \ (\x\?AB. \y\?AB. (f x < f y) = ((x, y) \ ?R))" unfolding bij_betw_def proof (intro exI conjI strip) show "inj_on f ?AB" proof (clarsimp simp: f_def inj_on_def split_def USigma_iff IJ usplit_def) fix x y assume \
: "\ * ordermap (A ` \) ?R (A (x - {Max x})) + ordermap (B ` J (x - {Max x})) ?R (B (Max x)) = \ * ordermap (A ` \) ?R (A (y - {Max y})) + ordermap (B ` J (y - {Max y})) ?R (B (Max y))" and x: "x - {Max x} \ \" and y: "y - {Max y} \ \" and mx: "Max x \ J (x - {Max x})" and "x = insert (Max x) x" and my: "Max y \ J (y - {Max y})" have "ordermap (A`\) ?R (A (x - {Max x})) = ordermap (A`\) ?R (A (y - {Max y}))" and B_eq: "ordermap (B ` J (x - {Max x})) ?R (B (Max x)) = ordermap (B ` J (y - {Max y})) ?R (B (Max y))" using mult_cancellation_lemma [OF \
] oB mx my x y by blast+ then have "A (x - {Max x}) = A (y - {Max y})" using x y by auto then have "x - {Max x} = y - {Max y}" by (metis x y inv_into_IA) then show "A (x - {Max x}) = A (y - {Max y}) \ B (Max x) = B (Max y)" using B_eq mx my by auto qed show "f ` ?AB = elts (\ * ?\)" proof show "f ` ?AB \ elts (\ * ?\)" using \Ord \\ apply (clarsimp simp add: f_def split_def USigma_iff IJ usplit_def) by (metis TC_small \ add_mult_less image_eqI ordermap_in_ordertype trans_llt wf_Ord_ordertype wf_llt) show "elts (\ * ?\) \ f ` ?AB" proof (clarsimp simp: f_def split_def image_iff USigma_iff IJ usplit_def Bex_def elim!: elts_multE split: prod.split) fix \ \ assume \: "\ \ elts \" and \: "\ \ elts ?\" have "\ \ ordermap (A ` \) (lenlex less_than) ` A ` \" by (meson \ ordermap_surj subset_iff) then obtain i where "i \ \" and yv: "\ = ordermap (A`\) ?R (A i)" by blast have "\ \ ordermap (B ` J i) (lenlex less_than) ` B ` J i" by (metis (no_types) \ \ \i \ \\ in_mono ordermap_surj) then obtain j where "j \ J i" and xu: "\ = ordermap (B`J i) ?R (B j)" by blast then have mji: "Max (insert j i) = j" by (meson IJ Max_insert2 \i \ \\ less_imp_le less_sets_def) have [simp]: "i - {j} = i" using IJ \i \ \\ \j \ J i\ less_setsD by fastforce show "\l. (\K. K - {Max K} \ \ \ Max K \ J (K - {Max K}) \ K = insert (Max K) K \ l = A (K - {Max K}) @ B (Max K)) \ \ * \ + \ = \ * ordermap (A ` \) ?R (take (length l - c) l) + ordermap (B ` J (inv_into \ A (take (length l - c) l))) ?R (drop (length l - c) l)" proof (intro conjI exI) let ?ji = "insert j i" show "A i @ B j = A (?ji - {Max ?ji}) @ B (Max ?ji)" by (auto simp: mji) qed (use \i \ \\ \j \ J i\ mji xu yv in auto) qed qed next fix p q assume "p \ ?AB" and "q \ ?AB" then obtain x y where peq: "p = A (x - {Max x}) @ B (Max x)" and qeq: "q = A (y - {Max y}) @ B (Max y)" and x: "x - {Max x} \ \" and y: "y - {Max y} \ \" and mx: "Max x \ J (x - {Max x})" and my: "Max y \ J (y - {Max y})" by (auto simp: USigma_iff IJ usplit_def) let ?mx = "x - {Max x}" let ?my = "y - {Max y}" show "(f p < f q) \ ((p, q) \ ?R)" proof assume "f p < f q" then consider "ordermap (A`\) ?R (A (x - {Max x})) < ordermap (A`\) ?R (A (y - {Max y}))" | "ordermap (A`\) ?R (A (x - {Max x})) = ordermap (A`\) ?R (A (y - {Max y}))" "ordermap (B`J (x - {Max x})) ?R (B (Max x)) < ordermap (B`J (y - {Max y})) ?R (B (Max y))" using x y mx my by (auto dest: mult_cancellation_less simp: f_def split_def peq qeq oB) then have "(A ?mx @ B (Max x), A ?my @ B (Max y)) \ ?R" proof cases case 1 then have "(A ?mx, A ?my) \ ?R" using x y by (force simp: Ord_mem_iff_lt intro: converse_ordermap_mono) then show ?thesis using x y mx my lenB lenlex_append1 by blast next case 2 then have "A ?mx = A ?my" using \?my \ \\ \?mx \ \\ by auto then have eq: "?mx = ?my" by (metis \?my \ \\ \?mx \ \\ inv_into_IA) then have "(B (Max x), B (Max y)) \ ?R" using mx my 2 by (force simp: Ord_mem_iff_lt intro: converse_ordermap_mono) with 2 show ?thesis by (simp add: eq irrefl_less_than) qed then show "(p,q) \ ?R" by (simp add: peq qeq f_def split_def sorted_list_of_set_Un AB) next assume pqR: "(p,q) \ ?R" then have \
: "(A ?mx @ B (Max x), A ?my @ B (Max y)) \ ?R" using peq qeq by blast then consider "(A ?mx, A ?my) \ ?R" | "A ?mx = A ?my \ (B (Max x), B (Max y)) \ ?R" proof (cases "(A ?mx, A ?my) \ ?R") case False have False if "(A ?my, A ?mx) \ ?R" by (metis \?my \ \\ \?mx \ \\ "\
" \(Max y) \ J ?my\ \(Max x) \ J ?mx\ lenB lenlex_append1 omega_sum_1_less order.asym that) then have "A ?mx = A ?my" by (meson False UNIV_I total_llt total_on_def) then show ?thesis using "\
" irrefl_less_than that(2) by auto qed (use that in blast) then have "\ * ordermap (A`\) ?R (A ?mx) + ordermap (B`J ?mx) ?R (B (Max x)) < \ * ordermap (A`\) ?R (A ?my) + ordermap (B`J ?my) ?R (B (Max y))" proof cases case 1 show ?thesis proof (rule add_mult_less_add_mult) show "ordermap (A`\) (lenlex less_than) (A ?mx) < ordermap (A`\) (lenlex less_than) (A ?my)" by (simp add: "1" \?my \ \\ \?mx \ \\ ordermap_mono_less) show "Ord (ordertype (A`\) ?R)" using wf_Ord_ordertype by blast show "ordermap (B ` J ?mx) ?R (B (Max x)) \ elts \" using Ord_less_TC_mem \Ord \\ \?mx \ \\ \(Max x) \ J ?mx\ oB by blast show "ordermap (B ` J ?my) ?R (B (Max y)) \ elts \" using Ord_less_TC_mem \Ord \\ \?my \ \\ \(Max y) \ J ?my\ oB by blast qed (use \?my \ \\ \?mx \ \\ \Ord \\ in auto) next case 2 with \?mx \ \\ show ?thesis using \(Max y) \ J ?my\ \(Max x) \ J ?mx\ ordermap_mono_less by (metis (no_types, hide_lams) Kirby.add_less_cancel_left TC_small image_iff inv_into_IA trans_llt wf_llt y) qed then show "f p < f q" using \?my \ \\ \?mx \ \\ \(Max y) \ J ?my\ \(Max x) \ J ?mx\ by (auto simp: peq qeq f_def split_def AB) qed qed qed auto qed auto subsubsection \The final part of 3.8, where two sequences are merged\ inductive merge :: "[nat list list,nat list list,nat list list,nat list list] \ bool" where NullNull: "merge [] [] [] []" | Null: "as \ [] \ merge as [] [concat as] []" | App: "\as1 \ []; bs1 \ []; concat as1 < concat bs1; concat bs1 < concat as2; merge as2 bs2 as bs\ \ merge (as1@as2) (bs1@bs2) (concat as1 # as) (concat bs1 # bs)" inductive_simps Null1 [simp]: "merge [] bs us vs" inductive_simps Null2 [simp]: "merge as [] us vs" lemma merge_single: "\concat as < concat bs; concat as \ []; concat bs \ []\ \ merge as bs [concat as] [concat bs]" using merge.App [of as bs "[]" "[]"] by (fastforce simp: less_list_def) lemma merge_length1_nonempty: assumes "merge as bs us vs" "as \ lists (- {[]})" shows "us \ lists (- {[]})" using assms by induction (auto simp: mem_lists_non_Nil) lemma merge_length2_nonempty: assumes "merge as bs us vs" "bs \ lists (- {[]})" shows "vs \ lists (- {[]})" using assms by induction (auto simp: mem_lists_non_Nil) lemma merge_length1_gt_0: assumes "merge as bs us vs" "as \ []" shows "length us > 0" using assms by induction auto lemma merge_length_le: assumes "merge as bs us vs" shows "length vs \ length us" using assms by induction auto lemma merge_length_le_Suc: assumes "merge as bs us vs" shows "length us \ Suc (length vs)" using assms by induction auto lemma merge_length_less2: assumes "merge as bs us vs" shows "length vs \ length as" using assms proof induction case (App as1 bs1 as2 bs2 as bs) then show ?case using length_greater_0_conv [of as1] by (simp, presburger) qed auto lemma merge_preserves: assumes "merge as bs us vs" shows "concat as = concat us \ concat bs = concat vs" using assms by induction auto lemma merge_interact: assumes "merge as bs us vs" "strict_sorted (concat as)" "strict_sorted (concat bs)" "bs \ lists (- {[]})" shows "strict_sorted (interact us vs)" using assms proof induction case (App as1 bs1 as2 bs2 as bs) then have bs: "concat bs1 < concat bs" "concat bs1 < concat as" and xx: "concat bs1 \ []" using merge_preserves strict_sorted_append_iff by fastforce+ then have "concat bs1 < interact as bs" unfolding less_list_def using App bs by (metis (no_types, lifting) Un_iff concat_append hd_in_set last_in_set merge_preserves set_interact sorted_wrt_append strict_sorted_append_iff) with App show ?case apply (simp add: strict_sorted_append_iff del: concat_eq_Nil_conv) by (metis hd_append2 less_list_def xx) qed auto lemma acc_lengths_merge1: assumes "merge as bs us vs" shows "list.set (acc_lengths k us) \ list.set (acc_lengths k as)" using assms proof (induction arbitrary: k) case (App as1 bs1 as2 bs2 as bs) then show ?case apply (simp add: acc_lengths_append strict_sorted_append_iff length_concat_acc_lengths) by (simp add: le_supI2 length_concat) qed (auto simp: length_concat_acc_lengths) lemma acc_lengths_merge2: assumes "merge as bs us vs" shows "list.set (acc_lengths k vs) \ list.set (acc_lengths k bs)" using assms proof (induction arbitrary: k) case (App as1 bs1 as2 bs2 as bs) then show ?case apply (simp add: acc_lengths_append strict_sorted_append_iff length_concat_acc_lengths) by (simp add: le_supI2 length_concat) qed (auto simp: length_concat_acc_lengths) lemma length_hd_le_concat: assumes "as \ []" shows "length (hd as) \ length (concat as)" by (metis (no_types) add.commute assms concat.simps(2) le_add2 length_append list.exhaust_sel) lemma length_hd_merge2: assumes "merge as bs us vs" shows "length (hd bs) \ length (hd vs)" using assms by induction (auto simp: length_hd_le_concat) lemma merge_less_sets_hd: assumes "merge as bs us vs" "strict_sorted (concat as)" "strict_sorted (concat bs)" "bs \ lists (- {[]})" shows "list.set (hd us) \ list.set (concat vs)" using assms proof induction case (App as1 bs1 as2 bs2 as bs) then have \
: "list.set (concat bs1) \ list.set (concat bs2)" by (force simp: dest: strict_sorted_imp_less_sets) have *: "list.set (concat as1) \ list.set (concat bs1)" using App by (metis concat_append strict_sorted_append_iff strict_sorted_imp_less_sets) then have "list.set (concat as1) \ list.set (concat bs)" using App \
less_sets_trans merge_preserves by (metis List.set_empty append_in_lists_conv le_zero_eq length_0_conv length_concat_ge) with * App.hyps show ?case by (fastforce simp: less_sets_UN1 less_sets_UN2 less_sets_Un2) qed auto lemma set_takeWhile: assumes "strict_sorted (concat as)" "as \ lists (- {[]})" shows "list.set (takeWhile (\x. x < y) as) = {x \ list.set as. x < y}" using assms proof (induction as) case (Cons a as) have "a < y" if a: "a < concat as" "strict_sorted a" "strict_sorted (concat as)" "x < y" "x \ []" "x \ list.set as" for x proof - have "last x \ list.set (concat as)" using set_concat that(5) that(6) by fastforce then have "last a < hd (concat as)" using Cons.prems that by (auto simp: less_list_def) also have "\ \ hd y" if "y \ []" using that a by (meson \last x \ list.set (concat as)\ dual_order.strict_trans less_list_def not_le sorted_hd_le strict_sorted_imp_sorted) finally show ?thesis by (simp add: less_list_def) qed then show ?case using Cons by (auto simp: strict_sorted_append_iff) qed auto proposition merge_exists: assumes "strict_sorted (concat as)" "strict_sorted (concat bs)" "as \ lists (- {[]})" "bs \ lists (- {[]})" "hd as < hd bs" "as \ []" "bs \ []" and disj: "\a b. \a \ list.set as; b \ list.set bs\ \ a bus vs. merge as bs us vs" using assms proof (induction "length as + length bs" arbitrary: as bs rule: less_induct) case (less as bs) obtain as1 as2 bs1 bs2 where A: "as1 \ []" "bs1 \ []" "concat as1 < concat bs1" "concat bs1 < concat as2" and B: "as = as1@as2" "bs = bs1@bs2" and C: "bs2 = [] \ (as2 \ [] \ hd as2 < hd bs2)" proof define as1 where "as1 \ takeWhile (\x. x < hd bs) as" define as2 where "as2 \ dropWhile (\x. x < hd bs) as" define bs1 where "bs1 \ if as2=[] then bs else takeWhile (\x. x < hd as2) bs" define bs2 where "bs2 \ if as2=[] then [] else dropWhile (\x. x < hd as2) bs" have as1: "as1 = takeWhile (\x. last x < hd (hd bs)) as" using less.prems by (auto simp: as1_def less_list_def cong: takeWhile_cong) have as2: "as2 = dropWhile (\x. last x < hd (hd bs)) as" using less.prems by (auto simp: as2_def less_list_def cong: dropWhile_cong) have hd_as2: "as2 \ [] \ \ hd as2 < hd bs" using as2_def hd_dropWhile by metis have hd_bs2: "bs2 \ [] \ \ hd bs2 < hd as2" using bs2_def hd_dropWhile by metis show "as1 \ []" by (simp add: as1_def less.prems takeWhile_eq_Nil_iff) show "bs1 \ []" by (metis as2 bs1_def hd_as2 hd_in_set less.prems(7) less.prems(8) set_dropWhileD takeWhile_eq_Nil_iff) show "bs2 = [] \ (as2 \ [] \ hd as2 < hd bs2)" by (metis as2_def bs2_def hd_bs2 less.prems(8) list.set_sel(1) set_dropWhileD) have AB: "list.set A \ list.set B" if "A \ list.set as1" "B \ list.set bs" for A B proof - have "A \ list.set as" using that by (metis as1 set_takeWhileD) then have "sorted A" by (metis concat.simps(2) concat_append less.prems(1) sorted_append split_list_last strict_sorted_imp_sorted) moreover have "sorted (hd bs)" by (metis concat.simps(2) hd_Cons_tl less.prems(2) less.prems(7) strict_sorted_append_iff strict_sorted_imp_sorted) ultimately show ?thesis using that less.prems apply (clarsimp simp add: as1_def set_takeWhile less_list_iff_less_sets less_sets_def) by (metis (full_types) UN_I hd_concat less_le_trans list.set_sel(1) set_concat sorted_hd_le strict_sorted_imp_sorted) qed show "as = as1@as2" by (simp add: as1_def as2_def) show "bs = bs1@bs2" by (simp add: bs1_def bs2_def) have "list.set (concat as1) \ list.set (concat bs1)" using AB set_takeWhileD by (fastforce simp: as1_def bs1_def less_sets_UN1 less_sets_UN2) then show "concat as1 < concat bs1" by (rule less_sets_imp_list_less) have "list.set (concat bs1) \ list.set (concat as2)" if "as2 \ []" proof (clarsimp simp add: bs1_def less_sets_UN1 less_sets_UN2 set_takeWhile less.prems) fix A B assume "A \ list.set as2" "B \ list.set bs" "B < hd as2" with that show "list.set B \ list.set A" using hd_as2 less.prems(1,2) apply (clarsimp simp add: less_sets_def less_list_def) apply (auto simp: as2_def) apply (simp flip: as2_def) by (smt (verit, ccfv_SIG) UN_I \as = as1 @ as2\ concat.simps(2) concat_append hd_concat in_set_conv_decomp_first le_less_trans less_le_trans set_concat sorted_append sorted_hd_le sorted_le_last strict_sorted_imp_sorted that) qed then show "concat bs1 < concat as2" by (simp add: bs1_def less_sets_imp_list_less) qed obtain cs ds where "merge as2 bs2 cs ds" proof (cases "as2 = [] \ bs2 = []") case True then show thesis using that C NullNull Null by metis next have \: "length as2 + length bs2 < length as + length bs" by (simp add: A B) case False moreover have "strict_sorted (concat as2)" "strict_sorted (concat bs2)" "as2 \ lists (- {[]})" "bs2 \ lists (- {[]})" "\a b. \a \ list.set as2; b \ list.set bs2\ \ a < b \ b < a" using B less.prems strict_sorted_append_iff by auto ultimately show ?thesis using C less.hyps [OF \] False that by force qed then obtain cs where "merge (as1 @ as2) (bs1 @ bs2) (concat as1 # cs) (concat bs1 # ds)" using A merge.App by blast then show ?case using B by blast qed subsubsection \Actual proof of Larson's Lemma 3.8\ proposition lemma_3_8: assumes "infinite N" obtains X where "X \ WW" "ordertype X (lenlex less_than) = \\\" "\u. u \ [X]\<^bsup>2\<^esup> \ \l. Form l u \ (l > 0 \ [enum N l] < inter_scheme l u \ List.set (inter_scheme l u) \ N)" proof - let ?LL = "lenlex less_than" define bf where "bf \ \M q. wfrec pair_less (\f (j,i). let R = (case prev j i of None \ M | Some u \ snd (f u)) in grab R (q j i))" have bf_rec: "bf M q (j,i) = (let R = (case prev j i of None \ M | Some u \ snd (bf M q u)) in grab R (q j i))" for M q j i by (subst (1) bf_def) (simp add: Let_def wfrec bf_def cut_apply prev_pair_less cong: conj_cong split: option.split) have "infinite (snd (bf M q u)) = infinite M \ fst (bf M q u) \ M \ snd (bf M q u) \ M" for M q u using wf_pair_less proof (induction u rule: wf_induct_rule) case (less u) then show ?case proof (cases u) case (Pair j i) with less.IH prev_pair_less show ?thesis apply (simp add: bf_rec [of M q j i] split: option.split) using fst_grab_subset snd_grab_subset by blast qed qed then have infinite_bf [simp]: "infinite (snd (bf M q u)) = infinite M" and bf_subset: "fst (bf M q u) \ M \ snd (bf M q u) \ M" for M q u by auto have bf_less_sets: "fst (bf M q ij) \ snd (bf M q ij)" if "infinite M" for M q ij using wf_pair_less proof (induction ij rule: wf_induct_rule) case (less u) then show ?case proof (cases u) case (Pair j i) with less_sets_grab show ?thesis by (simp add: bf_rec [of M q j i] less.IH prev_pair_less that split: option.split) qed qed have card_fst_bf: "finite (fst (bf M q (j,i))) \ card (fst (bf M q (j,i))) = q j i" if "infinite M" for M q j i by (simp add: that bf_rec [of M q j i] split: option.split) have bf_cong: "bf M q u = bf M q' u" if "snd u \ fst u" and eq: "\y x. \x\y; y\fst u\ \ q' y x = q y x" for M q q' u using wf_pair_less that proof (induction u rule: wf_induct_rule) case (less u) show ?case proof (cases u) case (Pair j i) with less.prems show ?thesis proof (clarsimp simp add: bf_rec [of M _ j i] split: option.split) fix j' i' assume *: "prev j i = Some (j',i')" then have **: "((j', i'), u) \ pair_less" by (simp add: Pair prev_pair_less) moreover have "i' < j'" using Pair less.prems by (simp add: prev_Some_less [OF *]) moreover have "\x y. \x \ y; y \ j'\ \ q' y x = q y x" using ** less.prems by (auto simp: pair_less_def Pair) ultimately show "grab (snd (bf M q (j',i'))) (q j i) = grab (snd (bf M q' (j',i'))) (q j i)" using less.IH by auto qed qed qed define ediff where "ediff \ \D:: nat \ nat set. \j i. enum (D j) (Suc i) - enum (D j) i" define F where "F \ \l (dl,a0::nat set,b0::nat \ nat \ nat set,M). let (d,Md) = grab (nxt M (enum N (Suc (2 * Suc l)))) (Suc l) in let (a,Ma) = grab Md (Min d) in let Gb = bf Ma (ediff (dl(l := d))) in let dl' = dl(l := d) in (dl', a, fst \ Gb, snd (Gb(l, l-1)))" define DF where "DF \ rec_nat (\i\{..<0}. {}, {}, \p. {}, N) F" have DF_simps: "DF 0 = (\i\{..<0}. {}, {}, \p. {}, N)" "DF (Suc l) = F l (DF l)" for l by (auto simp: DF_def) note cut_apply [simp] have inf [rule_format]: "\dl al bl L. DF l = (dl,al,bl,L) \ infinite L" for l by (induction l) (auto simp: DF_simps F_def Let_def grab_eqD infinite_nxtN assms split: prod.split) define \ where "\ \ \(dl, a, b, M). \l::nat. dl l \ a \ card a > 0 \ (\j\l. card (dl j) = Suc j) \ a \ \(range b) \ range b \ Collect finite \ a \ N \ \(range b) \ N \ infinite M \ b(l,l-1) \ M \ M \ N" have \_DF: "\ (DF (Suc l)) l" for l proof (induction l) case 0 show ?case using assms apply (clarsimp simp add: bf_rec F_def DF_simps \_def split: prod.split) apply (drule grab_eqD, blast dest: grab_eqD infinite_nxtN)+ apply (auto simp: less_sets_UN2 less_sets_grab card_fst_bf elim!: less_sets_weaken2) apply (metis card_1_singleton_iff Min_singleton greaterThan_iff insertI1 le0 nxt_subset_greaterThan subsetD) using nxt_subset snd_grab_subset bf_subset by blast+ next case (Suc l) then show ?case using assms unfolding Let_def DF_simps(2)[of "Suc l"] F_def \_def apply (clarsimp simp add: bf_rec DF_simps split: prod.split) apply (drule grab_eqD, metis grab_eqD infinite_nxtN)+ apply (safe, simp_all add: less_sets_UN2 less_sets_grab card_fst_bf card_Suc_eq_finite) apply (meson less_sets_weaken2) apply (metis Min_in gr0I greaterThan_iff insert_not_empty le_inf_iff less_asym nxt_def subsetD) apply (meson bf_subset less_sets_weaken2) apply (meson nxt_subset subset_eq) apply (meson bf_subset nxt_subset subset_eq) using bf_rec infinite_bf apply force using bf_less_sets bf_rec apply force by (metis bf_rec bf_subset nxt_subset subsetD) qed define d where "d \ \k. let (dk,ak,bk,M) = DF(Suc k) in dk k" define a where "a \ \k. let (dk,ak,bk,M) = DF(Suc k) in ak" define b where "b \ \k. let (dk,ak,bk,M) = DF(Suc k) in bk" define M where "M \ \k. let (dk,ak,bk,M) = DF k in M" have infinite_M [simp]: "infinite (M k)" for k by (auto simp: M_def inf split: prod.split) have M_Suc_subset: "M (Suc k) \ M k" for k apply (clarsimp simp add: Let_def M_def F_def DF_simps split: prod.split) apply (drule grab_eqD, blast dest: infinite_nxtN local.inf)+ using bf_subset nxt_subset by blast have Inf_M_Suc_ge: "Inf (M k) \ Inf (M (Suc k))" for k by (simp add: M_Suc_subset cInf_superset_mono infinite_imp_nonempty) have Inf_M_telescoping: "{Inf (M k)..} \ {Inf (M k')..}" if k': "k'\k" for k k' using that Inf_nat_def1 infinite_M unfolding Inf_nat_def atLeast_subset_iff by (metis M_Suc_subset finite.emptyI le_less_linear lift_Suc_antimono_le not_less_Least subsetD) have d_eq: "d k = fst (grab (nxt (M k) (enum N (Suc (2 * Suc k)))) (Suc k))" for k by (simp add: d_def M_def Let_def DF_simps F_def split: prod.split) then have finite_d [simp]: "finite (d k)" for k by simp then have d_ne [simp]: "d k \ {}" for k by (metis card.empty card_grab d_eq infinite_M infinite_nxtN nat.distinct(1)) have a_eq: "\M. a k = fst (grab M (Min (d k))) \ infinite M" for k apply (simp add: a_def d_def M_def Let_def DF_simps F_def split: prod.split) by (metis fst_conv grab_eqD infinite_nxtN local.inf) then have card_a: "card (a k) = Inf (d k)" for k by (metis cInf_eq_Min card_grab d_ne finite_d) have d_eq_dl: "d k = dl k" if "(dl,a,b,P) = DF l" "k < l" for k l dl a b P using that by (induction l arbitrary: dl a b P) (simp_all add: d_def DF_simps F_def Let_def split: prod.split_asm prod.split) have card_d [simp]: "card (d k) = Suc k" for k by (auto simp: d_eq infinite_nxtN) have d_ne [simp]: "d j \ {}" and a_ne [simp]: "a j \ {}" and finite_d [simp]: "finite (d j)" and finite_a [simp]: "finite (a j)" for j using \_DF [of "j"] by (auto simp: \_def a_def d_def card_gt_0_iff split: prod.split_asm) have da: "d k \ a k" for k using \_DF [of "k"] by (simp add: \_def a_def d_def split: prod.split_asm) have ab_same: "a k \ \(range(b k))" for k using \_DF [of "k"] by (simp add: \_def a_def b_def M_def split: prod.split_asm) have snd_bf_subset: "snd (bf M r (j,i)) \ snd (bf M r (j',i'))" if ji: "((j',i'), (j,i)) \ pair_less" "(j',i') \ IJ k" for M r k j i j' i' using wf_pair_less ji proof (induction rule: wf_induct_rule [where a= "(j,i)"]) case (less u) show ?case proof (cases u) case (Pair j i) then consider "prev j i = Some (j', i')" | x where "((j', i'), x) \ pair_less" "prev j i = Some x" using less.prems pair_less_prev by blast then show ?thesis proof cases case 2 with less.IH show ?thesis unfolding bf_rec Pair by (metis in_mono option.simps(5) prev_pair_less snd_grab_subset subsetI that(2)) qed (simp add: Pair bf_rec snd_grab_subset) qed qed have less_bf: "fst (bf M r (j',i')) \ fst (bf M r (j,i))" if ji: "((j',i'), (j,i)) \ pair_less" "(j',i') \ IJ k" and "infinite M" for M r k j i j' i' proof - consider "prev j i = Some (j', i')" | j'' i'' where "((j', i'), (j'',i'')) \ pair_less" "prev j i = Some (j'',i'')" by (metis pair_less_prev ji prod.exhaust_sel) then show ?thesis proof cases case 1 then show ?thesis using bf_less_sets bf_rec less_sets_fst_grab \infinite M\ by force next case 2 then have "fst (bf M r (j',i')) \ snd (bf M r (j'',i''))" by (meson bf_less_sets snd_bf_subset less_sets_weaken2 that) with 2 show ?thesis using bf_rec bf_subset less_sets_fst_grab \infinite M\ by auto qed qed have aM: "a k \ M (Suc k)" for k apply (clarsimp simp add: a_def M_def DF_simps F_def Let_def split: prod.split) by (meson bf_subset grab_eqD infinite_nxtN less_sets_weaken2 local.inf) then have "a k \ a (Suc k)" for k by (metis IntE card_d card.empty d_eq da fst_grab_subset less_sets_trans less_sets_weaken2 nat.distinct(1) nxt_def subsetI) then have aa: "a j \ a k" if "j b k (j,i)" if "k'\k" for k k' j i by (metis a_ne ab_same le_less less_sets_UN2 less_sets_trans rangeI that) have db: "d j \ b k (j,i)" if "j\k" for k j i by (meson a_ne ab da less_sets_trans that) have bMkk: "b k (k,k-1) \ M (Suc k)" for k using \_DF [of k] by (simp add: \_def b_def d_def M_def split: prod.split_asm) have b: "\P \ M k. infinite P \ (\j i. i\j \ j\k \ b k (j,i) = fst (bf P (ediff d) (j,i)))" for k proof (clarsimp simp: b_def DF_simps F_def Let_def split: prod.split) fix a a' d' dl bb P M' M'' assume gr: "grab M'' (Min d') = (a', M')" "grab (nxt P (enum N (Suc (Suc (Suc (2 * k)))))) (Suc k) = (d', M'')" and DF: "DF k = (dl, a, bb, P)" have deq: "d j = (if j = k then d' else dl j)" if "j\k" for j proof (cases "j < k") case True then show ?thesis by (metis DF d_eq_dl less_not_refl) next case False then show ?thesis using that DF gr by (auto simp: d_def DF_simps F_def Let_def split: prod.split) qed have "M' \ P" by (metis gr in_mono nxt_subset snd_conv snd_grab_subset subsetI) also have "P \ M k" using DF by (simp add: M_def) finally have "M' \ M k" . moreover have "infinite M'" using DF by (metis (mono_tags) finite_grab_iff gr infinite_nxtN local.inf snd_conv) moreover have "ediff (dl(k := d')) j i = ediff d j i" if "j\k" for j i by (simp add: deq that ediff_def) then have "bf M' (ediff (dl(k := d'))) (j,i) = bf M' (ediff d) (j,i)" if "i \ j" "j\k" for j i using bf_cong that by fastforce ultimately show "\P\M k. infinite P \ (\j i. i \ j \ j \ k \ fst (bf M' (ediff (dl(k := d'))) (j,i)) = fst (bf P (ediff d) (j,i)))" by auto qed have card_b: "card (b k (j,i)) = enum (d j) (Suc i) - enum (d j) i" if "j\k" for k j i \\there's a short proof of this from the previous result but it would need @{term"i\j"}\ proof (clarsimp simp: b_def DF_simps F_def Let_def split: prod.split) fix dl and a a' d':: "nat set" and bb M M' M'' assume gr: "grab M'' (Min d') = (a', M')" "grab (nxt M (enum N (Suc (Suc (Suc (2 * k)))))) (Suc k) = (d',M'')" and DF: "DF k = (dl, a, bb, M)" have "d j = (if j = k then d' else dl j)" proof (cases "j < k") case True then show ?thesis by (metis DF d_eq_dl less_not_refl) next case False then show ?thesis using that DF gr by (auto simp: d_def DF_simps F_def Let_def split: prod.split) qed then show "card (fst (bf M' (ediff (dl(k := d'))) (j,i))) = enum (d j) (Suc i) - enum (d j) i" using DF gr card_fst_bf grab_eqD infinite_nxtN local.inf ediff_def by auto qed have card_b_pos: "card (b k (j,i)) > 0" if "i < j" "j\k" for k j i by (simp add: card_b that finite_enumerate_step) have b_ne [simp]: "b k (j,i) \ {}" if "i < j" "j\k" for k j i using card_b_pos [OF that] less_imp_neq by fastforce+ have card_b_finite [simp]: "finite (b k u)" for k u using \_DF [of k] by (fastforce simp: \_def b_def) have bM: "b k (j,i) \ M (Suc k)" if "ik" for i j k proof - obtain M' where "M' \ M k" "infinite M'" and bk: "\j i. i\j \ j\k \ b k (j,i) = fst (bf M' (ediff d) (j,i))" using b by (metis (no_types, lifting)) show ?thesis proof (cases "j=k \ i = k-1") case False show ?thesis proof (rule less_sets_trans [OF _ bMkk]) show "b k (j,i) \ b k (k, k-1)" using that \infinite M'\ False by (force simp: bk pair_less_def IJ_def intro: less_bf) show "b k (k, k-1) \ {}" using b_ne that by auto qed qed (use bMkk in auto) qed have b_InfM: "\ (range (b k)) \ {\(M k)..}" for k proof (clarsimp simp add: \_def b_def M_def DF_simps F_def Let_def split: prod.split) fix r dl :: "nat \ nat set" and a b and d' a' M'' M' P and x j' i' :: nat assume gr: "grab M'' (Min d') = (a', M')" "grab (nxt P (enum N (Suc (Suc (Suc (2 * k)))))) (Suc k) = (d', M'')" and DF: "DF k = (dl, a, b, P)" and x: "x \ fst (bf M' (ediff (dl(k := d'))) (j', i'))" have "infinite P" using DF local.inf by blast then have "M' \ P" by (meson gr grab_eqD infinite_nxtN nxt_subset order.trans) with bf_subset show "\ P \ x" using Inf_nat_def x le_less_linear not_less_Least by fastforce qed have b_Inf_M_Suc: "b k (j,i) \ {Inf(M (Suc k))}" if "ik" for k j i using bMkk [of k] that by (metis Inf_nat_def1 bM finite.emptyI infinite_M less_setsD less_sets_singleton2) have bb_same: "b k (j',i') \ b k (j,i)" if "((j',i'), (j,i)) \ pair_less" "(j',i') \ IJ k" for k j i j' i' using that unfolding b_def DF_simps F_def Let_def by (auto simp: less_bf grab_eqD infinite_nxtN local.inf split: prod.split) have bb: "b k' (j',i') \ b k (j,i)" if j: "i' < j'" "j'\k'" and k: "k' {Inf(M (Suc k'))}" using Suc_lessD b_Inf_M_Suc nat_less_le j by blast show "b k (j,i) \ {Inf(M (Suc k'))..}" by (meson Inf_M_telescoping Suc_leI UnionI b_InfM rangeI subset_eq k) qed have M_subset_N: "M k \ N" for k proof (cases k) case (Suc k') with \_DF [of k'] show ?thesis by (auto simp: M_def Let_def \_def split: prod.split) qed (auto simp: M_def DF_simps) have a_subset_N: "a k \ N" for k using \_DF [of k] by (simp add: a_def \_def split: prod.split prod.split_asm) have d_subset_N: "d k \ N" for k using M_subset_N [of k] d_eq fst_grab_subset nxt_subset by blast have b_subset_N: "b k (j,i) \ N" for k j i using \_DF [of k] by (force simp: b_def \_def) define \:: "[nat,nat] \ nat set set" where "\ \ \j0 j. nsets {j0<..} j" have \_finite: "finite K" and \_card: "card K = j" if "K \ \ j0 j" for K j0 j using that by (auto simp add: \_def nsets_def) have \_enum: "j0 < enum K i" if "K \ \ j0 j" "i < card K" for K j0 j i using that by (auto simp: \_def nsets_def finite_enumerate_in_set subset_eq) have \_0 [simp]: "\ k 0 = {{}}" for k by (auto simp: \_def) have \_Suc: "\ j0 (Suc j) = USigma (\ j0 j) (\K. {Max (insert j0 K)<..})" (is "?lhs = ?rhs") for j j0 proof show "\ j0 (Suc j) \ USigma (\ j0 j) (\K. {Max (insert j0 K)<..})" unfolding \_def nsets_def USigma_def proof clarsimp fix K assume K: "K \ {j0<..}" "finite K" "card K = Suc j" then have "Max K \ K" by (metis Max_in card_0_eq nat.distinct(1)) then obtain i where "Max (insert j0 (K - {Max K})) < i" "K = insert i (K - {Max K})" using K by (simp add: subset_iff) (metis DiffE Max.coboundedI insertCI insert_Diff le_neq_implies_less) then show "\L\{j0<..}. finite L \ card L = j \ (\i\{Max (insert j0 L)<..}. K = insert i L)" using K by (metis \Max K \ K\ card_Diff_singleton_if diff_Suc_1 finite_Diff greaterThan_iff insert_subset) qed show "?rhs \ \ j0 (Suc j)" by (force simp: \_def nsets_def USigma_def) qed define BB where "BB \ \j0 j K. list_of (a j0 \ (\i \j. BB j j ` \ j j" have less_list_of: "BB j i K < list_of (b l (j,i))" if K: "K \ \ j i" "\j\K. j < l" and "i \ j" "j \ l" for j i K l unfolding BB_def proof (rule less_sets_imp_sorted_list_of_set) have "\i. i < card K \ b (enum K i) (j,i) \ b l (j, card K)" using that by (metis \_card \_enum \_finite bb finite_enumerate_in_set nat_less_le less_le_trans) then show "a j \ (\i b l (j,i)" using that unfolding \_def nsets_def by (auto simp: less_sets_Un1 less_sets_UN1 ab finite_enumerate_in_set subset_eq) qed auto have BB_Suc: "BB j0 (Suc j) K = usplit (\L k. BB j0 j L @ list_of (b k (j0, j))) K" if j: "j \ j0" and K: "K \ \ j0 (Suc j)" for j0 j K \\towards the ordertype proof\ proof - have Kj: "K \ {j0<..}" and [simp]: "finite K" and cardK: "card K = Suc j" using K by (auto simp: \_def nsets_def) have KMK: "K - {Max K} \ \ j0 j" using that by (simp add: \_Suc USigma_iff \_finite less_sets_def usplit_def) have "j0 < Max K" by (metis Kj Max_in cardK card_gt_0_iff greaterThan_iff subsetD zero_less_Suc) have MaxK: "Max K = enum K j" proof (rule Max_eqI) fix k assume "k \ K" with K cardK show "k \ enum K j" by (metis \finite K\ finite_enumerate_Ex finite_enumerate_mono_iff leI lessI not_less_eq) qed (auto simp: cardK finite_enumerate_in_set) have ene: "i enum (K - {enum K j}) i = enum K i" for i using finite_enumerate_Diff_singleton [OF \finite K\] by (simp add: cardK) have "BB j0 (Suc j) K = list_of ((a j0 \ (\x b (enum K j) (j0, j))" by (simp add: BB_def lessThan_Suc Un_ac) also have "\ = list_of ((a j0 \ (\i b (enum K j) (j0, j)" if "i enum K i" using that K by (metis \_enum cardK less_SucI less_imp_le_nat) show "enum K i < enum K j" by (simp add: cardK finite_enumerate_mono that) qed moreover have "a j0 \ b (enum K j) (j0, j)" using MaxK \j0 < Max K\ ab by auto ultimately show "a j0 \ (\x b (enum K j) (j0, j)" by (simp add: less_sets_Un1 less_sets_UN1) qed (auto simp: finite_UnI) also have "\ = BB j0 j (K - {Max K}) @ list_of (b (Max K) (j0, j))" by (simp add: BB_def MaxK ene) also have "\ = usplit (\L k. BB j0 j L @ list_of (b k (j0, j))) K" by (simp add: usplit_def) finally show ?thesis . qed have enum_d_0: "enum (d j) 0 = Inf (d j)" for j using enum_0_eq_Inf_finite by auto have Inf_b_less: "\(b k' (j',i')) < \(b k (j,i))" if j: "i' < j'" "i < j" "j'\k'" "j\k" and k: "k' (b k (k, k-1)) \ k-1" for k proof (induction k) case (Suc k) show ?case proof (cases "k=0") case False then have "\ (b k (k, k - 1)) < \ (b (Suc k) (Suc k, k))" using Inf_b_less by auto with Suc show ?thesis by simp qed auto qed auto have b_ge: "\ (b k (j,i)) \ k-1" if "k \ j" "j > i" for k j i proof - have "\ Suc (\ (b k (j, i))) < k" by (metis (no_types) Inf_b_less Suc_leI b_ge_k diff_Suc_1 lessI not_less that) then show ?thesis by simp qed have hd_b: "hd (list_of (b k (j,i))) = \ (b k (j,i))" if "i < j" "j \ k" for k j i using that by (simp add: hd_list_of cInf_eq_Min) have b_disjoint_less: "b (enum K i) (j0, i) \ b (enum K i') (j0, i') = {}" if K: "K \ {j0<..}" "finite K" "card K \ j0" "i' < j" "i < i'" "j \ j0" for i i' j j0 K proof (intro bb less_sets_imp_disjnt [unfolded disjnt_def]) show "i < j0" using that by linarith then show "j0 \ enum K i" by (meson K finite_enumerate_in_set greaterThan_iff less_imp_le_nat less_le_trans subsetD) show "enum K i < enum K i'" using K \j \ j0\ that by auto qed have b_disjoint: "b (enum K i) (j0, i) \ b (enum K i') (j0, i') = {}" if K: "K \ {j0<..}" "finite K" "card K \ j0" "i < j" "i' < j" "i \ i'" "j \ j0" for i i' j j0 K using that b_disjoint_less inf_commute neq_iff by metis have ot\: "ordertype ((\k. list_of (b k (j,i))) ` {Max (insert j K)<..}) ?LL = \" (is "?lhs = _") if K: "K \ \ j i" "j > i" for j i K proof - have Sucj: "Suc (Max (insert j K)) \ j" using \_finite that(1) le_Suc_eq by auto let ?N = "{Inf(b k (j,i))| k. Max (insert j K) < k}" have infN: "infinite ?N" proof (clarsimp simp add: infinite_nat_iff_unbounded_le) fix m show "\n\m. \k. n = \ (b k (j,i)) \ Max (insert j K) < k" using b_ge \j > i\ Sucj by (metis (no_types, lifting) diff_Suc_1 le_SucI le_trans less_Suc_eq_le nat_le_linear) qed have [simp]: "Max (insert j K) < k \ j < k \ (\a\K. a < k)" for k using that by (auto simp: \_finite) have "?lhs = ordertype ?N less_than" proof (intro ordertype_eqI strip) have "list_of (b k (j,i)) = list_of (b k' (j,i))" if "j \ k" "j \ k'" "hd (list_of (b k (j,i))) = hd (list_of (b k' (j,i)))" for k k' by (metis Inf_b_less \i < j\ hd_b nat_less_le not_le that) moreover have "\k' j' i'. hd (list_of (b k (j,i))) = \ (b k' (j', i')) \ i' < j' \ j' \ k'" if "j \ k" for k using that \i < j\ hd_b less_imp_le_nat by blast moreover have "\k'. hd (list_of (b k (j,i))) = \ (b k' (j,i)) \ j < k' \ (\a\K. a < k')" if "j < k" "\a\K. a < k" for k using that K hd_b less_imp_le_nat by blast moreover have "\ (b k (j,i)) \ hd ` (\k. list_of (b k (j,i))) ` {Max (insert j K)<..}" if "j < k" "\a\K. a < k" for k using that K by (auto simp: hd_b image_iff) ultimately show "bij_betw hd ((\k. list_of (b k (j,i))) ` {Max (insert j K)<..}) {\ (b k (j,i)) |k. Max (insert j K) < k}" by (auto simp: bij_betw_def inj_on_def) next fix ms ns assume "ms \ (\k. list_of (b k (j,i))) ` {Max (insert j K)<..}" and "ns \ (\k. list_of (b k (j,i))) ` {Max (insert j K)<..}" with that obtain k k' where ms: "ms = list_of (b k (j,i))" and ns: "ns = list_of (b k' (j,i))" and "j < k" "j < k'" and lt_k: "\a\K. a < k" and lt_k': "\a\K. a < k'" by (auto simp: \_finite) then have len_eq [simp]: "length ns = length ms" by (simp add: card_b) have nz: "length ns \ 0" using b_ne \i < j\ \j < k'\ ns by auto show "(hd ms, hd ns) \ less_than \ (ms, ns) \ ?LL" proof assume "(hd ms, hd ns) \ less_than" then show "(ms, ns) \ ?LL" using that nz by (fastforce simp: lenlex_def \_finite card_b intro: hd_lex) next assume \
: "(ms, ns) \ ?LL" then have "(list_of (b k' (j,i)), list_of (b k (j,i))) \ ?LL" using less_asym ms ns omega_sum_1_less by blast then show "(hd ms, hd ns) \ less_than" using \j < k\ \j < k'\ Inf_b_less [of i j i j] ms ns by (metis Cons_lenlex_iff \
len_eq b_ne card_b_finite diff_Suc_1 hd_Cons_tl hd_b length_Cons less_or_eq_imp_le less_than_iff linorder_neqE_nat sorted_list_of_set_eq_Nil_iff that(2)) qed qed auto also have "\ = \" using infN ordertype_nat_\ by blast finally show ?thesis . qed have ot\j: "ordertype (BB j0 j ` \ j0 j) ?LL = \\j" if "j \ j0" for j j0 using that proof (induction j) \\a difficult proof, but no hints in Larson's text\ case 0 then show ?case by (auto simp: XX_def) next case (Suc j) then have ih: "ordertype (BB j0 j ` \ j0 j) ?LL = \ \ j" by simp have "j \ j0" by (simp add: Suc.prems Suc_leD) have inj_BB: "inj_on (BB j0 j) ([{j0<..}]\<^bsup>j\<^esup>)" proof (clarsimp simp: inj_on_def BB_def nsets_def sorted_list_of_set_Un less_sets_UN2) fix X Y assume X: "X \ {j0<..}" and Y: "Y \ {j0<..}" and "finite X" "finite Y" and jeq: "j = card X" and "card Y = card X" and eq: "list_of (a j0 \ (\i (\in. \n < card X\ \ j0 \ enum X n" using X \finite X\ finite_enumerate_in_set less_imp_le_nat by blast have enumY: "\n. \n < card X\ \ j0 \ enum Y n" using subsetD [OF Y] by (metis \card Y = card X\ \finite Y\ finite_enumerate_in_set greaterThan_iff less_imp_le_nat) have smX: "strict_mono_sets {..i. b (enum X i) (j0, i))" and smY: "strict_mono_sets {..i. b (enum Y i) (j0, i))" using Suc.prems \card Y = card X\ \finite X\ \finite Y\ bb enumX enumY jeq by (auto simp: strict_mono_sets_def) have len_eq: "length ms = length ns" if "(ms, ns) \ list.set (zip (map (list_of \ (\i. b (enum X i) (j0,i))) (list_of {.. (\i. b (enum Y i) (j0,i))) (list_of {.. card X" for ms ns n using that by (induction n rule: nat.induct) (auto simp: card_b enumX enumY) have "concat (map (list_of \ (\i. b (enum X i) (j0, i))) (list_of {.. (\i. b (enum Y i) (j0, i))) (list_of {.. (\i. b (enum X i) (j0, i))) (list_of {.. (\i. b (enum Y i) (j0, i))) (list_of {.. (b (enum X i) (j0,i))" "Inf (b (enum Y i) (j0,i)) \ (b (enum Y i) (j0,i))" "i < j0" using Inf_nat_def1 Suc.prems b_ne enumX enumY jeq that by auto ultimately show ?thesis by (metis Inf_b_less enumX enumY leI nat_less_le that) qed then show "X = Y" by (simp add: \card Y = card X\ \finite X\ \finite Y\ finite_enum_ext) qed have BB_Suc': "BB j0 (Suc j) X = usplit (\L k. BB j0 j L @ list_of (b k (j0, j))) X" if "X \ USigma (\ j0 j) (\K. {Max (insert j0 K)<..})" for X using that by (simp add: USigma_iff \_finite less_sets_def usplit_def \_Suc BB_Suc \j \ j0\) have "ordertype (BB j0 (Suc j) ` \ j0 (Suc j)) ?LL = ordertype (usplit (\L k. BB j0 j L @ list_of (b k (j0, j))) ` USigma (\ j0 j) (\K. {Max (insert j0 K)<..})) ?LL" by (simp add: BB_Suc' \_Suc) also have "\ = \ * ordertype (BB j0 j ` \ j0 j) ?LL" proof (intro ordertype_append_image_IJ) fix L k assume "L \ \ j0 j" and "k \ {Max (insert j0 L)<..}" then have "j0 < k" and L: "\a. a \ L \ a < k" by (simp_all add: \_finite) then show "BB j0 j L < list_of (b k (j0, j))" by (simp add: \L \ \ j0 j\ \j \ j0\ \_finite less_list_of) next show "inj_on (BB j0 j) (\ j0 j)" by (simp add: \_def inj_BB) next fix L assume L: "L \ \ j0 j" then show "L \ {Max (insert j0 L)<..} \ finite L" by (simp add: \_finite less_sets_def) show "ordertype ((\i. list_of (b i (j0, j))) ` {Max (insert j0 L)<..}) ?LL = \" using L Suc.prems Suc_le_lessD ot\ by blast qed (auto simp: \_finite card_b) also have "\ = \ \ ord_of_nat (Suc j)" by (simp add: oexp_mult_commute ih) finally show ?case . qed define seqs where "seqs \ \j0 j K. list_of (a j0) # (map (list_of \ (\i. b (enum K i) (j0,i))) (list_of {.. lists (- {[]})" if K: "K \ \ j0 j" and "j \ j0" for K j j0 proof - have j0: "\i. i < card K \ j0 \ enum K i" and le_j0: "card K \ j0" using finite_enumerate_in_set that unfolding \_def nsets_def by fastforce+ show "BB j0 j K = concat (seqs j0 j K)" using that unfolding BB_def \_def nsets_def seqs_def by (fastforce simp: j0 ab bb less_sets_UN2 sorted_list_of_set_Un strict_mono_sets_def sorted_list_of_set_UN_lessThan) have "b (enum K i) (j0, i) \ {}" if "i < card K" for i using j0 le_j0 less_le_trans that by simp moreover have "card K = j" using K \_card by blast ultimately show "seqs j0 j K \ lists (- {[]})" by (clarsimp simp: seqs_def) (metis card_b_finite sorted_list_of_set_eq_Nil_iff) qed have BB_decomp: "\cs. BB j0 j K = concat cs \ cs \ lists (- {[]})" if K: "K \ \ j0 j" and "j \ j0" for K j j0 using BB_eq_concat_seqs seqs_ne K that(2) by blast have a_subset_M: "a k \ M k" for k apply (clarsimp simp: a_def M_def DF_simps F_def Let_def split: prod.split_asm) by (metis (no_types) fst_conv fst_grab_subset nxt_subset snd_conv snd_grab_subset subsetD) have ba_Suc: "b k (j,i) \ a (Suc k)" if "i < j" "j \ k" for i j k by (meson a_subset_M bM less_sets_weaken2 nat_less_le that) have ba: "b k (j,i) \ a r" if "i < j" "j \ k" "k < r" for i j k r by (metis Suc_lessI a_ne aa ba_Suc less_sets_trans that) have disjnt_ba: "disjnt (b k (j,i)) (a r)" if "i < j" "j \ k" for i j k r by (meson ab ba disjnt_sym less_sets_imp_disjnt not_le that) have bb_disjnt: "disjnt (b k (j,i)) (b l (r,q))" if "q < r" "i < j" "j \ k" "r \ l" "j < r" for i j q r k l proof (cases "k=l") case True with that show ?thesis by (force simp: pair_less_def IJ_def intro: bb_same less_sets_imp_disjnt) next case False with that show ?thesis by (metis bb less_sets_imp_disjnt disjnt_sym nat_neq_iff) qed have sum_card_b: "(\i {j0<..}" "finite K" "card K \ j0" and "j \ j0" for j0 j K using \j \ j0\ proof (induction j) case 0 then show ?case by auto next case (Suc j) then have "j < card K" using that(3) by linarith have dis: "disjnt (b (enum K j) (j0, j)) (\ij < card K\ by (force simp: finite_enumerate_in_set) have "(\ii = card (b (enum K j) (j0, j)) + enum (d j0) j - enum (d j0) 0" using \Suc j \ j0\ by (simp add: Suc.IH split: nat_diff_split) also have "\ = enum (d j0) (Suc j) - enum (d j0) 0" using j0_less Suc.prems card_b less_or_eq_imp_le by force finally show ?case . qed have card_UN_b: "card (\i {j0<..}" "finite K" "card K \ j0" and "j \ j0" for j0 j K using that by (simp add: card_UN_disjoint sum_card_b b_disjoint) have len_BB: "length (BB j j K) = enum (d j) j" if K: "K \ \ j j" and "j \ j" for j K proof - have dis_ab: "\i. i < j \ disjnt (a j) (b (enum K i) (j,i))" using K \_card \_enum ab less_sets_imp_disjnt nat_less_le by blast show ?thesis using K unfolding BB_def \_def nsets_def by (simp add: card_UN_b card_Un_disjnt dis_ab card_a cInf_le_finite finite_enumerate_in_set enum_0_eq_Inf_finite) qed have "d k \ d (Suc k)" for k by (metis aM a_ne d_eq da less_sets_fst_grab less_sets_trans less_sets_weaken2 nxt_subset) then have dd: "d k' \ d k" if "k' < k" for k' k by (meson UNIV_I d_ne less_sets_imp_strict_mono_sets strict_mono_sets_def that) show thesis proof show "(\ (range XX)) \ WW" by (auto simp: XX_def BB_def WW_def) show "ordertype (\ (range XX)) (?LL) = \ \ \" using ot\j by (simp add: XX_def ordertype_\\) next fix U assume U: "U \ [\ (range XX)]\<^bsup>2\<^esup>" then obtain x y where Ueq: "U = {x,y}" and len_xy: "length x \ length y" by (auto simp: lenlex_nsets_2_eq lenlex_length) show "\l. Form l U \ (0 < l \ [enum N l] < inter_scheme l U \ list.set (inter_scheme l U) \ N)" proof (cases "length x = length y") case True then show ?thesis using Form.intros(1) U Ueq by fastforce next case False then have xy: "length x < length y" using len_xy by auto obtain j r K L where K: "K \ \ j j" and xeq: "x = BB j j K" and ne: "BB j j K \ BB r r L" and L: "L \ \ r r" and yeq: "y = BB r r L" using U by (auto simp: Ueq XX_def) then have "length x = enum (d j) j" "length y = enum (d r) r" by (auto simp: len_BB) then have "j < r" using xy dd by (metis card_d finite_enumerate_in_set finite_d lessI less_asym less_setsD linorder_neqE_nat) then have aj_ar: "a j \ a r" using aa by auto have Ksub: "K \ {j<..}" and "finite K" "card K \ j" using K by (auto simp: \_def nsets_def) have Lsub: "L \ {r<..}" and "finite L" "card L \ r" using L by (auto simp: \_def nsets_def) have enumK: "enum K i > j" if "i < j" for i using K \_card \_enum that by blast have enumL: "enum L i > r" if "i < r" for i using L \_card \_enum that by blast have "list.set (acc_lengths w (seqs j0 j K)) \ (+) w ` d j0" if K: "K \ {j0<..}" "finite K" "card K \ j0" and "j \ j0" for j0 j K w using \j \ j0\ proof (induction j arbitrary: w) case 0 then show ?case by (simp add: seqs_def Inf_nat_def1 card_a) next case (Suc j) let ?db = "\ (d j0) + ((\i \ (d j0)" using Suc.prems card_d by (simp add: cInf_le_finite finite_enumerate_in_set) then have "?db = enum (d j0) (Suc j)" using Suc.prems that by (simp add: cInf_le_finite finite_enumerate_in_set sum_card_b card_b enum_d_0 \j0 < enum K j\ less_or_eq_imp_le) then have "?db \ d j0" using Suc.prems finite_enumerate_in_set by (auto simp: finite_enumerate_in_set) moreover have "list.set (acc_lengths w (seqs j0 j K)) \ (+) w ` d j0" by (simp add: Suc Suc_leD) then have "list.set (acc_lengths (w + \ (d j0)) (map (list_of \ (\i. b (enum K i) (j0,i))) (list_of {.. (+) w ` d j0" by (simp add: seqs_def card_a subset_insertI) ultimately show ?case by (simp add: seqs_def acc_lengths_append image_iff Inf_nat_def1 sum_sorted_list_of_set_map card_a) qed then have acc_lengths_subset_d: "list.set (acc_lengths 0 (seqs j0 j K)) \ d j0" if K: "K \ {j0<..}" "finite K" "card K \ j0" and "j \ j0" for j0 j K by (metis image_add_0 that) have "strict_sorted x" "strict_sorted y" by (auto simp: xeq yeq BB_def) have disjnt_xy: "disjnt (list.set x) (list.set y)" proof - have "disjnt (a j) (a r)" using \j < r\ aa less_sets_imp_disjnt by blast moreover have "disjnt (b (enum K i) (j,i)) (a r)" if "i < j" for i by (simp add: disjnt_ba enumK less_imp_le_nat that) moreover have "disjnt (a j) (b (enum L q) (r,q))" if "q < r" for q by (meson disjnt_ba disjnt_sym enumL less_imp_le_nat that) moreover have "disjnt (b (enum K i) (j,i)) (b (enum L q) (r,q))" if "i < j" "q < r" for i q by (meson \j < r\ bb_disjnt enumK enumL less_imp_le that) ultimately show ?thesis by (simp add: xeq yeq BB_def) qed have "\us vs. merge (seqs j j K) (seqs r r L) us vs" proof (rule merge_exists) show "strict_sorted (concat (seqs j j K))" using BB_eq_concat_seqs K \strict_sorted x\ xeq by auto show "strict_sorted (concat (seqs r r L))" using BB_eq_concat_seqs L \strict_sorted y\ yeq by auto show "seqs j j K \ lists (- {[]})" "seqs r r L \ lists (- {[]})" by (auto simp: K L seqs_ne) show "hd (seqs j j K) < hd (seqs r r L)" by (simp add: aj_ar less_sets_imp_list_less seqs_def) show "seqs j j K \ []" "seqs r r L \ []" using seqs_def by blast+ have less_bb: "b (enum K i) (j,i) \ b (enum L p) (r, p)" if "\ b (enum L p) (r, p) \ b (enum K i) (j,i)" and "i < j" "p < r" for i p by (metis IJ_iff \j < r\ bb bb_same enumK enumL less_imp_le_nat linorder_neqE_nat pair_lessI1 that) show "u < v \ v < u" if "u \ list.set (seqs j j K)" and "v \ list.set (seqs r r L)" for u v using that enumK enumL unfolding seqs_def apply (auto simp: seqs_def aj_ar intro!: less_bb less_sets_imp_list_less) apply (meson ab ba less_imp_le_nat not_le)+ done qed then obtain uus vvs where merge: "merge (seqs j j K) (seqs r r L) uus vvs" by metis then have "uus \ []" using merge_length1_gt_0 by (auto simp: seqs_def) then obtain u1 us where us: "u1#us = uus" by (metis neq_Nil_conv) define ku where "ku \ length (u1#us)" define ps where "ps \ acc_lengths 0 (u1#us)" have us_ne: "u1#us \ lists (- {[]})" using merge_length1_nonempty seqs_ne us merge us K by auto have xu_eq: "x = concat (u1#us)" using BB_eq_concat_seqs K merge merge_preserves us xeq by auto then have "strict_sorted u1" using \strict_sorted x\ strict_sorted_append_iff by auto have u_sub: "list.set ps \ list.set (acc_lengths 0 (seqs j j K))" using acc_lengths_merge1 merge ps_def us by blast have "vvs \ []" using merge BB_eq_concat_seqs L merge_preserves xy yeq by auto then obtain v1 vs where vs: "v1#vs = vvs" by (metis neq_Nil_conv) define kv where "kv \ length (v1#vs)" define qs where "qs \ acc_lengths 0 (v1#vs)" have vs_ne: "v1#vs \ lists (- {[]})" using L merge merge_length2_nonempty seqs_ne vs by auto have yv_eq: "y = concat (v1#vs)" using BB_eq_concat_seqs L merge merge_preserves vs yeq by auto then have "strict_sorted v1" using \strict_sorted y\ strict_sorted_append_iff by auto have v_sub: "list.set qs \ list.set (acc_lengths 0 (seqs r r L))" using acc_lengths_merge2 merge qs_def vs by blast have ss_concat_jj: "strict_sorted (concat (seqs j j K))" using BB_eq_concat_seqs K \strict_sorted x\ xeq by auto then obtain k: "0 < kv" "kv \ ku" "ku \ Suc kv" "kv \ Suc j" using us vs merge_length_le merge_length_le_Suc merge_length_less2 merge unfolding ku_def kv_def by fastforce define zs where "zs \ concat [ps,u1,qs,v1] @ interact us vs" have ss: "strict_sorted zs" proof - have ssp: "strict_sorted ps" unfolding ps_def by (meson strict_sorted_acc_lengths us_ne) have ssq: "strict_sorted qs" unfolding qs_def by (meson strict_sorted_acc_lengths vs_ne) have "d j \ list.set x" using da [of j] db [of j] K \_card \_enum nat_less_le by (auto simp: xeq BB_def less_sets_Un2 less_sets_UN2) then have ac_x: "acc_lengths 0 (seqs j j K) < x" by (meson Ksub \finite K\ \j \ card K\ acc_lengths_subset_d le_refl less_sets_imp_list_less less_sets_weaken1) then have "ps < x" by (meson Ksub \d j \ list.set x\ \finite K\ \j \ card K\ acc_lengths_subset_d le_refl less_sets_imp_list_less less_sets_weaken1 u_sub) then have "ps < u1" by (metis Nil_is_append_conv concat.simps(2) hd_append2 less_list_def xu_eq) have "d r \ list.set y" using da [of r] db [of r] L \_card \_enum nat_less_le by (auto simp: yeq BB_def less_sets_Un2 less_sets_UN2) then have "acc_lengths 0 (seqs r r L) < y" by (meson Lsub \finite L\ \r \ card L\ acc_lengths_subset_d le_refl less_sets_imp_list_less less_sets_weaken1) then have "qs < y" by (metis L Lsub \_card \d r \ list.set y\ \finite L\ acc_lengths_subset_d less_sets_imp_list_less less_sets_weaken1 order_refl v_sub) then have "qs < v1" by (metis concat.simps(2) gr_implies_not0 hd_append2 less_list_def list.size(3) xy yv_eq) have carda_v1: "card (a r) \ length v1" using length_hd_merge2 [OF merge] unfolding vs [symmetric] by (simp add: seqs_def) have ab_enumK: "\i. i < j \ a j \ b (enum K i) (j,i)" by (meson ab enumK le_trans less_imp_le_nat) have ab_enumL: "\q. q < r \ a j \ b (enum L q) (r,q)" by (meson \j < r\ ab enumL le_trans less_imp_le_nat) then have ay: "a j \ list.set y" by (auto simp: yeq BB_def less_sets_Un2 less_sets_UN2 aj_ar) have disjnt_hd_last_K_y: "disjnt {hd l..last l} (list.set y)" if l: "l \ list.set (seqs j j K)" for l proof (clarsimp simp add: yeq BB_def disjnt_iff Ball_def, intro conjI strip) fix u assume u: "u \ last l" and "hd l \ u" with l consider "u \ last (list_of (a j))" "hd (list_of (a j)) \ u" | i where "i last (list_of (b (enum K i) (j,i)))" "hd (list_of (b (enum K i) (j,i))) \ u" by (force simp: seqs_def) note l_cases = this then show "u \ a r" proof cases case 1 then show ?thesis by (metis a_ne aj_ar finite_a last_in_set leD less_setsD set_sorted_list_of_set sorted_list_of_set_eq_Nil_iff) next case 2 then show ?thesis by (metis enumK ab ba Inf_nat_def1 b_ne card_b_finite hd_b last_in_set less_asym less_setsD not_le set_sorted_list_of_set sorted_list_of_set_eq_Nil_iff) qed fix q assume "q < r" show "u \ b (enum L q) (r,q)" using l_cases proof cases case 1 then show ?thesis by (metis \q < r\ a_ne ab_enumL finite_a last_in_set leD less_setsD set_sorted_list_of_set sorted_list_of_set_eq_Nil_iff) next case 2 show ?thesis proof (cases "enum K i = enum L q") case True then show ?thesis using 2 bb_same [of concl: "enum L q" j i r q] \j < r\ u by (metis IJ_iff b_ne card_b_finite enumK last_in_set leD less_imp_le_nat less_setsD pair_lessI1 set_sorted_list_of_set sorted_list_of_set_eq_Nil_iff) next case False with 2 bb enumK enumL show ?thesis unfolding less_sets_def by (metis \q < r\ b_ne card_b_finite last_in_set leD less_imp_le_nat list.set_sel(1) nat_neq_iff set_sorted_list_of_set sorted_list_of_set_eq_Nil_iff) qed qed qed have u1_y: "list.set u1 \ list.set y" using vs yv_eq L \strict_sorted y\ merge merge_less_sets_hd merge_preserves seqs_ne ss_concat_jj us by fastforce have u1_subset_seqs: "list.set u1 \ list.set (concat (seqs j j K))" using merge_preserves [OF merge] us by auto have "b k (j,i) \ d (Suc k)" if "j\k" "i d k'" if "j\k" "i d (Suc k)" for k by (metis aM d_eq less_sets_fst_grab less_sets_weaken2 nxt_subset) then have ad: "a k \ d k'" if "k list.set u1" for n proof - obtain l where l: "l \ list.set (seqs j j K)" and n: "n \ list.set l" using n u1_subset_seqs by auto then consider "l = list_of (a j)" | i where "l = list_of (b (enum K i) (j,i))" "i < j" by (force simp: seqs_def) then show ?thesis proof cases case 1 then show ?thesis by (metis Inf_nat_def1 \j < r\ ad d_ne finite_a less_setsD n set_sorted_list_of_set) next case 2 then have "hd (list_of (b (enum K i) (j,i))) = Min (b (enum K i) (j,i))" by (meson b_ne card_b_finite enumK hd_list_of less_imp_le_nat) also have "\ \ n" using 2 n by (simp add: less_list_def disjnt_iff less_sets_def) also have f8: "n < hd y" using less_setsD that u1_y by (metis gr_implies_not0 list.set_sel(1) list.size(3) xy) finally have "l < y" using 2 disjnt_hd_last_K_y [OF l] by (simp add: disjnt_iff) (metis leI less_imp_le_nat less_list_def list.set_sel(1)) moreover have "last (list_of (b (enum K i) (j,i))) < hd (list_of (a r))" using \l < y\ L n by (auto simp: 2yeq BB_eq_concat_seqs seqs_def less_list_def) then have "enum K i < r" by (metis "2"(1) a_ne ab card_b_finite empty_iff finite.emptyI finite_a last_in_set leI less_asym less_setsD list.set_sel(1) n set_sorted_list_of_set) moreover have "j \ enum K i" by (simp add: "2"(2) enumK less_imp_le_nat) ultimately show ?thesis using 2 n bd [of j "enum K i" i r] Inf_nat_def1 less_setsD by force qed qed then have "last u1 < Inf (d r)" using \uus \ []\ us_ne by auto also have "\ \ length v1" using card_a carda_v1 by auto finally have "last u1 < length v1" . then have "u1 < qs" by (simp add: qs_def less_list_def) have "strict_sorted (interact (u1#us) (v1#vs))" using L \strict_sorted x\ \strict_sorted y\ merge merge_interact merge_preserves seqs_ne us vs xu_eq yv_eq by auto then have "strict_sorted (interact us vs)" "v1 < interact us vs" by (auto simp: strict_sorted_append_iff) moreover have "ps < u1 @ qs @ v1 @ interact us vs" using \ps < u1\ us_ne unfolding less_list_def by auto moreover have "u1 < qs @ v1 @ interact us vs" by (metis \u1 < qs\ \vvs \ []\ acc_lengths_eq_Nil_iff hd_append less_list_def qs_def vs) moreover have "qs < v1 @ interact us vs" using \qs < v1\ us_ne \last u1 < length v1\ vs_ne by (auto simp: less_list_def) ultimately show ?thesis by (simp add: zs_def strict_sorted_append_iff ssp ssq \strict_sorted u1\ \strict_sorted v1\) qed have ps_subset_d: "list.set ps \ d j" using K Ksub \_card \finite K\ acc_lengths_subset_d u_sub by blast have ps_less_u1: "ps < u1" proof - have "hd u1 = hd x" using us_ne by (auto simp: xu_eq) then have "hd u1 \ a j" by (simp add: xeq BB_eq_concat_seqs K seqs_def hd_append hd_list_of) then have "list.set ps \ {hd u1}" by (metis da ps_subset_d less_sets_def singletonD subset_iff) then show ?thesis by (metis less_hd_imp_less list.set(2) empty_set less_sets_imp_list_less) qed have qs_subset_d: "list.set qs \ d r" using L Lsub \_card \finite L\ acc_lengths_subset_d v_sub by blast have qs_less_v1: "qs < v1" proof - have "hd v1 = hd y" using vs_ne by (auto simp: yv_eq) then have "hd v1 \ a r" by (simp add: yeq BB_eq_concat_seqs L seqs_def hd_append hd_list_of) then have "list.set qs \ {hd v1}" by (metis da qs_subset_d less_sets_def singletonD subset_iff) then show ?thesis by (metis less_hd_imp_less list.set(2) empty_set less_sets_imp_list_less) qed have FB: "Form_Body ku kv x y zs" unfolding Form_Body.simps using ku_def kv_def ps_def qs_def ss us_ne vs_ne xu_eq xy yv_eq zs_def by blast then have "zs = (inter_scheme ((ku+kv) - Suc 0) {x,y})" by (simp add: Form_Body_imp_inter_scheme k) obtain l where "l \ 2 * (Suc j)" and l: "Form l U" and zs_eq_interact: "zs = inter_scheme l {x,y}" proof show "ku+kv-1 \ 2 * (Suc j)" using k by auto show "Form (ku+kv-1) U" proof (cases "ku=kv") case True then show ?thesis using FB Form.simps Ueq \0 < kv\ by (auto simp: mult_2) next case False then have "ku = Suc kv" using k by auto then show ?thesis using FB Form.simps Ueq \0 < kv\ by auto qed show "zs = inter_scheme (ku + kv - 1) {x, y}" using Form_Body_imp_inter_scheme by (simp add: FB k) qed then have "enum N l \ enum N (Suc (2 * Suc j))" by (simp add: assms less_imp_le_nat) also have "\ < Min (d j)" by (smt (verit, best) Min_gr_iff d_eq d_ne finite_d fst_grab_subset greaterThan_iff in_mono le_inf_iff nxt_def) finally have ls: "{enum N l} \ d j" by simp have "l > 0" by (metis l False Form_0_cases_raw Set.doubleton_eq_iff Ueq gr0I) show ?thesis unfolding Ueq proof (intro exI conjI impI) have zs_subset: "list.set zs \ list.set (acc_lengths 0 (seqs j j K)) \ list.set (acc_lengths 0 (seqs r r L)) \ list.set x \ list.set y" using u_sub v_sub by (auto simp: zs_def xu_eq yv_eq) also have "\ \ N" proof (simp, intro conjI) show "list.set (acc_lengths 0 (seqs j j K)) \ N" using d_subset_N Ksub \finite K\ \j \ card K\ acc_lengths_subset_d by blast show "list.set (acc_lengths 0 (seqs r r L)) \ N" using d_subset_N Lsub \finite L\ \r \ card L\ acc_lengths_subset_d by blast show "list.set x \ N" "list.set y \ N" by (simp_all add: xeq yeq BB_def a_subset_N UN_least b_subset_N) qed finally show "list.set (inter_scheme l {x, y}) \ N" using zs_eq_interact by blast have "[enum N l] < ps" using ps_subset_d ls by (metis empty_set less_sets_imp_list_less less_sets_weaken2 list.simps(15)) then show "[enum N l] < inter_scheme l {x, y}" by (simp add: zs_def less_list_def ps_def flip: zs_eq_interact) qed (use Ueq l in blast) qed qed qed subsection \The main partition theorem for @{term "\\\"}\ definition iso_ll where "iso_ll A B \ iso (lenlex less_than \ (A\A)) (lenlex less_than \ (B\B))" corollary ordertype_eq_ordertype_iso_ll: assumes "Field (Restr (lenlex less_than) A) = A" "Field (Restr (lenlex less_than) B) = B" shows "(ordertype A (lenlex less_than) = ordertype B (lenlex less_than)) \ (\f. iso_ll A B f)" proof - have "total_on A (lenlex less_than) \ total_on B (lenlex less_than)" by (meson UNIV_I total_lenlex total_on_def total_on_less_than) then show ?thesis by (simp add: assms wf_lenlex lenlex_transI iso_ll_def ordertype_eq_ordertype_iso_Restr) qed theorem partition_\\_aux: assumes "\ \ elts \" shows "partn_lst (lenlex less_than) WW [\\\,\] 2" (is "partn_lst ?R WW [\\\,\] 2") proof (cases "\ \ 1") case True then show ?thesis using strict_sorted_into_WW unfolding WW_def by (auto intro!: partn_lst_triv1[where i=1]) next case False obtain m where m: "\ = ord_of_nat m" using assms elts_\ by auto then have "m>1" using False by auto show ?thesis unfolding partn_lst_def proof clarsimp fix f assume f: "f \ [WW]\<^bsup>2\<^esup> \ {..: "?P0 \ ?P1" proof (rule disjCI) assume not1: "\ ?P1" have "\W'. ordertype W' ?R = \\n \ f ` [W']\<^bsup>2\<^esup> \ {0} \ W' \ WW_seg (n*m)" for n::nat proof - have fnm: "f \ [WW_seg (n*m)]\<^bsup>2\<^esup> \ {..\n, ord_of_nat m] 2" using ordertype_WW_seg [of "n*m"] by (simp add: partn_lst_VWF_imp_partn_lst [OF Theorem_3_2]) show ?thesis using partn_lst_E [OF * fnm, simplified] by (metis One_nat_def WW_seg_subset_WW less_2_cases m not1 nth_Cons_0 nth_Cons_Suc numeral_2_eq_2 subset_trans) qed then obtain W':: "nat \ nat list set" where otW': "\n. ordertype (W' n) ?R = \\n" and f_W': "\n. f ` [W' n]\<^bsup>2\<^esup> \ {0}" and seg_W': "\n. W' n \ WW_seg (n*m)" by metis define WW' where "WW' \ (\n. W' n)" have "WW' \ WW" using seg_W' WW_seg_subset_WW by (force simp: WW'_def) with f have f': "f \ [WW']\<^bsup>2\<^esup> \ {..\\" proof (rule antisym) have "ordertype WW' ?R \ ordertype WW ?R" by (simp add: \WW' \ WW\ lenlex_transI ordertype_mono wf_lenlex) with ordertype_WW show "ordertype WW' ?R \ \ \ \" by simp have "\ \ n \ ordertype (\ (range W')) ?R" for n::nat using oexp_Limit ordertype_\\ otW' by auto then show "\ \ \ \ ordertype WW' ?R" by (auto simp: elts_\ oexp_Limit ZFC_in_HOL.SUP_le_iff WW'_def) qed have FR_WW: "Field (Restr (lenlex less_than) WW) = WW" by (simp add: Limit_omega_oexp Limit_ordertype_imp_Field_Restr ordertype_WW) have FR_WW': "Field (Restr (lenlex less_than) WW') = WW'" by (simp add: Limit_omega_oexp Limit_ordertype_imp_Field_Restr ot') have FR_W: "Field (Restr (lenlex less_than) (WW_seg n)) = WW_seg n" if "n>0" for n by (simp add: Limit_omega_oexp ordertype_WW_seg that Limit_ordertype_imp_Field_Restr) have FR_W': "Field (Restr (lenlex less_than) (W' n)) = W' n" if "n>0" for n by (simp add: Limit_omega_oexp otW' that Limit_ordertype_imp_Field_Restr) have "\h. iso_ll (WW_seg n) (W' n) h" if "n>0" for n proof (subst ordertype_eq_ordertype_iso_ll [symmetric]) show "ordertype (WW_seg n) (lenlex less_than) = ordertype (W' n) (lenlex less_than)" by (simp add: ordertype_WW_seg otW') qed (auto simp: FR_W FR_W' that) then obtain h_seg where h_seg: "\n. n > 0 \ iso_ll (WW_seg n) (W' n) (h_seg n)" by metis define h where "h \ \l. if l=[] then [] else h_seg (length l) l" have bij_h_seg: "\n. n > 0 \ bij_betw (h_seg n) (WW_seg n) (W' n)" using h_seg by (simp add: iso_ll_def iso_iff2 FR_W FR_W') have len_h_seg: "length (h_seg (length l) l) = length l * m" if "length l > 0" "l \ WW" for l using bij_betwE [OF bij_h_seg] seg_W' that by (simp add: WW_seg_def subset_iff) have hlen: "length (h x) = length (h y) \ length x = length y" if "x \ WW" "y \ WW" for x y using that \1 < m\ h_def len_h_seg by force have h: "iso_ll WW WW' h" unfolding iso_ll_def iso_iff2 FR_WW FR_WW' proof (intro conjI strip) have W'_ne: "W' n \ {}" for n using otW' [of n] by auto then have "[] \ WW'" using seg_W' [of 0] by (auto simp: WW'_def WW_seg_def) let ?g = "\l. if l=[] then l else inv_into (WW_seg (length l div m)) (h_seg (length l div m)) l" have h_seg_iff: "\n a b. \a \ WW_seg n; b \ WW_seg n; n>0\ \ (a, b) \ lenlex less_than \ (h_seg n a, h_seg n b) \ lenlex less_than \ h_seg n a \ W' n \ h_seg n b \ W' n" using h_seg by (auto simp: iso_ll_def iso_iff2 FR_W FR_W') show "bij_betw h WW WW'" unfolding bij_betw_iff_bijections proof (intro exI conjI ballI) fix l assume "l \ WW" then have l: "l \ WW_seg (length l)" by (simp add: WW_seg_def) have "h l \ W' (length l)" proof (cases "l=[]") case True with seg_W' [of 0] W'_ne show ?thesis by (auto simp: WW_seg_def h_def) next case False then show ?thesis using bij_betwE bij_h_seg h_def l by fastforce qed show "h l \ WW'" using WW'_def \h l \ W' (length l)\ by blast show "?g (h l) = l" proof (cases "l=[]") case False then have "length l > 0" by auto then have "h_seg (length l) l \ []" using \1 < m\ \l \ WW\ len_h_seg by fastforce moreover have "bij_betw (h_seg (length l)) (WW_seg (length l)) (W' (length l))" using \0 < length l\ bij_h_seg by presburger ultimately show ?thesis using \l \ WW\ bij_betw_inv_into_left h_def l len_h_seg by fastforce qed (auto simp: h_def) next fix l assume "l \ WW'" then have l: "l \ W' (length l div m)" using WW_seg_def \1 < m\ seg_W' by (fastforce simp: WW'_def) show "?g l \ WW" proof (cases "l=[]") case False then have "l \ W' 0" using WW_seg_def seg_W' by fastforce with l have "inv_into (WW_seg (length l div m)) (h_seg (length l div m)) l \ WW_seg (length l div m)" by (metis Nat.neq0_conv bij_betwE bij_betw_inv_into bij_h_seg) then show ?thesis using False WW_seg_subset_WW by auto qed (auto simp: WW_def) show "h (?g l) = l" proof (cases "l=[]") case False then have "0 < length l div m" using WW_seg_def l seg_W' by fastforce then have "inv_into (WW_seg (length l div m)) (h_seg (length l div m)) l \ WW_seg (length l div m)" by (metis bij_betw_imp_surj_on bij_h_seg inv_into_into l) then show ?thesis using bij_h_seg [of "length l div m"] WW_seg_def \0 < length l div m\ bij_betw_inv_into_right l by (fastforce simp: h_def) qed (auto simp: h_def) qed fix a b assume "a \ WW" "b \ WW" show "(a, b) \ Restr (lenlex less_than) WW \ (h a, h b) \ Restr (lenlex less_than) WW'" (is "?lhs = ?rhs") proof assume L: ?lhs then consider "length a < length b" | "length a = length b" "(a, b) \ lex less_than" by (auto simp: lenlex_conv) then show ?rhs proof cases case 1 then have "length (h a) < length (h b)" using \1 < m\ \a \ WW\ \b \ WW\ h_def len_h_seg by auto then have "(h a, h b) \ lenlex less_than" by (auto simp: lenlex_conv) then show ?thesis using \a \ WW\ \b \ WW\ \bij_betw h WW WW'\ bij_betwE by fastforce next case 2 then have ab: "a \ WW_seg (length a)" "b \ WW_seg (length a)" using \a \ WW\ \b \ WW\ by (auto simp: WW_seg_def) have "length (h a) = length (h b)" using 2 \a \ WW\ \b \ WW\ h_def len_h_seg by force moreover have "(a, b) \ lenlex less_than" using L by blast then have "(h_seg (length a) a, h_seg (length a) b) \ lenlex less_than" using 2 ab h_seg_iff by blast ultimately show ?thesis using 2 \a \ WW\ \b \ WW\ \bij_betw h WW WW'\ bij_betwE h_def by fastforce qed next assume R: ?rhs then have R': "(h a, h b) \ lenlex less_than" by blast then consider "length a < length b" | "length a = length b" "(h a, h b) \ lex less_than" using \a \ WW\ \b \ WW\ \m > 1\ by (auto simp: lenlex_conv h_def len_h_seg split: if_split_asm) then show ?lhs proof cases case 1 then show ?thesis using omega_sum_less_iff \a \ WW\ \b \ WW\ by auto next case 2 then have ab: "a \ WW_seg (length a)" "b \ WW_seg (length a)" using \a \ WW\ \b \ WW\ by (auto simp: WW_seg_def) then have "(a, b) \ lenlex less_than" using bij_betwE [OF bij_h_seg] \a \ WW\ \b \ WW\ R' 2 by (simp add: h_def h_seg_iff split: if_split_asm) then show ?thesis using \a \ WW\ \b \ WW\ by blast qed qed qed let ?fh = "f \ image h" have "bij_betw h WW WW'" using h unfolding iso_ll_def iso_iff2 by (fastforce simp: FR_WW FR_WW') moreover have "{.. [WW]\<^bsup>2\<^esup> \ {0,1}" unfolding Pi_iff using bij_betwE f' bij_betw_nsets by (metis PiE comp_apply) have "f{x,y} = 0" if "x \ WW'" "y \ WW'" "length x = length y" "x \ y" for x y proof - obtain p q where "x \ W' p" and "y \ W' q" using WW'_def \x \ WW'\ \y \ WW'\ by blast then obtain n where "{x,y} \ [W' n]\<^bsup>2\<^esup>" using seg_W' \1 < m\ \length x = length y\ \x \ y\ by (auto simp: WW'_def WW_seg_def subset_iff) then show "f{x,y} = 0" using f_W' by blast qed then have fh_eq_0_eqlen: "?fh{x,y} = 0" if "x \ WW" "y \ WW" "length x = length y" "x\y" for x y using \bij_betw h WW WW'\ that hlen by (simp add: bij_betw_iff_bijections) metis have m_f_0: "\x\[M]\<^bsup>2\<^esup>. f x = 0" if "M \ WW" "card M = m" for M proof - have "finite M" using False m that by auto with not1 [simplified, rule_format, of M] f show ?thesis using that \1 < m\ apply (simp add: Pi_iff image_subset_iff finite_ordertype_eq_card m) by (metis less_2_cases nsets_mono numeral_2_eq_2 subset_iff) qed have m_fh_0: "\x\[M]\<^bsup>2\<^esup>. ?fh x = 0" if "M \ WW" "card M = m" for M proof - have "h ` M \ WW" using \WW' \ WW\ \bij_betw h WW WW'\ bij_betwE that(1) by fastforce moreover have "card (h ` M) = m" by (metis \bij_betw h WW WW'\ bij_betw_def bij_betw_subset card_image that) ultimately have "\x \ [h ` M]\<^bsup>2\<^esup>. f x = 0" by (metis m_f_0) then obtain Y where Y: "f (h ` Y) = 0" "Y \ M" and "finite (h ` Y)" "card (h ` Y) = 2" by (auto simp: nsets_def subset_image_iff) then have "card Y = 2" using \bij_betw h WW WW'\ \M \ WW\ by (metis bij_betw_def card_image inj_on_subset) with Y card.infinite[of Y] show ?thesis by (auto simp: nsets_def) qed obtain N j where "infinite N" and N: "\k u. \k > 0; u \ [WW]\<^bsup>2\<^esup>; Form k u; [enum N k] < inter_scheme k u; List.set (inter_scheme k u) \ N\ \ ?fh u = j k" using lemma_3_6 [OF fh] by blast have infN': "infinite (enum N ` {k<..})" for k by (simp add: \infinite N\ enum_works finite_image_iff infinite_Ioi strict_mono_imp_inj_on) have j_0: "j k = 0" if "k>0" for k proof - obtain M where M: "M \ [WW]\<^bsup>m\<^esup>" and MF: "\u. u \ [M]\<^bsup>2\<^esup> \ Form k u" and Mi: "\u. u \ [M]\<^bsup>2\<^esup> \ List.set (inter_scheme k u) \ enum N ` {k<..}" using lemma_3_7 [OF infN' \k > 0\] by metis obtain u where u: "u \ [M]\<^bsup>2\<^esup>" "?fh u = 0" using m_fh_0 [of M] M [unfolded nsets_def] by force moreover have \
: "Form k u" "List.set (inter_scheme k u) \ enum N ` {k<..}" by (simp_all add: MF Mi \u \ [M]\<^bsup>2\<^esup>\) then have "hd (inter_scheme k u) \ enum N ` {k<..}" using hd_in_set inter_scheme_simple that by blast then have "[enum N k] < inter_scheme k u" using strict_mono_enum [OF \infinite N\] by (auto simp: less_list_def strict_mono_def) moreover have "u \ [WW]\<^bsup>2\<^esup>" using M u by (auto simp: nsets_def) moreover have "enum N ` {k<..} \ N" using \infinite N\ range_enum by auto ultimately show ?thesis using N \
that by auto qed obtain X where "X \ WW" and otX: "ordertype X (lenlex less_than) = \\\" and X: "\u. u \ [X]\<^bsup>2\<^esup> \ \l. Form l u \ (l > 0 \ [enum N l] < inter_scheme l u \ List.set (inter_scheme l u) \ N)" using lemma_3_8 [OF \infinite N\] ot' by blast have 0: "?fh ` [X]\<^bsup>2\<^esup> \ {0}" proof clarsimp fix u assume u: "u \ [X]\<^bsup>2\<^esup>" obtain l where "Form l u" and l: "l > 0 \ [enum N l] < inter_scheme l u \ List.set (inter_scheme l u) \ N" using u X by blast have "?fh u = 0" proof (cases "l = 0") case True then show ?thesis by (metis Form_0_cases_raw \Form l u\ \X \ WW\ doubleton_in_nsets_2 fh_eq_0_eqlen subset_iff u) next case False then obtain "[enum N l] < inter_scheme l u" "List.set (inter_scheme l u) \ N" "j l = 0" using Nat.neq0_conv j_0 l by blast with False show ?thesis using \X \ WW\ N inter_scheme \Form l u\ doubleton_in_nsets_2 u by (auto simp: nsets_def) qed then show "f (h ` u) = 0" by auto qed show ?P0 proof (intro exI conjI) show "h ` X \ WW" using \WW' \ WW\ \X \ WW\ \bij_betw h WW WW'\ bij_betw_imp_surj_on by fastforce show "ordertype (h ` X) (lenlex less_than) = \ \ \" proof (subst ordertype_inc_eq) show "(h x, h y) \ lenlex less_than" if "x \ X" "y \ X" "(x, y) \ lenlex less_than" for x y using that h \X \ WW\ by (auto simp: FR_WW FR_WW' iso_iff2 iso_ll_def) qed (use otX in auto) show "f ` [h ` X]\<^bsup>2\<^esup> \ {0}" proof (clarsimp simp: image_subset_iff nsets_def) fix Y assume "Y \ h ` X" and "finite Y" and "card Y = 2" have "inv_into WW h ` Y \ X" using \X \ WW\ \Y \ h ` X\ \bij_betw h WW WW'\ bij_betw_inv_into_LEFT by blast moreover have "finite (inv_into WW h ` Y)" using \finite Y\ by blast moreover have "card (inv_into WW h ` Y) = 2" by (metis \X \ WW\ \Y \ h ` X\ \card Y = 2\ card_image inj_on_inv_into subset_image_iff subset_trans) ultimately have "f (h ` inv_into WW h ` Y) = 0" using 0 by (auto simp: image_subset_iff nsets_def) then show "f Y = 0" by (metis \X \ WW\ \Y \ h ` X\ image_inv_into_cancel image_mono order_trans) qed qed qed then show "\iH\WW. ordertype H ?R = [\\\, \] ! i \ f ` [H]\<^bsup>2\<^esup> \ {i}" by (metis One_nat_def lessI nth_Cons_0 nth_Cons_Suc zero_less_Suc) qed qed text \Theorem 3.1 of Jean A. Larson, ibid.\ theorem partition_\\: "\ \ elts \ \ partn_lst_VWF (\\\) [\\\,\] 2" using partn_lst_imp_partn_lst_VWF_eq [OF partition_\\_aux] ordertype_WW by auto end diff --git a/thys/Progress_Tracking/Combined.thy b/thys/Progress_Tracking/Combined.thy --- a/thys/Progress_Tracking/Combined.thy +++ b/thys/Progress_Tracking/Combined.thy @@ -1,1221 +1,1225 @@ section \Combined Progress Tracking Protocol\label{sec:combined}\ (*<*) theory Combined imports Exchange Propagate begin (*>*) lemma fold_invar: assumes "finite M" and "P z" and "\z. \x\M. P z \ P (f x z)" and "comp_fun_commute f" shows "P (Finite_Set.fold f z M)" - using assms by (induct M arbitrary: z rule: finite_induct) (auto simp: comp_fun_commute.fold_insert) +proof - + interpret commute: comp_fun_commute f + by (fact \comp_fun_commute f\) + from assms show ?thesis + by (induct M arbitrary: z rule: finite_induct) auto +qed subsection\Could-result-in Relation\ context dataflow_topology begin definition cri_less_eq :: "('loc \ 't) \ ('loc \ 't) \ bool" ("_\\<^sub>p_" [51,51] 50) where "cri_less_eq = (\(loc1,t1) (loc2,t2). (\s. s \\<^sub>A path_summary loc1 loc2 \ results_in t1 s \ t2))" definition cri_less :: "('loc \ 't) \ ('loc \ 't) \ bool" ("_<\<^sub>p_" [51,51] 50) where "cri_less x y = (x \\<^sub>p y \ x \ y)" lemma cri_asym1: "x <\<^sub>p y \ \ y <\<^sub>p x" for x y apply (cases x; cases y) proof (rule ccontr) fix loc1 t1 loc2 t2 assume as: "\ (x <\<^sub>p y \ \ y <\<^sub>p x)" "x = (loc1, t1)" "y = (loc2, t2)" then have as1: "loc1 \ loc2" unfolding cri_less_def cri_less_eq_def by clarsimp (metis add.right_neutral order.antisym order.trans le_plus(2) results_in_mono(2) results_in_zero) from as obtain s1 where s1: "s1 \\<^sub>A path_summary loc1 loc2" "results_in t1 s1 \ t2" using cri_less_def cri_less_eq_def by auto then obtain path1 where path1: "flow.path loc1 loc2 path1" "s1 = flow.sum_path_weights path1" "path1 \ []" using as1 flow.path_weight_conv_path flow.path0E by blast from as obtain s2 where s2: "s2 \\<^sub>A path_summary loc2 loc1" "results_in t2 s2 \ t1" using cri_less_def cri_less_eq_def by auto then obtain path2 where path2: "flow.path loc2 loc1 path2" "s2 = flow.sum_path_weights path2" "path2 \ []" using as1 flow.path_weight_conv_path flow.path0E by blast with path1 have path_total: "flow.path loc1 loc1 (path1@path2)" using flow.path_trans path2(3) by blast then obtain s3 where s3: "s3 = flow.sum_path_weights (path1@path2)" by auto then have s3_sum: "s3 = followed_by s1 s2" using path1 path2 flow.sum_weights_append by auto have "\ t1 < results_in t1 s3" using s3_sum s1(2) s2(2) followed_by_summary by (metis le_less_trans less_le_not_le results_in_mono(1)) then show False using path_total no_zero_cycle s3 path1(3) path2(3) by blast qed lemma cri_asym2: "x <\<^sub>p y \ x \ y" by (simp add: cri_less_def) sublocale cri: order cri_less_eq cri_less apply standard subgoal using cri_asym1 cri_asym2 cri_less_def by blast subgoal for x unfolding cri_less_eq_def using flow.path_weight_refl results_in_zero by fastforce subgoal for x y z apply (cases x; cases y; cases z) proof - fix a b aa ba ab bb assume as: "x \\<^sub>p y" "y \\<^sub>p z" "x = (a, b)" "y = (aa, ba)" "z = (ab, bb)" then obtain s1 where s1: "s1 \\<^sub>A path_summary a aa" "results_in b s1 \ ba" using cri_less_eq_def by auto from as(2,4,5) obtain s2 where s2: "s2 \\<^sub>A path_summary aa ab" "results_in ba s2 \ bb" using cri_less_eq_def by auto with s1 obtain s3 where s3: "s3 \\<^sub>A path_summary a ab" "s3 \ followed_by s1 s2" using flow.path_weight_elem_trans by blast with s1 s2 have "results_in b s3 \ bb" proof - have "\s. results_in (results_in b s1) s \ results_in ba s" by (meson results_in_mono(1) s1(2)) then show ?thesis by (metis (no_types) results_in_mono(2) followed_by_summary order_trans s2(2) s3(2)) qed with as s3 show ?thesis unfolding cri_less_eq_def by blast qed using cri_asym1 cri_asym2 cri_less_def by blast lemma wf_cri: "wf {(l, l'). (l, t) <\<^sub>p (l', t)}" by (rule finite_acyclic_wf) (auto intro: cri.acyclicI_order[THEN acyclic_converse[THEN iffD1]]) end subsection\Specification\ subsubsection\Configuration\ record ('p::finite, 't::order, 'loc) configuration = exchange_config :: "('p, ('loc \ 't)) Exchange.configuration" prop_config :: "'p \ ('loc, 't) Propagate.configuration" init :: "'p \ bool" (* True = initialization finished *) type_synonym ('p, 't, 'loc) computation = "('p, 't, 'loc) configuration stream" context dataflow_topology begin definition the_cm where "the_cm c loc t n = (THE c'. next_change_multiplicity' c c' loc t n)" text\@{term the_cm} is not commutative in general, only if the necessary conditions hold. It can be converted to @{term apply_cm} for which we prove @{term comp_fun_commute}.\ definition apply_cm where "apply_cm c loc t n = (let new_pointstamps = (\loc'. (if loc' = loc then update_zmultiset (c_pts c loc') t n else c_pts c loc')) in c \ c_pts := new_pointstamps \ \ c_work := (\loc'. c_work c loc' + frontier_changes (new_pointstamps loc') (c_pts c loc'))\)" definition cm_all' where "cm_all' c0 \ = Finite_Set.fold (\(loc, t) c. apply_cm c loc t (zcount \ (loc,t))) c0 (set_zmset \)" definition cm_all where "cm_all c0 \ = Finite_Set.fold (\(loc, t) c. the_cm c loc t (zcount \ (loc,t))) c0 (set_zmset \)" definition "propagate_all c0 = while_option (\c. \loc. (c_work c loc) \ {#}\<^sub>z) (\c. SOME c'. \loc t. next_propagate' c c' loc t) c0" subsubsection\Initial state and state transitions\ definition InitConfig :: "('p::finite, 't::order, 'loc) configuration \ bool" where "InitConfig c = ((\p. init c p = False) \ cri.init_config (exchange_config c) \ (\p loc t. zcount (c_pts (prop_config c p) loc) t = zcount (c_glob (exchange_config c) p) (loc, t)) \ (\w. init_config (prop_config c w)))" definition NextPerformOp' :: "('p::finite, 't::order, 'loc) configuration \ ('p, 't, 'loc) configuration \ 'p \ ('loc \ 't) multiset \ ('p \ ('loc \ 't)) multiset \ ('loc \ 't) multiset \ bool" where "NextPerformOp' c0 c1 p \neg \mint_msg \mint_self = ( cri.next_performop' (exchange_config c0) (exchange_config c1) p \neg \mint_msg \mint_self \ unchanged prop_config c0 c1 \ unchanged init c0 c1)" abbreviation NextPerformOp where "NextPerformOp c0 c1 \ \p \neg \mint_msg \mint_self. NextPerformOp' c0 c1 p \neg \mint_msg \mint_self" definition NextRecvCap' :: "('p::finite, 't::order, 'loc) configuration \ ('p, 't, 'loc) configuration \ 'p \ 'loc \ 't \ bool" where "NextRecvCap' c0 c1 p t = ( cri.next_recvcap' (exchange_config c0) (exchange_config c1) p t \ unchanged prop_config c0 c1 \ unchanged init c0 c1)" abbreviation NextRecvCap where "NextRecvCap c0 c1 \ \p t. NextRecvCap' c0 c1 p t" definition NextSendUpd' :: "('p::finite, 't::order, 'loc) configuration \ ('p, 't, 'loc) configuration \ 'p \ ('loc \ 't) set \ bool" where "NextSendUpd' c0 c1 p tt = ( cri.next_sendupd' (exchange_config c0) (exchange_config c1) p tt \ unchanged prop_config c0 c1 \ unchanged init c0 c1)" abbreviation NextSendUpd where "NextSendUpd c0 c1 \ \p tt. NextSendUpd' c0 c1 p tt" definition NextRecvUpd' :: "('p::finite, 't::order, 'loc) configuration \ ('p, 't, 'loc) configuration \ 'p \ 'p \ bool" where "NextRecvUpd' c0 c1 p q = ( init c0 q \ \Once init is set we are guaranteed that the CM transitions' premises are satisfied\ \ cri.next_recvupd' (exchange_config c0) (exchange_config c1) p q \ unchanged init c0 c1 \ (\p'. prop_config c1 p' = (if p' = q then cm_all (prop_config c0 q) (hd (c_msg (exchange_config c0) p q)) else prop_config c0 p')))" abbreviation NextRecvUpd where "NextRecvUpd c0 c1 \ \p q. NextRecvUpd' c0 c1 p q" definition NextPropagate' :: "('p::finite, 't::order, 'loc) configuration \ ('p, 't, 'loc) configuration \ 'p \ bool" where "NextPropagate' c0 c1 p = ( unchanged exchange_config c0 c1 \ init c1 = (init c0)(p := True) \ (\p'. Some (prop_config c1 p') = (if p' = p then propagate_all (prop_config c0 p') else Some (prop_config c0 p'))))" abbreviation NextPropagate where "NextPropagate c0 c1 \ \p. NextPropagate' c0 c1 p" definition "Next'" where "Next' c0 c1 = (NextPerformOp c0 c1 \ NextSendUpd c0 c1 \ NextRecvUpd c0 c1 \ NextPropagate c0 c1 \ NextRecvCap c0 c1 \ c1 = c0)" abbreviation "Next" where "Next s \ Next' (shd s) (shd (stl s))" definition FullSpec :: "('p :: finite, 't :: order, 'loc) computation \ bool" where "FullSpec s = (holds InitConfig s \ alw Next s)" lemma NextPerformOpD: assumes "NextPerformOp' c0 c1 p \neg \mint_msg \mint_self" shows "cri.next_performop' (exchange_config c0) (exchange_config c1) p \neg \mint_msg \mint_self" "unchanged prop_config c0 c1" "unchanged init c0 c1" using assms unfolding NextPerformOp'_def by simp_all lemma NextSendUpdD: assumes "NextSendUpd' c0 c1 p tt" shows "cri.next_sendupd' (exchange_config c0) (exchange_config c1) p tt" "unchanged prop_config c0 c1" "unchanged init c0 c1" using assms unfolding NextSendUpd'_def by simp_all lemma NextRecvUpdD: assumes "NextRecvUpd' c0 c1 p q" shows "init c0 q" "cri.next_recvupd' (exchange_config c0) (exchange_config c1) p q" "unchanged init c0 c1" "(\p'. prop_config c1 p' = (if p' = q then cm_all (prop_config c0 q) (hd (c_msg (exchange_config c0) p q)) else prop_config c0 p'))" using assms unfolding NextRecvUpd'_def by simp_all lemma NextPropagateD: assumes "NextPropagate' c0 c1 p" shows "unchanged exchange_config c0 c1" "init c1 = (init c0)(p := True)" "(\p'. Some (prop_config c1 p') = (if p' = p then propagate_all (prop_config c0 p') else Some (prop_config c0 p')))" using assms unfolding NextPropagate'_def by simp_all lemma NextRecvCapD: assumes "NextRecvCap' c0 c1 p t" shows "cri.next_recvcap' (exchange_config c0) (exchange_config c1) p t" "unchanged prop_config c0 c1" "unchanged init c0 c1" using assms unfolding NextRecvCap'_def by simp_all subsection\Auxiliary Lemmas\ subsubsection\Auxiliary Lemmas for CM Conversion\ lemma apply_cm_is_cm: "\t'. t' \\<^sub>A frontier (c_imp c loc) \ t' \ t \ n \ 0 \ next_change_multiplicity' c (apply_cm c loc t n) loc t n" by (auto simp: next_change_multiplicity'_def apply_cm_def intro!: Propagate.configuration.equality) lemma update_zmultiset_commute: "update_zmultiset (update_zmultiset M t' n') t n = update_zmultiset (update_zmultiset M t n) t' n'" by transfer (auto simp: equiv_zmset_def split: if_splits) lemma apply_cm_commute: "apply_cm (apply_cm c loc t n) loc' t' n' = apply_cm (apply_cm c loc' t' n') loc t n" unfolding apply_cm_def Let_def by (auto intro!: Propagate.configuration.equality simp: update_zmultiset_commute) lemma comp_fun_commute_apply_cm[simp]: "comp_fun_commute (\(loc, t) c. apply_cm c loc t (f loc t))" by (auto intro!: Propagate.configuration.equality simp: update_zmultiset_commute comp_fun_commute_def o_def apply_cm_commute) lemma ex_cm_imp_conds: assumes "\c'. next_change_multiplicity' c c' loc t n" shows "\t'. t' \\<^sub>A frontier (c_imp c loc) \ t' \ t" "n \ 0" using assms by (auto simp: next_change_multiplicity'_def) lemma the_cm_eq_apply_cm: assumes "\c'. next_change_multiplicity' c c' loc t n" shows "the_cm c loc t n = apply_cm c loc t n" proof - from assms have cond: "\t'. t' \\<^sub>A frontier (c_imp c loc) \ t' \ t" "n \ 0" using ex_cm_imp_conds by blast+ show ?thesis unfolding the_cm_def apply (rule the1_equality) apply (rule next_change_multiplicity'_unique[OF cond(2,1)]) unfolding apply_cm_def next_change_multiplicity'_def using cond apply (auto intro!: Propagate.configuration.equality) done qed lemma apply_cm_preserves_cond: assumes "\(loc,t)\set_zmset \. \t'. t' \\<^sub>A frontier (c_imp c0 loc) \ t' \ t" shows "\(loc,t)\set_zmset \. \t'. t' \\<^sub>A frontier (c_imp (apply_cm c0 loc' t'' n) loc) \ t' \ t" using assms by (auto simp: apply_cm_def) lemma cm_all_eq_cm_all': assumes "\(loc,t)\set_zmset \. \t'. t' \\<^sub>A frontier (c_imp c0 loc) \ t' \ t" shows "cm_all c0 \ = cm_all' c0 \" unfolding cm_all_def cm_all'_def apply (rule fold_closed_eq[where B = "{c. \(loc,t)\set_zmset \. \t'. t' \\<^sub>A frontier (c_imp c loc) \ t' \ t}"]) subgoal for a \ apply (cases a) apply simp apply (rule the_cm_eq_apply_cm) apply (rule ex1_implies_ex) apply (rule next_change_multiplicity'_unique) apply auto done subgoal for a \ apply (cases a) apply simp apply (rule apply_cm_preserves_cond) apply auto done subgoal using assms by simp done lemma cm_eq_the_cm: assumes "next_change_multiplicity' c c' loc t n" shows "the_cm c loc t n = c'" proof - from assms have cond: "\u. u \\<^sub>A frontier (c_imp c loc) \ u \ t" "n \ 0" unfolding next_change_multiplicity'_def by auto then show ?thesis using next_change_multiplicity'_unique[OF cond(2,1)] assms the_cm_def by auto qed lemma zcount_ps_apply_cm: "zcount (c_pts (apply_cm c loc t n) loc') t' = zcount (c_pts c loc') t' + (if loc = loc' \ t = t' then n else 0)" by (auto simp: apply_cm_def zcount_update_zmultiset) lemma zcount_pointstamps_update: "zcount (c_pts (c\c_pts:=M\) loc) x = zcount (M loc) x" by auto lemma nop: "loc1 \ loc2 \ t1 \ t2 \ zcount (c_pts (apply_cm c loc2 t2 (zcount \ (loc2, t2))) loc1) t1 = zcount (c_pts c loc1) t1" apply (simp add: apply_cm_def) using zcount_update_zmultiset by (simp add: zcount_update_zmultiset) lemma fold_nop: "zcount (c_pts (Finite_Set.fold (\(loc', t') c. apply_cm c loc' t' (zcount \' (loc', t'))) c (set_zmset \ - {(loc, t)})) loc) t = zcount (c_pts c loc) t" proof - have "finite (set_zmset \- {(loc, t)})" using finite_set_zmset by blast then show ?thesis proof (induct "set_zmset \ - {(loc, t)}" arbitrary: \ rule: finite_induct) case empty then show ?case using fold_empty by auto next let ?f = "(\(loc', t') c. apply_cm c loc' t' (zcount \' (loc',t')))" case (insert x F) obtain loc' t' where x_pair: "x = (loc', t')" by (meson surj_pair) from insert have nonmember: "x \ (loc, t)" by auto then have finite_s: "finite F" using insert by auto - have commute: "comp_fun_commute ?f" + interpret commute: comp_fun_commute ?f by (simp add: comp_fun_commute_def comp_def apply_cm_commute) - with finite_s have step1: + from finite_s have step1: "Finite_Set.fold ?f c (insert x F) = ?f x (Finite_Set.fold ?f c F)" - by (metis (mono_tags, lifting) comp_fun_commute.fold_insert insert.hyps(1) insert.hyps(2)) + by (metis (mono_tags, lifting) commute.fold_insert insert.hyps(1) insert.hyps(2)) from nonmember have step2: "zcount (c_pts (?f x (Finite_Set.fold ?f c F)) loc) t = zcount (c_pts (Finite_Set.fold ?f c F) loc) t" using x_pair by (metis (mono_tags, lifting) case_prod_conv nop) with step1 and x_pair have final: "zcount (c_pts (Finite_Set.fold ?f c (insert x F)) loc) t = zcount (c_pts (Finite_Set.fold ?f c F) loc) t" by simp from insert(2,4) obtain \2 where \2: "\2 = filter_zmset (\y. y\x) \" by blast then have "F = set_zmset \2 - {(loc, t)}" proof - from \2 have *: "x \#\<^sub>z \2" by (simp add: not_in_iff_zmset) from insert(4) and nonmember have **: "x \#\<^sub>z \" by blast from \2 ** have ***: "\y. y \#\<^sub>z \ \ (y \#\<^sub>z \2 \ y = x)" by (metis (mono_tags, lifting) count_filter_zmset zcount_ne_zero_iff) with *** have "\y. (y \ set_zmset \ = (y \ (set_zmset \2 \ {x})))" by blast then have "set_zmset \ = set_zmset \2 \ {x}" by (auto simp add: set_eq_iff) with insert(2,3,4) * show ?thesis by (metis (mono_tags, lifting) Diff_insert Diff_insert2 Diff_insert_absorb Un_empty_right Un_insert_right) qed with final insert show ?case by metis qed qed lemma zcount_pointstamps_cm_all': "zcount (c_pts (cm_all' c \) loc) x = zcount (c_pts c loc) x + zcount \ (loc,x)" proof - let ?f = "(\(loc', t') c. apply_cm c loc' t' (zcount \ (loc',t')))" have ?thesis proof (cases "zcount \ (loc,x) = 0") case case_nonmember: True then have set_zmset\: "set_zmset \ - {(loc, x)} = set_zmset \" using zcount_eq_zero_iff by fastforce have "zcount (c_pts (cm_all' c \) loc) x = zcount (c_pts c loc) x" unfolding cm_all'_def apply (subst set_zmset\[symmetric]) apply (simp add: fold_nop) done with case_nonmember show ?thesis by auto next case case_member: False then have fold_rec: "Finite_Set.fold ?f c (set_zmset \) = ?f (loc, x) (Finite_Set.fold ?f c (set_zmset \ - {(loc, x)}))" proof - + interpret commute: comp_fun_commute "(\(loc, t) c. apply_cm c loc t (f loc t))" for f + by (fact comp_fun_commute_apply_cm) have "(loc, x) \#\<^sub>z \" by (meson case_member zcount_inI) then show ?thesis - using comp_fun_commute_apply_cm - apply (intro comp_fun_commute.fold_rec) - apply simp_all - done + by (intro commute.fold_rec) simp_all qed have "zcount (c_pts (Finite_Set.fold ?f c (set_zmset \ - {(loc, x)})) loc) x = zcount (c_pts c loc) x" by (simp add: fold_nop) then have "zcount (c_pts (Finite_Set.fold ?f c (set_zmset \)) loc) x = zcount (c_pts (?f (loc, x) c) loc) x" using fold_nop fold_rec by (simp add: zcount_ps_apply_cm) then show ?thesis by (simp add: zcount_ps_apply_cm cm_all'_def) qed then show ?thesis .. qed lemma implications_apply_cm[simp]: "c_imp (apply_cm c loc t n) = c_imp c" by (auto simp: apply_cm_def) lemma implications_cm_all[simp]: "c_imp (cm_all' c \) = c_imp c" unfolding cm_all'_def Let_def apply (rule fold_invar[OF finite_set_zmset]) apply auto done lemma lift_cm_inv_cm_all': assumes "(\c0 c1 loc t n. P c0 \ next_change_multiplicity' c0 c1 loc t n \ P c1)" and "\(loc,t)\#\<^sub>z\. \t'. t' \\<^sub>A frontier (c_imp c0 loc) \ t' \ t" and "P c0" shows "P (cm_all' c0 \)" proof - let ?cond_invar = "\c. \(loc, t)\#\<^sub>z\. \t'. t' \\<^sub>A frontier (c_imp c loc) \ t' \ t" let ?invar = "\c. ?cond_invar c \ P c" show ?thesis unfolding cm_all'_def apply (rule conjunct2[OF fold_invar[OF finite_set_zmset, of ?invar]]) using assms(2,3) apply simp subgoal apply safe apply auto [] apply (rule assms(1)) apply simp apply (rule apply_cm_is_cm) apply auto done apply simp done qed lemma lift_cm_inv_cm_all: assumes "\c0 c1 loc t n. P c0 \ next_change_multiplicity' c0 c1 loc t n \ P c1" and "\(loc,t)\#\<^sub>z\. \t'. t' \\<^sub>A frontier (c_imp c0 loc) \ t' \ t" and "P c0" shows "P (cm_all c0 \)" apply (subst cm_all_eq_cm_all') using assms(2) apply simp using assms apply (rule lift_cm_inv_cm_all') apply simp_all done (* Finds a minimal timestamp - location pair in a non-empty zmset (e.g. in c_work) *) lemma obtain_min_worklist: assumes "(a (loc'::(_ :: finite))::(('t :: order) zmultiset)) \ {#}\<^sub>z" obtains loc t where "t \#\<^sub>z a loc" and "\t' loc'. t' \#\<^sub>z a loc' \ \ t' < t" using assms proof atomize_elim obtain f where f: "f = minimal_antichain (\loc'. set_zmset (a loc'))" by blast from assms obtain x' where x': "x' \ (\loc'. set_zmset (a loc'))" using pos_zcount_in_zmset by fastforce from assms have "finite (\loc'. set_zmset (a loc'))" using finite_UNIV by blast with x' have "f \ {}" unfolding f using minimal_antichain_nonempty by blast then obtain t where t: "t \ f" "(\s\f. \ s < t)" using ex_min_if_finite f minimal_antichain_def by fastforce with f have thesis1: "\t' loc'. t' \#\<^sub>z a loc' \ \ t' < t" "\loc. t \#\<^sub>z a loc" by (simp add: minimal_antichain_def)+ then show "\t loc. t \#\<^sub>z a loc \ (\t' loc'. t' \#\<^sub>z a loc' \ \ t' < t)" by blast qed lemma propagate_pointstamps_eq: assumes "c_work c loc \ {#}\<^sub>z" shows "c_pts c = c_pts (SOME c'. \loc t. next_propagate' c c' loc t)" proof - from assms obtain loc' t where loc't: "t \#\<^sub>z c_work c loc'" "\t' loc'. t' \#\<^sub>z c_work c loc' \ \ t' < t" apply (rule obtain_min_worklist[of "c_work c" "loc"]) by blast let ?imps = "\locX. if locX = loc' then c_imp c locX + {#t' \#\<^sub>z c_work c locX. t' = t#} else c_imp c locX" let ?wl = "\locX. if locX = loc' then {#t' \#\<^sub>z c_work c locX. t' \ t#} else c_work c locX + after_summary (frontier_changes (?imps loc') (c_imp c loc')) (summary loc' locX)" let ?c = "c\c_imp := ?imps, c_work := ?wl\" from loc't assms have propagate: "\c'. \loc t. next_propagate' c c' loc t" by (intro exI[of _ ?c] exI[of _ loc'] exI[of _ t]) (auto simp add: next_propagate'_def intro!: Propagate.configuration.equality) { fix c' loc t assume "next_propagate' c c' loc t" then have "c_pts c = c_pts c'" by (simp add: next_propagate'_def) } with propagate show ?thesis by (simp add: verit_sko_ex') qed lemma propagate_all_imp_InvGlobPointstampsEq: "Some c1 = propagate_all c0 \ c_pts c0 = c_pts c1" unfolding propagate_all_def using while_option_rule[where P="(\c. c_pts c0 = c_pts c)" and b="(\c. \loc. c_work c loc \ {#}\<^sub>z)" and c="(\c. SOME c'. \loc t. next_propagate' c c' loc t)"] propagate_pointstamps_eq by (metis (no_types, lifting)) lemma exists_next_propagate': assumes "c_work c loc \ {#}\<^sub>z" shows "\c' loc t. next_propagate' c c' loc t" proof - from assms obtain loc' t where loc't: "t \#\<^sub>z c_work c loc'" "\t' loc'. t' \#\<^sub>z c_work c loc' \ \ t' < t" by (rule obtain_min_worklist) let ?imps = "\locX. if locX = loc' then c_imp c locX + {#t' \#\<^sub>z c_work c locX. t' = t#} else c_imp c locX" let ?wl = "\locX. if locX = loc' then {#t' \#\<^sub>z c_work c locX. t' \ t#} else c_work c locX + after_summary (frontier_changes (?imps loc') (c_imp c loc')) (summary loc' locX)" let ?c = "c\c_imp := ?imps, c_work := ?wl\" from loc't assms show ?thesis by (intro exI[of _ ?c] exI[of _ loc'] exI[of _ t]) (auto simp: next_propagate'_def intro!: Propagate.configuration.equality) qed lemma lift_propagate_inv_propagate_all: assumes "(\c0 c1 loc t. P c0 \ next_propagate' c0 c1 loc t \ P c1)" and "P c0" and "propagate_all c0 = Some c1" shows "P c1" apply (rule while_option_rule[of P _, rotated, OF assms(3)[unfolded propagate_all_def], OF assms(2)]) apply clarify subgoal for c loc apply (drule exists_next_propagate') apply (simp add: assms(1) verit_sko_ex') done done subsection\Exchange is a Subsystem of Tracker\ text\Steps in the Tracker are valid steps in the Exchange protocol.\ lemma next_imp_exchange_next: "Next' c0 c1 \ cri.next' (exchange_config c0) (exchange_config c1)" unfolding Next'_def cri.next'_def NextPerformOp'_def NextRecvUpd'_def NextSendUpd'_def NextPropagate'_def NextRecvCap'_def by auto lemma alw_next_imp_exchange_next: "alw Next s \ alw cri.next (smap exchange_config s)" by (coinduction arbitrary: s rule: alw.coinduct) (auto dest: alwD next_imp_exchange_next) text\Any Tracker trace is a valid Exchange trace\ lemma spec_imp_exchange_spec: "FullSpec s \ cri.spec (smap exchange_config s)" unfolding cri.spec_def FullSpec_def by (auto simp: InitConfig_def intro: alw_next_imp_exchange_next) lemma lift_exchange_invariant: assumes "\s. cri.spec s \ alw (holds P) s" shows "FullSpec s \ alw (\s. P (exchange_config (shd s))) s" proof - assume "FullSpec s" note spec_imp_exchange_spec[OF this] note assms[OF this] then show ?thesis by (auto simp: alw_holds_smap_conv_comp) qed text\Lifted Exchange invariants\ lemmas exch_alw_InvCapsNonneg = lift_exchange_invariant[OF cri.alw_InvCapsNonneg, unfolded atomize_imp, simplified, folded atomize_imp] and exch_alw_InvRecordCount = lift_exchange_invariant[OF cri.alw_InvRecordCount, simplified atomize_imp, simplified, folded atomize_imp] and exch_alw_InvRecordsNonneg = lift_exchange_invariant[OF cri.alw_InvRecordsNonneg, simplified atomize_imp, simplified, folded atomize_imp] and exch_alw_InvGlobVacantImpRecordsVacant = lift_exchange_invariant[OF cri.alw_InvGlobVacantImpRecordsVacant, simplified atomize_imp, simplified, folded atomize_imp] and exch_alw_InvGlobNonposImpRecordsNonpos = lift_exchange_invariant[OF cri.alw_InvGlobNonposImpRecordsNonpos, simplified atomize_imp, simplified, folded atomize_imp] and exch_alw_InvJustifiedGII = lift_exchange_invariant[OF cri.alw_InvJustifiedGII, simplified atomize_imp, simplified, folded atomize_imp] and exch_alw_InvJustifiedII = lift_exchange_invariant[OF cri.alw_InvJustifiedII, simplified atomize_imp, simplified, folded atomize_imp] and exch_alw_InvGlobNonposEqVacant = lift_exchange_invariant[OF cri.alw_InvGlobNonposEqVacant, simplified atomize_imp, simplified, folded atomize_imp] and exch_alw_InvMsgInGlob = lift_exchange_invariant[OF cri.alw_InvMsgInGlob, simplified atomize_imp, simplified, folded atomize_imp] and exch_alw_InvTempJustified = lift_exchange_invariant[OF cri.alw_InvTempJustified, simplified atomize_imp, simplified, folded atomize_imp] subsection\Definitions\ (* First variant of global safe *) definition safe_combined :: "('p::finite, 't::order, 'loc) configuration \ bool" where "safe_combined c \ \loc1 loc2 t s p. zcount (cri.records (exchange_config c)) (loc1, t) > 0 \ s \\<^sub>A path_summary loc1 loc2 \ init c p \ (\t'. t' \\<^sub>A frontier (c_imp (prop_config c p) loc2) \ t' \ results_in t s)" (* Second variant of global safe *) definition safe_combined2 :: "('p::finite, 't::order, 'loc) configuration \ bool" where "safe_combined2 c \ \loc1 loc2 t s p1 p2. zcount (c_caps (exchange_config c) p1) (loc1, t) > 0 \ s \\<^sub>A path_summary loc1 loc2 \ init c p2 \ (\t'. t' \\<^sub>A frontier (c_imp (prop_config c p2) loc2) \ t' \ results_in t s)" definition InvGlobPointstampsEq :: "('p :: finite, 't :: order, 'loc) configuration \ bool" where "InvGlobPointstampsEq c = ( (\p loc t. zcount (c_pts (prop_config c p) loc) t = zcount (c_glob (exchange_config c) p) (loc, t)))" lemma safe_combined_implies_safe_combined2: assumes "cri.InvCapsNonneg (exchange_config c)" and "safe_combined c" shows "safe_combined2 c" unfolding safe_combined2_def apply clarify subgoal for loc1 loc2 t s p1 p2 apply (rule assms(2)[unfolded safe_combined_def, rule_format, of loc1 t s loc2 p2]) apply (simp add: cri.records_def) apply (rule add_pos_nonneg) apply (subst zcount_sum) apply (rule sum_pos[where y = p1]) using assms(1) apply (simp_all add: cri.InvCapsNonneg_def) done done subsection\Propagate is a Subsystem of Tracker\ subsubsection\CM Conditions\ definition InvMsgCMConditions where "InvMsgCMConditions c = (\p q. init c q \ c_msg (exchange_config c) p q \ [] \ (\(loc,t) \#\<^sub>z (hd (c_msg (exchange_config c) p q)). \t'. t' \\<^sub>A frontier (c_imp (prop_config c q) loc) \ t' \ t))" text\Pointstamps in incoming messages all satisfy the CM premise, which is required during NextRecvUpd' steps.\ lemma msg_is_cm_safe: fixes c :: "('p::finite, 't::order, 'loc) configuration" assumes "safe (prop_config c q)" and "InvGlobPointstampsEq c" and "cri.InvMsgInGlob (exchange_config c)" and "c_msg (exchange_config c) p q \ []" shows "\(loc,t) \#\<^sub>z (hd (c_msg (exchange_config c) p q)). \t'. t' \\<^sub>A frontier (c_imp (prop_config c q) loc) \ t' \ t" using assms(3)[unfolded cri.InvMsgInGlob_def, rule_format, OF assms(4)] assms(1)[unfolded safe_def, rule_format] apply (clarsimp simp: cri_less_eq_def assms(2)[unfolded InvGlobPointstampsEq_def, rule_format, symmetric]) using order_trans apply blast done subsubsection\Propagate Safety and InvGlobPointstampsEq\ text\To be able to use the @{thm[source] msg_is_cm_safe} lemma at all times and show that Propagate is a subsystem we need to prove that the specification implies Propagate's safe and the InvGlobPointstampsEq. Both of these depend on the CM conditions being satisfied during the NextRecvUpd' step and the safety proof additionally depends on other Propagate invariants, which means that we need to prove all of these jointly.\ abbreviation prop_invs where "prop_invs c \ inv_implications_nonneg c \ inv_imps_work_sum c" abbreviation prop_safe where "prop_safe c \ impl_safe c \ safe c" definition inv_init_imp_prop_safe where "inv_init_imp_prop_safe c = (\p. init c p \ prop_safe (prop_config c p))" lemma NextRecvUpd'_preserves_prop_safe: assumes "prop_safe (prop_config c0 q)" and "InvGlobPointstampsEq c0" and "cri.InvMsgInGlob (exchange_config c0)" and "NextRecvUpd' c0 c1 p q" shows "prop_safe (prop_config c1 q)" proof - have safe: "safe (prop_config c0 q)" using assms(1) by blast note recvupd_change = cri.next_recvupdD(1)[OF NextRecvUpdD(2)[OF assms(4)]] note cm_conds = msg_is_cm_safe[OF safe assms(2,3) recvupd_change] have safes: "prop_safe c0 \ next_change_multiplicity' c0 c1 loc t n \ prop_safe c1" for c0 c1 loc t n using cm_preserves_safe cm_preserves_impl_safe by auto show "prop_safe (prop_config c1 q)" using lift_cm_inv_cm_all[rotated, OF cm_conds, of prop_safe, OF assms(1)] safes NextRecvUpdD(4)[OF assms(4)] by metis qed lemma NextRecvUpd'_preserves_InvGlobPointstampsEq: assumes "impl_safe (prop_config c0 q) \ safe (prop_config c0 q)" and "InvGlobPointstampsEq c0" and "cri.InvMsgInGlob (exchange_config c0)" and "NextRecvUpd' c0 c1 p q" shows "InvGlobPointstampsEq c1" proof - have safe: "safe (prop_config c0 q)" using assms(1) by blast note recvupd_change = cri.next_recvupdD(1)[OF NextRecvUpdD(2)[OF assms(4)]] note cm_conds = msg_is_cm_safe[OF safe assms(2,3) recvupd_change] show "InvGlobPointstampsEq c1" using assms(2,4) cm_conds unfolding NextRecvUpd'_def cri.next_recvupd'_def Let_def InvGlobPointstampsEq_def by (clarsimp simp: zcount_pointstamps_cm_all' cm_all_eq_cm_all')+ qed \ \Whenever some worker p propagates it ends up in a Propagate-safe state\ lemma NextPropagate'_causes_safe: assumes "NextPropagate' c0 c1 p" and "inv_imps_work_sum (prop_config c1 p)" and "inv_implications_nonneg (prop_config c1 p)" shows "safe (prop_config c1 p)" "impl_safe (prop_config c1 p)" proof - from assms(1) have "Some (prop_config c1 p) = propagate_all (prop_config c0 p)" by (simp add: NextPropagate'_def) then have wl: "c_work (prop_config c1 p) loc = {#}\<^sub>z" for loc unfolding propagate_all_def by (subst (asm) eq_commute) (auto dest: while_option_stop) show "safe (prop_config c1 p)" "impl_safe (prop_config c1 p)" by (rule safe[OF assms(2,3) wl]) (rule impl_safe[OF assms(2,3) wl]) qed \ \NextPropagate' preserves Propagate-safe at all workers\ lemma NextPropagate'_preserves_safe: assumes "NextPropagate' c0 c1 q" and "inv_imps_work_sum (prop_config c1 p)" and "inv_implications_nonneg (prop_config c1 p)" and "safe (prop_config c0 p)" shows "safe (prop_config c1 p)" apply (cases "p=q") subgoal using assms(1-3) by (auto intro: NextPropagate'_causes_safe) subgoal using assms(1,4) by (auto dest: spec[of _ p] simp: NextPropagate'_def) done lemma NextPropagate'_preserves_impl_safe: assumes "NextPropagate' c0 c1 q" and "inv_imps_work_sum (prop_config c1 p)" and "inv_implications_nonneg (prop_config c1 p)" and "impl_safe (prop_config c0 p)" shows "impl_safe (prop_config c1 p)" apply (cases "p=q") subgoal using assms(1-3) by (auto intro: NextPropagate'_causes_safe) subgoal using assms(1,4) by (auto dest: spec[of _ p] simp: NextPropagate'_def) done lemma NextRecvUpd'_preserves_inv_init_imp_prop_safe: assumes "cri.InvMsgInGlob (exchange_config c0)" and "inv_init_imp_prop_safe c0" and "InvGlobPointstampsEq c0" and "NextRecvUpd' c0 c1 p q" shows "inv_init_imp_prop_safe c1" using assms(2) unfolding inv_init_imp_prop_safe_def apply clarify subgoal for p apply (cases "p=q") subgoal apply (drule spec[of _p]) apply (simp add: NextRecvUpdD(1)[OF assms(4)]) apply (drule NextRecvUpd'_preserves_prop_safe[OF _ assms(3,1,4)]) apply simp done subgoal using NextRecvUpdD(3,4)[OF assms(4)] by fastforce done done lemma NextRecvUpd'_preserves_prop_invs: assumes "cri.InvMsgInGlob (exchange_config c0)" and "inv_init_imp_prop_safe c0" and "\p. prop_invs (prop_config c0 p)" and "InvGlobPointstampsEq c0" and "NextRecvUpd' c0 c1 p q" shows "\p. prop_invs (prop_config c1 p)" proof - have safe: "safe (prop_config c0 q)" using NextRecvUpdD(1) assms(2,5) inv_init_imp_prop_safe_def by blast note recvupd_change = cri.next_recvupdD(1)[OF NextRecvUpdD(2)[OF assms(5)]] note cm_conds = msg_is_cm_safe[OF safe assms(4,1) recvupd_change] have invs: "prop_invs c0 \ next_change_multiplicity' c0 c1 loc t n \ prop_invs c1" for c0 c1 loc t n using cm_preserves_inv_imps_work_sum cm_preserves_inv_implications_nonneg by auto show "\q. prop_invs (prop_config c1 q)" apply rule subgoal for q' apply (cases "q'=q") subgoal using lift_cm_inv_cm_all[rotated, OF cm_conds, of prop_invs, OF assms(3)[rule_format]] invs NextRecvUpdD(4)[OF assms(5)] by metis subgoal using NextRecvUpdD(4) assms(3) assms(5) by fastforce done done qed lemma NextPropagate'_preserves_prop_invs: assumes "prop_invs (prop_config c0 q)" and "NextPropagate' c0 c1 p" shows "prop_invs (prop_config c1 q)" apply (cases "p=q") subgoal using assms(1) lift_propagate_inv_propagate_all[ of prop_invs, rotated 2, OF NextPropagateD(3)[OF assms(2), rule_format, of p, simplified, symmetric]] by (simp add: iiws_imp_iipwn p_preserves_inv_implications_nonneg p_preserves_inv_imps_work_sum) subgoal by (metis NextPropagateD(3) assms(1) assms(2) option.simps(1)) done lemma NextPropagate'_preserves_inv_init_imp_prop_safe: assumes "prop_invs (prop_config c0 p)" and "inv_init_imp_prop_safe c0" and "NextPropagate' c0 c1 p" shows "inv_init_imp_prop_safe c1" using assms(2) unfolding inv_init_imp_prop_safe_def apply clarsimp subgoal for p' apply (cases "p'=p") subgoal using NextPropagate'_preserves_prop_invs[OF assms(1,3)] using NextPropagate'_causes_safe(1,2)[OF assms(3)] by blast subgoal using NextPropagateD(2,3)[OF assms(3)] by (auto dest: spec[of _ p']) done done lemma Next'_preserves_invs: assumes "cri.InvMsgInGlob (exchange_config c0)" and "inv_init_imp_prop_safe c0" and "InvGlobPointstampsEq c0" and "Next' c0 c1" and "\p. prop_invs (prop_config c0 p)" shows "inv_init_imp_prop_safe c1" "\p. prop_invs (prop_config c1 p)" "InvGlobPointstampsEq c1" subgoal using assms(4) unfolding Next'_def using assms(2) apply (elim disjE) subgoal unfolding inv_init_imp_prop_safe_def using NextPerformOpD(2,3) by fastforce subgoal unfolding inv_init_imp_prop_safe_def using NextSendUpdD(2,3) by fastforce subgoal using NextRecvUpd'_preserves_inv_init_imp_prop_safe[OF assms(1,2,3)] by blast subgoal using NextPropagate'_preserves_inv_init_imp_prop_safe assms(5) by blast subgoal unfolding inv_init_imp_prop_safe_def using NextRecvCapD(2,3) by fastforce subgoal by simp done subgoal using assms(4) unfolding Next'_def using assms(5) apply (elim disjE) subgoal using NextPerformOpD(2) by fastforce subgoal using NextSendUpdD(2) by fastforce subgoal using assms(1,2,3) NextRecvUpd'_preserves_prop_invs by blast subgoal using NextPropagate'_preserves_prop_invs by blast subgoal unfolding inv_init_imp_prop_safe_def using NextRecvCapD(2,3) by fastforce subgoal by simp done subgoal using assms(4) unfolding Next'_def using assms(3) apply (elim disjE) subgoal by (metis InvGlobPointstampsEq_def NextPerformOpD(1,2) cri.next_performopD(7)) subgoal by (metis InvGlobPointstampsEq_def NextSendUpdD(1,2) cri.next_sendupdD(5)) subgoal using NextRecvUpdD(1) NextRecvUpd'_preserves_InvGlobPointstampsEq assms(1,2) inv_init_imp_prop_safe_def by blast subgoal unfolding NextPropagate'_def InvGlobPointstampsEq_def using propagate_all_imp_InvGlobPointstampsEq by (metis option.inject) subgoal by (metis InvGlobPointstampsEq_def NextRecvCapD(1) NextRecvCapD(2) cri.next_recvcapD(4)) subgoal by simp done done lemma init_imp_InvGlobPointstampsEq: "InitConfig c \ InvGlobPointstampsEq c" by (simp add: InitConfig_def cri.init_config_def InvGlobPointstampsEq_def) lemma init_imp_inv_init_imp_prop_safe: "InitConfig c \ inv_init_imp_prop_safe c" by (simp add: inv_init_imp_prop_safe_def InitConfig_def) lemma init_imp_prop_invs: "InitConfig c \ \p. prop_invs (prop_config c p)" by (simp add: InitConfig_def init_imp_inv_implications_nonneg init_imp_inv_imps_work_sum) abbreviation all_invs where "all_invs c \ InvGlobPointstampsEq c \ inv_init_imp_prop_safe c \ (\p. prop_invs (prop_config c p))" lemma alw_Next'_alw_invs: assumes "holds all_invs s" and "alw (holds (\c. cri.InvMsgInGlob (exchange_config c))) s" and "alw Next s" shows "alw (holds all_invs) s" using assms apply (coinduction arbitrary: s) apply clarsimp apply safe apply (metis (mono_tags, lifting) alw_holds2 Next'_preserves_invs(3) alwD) apply (metis (mono_tags, lifting) alw_holds2 Next'_preserves_invs(1) alwD) apply (metis (mono_tags, lifting) alw_holds2 Next'_preserves_invs(2) alwD) apply (metis (mono_tags, lifting) alw_holds2 Next'_preserves_invs(2) alwD) apply auto done lemma alw_invs: "FullSpec s \ alw (holds all_invs) s" apply (frule exch_alw_InvMsgInGlob) unfolding FullSpec_def apply clarsimp apply (frule init_imp_InvGlobPointstampsEq) apply (frule init_imp_inv_init_imp_prop_safe) apply (drule init_imp_prop_invs) by (simp add: alw_Next'_alw_invs alw_mono) lemma alw_InvGlobPointstampsEq: "FullSpec s \ alw (holds InvGlobPointstampsEq) s" using alw_invs alw_mono holds_mono by blast lemma alw_inv_init_imp_prop_safe: "FullSpec s \ alw (holds inv_init_imp_prop_safe) s" using alw_invs alw_mono holds_mono by blast lemma alw_holds_conv_shd: "alw (holds \) s = alw (\s. \ (shd s)) s" by (simp add: alw_iff_sdrop) lemma alw_prop_invs: "FullSpec s \ alw (holds (\c. \p. prop_invs (prop_config c p))) s" by (auto intro: alw_mono[of "holds all_invs" s "holds (\c. \p. prop_invs (prop_config c p))"] dest: alw_invs) lemma nrec_pts_delayed: assumes "cri.InvGlobNonposImpRecordsNonpos (exchange_config c)" and "zcount (cri.records (exchange_config c)) x > 0" shows "\x'. x' \\<^sub>p x \ zcount (c_glob (exchange_config c) p) x' > 0" proof - from assms have r: "\p. \ cri.nonpos_upto (c_glob (exchange_config c) p) x" unfolding cri.InvGlobNonposImpRecordsNonpos_def cri.nonpos_upto_def by (metis linorder_not_less cri.order.order_iff_strict) show ?thesis using r[rule_format, of p] by (auto simp: cri.nonpos_upto_def not_le) qed lemma help_lemma: assumes "0 < zcount (c_pts (prop_config c p) loc0) t0" and "(loc0, t0) \\<^sub>p (loc1, t1)" and "s2 \\<^sub>A path_summary loc1 loc2" and "safe (prop_config c p)" shows "\ t2. (t2 \ results_in t1 s2 \ t2 \\<^sub>A frontier (c_imp (prop_config c p) loc2))" proof - from assms(2) obtain s1 where s1: "s1 \\<^sub>A path_summary loc0 loc1" "results_in t0 s1 \ t1" by (auto simp add: cri_less_eq_def) from s1(1) assms(3) obtain s_full where s_full: "s_full \\<^sub>A path_summary loc0 loc2" "s_full \ followed_by s1 s2" using flow.path_weight_elem_trans by blast from s_full(1) assms(1,4) obtain t2 where t2: "t2 \\<^sub>A frontier (c_imp (prop_config c p) loc2)" "t2 \ results_in t0 s_full" unfolding safe_def by blast from t2(2) and s_full(2) have "t2 \ results_in (results_in t0 s1) s2" by (metis followed_by_summary order_trans results_in_mono(2)) with s1(2) have "t2 \ results_in t1 s2" by (meson order.trans results_in_mono(1)) with t2(1) show ?thesis by auto qed \ \Lift an invariant's preservation proof over @{term next_propagate'} to NextPropagate' transitions\ lemma lift_prop_inv_NextPropagate': assumes "(\c0 c1 loc t. P c0 \ next_propagate' c0 c1 loc t \ P c1)" shows "P (prop_config c0 p') \ NextPropagate' c0 c1 p \ P (prop_config c1 p')" proof - assume pc0: "P (prop_config c0 p')" assume np: "NextPropagate' c0 c1 p" have n_p: "(\c0 c1. P c0 \ next_propagate c0 c1 \ P c1)" using assms by auto let ?f = "\c. SOME c'. next_propagate c c'" let ?b = "\c. \loc. c_work c loc \ {#}\<^sub>z" from np have pc1: "Some (prop_config c1 p) = propagate_all (prop_config c0 p)" by (simp add: NextPropagate'_def) show ?thesis apply (cases "p'=p") subgoal apply (rule while_option_rule[of P ?b ?f "prop_config c0 p"]) apply (rule n_p) apply assumption apply (rule iffD1[OF verit_sko_ex]) apply (elim exE) apply (rule exists_next_propagate') apply assumption using pc1 apply (simp add: propagate_all_def) using pc0 apply simp done subgoal using np pc0 by (auto simp: NextPropagate'_def dest!: spec[of _ p']) done qed subsubsection\Propagate is a Subsystem\ lemma NextRecvUpd'_next': assumes "safe (prop_config c0 q)" and "InvGlobPointstampsEq c0" and "cri.InvMsgInGlob (exchange_config c0)" and "NextRecvUpd' c0 c1 p q" shows "next'\<^sup>+\<^sup>+ (prop_config c0 q') (prop_config c1 q')" apply (subst NextRecvUpdD(4)[OF assms(4), rule_format]) apply simp apply safe subgoal apply (subst cm_all_eq_cm_all') apply clarsimp apply (drule assms(3)[unfolded cri.InvMsgInGlob_def, rule_format, OF cri.next_recvupdD(1)[OF NextRecvUpdD(2)[OF assms(4)]]]) apply clarsimp subgoal for loc t loc' t' apply (subst (asm) assms(2)[unfolded InvGlobPointstampsEq_def, rule_format, symmetric]) apply (clarsimp simp: cri_less_eq_def) subgoal for s using assms(1)[unfolded safe_def, rule_format, of loc' t' s loc] apply - apply (drule meta_mp) apply simp apply clarsimp subgoal for t'' apply (clarsimp intro!: exI[of _ t'']) using order_trans apply blast done done done apply (rule lift_cm_inv_cm_all') apply (rule tranclp.intros(2)) apply (auto simp: next'_def) [2] apply clarsimp apply (drule assms(3)[unfolded cri.InvMsgInGlob_def, rule_format, OF cri.next_recvupdD(1)[OF NextRecvUpdD(2)[OF assms(4)]]]) apply clarsimp subgoal for loc t loc' t' apply (subst (asm) assms(2)[unfolded InvGlobPointstampsEq_def, rule_format, symmetric]) apply (clarsimp simp: cri_less_eq_def) subgoal for s using assms(1)[unfolded safe_def, rule_format, of loc' t' s loc] apply - apply (drule meta_mp) apply simp apply clarsimp subgoal for t'' apply (clarsimp intro!: exI[of _ t'']) using order_trans apply blast done done done apply (auto simp: next'_def) done apply (auto simp: next'_def) done lemma NextPropagate'_next': assumes "NextPropagate' c0 c1 p" shows "next'\<^sup>+\<^sup>+ (prop_config c0 q) (prop_config c1 q)" apply (cases "p=q") subgoal apply (rule lift_propagate_inv_propagate_all[of _ "prop_config c0 p"]) apply (rule tranclp.intros(2)) apply (auto simp: next'_def NextPropagateD(3)[OF assms, rule_format]) done subgoal by (metis NextPropagateD(3) assms next'_def option.simps(1) tranclp.intros(1)) done lemma next_imp_propagate_next: assumes "inv_init_imp_prop_safe c0" and "InvGlobPointstampsEq c0" and "cri.InvMsgInGlob (exchange_config c0)" shows "Next' c0 c1 \ next'\<^sup>+\<^sup>+ (prop_config c0 p) (prop_config c1 p)" unfolding Next'_def NextPerformOp'_def NextSendUpd'_def NextRecvCap'_def apply safe subgoal by (auto simp: next'_def) subgoal by (auto simp: next'_def) subgoal for p' q using assms(1)[unfolded inv_init_imp_prop_safe_def, rule_format, of q] apply - apply (drule meta_mp) apply (rule NextRecvUpdD(1)) apply simp apply (cases "q=p") apply (auto dest!: NextRecvUpd'_next'[rotated, OF assms(2-)]) [] apply (auto simp add: NextRecvUpd'_def next'_def) done subgoal by (rule NextPropagate'_next') subgoal by (auto simp: next'_def) subgoal by (auto simp: next'_def) done lemma alw_next_imp_propagate_next: assumes "alw (holds inv_init_imp_prop_safe) s" and "alw (holds InvGlobPointstampsEq) s" and "alw (holds cri.InvMsgInGlob) (smap exchange_config s)" and "alw Next s" shows "alw (relates (next'\<^sup>+\<^sup>+)) (smap (\s. prop_config s p) s)" using assms by (coinduction arbitrary: s rule: alw.coinduct) (auto intro!: next_imp_propagate_next simp: relates_def alw_holds_smap_conv_comp) text\Any Tracker trace is a valid Propagate trace (using the transitive closure of next, since tracker may take multiple propagate steps at once).\ lemma spec_imp_propagate_spec: "FullSpec s \ (holds init_config aand alw (relates (next'\<^sup>+\<^sup>+))) (smap (\c. prop_config c p) s)" apply (frule alw_inv_init_imp_prop_safe) apply (frule alw_InvGlobPointstampsEq) apply (frule spec_imp_exchange_spec) apply (drule cri.alw_InvMsgInGlob) apply (auto intro!: alw_next_imp_propagate_next simp: FullSpec_def InitConfig_def) done subsection\Safety Proofs\ lemma safe_satisfied: assumes "cri.InvGlobNonposImpRecordsNonpos (exchange_config c)" and "inv_init_imp_prop_safe c" and "InvGlobPointstampsEq c" shows "safe_combined c" proof - { fix loc1 loc2 t s p assume as: "0 < zcount (cri.records (exchange_config c)) (loc1, t)" "s \\<^sub>A path_summary loc1 loc2" "init c p" obtain loc0 t0 where delayed: "(loc0, t0) \\<^sub>p (loc1, t)" "0 < zcount (c_glob (exchange_config c) p) (loc0, t0)" using nrec_pts_delayed[OF assms(1) as(1)] by fast with as(2,3) assms(2) have "\t2. t2 \ results_in t s \ t2 \\<^sub>A frontier (c_imp (prop_config c p) loc2)" using help_lemma delayed by (metis InvGlobPointstampsEq_def assms(3) inv_init_imp_prop_safe_def) } then show ?thesis unfolding safe_combined_def by blast qed lemma alw_safe_combined: "FullSpec s \ alw (holds safe_combined) s" apply (frule alw_inv_init_imp_prop_safe) apply (frule exch_alw_InvGlobNonposImpRecordsNonpos) apply (drule alw_InvGlobPointstampsEq) apply (coinduction arbitrary: s rule: alw.coinduct) apply clarsimp apply (rule conjI) apply (metis alwD alw_holds2 safe_satisfied) apply (rule disjI1) apply blast done lemma alw_safe_combined2: "FullSpec s \ alw (holds safe_combined2) s" apply (frule exch_alw_InvCapsNonneg) apply (drule alw_safe_combined) apply (simp add: alw_iff_sdrop safe_combined_implies_safe_combined2) done lemma alw_implication_frontier_eq_implied_frontier: "FullSpec s \ alw (holds (\c. worklists_vacant_to (prop_config c p) b \ b \\<^sub>A frontier (c_imp (prop_config c p) loc) \ b \\<^sub>A implied_frontier (c_pts (prop_config c p)) loc)) s" by (drule alw_prop_invs) (auto simp: implication_frontier_iff_implied_frontier_vacant all_imp_alw elim: alw_mp) end (*<*) end (*>*) \ No newline at end of file diff --git a/thys/Regex_Equivalence/Deriv_PDeriv.thy b/thys/Regex_Equivalence/Deriv_PDeriv.thy --- a/thys/Regex_Equivalence/Deriv_PDeriv.thy +++ b/thys/Regex_Equivalence/Deriv_PDeriv.thy @@ -1,304 +1,304 @@ (* Author: Tobias Nipkow, Dmitriy Traytel *) section "Connection Between Derivatives and Partial Derivatives" (*<*) theory Deriv_PDeriv imports Derivatives_Finite begin (*>*) lemma pderiv_not_is_Zero_is_Plus[simp]: "\x \ pderiv a r. \ is_Zero x \ \ is_Plus x" by (induct r) auto lemma finite_pderiv[simp]: "finite (pderiv a r)" by (induct r) auto lemma PLUS_inject: "\\x \ set xs \ set ys. \ is_Zero x \ \ is_Plus x; sorted xs; sorted ys\ \ (PLUS xs = PLUS ys) \ xs = ys" proof (induct xs arbitrary: ys rule: list_singleton_induct) case nil then show ?case by (induct ys rule: list_singleton_induct) auto next case single then show ?case by (induct ys rule: list_singleton_induct) auto next case cons then show ?case by (induct ys rule: list_singleton_induct) auto qed lemma sorted_list_of_set_inject: "\finite R; finite S\ \ (sorted_list_of_set R = sorted_list_of_set S) \ R = S" proof (induct R arbitrary: S rule: finite_linorder_min_induct) case empty then show ?case proof (induct S rule: finite_linorder_min_induct) case (insert b S) then show ?case by simp (metis insort_not_Nil) qed simp next case (insert a R) from this(4,1-3) show ?case proof (induct S rule: finite_linorder_min_induct) case (insert b S) show ?case proof assume "sorted_list_of_set (insert a R) = sorted_list_of_set (insert b S)" with insert(1,2,4,5) have "insort a (sorted_list_of_set R) = insort b (sorted_list_of_set S)" - by (elim box_equals[OF _ sorted_list_of_set.insert sorted_list_of_set.insert]) auto + by fastforce with insert(2,5) have "a # sorted_list_of_set R = b # sorted_list_of_set S" apply (cases "sorted_list_of_set R" "sorted_list_of_set S" rule: list.exhaust[case_product list.exhaust]) apply (auto split: if_splits simp add: not_le) using insort_not_Nil apply metis using insert.prems(1) set_sorted_list_of_set apply fastforce using insert.prems(1) set_sorted_list_of_set apply fastforce using insert.prems(1) set_sorted_list_of_set apply fastforce using insert.hyps(1) set_sorted_list_of_set apply fastforce using insert.hyps(1) set_sorted_list_of_set apply fastforce using insert.hyps(1) set_sorted_list_of_set apply fastforce using insert.hyps(1) set_sorted_list_of_set apply fastforce using insert.hyps(1) set_sorted_list_of_set apply fastforce done with insert show "insert a R = insert b S" by auto next assume "insert a R = insert b S" then show "sorted_list_of_set (insert a R) = sorted_list_of_set (insert b S)" by simp qed qed simp qed lemma flatten_PLUS_inject: "\\x \ R \ S. \ is_Zero x \ \ is_Plus x; finite R; finite S\ \ (flatten PLUS R = flatten PLUS S) = (R = S)" by (rule trans[OF PLUS_inject sorted_list_of_set_inject]) auto primrec pset where "pset Zero = {}" | "pset One = {One}" | "pset (Atom a) = {Atom a}" | "pset (Plus r s) = pset r \ pset s" | "pset (Times r s) = Timess (pset r) s" | "pset (Star r) = {Star r}" lemma pset_not_is_Zero_is_Plus[simp]: "\x \ pset r. \ is_Zero x \ \ is_Plus x" by (induct r) auto lemma finite_pset[simp]: "finite (pset r)" by (induct r) auto lemma pset_deriv: "pset (deriv a r) = pderiv a r" by (induct r) auto definition pnorm where "pnorm = flatten PLUS o pset" lemma pnorm_deriv_eq_iff_pderiv_eq: "pnorm (deriv a r) = pnorm (deriv a s) \ pderiv a r = pderiv a s" unfolding pnorm_def o_apply pset_deriv by (rule flatten_PLUS_inject) auto fun pnPlus :: "'a::linorder rexp \ 'a rexp \ 'a rexp" where "pnPlus Zero r = r" | "pnPlus r Zero = r" | "pnPlus (Plus r s) t = pnPlus r (pnPlus s t)" | "pnPlus r (Plus s t) = (if r = s then (Plus s t) else if le_rexp r s then Plus r (Plus s t) else Plus s (pnPlus r t))" | "pnPlus r s = (if r = s then r else if le_rexp r s then Plus r s else Plus s r)" fun pnTimes :: "'a::linorder rexp \ 'a rexp \ 'a rexp" where "pnTimes Zero r = Zero" | "pnTimes (Plus r s) t = pnPlus (pnTimes r t) (pnTimes s t)" | "pnTimes r s = Times r s" primrec pnorm_alt :: "'a::linorder rexp \ 'a rexp" where "pnorm_alt Zero = Zero" | "pnorm_alt One = One" | "pnorm_alt (Atom a) = Atom a" | "pnorm_alt (Plus r s) = pnPlus (pnorm_alt r) (pnorm_alt s)" | "pnorm_alt (Times r s) = pnTimes (pnorm_alt r) s" | "pnorm_alt (Star r) = Star r" lemma pset_pnPlus: "pset (pnPlus r s) = pset (Plus r s)" by (induct r s rule: pnPlus.induct) auto lemma pset_pnTimes: "pset (pnTimes r s) = pset (Times r s)" by (induct r s rule: pnTimes.induct) (auto simp: pset_pnPlus) lemma pset_pnorm_alt_Times: "s \ pset r \ pnTimes (pnorm_alt s) t = Times (pnorm_alt s) t" by (induct r arbitrary: s t) auto lemma pset_pnorm_alt: "pset (pnorm_alt r) = pnorm_alt ` pset r" by (induct r) (auto simp: pset_pnPlus pset_pnTimes pset_pnorm_alt_Times image_iff) lemma pset_pnTimes_Times: "s \ pset r \ pnTimes s t = Times s t" by (induct r arbitrary: s t) auto lemma pset_pnorm_alt_id: "s \ pset r \ pnorm_alt s = s" by (induct r arbitrary: s) (auto simp: pset_pnTimes_Times) lemma pnorm_alt_image_pset: "pnorm_alt ` pset r = pset r" by (induction r) (auto, auto simp add: pset_pnorm_alt_id pset_pnTimes_Times image_iff) lemma pnorm_pnorm_alt: "pnorm (pnorm_alt r) = pnorm r" by (induct r) (auto simp: pnorm_def pset_pnPlus pset_pnTimes pset_pnorm_alt pnorm_alt_image_pset) lemma pnPlus_singleton_PLUS: "\xs \ []; sorted xs; distinct xs; \x \ {x} \ set xs. \is_Zero x \ \is_Plus x\ \ pnPlus x (PLUS xs) = (if x \ set xs then PLUS xs else PLUS (insort x xs))" proof (induct xs rule: list_singleton_induct) case (single y) thus ?case unfolding is_Zero_def is_Plus_def apply (cases x y rule: linorder_cases) apply (induct x y rule: pnPlus.induct) apply (auto simp: less_rexp_def less_eq_rexp_def) apply (cases y) apply auto apply (induct x y rule: pnPlus.induct) apply (auto simp: less_rexp_def less_eq_rexp_def) apply (induct x y rule: pnPlus.induct) apply (auto simp: less_rexp_def less_eq_rexp_def) done next case (cons y1 y2 ys) thus ?case unfolding is_Zero_def is_Plus_def apply (cases x) apply (metis UnCI insertI1) apply simp apply (metis antisym less_eq_rexp_def) apply simp apply (metis antisym less_eq_rexp_def) apply (metis UnCI insertI1) apply simp apply (metis antisym less_eq_rexp_def) apply simp apply (metis antisym less_eq_rexp_def) done qed simp lemma pnPlus_PlusL[simp]: "t \ Zero \ pnPlus (Plus r s) t = pnPlus r (pnPlus s t)" by (induct t) auto lemma pnPlus_ZeroR[simp]: "pnPlus r Zero = r" by (induct r) auto lemma PLUS_eq_Zero: "PLUS xs = Zero \ xs = [] \ xs = [Zero]" by (induct xs rule: list_singleton_induct) auto lemma pnPlus_PLUS: "\xs1 \ []; xs2 \ []; \x \ set (xs1 @ xs2). \is_Zero x \ \is_Plus x; sorted xs2; distinct xs2\\ pnPlus (PLUS xs1) (PLUS xs2) = flatten PLUS (set (xs1 @ xs2))" proof (induct xs1 arbitrary: xs2 rule: list_singleton_induct) case (single x1) thus ?case apply (auto intro!: trans[OF pnPlus_singleton_PLUS] - simp: insert_absorb simp del: sorted_list_of_set_insert) + simp: insert_absorb simp del: sorted_list_of_set_insert_remove) apply (metis List.finite_set finite_sorted_distinct_unique sorted_list_of_set) apply (rule arg_cong[of _ _ PLUS]) apply (metis remdups_id_iff_distinct sorted_list_of_set_sort_remdups sorted_sort_id) done next case (cons x11 x12 xs1) then show ?case unfolding rexp_of_list.simps apply (subst pnPlus_PlusL) apply (unfold PLUS_eq_Zero) [] apply (metis in_set_conv_decomp rexp.disc(1)) apply (subst cons(1)) - apply (simp_all del: sorted_list_of_set_insert) + apply (simp_all del: sorted_list_of_set_insert_remove) apply (rule trans[OF pnPlus_singleton_PLUS]) - apply (simp_all add: sorted_insort set_insort_key del: sorted_list_of_set_insert) + apply (simp_all add: sorted_insort set_insort_key del: sorted_list_of_set_insert_remove) apply safe unfolding insert_commute[of x11] apply (auto simp only: Un_insert_left[of x11, symmetric] insert_absorb) [] apply (auto simp only: Un_insert_right[of _ x11, symmetric] insert_absorb) [] done qed simp lemma pnPlus_flatten_PLUS: "\X1 \ {}; X2 \ {}; finite X1; finite X2; \x \ X1 \ X2. \is_Zero x \ \is_Plus x\\ pnPlus (flatten PLUS X1) (flatten PLUS X2) = flatten PLUS (X1 \ X2)" by (rule trans[OF pnPlus_PLUS]) auto lemma pnPlus_pnorm: "pnPlus (pnorm r) (pnorm s) = pnorm (Plus r s)" by (cases "pset r = {} \ pset s = {}") (auto simp: pnorm_def pset_pnPlus pset_pnorm_alt intro: pnPlus_flatten_PLUS) lemma pnTimes_not_Zero_or_Plus[simp]: "\\ is_Zero x; \ is_Plus x\ \ pnTimes x r = Times x r" by (cases x) auto lemma pnTimes_PLUS: "\xs \ []; \x \ set xs. \is_Zero x \ \is_Plus x\\ pnTimes (PLUS xs) r = flatten PLUS (Timess (set xs) r)" proof (induct xs arbitrary: r rule: list_singleton_induct) case (cons x y xs) then show ?case unfolding rexp_of_list.simps pnTimes.simps apply (subst pnTimes_not_Zero_or_Plus) - apply (simp_all add: sorted_insort set_insort_key del: sorted_list_of_set_insert) + apply (simp_all add: sorted_insort set_insort_key del: sorted_list_of_set_insert_remove) apply (subst pnPlus_singleton_PLUS) - apply (simp_all add: sorted_insort set_insort_key del: sorted_list_of_set_insert) + apply (simp_all add: sorted_insort set_insort_key del: sorted_list_of_set_insert_remove) unfolding insert_commute[of "Times y r"] - apply (simp del: sorted_list_of_set_insert) + apply (simp del: sorted_list_of_set_insert_remove) apply safe apply (subst insert_absorb[of "Times x r"]) apply simp_all done qed auto lemma pnTimes_flatten_PLUS: "\X1 \ {}; finite X1; \x \ X1. \is_Zero x \ \is_Plus x\\ pnTimes (flatten PLUS X1) r = flatten PLUS (Timess X1 r)" by (rule trans[OF pnTimes_PLUS]) auto lemma pnTimes_pnorm: "pnTimes (pnorm r1) r2 = pnorm (Times r1 r2)" by (cases "pset r1 = {}") (auto simp: pnorm_def pset_pnTimes pset_pnorm_alt intro: pnTimes_flatten_PLUS) lemma pnorm_alt[symmetric]: "pnorm_alt r = pnorm r" by (induct r) (simp_all only: pnorm_alt.simps pnPlus_pnorm pnTimes_pnorm, auto simp: pnorm_def) lemma insort_eq_Cons: "\\a \ set xs. b < a; sorted xs\ \ insort b xs = b # xs" by (cases xs) auto lemma pderiv_PLUS: "pderiv a (PLUS (x # xs)) = pderiv a x \ pderiv a (PLUS xs)" by (cases xs) auto lemma pderiv_set_flatten_PLUS: "finite X \ pderiv (a :: 'a :: linorder) (flatten PLUS X) = pderiv_set a X" proof (induction X rule: finite_linorder_min_induct) case (insert b X) then have "b \ X" by auto then have "pderiv a (flatten PLUS (insert b X)) = pderiv a b \ pderiv a (flatten PLUS X)" by (simp add: pderiv_PLUS insort_eq_Cons insert.hyps) also from insert.IH have "\ = pderiv a b \ pderiv_set a X" by simp finally show ?case by simp qed simp lemma fold_pderiv_set_flatten_PLUS: "\w \ []; finite X\ \ fold pderiv_set w {flatten PLUS X} = fold pderiv_set w X" by (induct w arbitrary: X) (simp_all add: pderiv_set_flatten_PLUS) lemma fold_pnorm_deriv: "fold (\a r. pnorm (deriv a r)) w s = flatten PLUS (fold pderiv_set w {s})" proof (induction w arbitrary: s) case (Cons x w) then show ?case proof (cases "w = []") case False show ?thesis using fold_pderiv_set_flatten_PLUS[OF False] Cons.IH by (auto simp: pnorm_def pset_deriv) qed (simp add: pnorm_def pset_deriv) qed simp primrec pnderiv :: "'a :: linorder \ 'a rexp \ 'a rexp" where "pnderiv c (Zero) = Zero" | "pnderiv c (One) = Zero" | "pnderiv c (Atom c') = (if c = c' then One else Zero)" | "pnderiv c (Plus r1 r2) = pnPlus (pnderiv c r1) (pnderiv c r2)" | "pnderiv c (Times r1 r2) = (if nullable r1 then pnPlus (pnTimes (pnderiv c r1) r2) (pnderiv c r2) else pnTimes (pnderiv c r1) r2)" | "pnderiv c (Star r) = pnTimes (pnderiv c r) (Star r)" lemma pnderiv_alt[code]: "pnderiv a r = pnorm (deriv a r)" by (induct r) (auto simp: pnorm_alt) lemma pnderiv_pderiv: "pnderiv a r = flatten PLUS (pderiv a r)" unfolding pnderiv_alt pnorm_def o_apply pset_deriv .. (*<*) end (*>*) diff --git a/thys/Regex_Equivalence/Derivatives_Finite.thy b/thys/Regex_Equivalence/Derivatives_Finite.thy --- a/thys/Regex_Equivalence/Derivatives_Finite.thy +++ b/thys/Regex_Equivalence/Derivatives_Finite.thy @@ -1,503 +1,503 @@ (* Author: Dmitriy Traytel, ported by Tobias Nipkow *) section \Finiteness of Derivatives Modulo ACI\ (*<*) (* split into Norm and Fin theories *) theory Derivatives_Finite imports "Regular-Sets.Derivatives" begin (*>*) text \Lifting constructors to lists\ fun rexp_of_list where "rexp_of_list OP N [] = N" | "rexp_of_list OP N [r] = r" | "rexp_of_list OP N (r # rs) = OP r (rexp_of_list OP N rs)" abbreviation "PLUS \ rexp_of_list Plus Zero" abbreviation "TIMES \ rexp_of_list Times One" lemma list_singleton_induct [case_names nil single cons]: assumes "P []" and "\x. P [x]" and "\x y xs. P (y # xs) \ P (x # (y # xs))" shows "P xs" using assms by induction_schema (pat_completeness, lexicographic_order) subsection \ACI normalization\ fun toplevel_summands :: "'a rexp \ 'a rexp set" where "toplevel_summands (Plus r s) = toplevel_summands r \ toplevel_summands s" | "toplevel_summands r = {r}" abbreviation "flatten LISTOP X \ LISTOP (sorted_list_of_set X)" lemma toplevel_summands_nonempty[simp]: "toplevel_summands r \ {}" by (induct r) auto lemma toplevel_summands_finite[simp]: "finite (toplevel_summands r)" by (induct r) auto primrec ACI_norm :: "('a::linorder) rexp \ 'a rexp" ("\_\") where "\Zero\ = Zero" | "\One\ = One" | "\Atom a\ = Atom a" | "\Plus r s\ = flatten PLUS (toplevel_summands (Plus \r\ \s\))" | "\Times r s\ = Times \r\ \s\" | "\Star r\ = Star \r\" lemma Plus_toplevel_summands: "Plus r s \ toplevel_summands t \ False" by (induction t) auto lemma toplevel_summands_not_Plus[simp]: "(\r s. x \ Plus r s) \ toplevel_summands x = {x}" by (induction x) auto lemma toplevel_summands_PLUS_strong: "\xs \ []; list_all (\x. \(\r s. x = Plus r s)) xs\ \ toplevel_summands (PLUS xs) = set xs" by (induct xs rule: list_singleton_induct) auto lemma toplevel_summands_flatten: "\X \ {}; finite X; \x \ X. \(\r s. x = Plus r s)\ \ toplevel_summands (flatten PLUS X) = X" using toplevel_summands_PLUS_strong[of "sorted_list_of_set X"] unfolding list_all_iff by fastforce lemma ACI_norm_Plus: "\r\ = Plus s t \ \s t. r = Plus s t" by (induction r) auto lemma toplevel_summands_flatten_ACI_norm_image: "toplevel_summands (flatten PLUS (ACI_norm ` toplevel_summands r)) = ACI_norm ` toplevel_summands r" by (intro toplevel_summands_flatten) (auto dest!: ACI_norm_Plus intro: Plus_toplevel_summands) lemma toplevel_summands_flatten_ACI_norm_image_Union: "toplevel_summands (flatten PLUS (ACI_norm ` toplevel_summands r \ ACI_norm ` toplevel_summands s)) = ACI_norm ` toplevel_summands r \ ACI_norm ` toplevel_summands s" by (intro toplevel_summands_flatten) (auto dest!: ACI_norm_Plus[OF sym] intro: Plus_toplevel_summands) lemma toplevel_summands_ACI_norm: "toplevel_summands \r\ = ACI_norm ` toplevel_summands r" by (induction r) (auto simp: toplevel_summands_flatten_ACI_norm_image_Union) lemma ACI_norm_flatten: "\r\ = flatten PLUS (ACI_norm ` toplevel_summands r)" by (induction r) (auto simp: image_Un toplevel_summands_flatten_ACI_norm_image) theorem ACI_norm_idem[simp]: "\\r\\ = \r\" proof (induct r) case (Plus r s) have "\\Plus r s\\ = \flatten PLUS (toplevel_summands \r\ \ toplevel_summands \s\)\" (is "_ = \flatten PLUS ?U\") by simp also have "\ = flatten PLUS (ACI_norm ` toplevel_summands (flatten PLUS ?U))" by (simp only: ACI_norm_flatten) also have "toplevel_summands (flatten PLUS ?U) = ?U" by (intro toplevel_summands_flatten) (auto intro: Plus_toplevel_summands) also have "flatten PLUS (ACI_norm ` ?U) = flatten PLUS (toplevel_summands \r\ \ toplevel_summands \s\)" by (simp only: image_Un toplevel_summands_ACI_norm[symmetric] Plus) finally show ?case by simp qed auto subsection "Atoms" lemma atoms_toplevel_summands: "atoms s = (\r\toplevel_summands s. atoms r)" by (induct s) auto lemma wf_PLUS: "atoms (PLUS xs) \ \ \ (\r \ set xs. atoms r \ \)" by (induct xs rule: list_singleton_induct) auto lemma atoms_PLUS: "atoms (PLUS xs) = (\r \ set xs. atoms r)" by (induct xs rule: list_singleton_induct) auto lemma atoms_flatten_PLUS: "finite X \ atoms (flatten PLUS X) = (\r \ X. atoms r)" using wf_PLUS[of "sorted_list_of_set X"] by auto theorem atoms_ACI_norm: "atoms \r\ = atoms r" proof (induct r) case (Plus r1 r2) thus ?case using atoms_toplevel_summands[of "\r1\"] atoms_toplevel_summands[of "\r2\"] by(simp add: atoms_flatten_PLUS ball_Un Un_commute) qed auto subsection "Language" lemma toplevel_summands_lang: "r \ toplevel_summands s \ lang r \ lang s" by (induct s) auto lemma toplevel_summands_lang_UN: "lang s = (\r\toplevel_summands s. lang r)" by (induct s) auto lemma toplevel_summands_in_lang: "w \ lang s = (\r\toplevel_summands s. w \ lang r)" by (induct s) auto lemma lang_PLUS: "lang (PLUS xs) = (\r \ set xs. lang r)" by (induct xs rule: list_singleton_induct) auto lemma lang_PLUS_map[simp]: "lang (PLUS (map f xs)) = (\a\set xs. lang (f a))" by (induct xs rule: list_singleton_induct) auto lemma lang_flatten_PLUS[simp]: "finite X \ lang (flatten PLUS X) = (\r \ X. lang r)" using lang_PLUS[of "sorted_list_of_set X"] by fastforce theorem lang_ACI_norm[simp]: "lang \r\ = lang r" proof (induct r) case (Plus r1 r2) moreover from Plus[symmetric] have "lang (Plus r1 r2) \ lang \Plus r1 r2\" using toplevel_summands_in_lang[of _ "\r1\"] toplevel_summands_in_lang[of _ "\r2\"] by auto ultimately show ?case by (fastforce dest!: toplevel_summands_lang) qed auto subsection \Finiteness of ACI-Equivalent Derivatives\ lemma toplevel_summands_deriv: "toplevel_summands (deriv as r) = (\s\toplevel_summands r. toplevel_summands (deriv as s))" by (induction r) (auto simp: Let_def) lemma derivs_Zero[simp]: "derivs xs Zero = Zero" by (induction xs) auto lemma derivs_One: "derivs xs One \ {Zero, One}" by (induction xs) auto lemma derivs_Atom: "derivs xs (Atom as) \ {Zero, One, Atom as}" proof (induction xs) case Cons thus ?case by (auto intro: insertE[OF derivs_One]) qed simp lemma derivs_Plus: "derivs xs (Plus r s) = Plus (derivs xs r) (derivs xs s)" by (induction xs arbitrary: r s) auto lemma derivs_PLUS: "derivs xs (PLUS ys) = PLUS (map (derivs xs) ys)" by (induction ys rule: list_singleton_induct) (auto simp: derivs_Plus) lemma toplevel_summands_derivs_Times: "toplevel_summands (derivs xs (Times r s)) \ {Times (derivs xs r) s} \ {r'. \ys zs. r' \ toplevel_summands (derivs ys s) \ ys \ [] \ zs @ ys = xs}" proof (induction xs arbitrary: r s) case (Cons x xs) thus ?case by (auto simp: Let_def derivs_Plus) (fastforce intro: exI[of _ "x#xs"])+ qed simp lemma toplevel_summands_derivs_Star_nonempty: "xs \ [] \ toplevel_summands (derivs xs (Star r)) \ {Times (derivs ys r) (Star r) | ys. \zs. ys \ [] \ zs @ ys = xs}" proof (induction xs rule: length_induct) case (1 xs) then obtain y ys where "xs = y # ys" by (cases xs) auto thus ?case using spec[OF 1(1)] by (auto dest!: subsetD[OF toplevel_summands_derivs_Times] intro: exI[of _ "y#ys"]) (auto elim!: impE dest!: meta_spec subsetD) qed lemma toplevel_summands_derivs_Star: "toplevel_summands (derivs xs (Star r)) \ {Star r} \ {Times (derivs ys r) (Star r) | ys. \zs. ys \ [] \ zs @ ys = xs}" by (cases "xs = []") (auto dest!: toplevel_summands_derivs_Star_nonempty) lemma toplevel_summands_PLUS: "xs \ [] \ toplevel_summands (PLUS (map f xs)) = (\r \ set xs. toplevel_summands (f r))" by (induction xs rule: list_singleton_induct) simp_all lemma ACI_norm_toplevel_summands_Zero: "toplevel_summands r \ {Zero} \ \r\ = Zero" by (subst ACI_norm_flatten) (auto dest: subset_singletonD) lemma finite_ACI_norm_toplevel_summands: "finite {f \s\ |s. toplevel_summands s \ B}" if "finite B" proof - have *: "{f \s\ | s. toplevel_summands s \ B} \ (f \ flatten PLUS \ (`) ACI_norm) ` Pow B" by (subst ACI_norm_flatten) auto with that show ?thesis by (rule finite_surj [OF iffD2 [OF finite_Pow_iff]]) qed theorem finite_derivs: "finite {\derivs xs r\ | xs . True}" proof (induct r) case Zero show ?case by simp next case One show ?case by (rule finite_surj[of "{Zero, One}"]) (blast intro: insertE[OF derivs_One])+ next case (Atom as) show ?case by (rule finite_surj[of "{Zero, One, Atom as}"]) (blast intro: insertE[OF derivs_Atom])+ next case (Plus r s) show ?case by (auto simp: derivs_Plus intro!: finite_surj[OF finite_cartesian_product[OF Plus]]) next case (Times r s) hence "finite (\ (toplevel_summands ` {\derivs xs s\ | xs . True}))" by auto moreover have "{\r'\ |r'. \ys. r' \ toplevel_summands (derivs ys s)} = {r'. \ys. r' \ toplevel_summands \derivs ys s\}" by (auto simp: toplevel_summands_ACI_norm) ultimately have fin: "finite {\r'\ |r'. \ys. r' \ toplevel_summands (derivs ys s)}" by (fastforce intro: finite_subset[of _ "\ (toplevel_summands ` {\derivs xs s\ | xs . True})"]) let ?X = "\xs. {Times (derivs ys r) s | ys. True} \ {r'. r' \ (\ys. toplevel_summands (derivs ys s))}" show ?case proof (simp only: ACI_norm_flatten, rule finite_surj[of "{X. \xs. X \ ACI_norm ` ?X xs}" _ "flatten PLUS"]) show "finite {X. \xs. X \ ACI_norm ` ?X xs}" using fin by (fastforce simp: image_Un elim: finite_subset[rotated] intro: finite_surj[OF Times(1), of _ "\r. Times r \s\"]) qed (fastforce dest!: subsetD[OF toplevel_summands_derivs_Times] intro!: imageI) next case (Star r) let ?f = "\f r'. Times r' (Star (f r))" let ?X = "{Star r} \ ?f id ` {r'. r' \ {derivs ys r|ys. True}}" show ?case proof (simp only: ACI_norm_flatten, rule finite_surj[of "{X. X \ ACI_norm ` ?X}" _ "flatten PLUS"]) have *: "\X. ACI_norm ` ?f (\x. x) ` X = ?f ACI_norm ` ACI_norm ` X" by (auto simp: image_def) show "finite {X. X \ ACI_norm ` ?X}" by (rule finite_Collect_subsets) (auto simp: * intro!: finite_imageI[of _ "?f ACI_norm"] intro: finite_subset[OF _ Star]) qed (fastforce dest!: subsetD[OF toplevel_summands_derivs_Star] intro!: imageI) qed subsection \Deriving preserves ACI-equivalence\ lemma ACI_norm_PLUS: "list_all2 (\r s. \r\ = \s\) xs ys \ \PLUS xs\ = \PLUS ys\" proof (induct rule: list_all2_induct) case (Cons x xs y ys) hence "length xs = length ys" by (elim list_all2_lengthD) thus ?case using Cons by (induct xs ys rule: list_induct2) auto qed simp lemma toplevel_summands_ACI_norm_deriv: "(\a\toplevel_summands r. toplevel_summands \deriv as \a\\) = toplevel_summands \deriv as \r\\" proof (induct r) case (Plus r1 r2) thus ?case unfolding toplevel_summands.simps toplevel_summands_ACI_norm toplevel_summands_deriv[of as "\Plus r1 r2\"] image_Un Union_Un_distrib by (simp add: image_UN) qed (auto simp: Let_def) lemma toplevel_summands_nullable: "nullable s = (\r\toplevel_summands s. nullable r)" by (induction s) auto lemma nullable_PLUS: "nullable (PLUS xs) = (\r \ set xs. nullable r)" by (induction xs rule: list_singleton_induct) auto theorem ACI_norm_nullable: "nullable \r\ = nullable r" proof (induction r) case (Plus r1 r2) thus ?case using toplevel_summands_nullable by (auto simp: nullable_PLUS) qed auto theorem ACI_norm_deriv: "\deriv as \r\\ = \deriv as r\" proof (induction r arbitrary: as) case (Plus r1 r2) thus ?case unfolding deriv.simps ACI_norm_flatten[of "deriv as \Plus r1 r2\"] toplevel_summands_deriv[of as "\Plus r1 r2\"] image_Un image_UN by (auto simp: toplevel_summands_ACI_norm toplevel_summands_flatten_ACI_norm_image_Union) (auto simp: toplevel_summands_ACI_norm[symmetric] toplevel_summands_ACI_norm_deriv) qed (simp_all add: ACI_norm_nullable) corollary deriv_preserves: "\r\ = \s\ \ \deriv as r\ = \deriv as s\" by (rule box_equals[OF _ ACI_norm_deriv ACI_norm_deriv]) (erule arg_cong) lemma derivs_snoc[simp]: "derivs (xs @ [x]) r = (deriv x (derivs xs r))" by (induction xs arbitrary: r) auto theorem ACI_norm_derivs: "\derivs xs \r\\ = \derivs xs r\" proof (induction xs arbitrary: r rule: rev_induct) case (snoc x xs) thus ?case using ACI_norm_deriv[of x "derivs xs r"] ACI_norm_deriv[of x "derivs xs \r\"] by auto qed simp subsection \Alternative ACI defintions\ text \Not necessary but conceptually nicer (and seems also to be faster?!)\ fun ACI_nPlus :: "'a::linorder rexp \ 'a rexp \ 'a rexp" where "ACI_nPlus (Plus r1 r2) s = ACI_nPlus r1 (ACI_nPlus r2 s)" | "ACI_nPlus r (Plus s1 s2) = (if r = s1 then Plus s1 s2 else if r < s1 then Plus r (Plus s1 s2) else Plus s1 (ACI_nPlus r s2))" | "ACI_nPlus r s = (if r = s then r else if r < s then Plus r s else Plus s r)" primrec ACI_norm_alt where "ACI_norm_alt Zero = Zero" | "ACI_norm_alt One = One" | "ACI_norm_alt (Atom a) = Atom a" | "ACI_norm_alt (Plus r s) = ACI_nPlus (ACI_norm_alt r) (ACI_norm_alt s)" | "ACI_norm_alt (Times r s) = Times (ACI_norm_alt r) (ACI_norm_alt s)" | "ACI_norm_alt (Star r) = Star (ACI_norm_alt r)" lemma toplevel_summands_ACI_nPlus: "toplevel_summands (ACI_nPlus r s) = toplevel_summands (Plus r s)" by (induct r s rule: ACI_nPlus.induct) auto lemma toplevel_summands_ACI_norm_alt: "toplevel_summands (ACI_norm_alt r) = ACI_norm_alt ` toplevel_summands r" by (induct r) (auto simp: toplevel_summands_ACI_nPlus) lemma ACI_norm_alt_Plus: "ACI_norm_alt r = Plus s t \ \s t. r = Plus s t" by (induct r) auto lemma toplevel_summands_flatten_ACI_norm_alt_image: "toplevel_summands (flatten PLUS (ACI_norm_alt ` toplevel_summands r)) = ACI_norm_alt ` toplevel_summands r" by (intro toplevel_summands_flatten) (auto dest!: ACI_norm_alt_Plus intro: Plus_toplevel_summands) lemma ACI_norm_ACI_norm_alt: "\ACI_norm_alt r\ = \r\" proof (induction r) case (Plus r s) show ?case using ACI_norm_flatten [of r] ACI_norm_flatten [of s] by (auto simp add: toplevel_summands_ACI_nPlus) (metis ACI_norm_flatten Plus.IH(1) Plus.IH(2) image_Un toplevel_summands.simps(1) toplevel_summands_ACI_nPlus toplevel_summands_ACI_norm) qed auto lemma ACI_nPlus_singleton_PLUS: "\xs \ []; sorted xs; distinct xs; \x \ {x} \ set xs. \(\r s. x = Plus r s)\ \ ACI_nPlus x (PLUS xs) = (if x \ set xs then PLUS xs else PLUS (insort x xs))" proof (induct xs rule: list_singleton_induct) case (single y) thus ?case by (cases x y rule: linorder_cases) (induct x y rule: ACI_nPlus.induct, auto)+ next case (cons y1 y2 ys) thus ?case by (cases x) (auto) qed simp lemma ACI_nPlus_PLUS: "\xs1 \ []; xs2 \ []; \x \ set (xs1 @ xs2). \(\r s. x = Plus r s); sorted xs2; distinct xs2\\ ACI_nPlus (PLUS xs1) (PLUS xs2) = flatten PLUS (set (xs1 @ xs2))" proof (induct xs1 arbitrary: xs2 rule: list_singleton_induct) case (single x1) thus ?case - apply (auto intro!: trans[OF ACI_nPlus_singleton_PLUS] simp del: sorted_list_of_set_insert) + apply (auto intro!: trans[OF ACI_nPlus_singleton_PLUS] simp del: sorted_list_of_set_insert_remove) apply (simp only: insert_absorb) apply (metis List.finite_set finite_sorted_distinct_unique sorted_list_of_set) apply (rule arg_cong[of _ _ PLUS]) apply (metis remdups_id_iff_distinct sorted_list_of_set_sort_remdups sorted_sort_id) done next case (cons x11 x12 xs1) thus ?case - apply (simp del: sorted_list_of_set_insert) + apply (simp del: sorted_list_of_set_insert_remove) apply (rule trans[OF ACI_nPlus_singleton_PLUS]) - apply (auto simp del: sorted_list_of_set_insert simp add: insert_commute[of x11]) + apply (auto simp del: sorted_list_of_set_insert_remove simp add: insert_commute[of x11]) apply (auto simp only: Un_insert_left[of x11, symmetric] insert_absorb) [] apply (auto simp only: Un_insert_right[of _ x11, symmetric] insert_absorb) [] apply (auto simp add: insert_commute[of x12]) done qed simp lemma ACI_nPlus_flatten_PLUS: "\X1 \ {}; X2 \ {}; finite X1; finite X2; \x \ X1 \ X2. \(\r s. x = Plus r s)\\ ACI_nPlus (flatten PLUS X1) (flatten PLUS X2) = flatten PLUS (X1 \ X2)" by (rule trans[OF ACI_nPlus_PLUS]) auto lemma ACI_nPlus_ACI_norm [simp]: "ACI_nPlus \r\ \s\ = \Plus r s\" by (auto simp: image_Un Un_assoc ACI_norm_flatten [of r] ACI_norm_flatten [of s] ACI_norm_flatten [of "Plus r s"] toplevel_summands_flatten_ACI_norm_image intro!: trans [OF ACI_nPlus_flatten_PLUS]) (metis ACI_norm_Plus Plus_toplevel_summands)+ lemma ACI_norm_alt: "ACI_norm_alt r = \r\" by (induct r) auto declare ACI_norm_alt[symmetric, code] inductive ACI where ACI_refl: "ACI r r" | ACI_sym: "ACI r s \ ACI s r" | ACI_trans: "ACI r s \ ACI s t \ ACI r t" | ACI_Plus_cong: "\ACI r1 s1; ACI r2 s2\ \ ACI (Plus r1 r2) (Plus s1 s2)" | ACI_Times_cong: "\ACI r1 s1; ACI r2 s2\ \ ACI (Times r1 r2) (Times s1 s2)" | ACI_Star_cong: "ACI r s \ ACI (Star r) (Star s)" | ACI_assoc: "ACI (Plus (Plus r s) t) (Plus r (Plus s t))" | ACI_comm: "ACI (Plus r s) (Plus s r)" | ACI_idem: "ACI (Plus r r) r" lemma ACI_atoms: "ACI r s \ atoms r = atoms s" by (induct rule: ACI.induct) auto lemma ACI_nullable: "ACI r s \ nullable r = nullable s" by (induct rule: ACI.induct) auto lemma ACI_lang: "ACI r s \ lang r = lang s" by (induct rule: ACI.induct) auto lemma ACI_deriv: "ACI r s \ ACI (deriv a r) (deriv a s)" proof (induct arbitrary: a rule: ACI.induct) case (ACI_Times_cong r1 s1 r2 s2) thus ?case by (auto simp: Let_def intro: ACI.intros dest: ACI_nullable) (metis ACI.ACI_Times_cong ACI_Plus_cong) qed (auto intro: ACI.intros) lemma ACI_Plus_assocI[intro]: "ACI (Plus r1 r2) s2 \ ACI (Plus r1 (Plus s1 r2)) (Plus s1 s2)" "ACI (Plus r1 r2) s2 \ ACI (Plus r1 (Plus r2 s1)) (Plus s1 s2)" by (metis ACI_assoc ACI_comm ACI_Plus_cong ACI_refl ACI_trans)+ lemma ACI_Plus_idemI[intro]: "\ACI r s1; ACI r s2\ \ ACI r (Plus s1 s2)" by (metis ACI_Plus_cong ACI_idem ACI_sym ACI_trans) lemma ACI_Plus_idemI'[intro]: "\ACI r1 s1; ACI (Plus r1 r2) s2\ \ ACI (Plus r1 r2) (Plus s1 s2)" by (rule ACI_trans[OF ACI_Plus_cong[OF ACI_sym[OF ACI_idem] ACI_refl] ACI_trans[OF ACI_assoc ACI_trans[OF ACI_Plus_cong ACI_refl]]]) lemma ACI_ACI_nPlus: "\ACI r1 s1; ACI r2 s2\ \ ACI (ACI_nPlus r1 r2) (Plus s1 s2)" proof (induct r1 r2 arbitrary: s1 s2 rule: ACI_nPlus.induct) case 1 from 1(2)[OF ACI_refl 1(1)[OF ACI_refl 1(4)]] 1(3) show ?case by (auto intro: ACI_comm ACI_trans) next case ("2_1" r1 r2) with ACI_Plus_cong[OF ACI_refl "2_1"(1)[OF _ _ "2_1"(2) ACI_refl], of r1] show ?case by (auto intro: ACI.intros) next case ("2_2" r1 r2) with ACI_Plus_cong[OF ACI_refl "2_2"(1)[OF _ _ "2_2"(2) ACI_refl], of r1] show ?case by (auto intro: ACI.intros) next case ("2_3" _ r1 r2) with ACI_Plus_cong[OF ACI_refl "2_3"(1)[OF _ _ "2_3"(2) ACI_refl], of r1] show ?case by (auto intro: ACI.intros) next case ("2_4" _ _ r1 r2) with ACI_Plus_cong[OF ACI_refl "2_4"(1)[OF _ _ "2_4"(2) ACI_refl], of r1] show ?case by (auto intro: ACI.intros) next case ("2_5" _ r1 r2) with ACI_Plus_cong[OF ACI_refl "2_5"(1)[OF _ _ "2_5"(2) ACI_refl], of r1] show ?case by (auto intro: ACI.intros) qed (auto intro: ACI.intros) lemma ACI_ACI_norm: "ACI \r\ r" unfolding ACI_norm_alt[symmetric] by (induct r) (auto intro: ACI.intros simp: ACI_ACI_nPlus) lemma ACI_norm_eqI: "ACI r s \ \r\ = \s\" by (induct rule: ACI.induct) (auto simp: toplevel_summands_ACI_norm ACI_norm_flatten[symmetric] toplevel_summands_flatten_ACI_norm_image_Union ac_simps) lemma ACI_I: "\r\ = \s\ \ ACI r s" by (metis ACI_ACI_norm ACI_sym ACI_trans) lemma ACI_decidable: "ACI r s = (\r\ = \s\)" by (metis ACI_I ACI_norm_eqI) (*<*) end (*>*) diff --git a/thys/SC_DOM_Components/Core_DOM_DOM_Components.thy b/thys/SC_DOM_Components/Core_DOM_DOM_Components.thy --- a/thys/SC_DOM_Components/Core_DOM_DOM_Components.thy +++ b/thys/SC_DOM_Components/Core_DOM_DOM_Components.thy @@ -1,2351 +1,2351 @@ (*********************************************************************************** * Copyright (c) 2016-2020 The University of Sheffield, UK * 2019-2020 University of Exeter, UK * * 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. * * 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 HOLDER 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. * * SPDX-License-Identifier: BSD-2-Clause ***********************************************************************************) section \Core SC DOM Components\ theory Core_DOM_DOM_Components imports Core_SC_DOM.Core_DOM begin subsection \Components\ locale l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_root_node_defs get_root_node get_root_node_locs + l_to_tree_order_defs to_tree_order for get_root_node :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr) prog" and get_root_node_locs :: "((_) heap \ (_) heap \ bool) set" and to_tree_order :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr list) prog" begin definition a_get_dom_component :: "(_) object_ptr \ (_, (_) object_ptr list) dom_prog" where "a_get_dom_component ptr = do { root \ get_root_node ptr; to_tree_order root }" definition a_is_strongly_dom_component_safe :: "(_) object_ptr set \ (_) object_ptr set \ (_) heap \ (_) heap \ bool" where "a_is_strongly_dom_component_safe S\<^sub>a\<^sub>r\<^sub>g S\<^sub>r\<^sub>e\<^sub>s\<^sub>u\<^sub>l\<^sub>t h h' = ( let removed_pointers = fset (object_ptr_kinds h) - fset (object_ptr_kinds h') in let added_pointers = fset (object_ptr_kinds h') - fset (object_ptr_kinds h) in let arg_components = (\ptr \ (\ptr \ S\<^sub>a\<^sub>r\<^sub>g. set |h \ a_get_dom_component ptr|\<^sub>r) \ fset (object_ptr_kinds h). set |h \ a_get_dom_component ptr|\<^sub>r) in let arg_components' = (\ptr \ (\ptr \ S\<^sub>a\<^sub>r\<^sub>g. set |h \ a_get_dom_component ptr|\<^sub>r) \ fset (object_ptr_kinds h'). set |h' \ a_get_dom_component ptr|\<^sub>r) in removed_pointers \ arg_components \ added_pointers \ arg_components' \ S\<^sub>r\<^sub>e\<^sub>s\<^sub>u\<^sub>l\<^sub>t \ arg_components' \ (\outside_ptr \ fset (object_ptr_kinds h) \ fset (object_ptr_kinds h') - (\ptr \ S\<^sub>a\<^sub>r\<^sub>g. set |h \ a_get_dom_component ptr|\<^sub>r). preserved (get_M outside_ptr id) h h'))" definition a_is_weakly_dom_component_safe :: "(_) object_ptr set \ (_) object_ptr set \ (_) heap \ (_) heap \ bool" where "a_is_weakly_dom_component_safe S\<^sub>a\<^sub>r\<^sub>g S\<^sub>r\<^sub>e\<^sub>s\<^sub>u\<^sub>l\<^sub>t h h' = ( let removed_pointers = fset (object_ptr_kinds h) - fset (object_ptr_kinds h') in let added_pointers = fset (object_ptr_kinds h') - fset (object_ptr_kinds h) in let arg_components = (\ptr \ (\ptr \ S\<^sub>a\<^sub>r\<^sub>g. set |h \ a_get_dom_component ptr|\<^sub>r) \ fset (object_ptr_kinds h). set |h \ a_get_dom_component ptr|\<^sub>r) in let arg_components' = (\ptr \ (\ptr \ S\<^sub>a\<^sub>r\<^sub>g. set |h \ a_get_dom_component ptr|\<^sub>r) \ fset (object_ptr_kinds h'). set |h' \ a_get_dom_component ptr|\<^sub>r) in removed_pointers \ arg_components \ S\<^sub>r\<^sub>e\<^sub>s\<^sub>u\<^sub>l\<^sub>t \ arg_components' \ added_pointers \ (\outside_ptr \ fset (object_ptr_kinds h) \ fset (object_ptr_kinds h') - (\ptr \ S\<^sub>a\<^sub>r\<^sub>g. set |h \ a_get_dom_component ptr|\<^sub>r). preserved (get_M outside_ptr id) h h'))" lemma "a_is_strongly_dom_component_safe S\<^sub>a\<^sub>r\<^sub>g S\<^sub>r\<^sub>e\<^sub>s\<^sub>u\<^sub>l\<^sub>t h h' \ a_is_weakly_dom_component_safe S\<^sub>a\<^sub>r\<^sub>g S\<^sub>r\<^sub>e\<^sub>s\<^sub>u\<^sub>l\<^sub>t h h'" by(auto simp add: a_is_strongly_dom_component_safe_def a_is_weakly_dom_component_safe_def Let_def) definition is_document_component :: "(_) object_ptr list \ bool" where "is_document_component c = is_document_ptr_kind (hd c)" definition is_disconnected_component :: "(_) object_ptr list \ bool" where "is_disconnected_component c = is_node_ptr_kind (hd c)" end global_interpretation l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_root_node get_root_node_locs to_tree_order defines get_dom_component = a_get_dom_component and is_strongly_dom_component_safe = a_is_strongly_dom_component_safe and is_weakly_dom_component_safe = a_is_weakly_dom_component_safe . locale l_get_dom_component_defs = fixes get_dom_component :: "(_) object_ptr \ (_, (_) object_ptr list) dom_prog" fixes is_strongly_dom_component_safe :: "(_) object_ptr set \ (_) object_ptr set \ (_) heap \ (_) heap \ bool" fixes is_weakly_dom_component_safe :: "(_) object_ptr set \ (_) object_ptr set \ (_) heap \ (_) heap \ bool" locale l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_to_tree_order_wf + l_get_dom_component_defs + l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_get_ancestors + l_get_ancestors_wf + l_get_root_node + l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent + l_get_parent_wf + l_get_element_by + l_to_tree_order_wf_get_root_node_wf + (* l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M _ _ _ get_child_nodes + l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M _ _ get_child_nodes+ l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M _ _ "l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.a_to_tree_order get_child_nodes" for get_child_nodes :: "(_::linorder) object_ptr \ (_, (_) node_ptr list) dom_prog" *) assumes get_dom_component_impl: "get_dom_component = a_get_dom_component" assumes is_strongly_dom_component_safe_impl: "is_strongly_dom_component_safe = a_is_strongly_dom_component_safe" assumes is_weakly_dom_component_safe_impl: "is_weakly_dom_component_safe = a_is_weakly_dom_component_safe" begin lemmas get_dom_component_def = a_get_dom_component_def[folded get_dom_component_impl] lemmas is_strongly_dom_component_safe_def = a_is_strongly_dom_component_safe_def[folded is_strongly_dom_component_safe_impl] lemmas is_weakly_dom_component_safe_def = a_is_weakly_dom_component_safe_def[folded is_weakly_dom_component_safe_impl] lemma get_dom_component_ptr_in_heap: assumes "h \ ok (get_dom_component ptr)" shows "ptr |\| object_ptr_kinds h" using assms get_root_node_ptr_in_heap by(auto simp add: get_dom_component_def) lemma get_dom_component_ok: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "ptr |\| object_ptr_kinds h" shows "h \ ok (get_dom_component ptr)" using assms apply(auto simp add: get_dom_component_def a_get_root_node_def intro!: bind_is_OK_pure_I)[1] using get_root_node_ok to_tree_order_ok get_root_node_ptr_in_heap apply blast by (simp add: local.get_root_node_root_in_heap local.to_tree_order_ok) lemma get_dom_component_ptr: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_dom_component ptr \\<^sub>r c" shows "ptr \ set c" proof(insert assms(1) assms(4), induct ptr rule: heap_wellformed_induct_rev ) case (step child) then show ?case proof (cases "is_node_ptr_kind child") case True obtain node_ptr where node_ptr: "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr = child" using \is_node_ptr_kind child\ node_ptr_casts_commute3 by blast have "child |\| object_ptr_kinds h" using \h \ get_dom_component child \\<^sub>r c\ get_dom_component_ptr_in_heap by fast with node_ptr have "node_ptr |\| node_ptr_kinds h" by auto then obtain parent_opt where parent: "h \ get_parent node_ptr \\<^sub>r parent_opt" using get_parent_ok \type_wf h\ \known_ptrs h\ by fast then show ?thesis proof (induct parent_opt) case None then have "h \ get_root_node (cast node_ptr) \\<^sub>r cast node_ptr" by (simp add: local.get_root_node_no_parent) then show ?case using \type_wf h\ \known_ptrs h\ node_ptr step(2) apply(auto simp add: get_dom_component_def a_get_root_node_def elim!: bind_returns_result_E2)[1] using to_tree_order_ptr_in_result returns_result_eq by fastforce next case (Some parent_ptr) then have "h \ get_dom_component parent_ptr \\<^sub>r c" using step(2) node_ptr \type_wf h\ \known_ptrs h\ \heap_is_wellformed h\ get_root_node_parent_same by(auto simp add: get_dom_component_def elim!: bind_returns_result_E2 intro!: bind_pure_returns_result_I) then have "parent_ptr \ set c" using step node_ptr Some by blast then show ?case using \type_wf h\ \known_ptrs h\ \heap_is_wellformed h\ step(2) node_ptr Some apply(auto simp add: get_dom_component_def elim!: bind_returns_result_E2)[1] using to_tree_order_parent by blast qed next case False then show ?thesis using \type_wf h\ \known_ptrs h\ step(2) apply(auto simp add: get_dom_component_def elim!: bind_returns_result_E2)[1] by (metis is_OK_returns_result_I local.get_root_node_not_node_same local.get_root_node_ptr_in_heap local.to_tree_order_ptr_in_result returns_result_eq) qed qed lemma get_dom_component_parent_inside: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_dom_component ptr \\<^sub>r c" assumes "cast node_ptr \ set c" assumes "h \ get_parent node_ptr \\<^sub>r Some parent" shows "parent \ set c" proof - have "parent |\| object_ptr_kinds h" using assms(6) get_parent_parent_in_heap by blast obtain root_ptr where root_ptr: "h \ get_root_node ptr \\<^sub>r root_ptr" and c: "h \ to_tree_order root_ptr \\<^sub>r c" using assms(4) by (metis (no_types, hide_lams) bind_returns_result_E2 get_dom_component_def get_root_node_pure) then have "h \ get_root_node (cast node_ptr) \\<^sub>r root_ptr" using assms(1) assms(2) assms(3) assms(5) to_tree_order_same_root by blast then have "h \ get_root_node parent \\<^sub>r root_ptr" using assms(6) get_root_node_parent_same by blast then have "h \ get_dom_component parent \\<^sub>r c" using c get_dom_component_def by auto then show ?thesis using assms(1) assms(2) assms(3) get_dom_component_ptr by blast qed lemma get_dom_component_subset: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_dom_component ptr \\<^sub>r c" assumes "ptr' \ set c" shows "h \ get_dom_component ptr' \\<^sub>r c" proof(insert assms(1) assms(5), induct ptr' rule: heap_wellformed_induct_rev ) case (step child) then show ?case proof (cases "is_node_ptr_kind child") case True obtain node_ptr where node_ptr: "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr = child" using \is_node_ptr_kind child\ node_ptr_casts_commute3 by blast have "child |\| object_ptr_kinds h" using to_tree_order_ptrs_in_heap assms(1) assms(2) assms(3) assms(4) step(2) unfolding get_dom_component_def by (meson bind_returns_result_E2 get_root_node_pure) with node_ptr have "node_ptr |\| node_ptr_kinds h" by auto then obtain parent_opt where parent: "h \ get_parent node_ptr \\<^sub>r parent_opt" using get_parent_ok \type_wf h\ \known_ptrs h\ by fast then show ?thesis proof (induct parent_opt) case None then have "h \ get_root_node child \\<^sub>r child" using assms(1) get_root_node_no_parent node_ptr by blast then show ?case using \type_wf h\ \known_ptrs h\ node_ptr step(2) assms(4) assms(1) by (metis (no_types) bind_pure_returns_result_I2 bind_returns_result_E2 get_dom_component_def get_root_node_pure is_OK_returns_result_I returns_result_eq to_tree_order_same_root) next case (Some parent_ptr) then have "h \ get_dom_component parent_ptr \\<^sub>r c" using step get_dom_component_parent_inside assms node_ptr by blast then show ?case using Some node_ptr apply(auto simp add: get_dom_component_def elim!: bind_returns_result_E2 del: bind_pure_returns_result_I intro!: bind_pure_returns_result_I)[1] using get_root_node_parent_same by blast qed next case False then have "child |\| object_ptr_kinds h" using assms(1) assms(4) step(2) by (metis (no_types, lifting) assms(2) assms(3) bind_returns_result_E2 get_root_node_pure get_dom_component_def to_tree_order_ptrs_in_heap) then have "h \ get_root_node child \\<^sub>r child" using assms(1) False get_root_node_not_node_same by blast then show ?thesis using assms(1) assms(2) assms(3) assms(4) step.prems by (metis (no_types) False \child |\| object_ptr_kinds h\ bind_pure_returns_result_I2 bind_returns_result_E2 get_dom_component_def get_root_node_ok get_root_node_pure returns_result_eq to_tree_order_node_ptrs) qed qed lemma get_dom_component_to_tree_order_subset: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ to_tree_order ptr \\<^sub>r nodes" assumes "h \ get_dom_component ptr \\<^sub>r c" shows "set nodes \ set c" using assms apply(auto simp add: get_dom_component_def elim!: bind_returns_result_E2)[1] by (meson to_tree_order_subset assms(5) contra_subsetD get_dom_component_ptr) lemma get_dom_component_to_tree_order: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_dom_component ptr \\<^sub>r c" assumes "h \ to_tree_order ptr' \\<^sub>r to" assumes "ptr \ set to" shows "h \ get_dom_component ptr' \\<^sub>r c" by (metis (no_types, hide_lams) assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) get_dom_component_ok get_dom_component_subset get_dom_component_to_tree_order_subset is_OK_returns_result_E local.to_tree_order_ptr_in_result local.to_tree_order_ptrs_in_heap select_result_I2 subsetCE) lemma get_dom_component_root_node_same: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_dom_component ptr \\<^sub>r c" assumes "h \ get_root_node ptr \\<^sub>r root_ptr" assumes "x \ set c" shows "h \ get_root_node x \\<^sub>r root_ptr" proof(insert assms(1) assms(6), induct x rule: heap_wellformed_induct_rev ) case (step child) then show ?case proof (cases "is_node_ptr_kind child") case True obtain node_ptr where node_ptr: "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr = child" using \is_node_ptr_kind child\ node_ptr_casts_commute3 by blast have "child |\| object_ptr_kinds h" using to_tree_order_ptrs_in_heap assms(1) assms(2) assms(3) assms(4) step(2) unfolding get_dom_component_def by (meson bind_returns_result_E2 get_root_node_pure) with node_ptr have "node_ptr |\| node_ptr_kinds h" by auto then obtain parent_opt where parent: "h \ get_parent node_ptr \\<^sub>r parent_opt" using get_parent_ok \type_wf h\ \known_ptrs h\ by fast then show ?thesis proof (induct parent_opt) case None then have "h \ get_root_node child \\<^sub>r child" using assms(1) get_root_node_no_parent node_ptr by blast then show ?case using \type_wf h\ \known_ptrs h\ node_ptr step(2) assms(4) assms(1) assms(5) by (metis (no_types) \child |\| object_ptr_kinds h\ bind_pure_returns_result_I get_dom_component_def get_dom_component_ptr get_dom_component_subset get_root_node_pure is_OK_returns_result_E returns_result_eq to_tree_order_ok to_tree_order_same_root) next case (Some parent_ptr) then have "h \ get_dom_component parent_ptr \\<^sub>r c" using step get_dom_component_parent_inside assms node_ptr by (meson get_dom_component_subset) then show ?case using Some node_ptr apply(auto simp add: get_dom_component_def elim!: bind_returns_result_E2)[1] using get_root_node_parent_same using \h \ get_dom_component parent_ptr \\<^sub>r c\ assms(1) assms(2) assms(3) get_dom_component_ptr step.hyps by blast qed next case False then have "child |\| object_ptr_kinds h" using assms(1) assms(4) step(2) by (metis (no_types, lifting) assms(2) assms(3) bind_returns_result_E2 get_root_node_pure get_dom_component_def to_tree_order_ptrs_in_heap) then have "h \ get_root_node child \\<^sub>r child" using assms(1) False get_root_node_not_node_same by auto then show ?thesis using assms(1) assms(2) assms(3) assms(4) step.prems assms(5) by (metis (no_types, hide_lams) bind_returns_result_E2 get_dom_component_def get_root_node_pure returns_result_eq to_tree_order_same_root) qed qed lemma get_dom_component_no_overlap: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_dom_component ptr \\<^sub>r c" assumes "h \ get_dom_component ptr' \\<^sub>r c'" shows "set c \ set c' = {} \ c = c'" proof (rule ccontr, auto) fix x assume 1: "c \ c'" and 2: "x \ set c" and 3: "x \ set c'" obtain root_ptr where root_ptr: "h \ get_root_node ptr \\<^sub>r root_ptr" using assms(4) unfolding get_dom_component_def by (meson bind_is_OK_E is_OK_returns_result_I) moreover obtain root_ptr' where root_ptr': "h \ get_root_node ptr' \\<^sub>r root_ptr'" using assms(5) unfolding get_dom_component_def by (meson bind_is_OK_E is_OK_returns_result_I) ultimately have "root_ptr \ root_ptr'" using 1 assms unfolding get_dom_component_def by (meson bind_returns_result_E3 get_root_node_pure returns_result_eq) moreover have "h \ get_root_node x \\<^sub>r root_ptr" using 2 root_ptr get_dom_component_root_node_same assms by blast moreover have "h \ get_root_node x \\<^sub>r root_ptr'" using 3 root_ptr' get_dom_component_root_node_same assms by blast ultimately show False using select_result_I2 by force qed lemma get_dom_component_separates_tree_order: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_dom_component ptr \\<^sub>r c" assumes "h \ to_tree_order ptr \\<^sub>r to" assumes "h \ get_dom_component ptr' \\<^sub>r c'" assumes "ptr' \ set c" shows "set to \ set c' = {}" proof - have "c \ c'" using assms get_dom_component_ptr by blast then have "set c \ set c' = {}" using assms get_dom_component_no_overlap by blast moreover have "set to \ set c" using assms get_dom_component_to_tree_order_subset by blast ultimately show ?thesis by blast qed lemma get_dom_component_separates_tree_order_general: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_dom_component ptr \\<^sub>r c" assumes "h \ to_tree_order ptr'' \\<^sub>r to''" assumes "ptr'' \ set c" assumes "h \ get_dom_component ptr' \\<^sub>r c'" assumes "ptr' \ set c" shows "set to'' \ set c' = {}" proof - obtain root_ptr where root_ptr: "h \ get_root_node ptr \\<^sub>r root_ptr" using assms(4) by (metis bind_is_OK_E get_dom_component_def is_OK_returns_result_I) then have c: "h \ to_tree_order root_ptr \\<^sub>r c" using assms(4) unfolding get_dom_component_def by (simp add: bind_returns_result_E3) with root_ptr show ?thesis using assms get_dom_component_separates_tree_order get_dom_component_subset by meson qed end interpretation i_get_dom_component?: l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name by(auto simp add: l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def get_dom_component_def is_strongly_dom_component_safe_def is_weakly_dom_component_safe_def instances) declare l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_child\_nodes\ locale l_get_dom_component_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_child_nodes_is_strongly_dom_component_safe: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_dom_component ptr \\<^sub>r c" assumes "h \ get_child_nodes ptr' \\<^sub>r children" assumes "child \ set children" shows "cast child \ set c \ ptr' \ set c" proof assume 1: "cast child \ set c" obtain root_ptr where root_ptr: "h \ get_root_node ptr \\<^sub>r root_ptr" by (metis assms(4) bind_is_OK_E get_dom_component_def is_OK_returns_result_I) have "h \ get_root_node (cast child) \\<^sub>r root_ptr" using "1" assms(1) assms(2) assms(3) assms(4) get_dom_component_root_node_same root_ptr by blast then have "h \ get_root_node ptr' \\<^sub>r root_ptr" using assms(1) assms(2) assms(3) assms(5) assms(6) local.child_parent_dual local.get_root_node_parent_same by blast then have "h \ get_dom_component ptr' \\<^sub>r c" using "1" assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) local.child_parent_dual local.get_dom_component_parent_inside local.get_dom_component_subset by blast then show "ptr' \ set c" using assms(1) assms(2) assms(3) get_dom_component_ptr by blast next assume 1: "ptr' \ set c" obtain root_ptr where root_ptr: "h \ get_root_node ptr \\<^sub>r root_ptr" by (metis assms(4) bind_is_OK_E get_dom_component_def is_OK_returns_result_I) have "h \ get_root_node ptr' \\<^sub>r root_ptr" using "1" assms(1) assms(2) assms(3) assms(4) get_dom_component_root_node_same root_ptr by blast then have "h \ get_root_node (cast child) \\<^sub>r root_ptr" using assms(1) assms(2) assms(3) assms(5) assms(6) local.child_parent_dual local.get_root_node_parent_same by blast then have "h \ get_dom_component ptr' \\<^sub>r c" using "1" assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) local.child_parent_dual local.get_dom_component_parent_inside local.get_dom_component_subset by blast then show "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child \ set c" by (smt \h \ get_root_node (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r root_ptr\ assms(1) assms(2) assms(3) assms(5) assms(6) disjoint_iff_not_equal is_OK_returns_result_E is_OK_returns_result_I l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_dom_component_no_overlap local.child_parent_dual local.get_dom_component_ok local.get_dom_component_parent_inside local.get_dom_component_ptr local.get_root_node_ptr_in_heap local.l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms) qed lemma get_child_nodes_get_dom_component: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_dom_component ptr \\<^sub>r c" assumes "h \ get_child_nodes ptr \\<^sub>r children" shows "cast ` set children \ set c" using assms get_child_nodes_is_strongly_dom_component_safe using local.get_dom_component_ptr by auto lemma assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_child_nodes ptr \\<^sub>r children" assumes "h \ get_child_nodes ptr \\<^sub>h h'" shows "is_strongly_dom_component_safe {ptr} (cast ` set children) h h'" proof - have "h = h'" using assms(5) by (meson local.get_child_nodes_pure pure_returns_heap_eq) then show ?thesis using assms apply(auto simp add: is_strongly_dom_component_safe_def Let_def preserved_def)[1] by (smt IntI finite_set_in get_child_nodes_is_strongly_dom_component_safe is_OK_returns_result_E is_OK_returns_result_I local.get_child_nodes_ptr_in_heap local.get_dom_component_impl local.get_dom_component_ok local.get_dom_component_ptr select_result_I2) qed end interpretation i_get_dom_component_get_child_nodes?: l_get_dom_component_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs by(auto simp add: l_get_dom_component_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_dom_component_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_parent\ locale l_get_dom_component_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_parent_is_strongly_dom_component_safe_step: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_dom_component ptr \\<^sub>r c" assumes "h \ get_parent ptr' \\<^sub>r Some parent" shows "parent \ set c \ cast ptr' \ set c" by (meson assms(1) assms(2) assms(3) assms(4) assms(5) is_OK_returns_result_E l_to_tree_order_wf.to_tree_order_parent local.get_dom_component_parent_inside local.get_dom_component_subset local.get_dom_component_to_tree_order_subset local.get_parent_parent_in_heap local.l_to_tree_order_wf_axioms local.to_tree_order_ok local.to_tree_order_ptr_in_result subsetCE) lemma get_parent_is_strongly_dom_component_safe: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_parent node_ptr \\<^sub>r Some parent" assumes "h \ get_parent node_ptr \\<^sub>h h'" shows "is_strongly_dom_component_safe {cast node_ptr} {parent} h h'" proof - have "h = h'" using assms(5) by (meson local.get_parent_pure pure_returns_heap_eq) then show ?thesis using assms apply(auto simp add: is_strongly_dom_component_safe_def Let_def preserved_def)[1] by (metis IntI finite_set_in local.get_dom_component_impl local.get_dom_component_ok local.get_dom_component_parent_inside local.get_dom_component_ptr local.get_parent_parent_in_heap local.known_ptrs_known_ptr local.parent_child_rel_child_in_heap local.parent_child_rel_parent returns_result_select_result) qed end interpretation i_get_dom_component_get_parent?: l_get_dom_component_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs by(auto simp add: l_get_dom_component_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_dom_component_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_root\_node\ locale l_get_dom_component_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_root_node_is_strongly_dom_component_safe_step: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_dom_component ptr \\<^sub>r c" assumes "h \ get_root_node ptr' \\<^sub>r root" shows "root \ set c \ ptr' \ set c" proof assume 1: "root \ set c" obtain root_ptr where root_ptr: "h \ get_root_node ptr \\<^sub>r root_ptr" by (metis assms(4) bind_is_OK_E get_dom_component_def is_OK_returns_result_I) have "h \ get_root_node root \\<^sub>r root_ptr" using "1" assms(1) assms(2) assms(3) assms(4) get_dom_component_root_node_same root_ptr by blast moreover have "h \ get_root_node ptr' \\<^sub>r root_ptr" by (metis (no_types, lifting) calculation assms(1) assms(2) assms(3) assms(5) is_OK_returns_result_E local.get_root_node_root_in_heap local.to_tree_order_ok local.to_tree_order_ptr_in_result local.to_tree_order_same_root select_result_I2) ultimately have "h \ get_dom_component ptr' \\<^sub>r c" apply(auto simp add: get_dom_component_def)[1] by (smt "1" assms(1) assms(2) assms(3) assms(4) bind_pure_returns_result_I bind_returns_result_E3 local.get_dom_component_def local.get_dom_component_subset local.get_root_node_pure) then show "ptr' \ set c" using assms(1) assms(2) assms(3) get_dom_component_ptr by blast next assume 1: "ptr' \ set c" obtain root_ptr where root_ptr: "h \ get_root_node ptr \\<^sub>r root_ptr" by (metis assms(4) bind_is_OK_E get_dom_component_def is_OK_returns_result_I) have "h \ get_root_node ptr' \\<^sub>r root_ptr" using "1" assms(1) assms(2) assms(3) assms(4) get_dom_component_root_node_same root_ptr by blast then have "h \ get_root_node root \\<^sub>r root_ptr" by (metis assms(1) assms(2) assms(3) assms(5) is_OK_returns_result_E local.get_root_node_root_in_heap local.to_tree_order_ok local.to_tree_order_ptr_in_result local.to_tree_order_same_root returns_result_eq) then have "h \ get_dom_component ptr' \\<^sub>r c" using "1" assms(1) assms(2) assms(3) assms(4) local.get_dom_component_subset by blast then show "root \ set c" using assms(5) bind_returns_result_E3 local.get_dom_component_def local.to_tree_order_ptr_in_result by fastforce qed lemma get_root_node_is_strongly_dom_component_safe: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_root_node ptr \\<^sub>r root" assumes "h \ get_root_node ptr \\<^sub>h h'" shows "is_strongly_dom_component_safe {ptr} {root} h h'" proof - have "h = h'" using assms(5) by (meson local.get_root_node_pure pure_returns_heap_eq) then show ?thesis using assms apply(auto simp add: is_strongly_dom_component_safe_def Let_def preserved_def)[1] by (metis (no_types, lifting) IntI finite_set_in get_root_node_is_strongly_dom_component_safe_step is_OK_returns_result_I local.get_dom_component_impl local.get_dom_component_ok local.get_dom_component_ptr local.get_root_node_ptr_in_heap returns_result_select_result) qed end interpretation i_get_dom_component_get_root_node?: l_get_dom_component_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs by(auto simp add: l_get_dom_component_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_dom_component_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_element\_by\_id\ locale l_get_dom_component_get_element_by_id\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_element_by_id_is_strongly_dom_component_safe_step: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_dom_component ptr \\<^sub>r c" assumes "h \ get_element_by_id ptr' idd \\<^sub>r Some result" shows "cast result \ set c \ ptr' \ set c" proof assume 1: "cast result \ set c" obtain root_ptr where root_ptr: "h \ get_root_node ptr \\<^sub>r root_ptr" by (metis assms(4) bind_is_OK_E get_dom_component_def is_OK_returns_result_I) then have "h \ to_tree_order root_ptr \\<^sub>r c" using \h \ get_dom_component ptr \\<^sub>r c\ by (simp add: get_dom_component_def bind_returns_result_E3) obtain to' where to': "h \ to_tree_order ptr' \\<^sub>r to'" using \h \ get_element_by_id ptr' idd \\<^sub>r Some result\ apply(simp add: get_element_by_id_def first_in_tree_order_def) by (meson bind_is_OK_E is_OK_returns_result_I) then have "cast result \ set to'" using \h \ get_element_by_id ptr' idd \\<^sub>r Some result\ get_element_by_id_result_in_tree_order by blast have "h \ get_root_node (cast result) \\<^sub>r root_ptr" using "1" assms(1) assms(2) assms(3) assms(4) get_dom_component_root_node_same root_ptr by blast then have "h \ get_root_node ptr' \\<^sub>r root_ptr" using \cast result \ set to'\ \h \ to_tree_order ptr' \\<^sub>r to'\ using "1" assms(1) assms(2) assms(3) assms(4) get_dom_component_ptr get_dom_component_root_node_same get_dom_component_subset get_dom_component_to_tree_order by blast then have "h \ get_dom_component ptr' \\<^sub>r c" using \h \ to_tree_order root_ptr \\<^sub>r c\ using get_dom_component_def by auto then show "ptr' \ set c" using assms(1) assms(2) assms(3) get_dom_component_ptr by blast next assume "ptr' \ set c" moreover obtain to' where to': "h \ to_tree_order ptr' \\<^sub>r to'" by (meson assms(1) assms(2) assms(3) assms(4) calculation get_dom_component_ptr_in_heap get_dom_component_subset is_OK_returns_result_E is_OK_returns_result_I to_tree_order_ok) ultimately have "set to' \ set c" using assms(1) assms(2) assms(3) assms(4) get_dom_component_subset get_dom_component_to_tree_order_subset by blast moreover have "cast result \ set to'" using assms(5) get_element_by_id_result_in_tree_order to' by blast ultimately show "cast result \ set c" by blast qed lemma get_element_by_id_is_strongly_dom_component_safe: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_element_by_id ptr idd \\<^sub>r Some result" assumes "h \ get_element_by_id ptr idd \\<^sub>h h'" shows "is_strongly_dom_component_safe {ptr} {cast result} h h'" proof - have "h = h'" using assms(5) by(auto simp add: preserved_def get_element_by_id_def first_in_tree_order_def elim!: bind_returns_heap_E2 intro!: map_filter_M_pure bind_pure_I split: option.splits list.splits) have "ptr |\| object_ptr_kinds h" using assms(4) apply(auto simp add: get_element_by_id_def)[1] by (metis (no_types, lifting) assms(1) assms(2) assms(3) bind_is_OK_E is_OK_returns_result_I local.first_in_tree_order_def local.to_tree_order_ptr_in_result local.to_tree_order_ptrs_in_heap) obtain to where to: "h \ to_tree_order ptr \\<^sub>r to" by (meson \ptr |\| object_ptr_kinds h\ assms(1) assms(2) assms(3) is_OK_returns_result_E local.to_tree_order_ok) then have "cast result \ set to" using assms(4) local.get_element_by_id_result_in_tree_order by auto obtain c where c: "h \ a_get_dom_component ptr \\<^sub>r c" using \ptr |\| object_ptr_kinds h\ assms(1) assms(2) assms(3) local.get_dom_component_impl local.get_dom_component_ok by blast then show ?thesis using assms \h = h'\ apply(auto simp add: is_strongly_dom_component_safe_def Let_def preserved_def get_element_by_id_def first_in_tree_order_def elim!: bind_returns_result_E2 intro!: map_filter_M_pure bind_pure_I split: option.splits list.splits)[1] by (metis (no_types, lifting) Int_iff \ptr |\| object_ptr_kinds h\ assms(4) finite_set_in get_element_by_id_is_strongly_dom_component_safe_step local.get_dom_component_impl local.get_dom_component_ptr select_result_I2) qed end interpretation i_get_dom_component_get_element_by_id?: l_get_dom_component_get_element_by_id\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs by(auto simp add: l_get_dom_component_get_element_by_id\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_dom_component_get_element_by_id\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_elements\_by\_class\_name\ locale l_get_dom_component_get_elements_by_class_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_elements_by_class_name_is_strongly_dom_component_safe_step: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_dom_component ptr \\<^sub>r c" assumes "h \ get_elements_by_class_name ptr' idd \\<^sub>r results" assumes "result \ set results" shows "cast result \ set c \ ptr' \ set c" proof assume 1: "cast result \ set c" obtain root_ptr where root_ptr: "h \ get_root_node ptr \\<^sub>r root_ptr" by (metis assms(4) bind_is_OK_E get_dom_component_def is_OK_returns_result_I) then have "h \ to_tree_order root_ptr \\<^sub>r c" using \h \ get_dom_component ptr \\<^sub>r c\ by (simp add: get_dom_component_def bind_returns_result_E3) obtain to' where to': "h \ to_tree_order ptr' \\<^sub>r to'" using assms(5) apply(simp add: get_elements_by_class_name_def first_in_tree_order_def) by (meson bind_is_OK_E is_OK_returns_result_I) then have "cast result \ set to'" using assms get_elements_by_class_name_result_in_tree_order by blast have "h \ get_root_node (cast result) \\<^sub>r root_ptr" using "1" assms(1) assms(2) assms(3) assms(4) get_dom_component_root_node_same root_ptr by blast then have "h \ get_root_node ptr' \\<^sub>r root_ptr" using \cast result \ set to'\ \h \ to_tree_order ptr' \\<^sub>r to'\ using "1" assms(1) assms(2) assms(3) assms(4) get_dom_component_ptr get_dom_component_root_node_same get_dom_component_subset get_dom_component_to_tree_order by blast then have "h \ get_dom_component ptr' \\<^sub>r c" using \h \ to_tree_order root_ptr \\<^sub>r c\ using get_dom_component_def by auto then show "ptr' \ set c" using assms(1) assms(2) assms(3) get_dom_component_ptr by blast next assume "ptr' \ set c" moreover obtain to' where to': "h \ to_tree_order ptr' \\<^sub>r to'" by (meson assms(1) assms(2) assms(3) assms(4) calculation get_dom_component_ptr_in_heap get_dom_component_subset is_OK_returns_result_E is_OK_returns_result_I to_tree_order_ok) ultimately have "set to' \ set c" using assms(1) assms(2) assms(3) assms(4) get_dom_component_subset get_dom_component_to_tree_order_subset by blast moreover have "cast result \ set to'" using assms get_elements_by_class_name_result_in_tree_order to' by blast ultimately show "cast result \ set c" by blast qed lemma get_elements_by_class_name_pure [simp]: "pure (get_elements_by_class_name ptr name) h" by(auto simp add: get_elements_by_class_name_def intro!: bind_pure_I map_filter_M_pure split: option.splits) lemma get_elements_by_class_name_is_strongly_dom_component_safe: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_elements_by_class_name ptr name \\<^sub>r results" assumes "h \ get_elements_by_class_name ptr name \\<^sub>h h'" shows "is_strongly_dom_component_safe {ptr} (cast ` set results) h h'" proof - have "h = h'" using assms(5) by (meson get_elements_by_class_name_pure pure_returns_heap_eq) have "ptr |\| object_ptr_kinds h" using assms(4) apply(auto simp add: get_elements_by_class_name_def)[1] by (metis (no_types, lifting) assms(1) assms(2) assms(3) bind_is_OK_E is_OK_returns_result_I local.first_in_tree_order_def local.to_tree_order_ptr_in_result local.to_tree_order_ptrs_in_heap) obtain to where to: "h \ to_tree_order ptr \\<^sub>r to" by (meson \ptr |\| object_ptr_kinds h\ assms(1) assms(2) assms(3) is_OK_returns_result_E local.to_tree_order_ok) then have "cast ` set results \ set to" using assms(4) local.get_elements_by_class_name_result_in_tree_order by auto obtain c where c: "h \ a_get_dom_component ptr \\<^sub>r c" using \ptr |\| object_ptr_kinds h\ assms(1) assms(2) assms(3) local.get_dom_component_impl local.get_dom_component_ok by blast then show ?thesis using assms \h = h'\ apply(auto simp add: is_strongly_dom_component_safe_def Let_def preserved_def get_elements_by_class_name_def first_in_tree_order_def elim!: bind_returns_result_E2 intro!: map_filter_M_pure bind_pure_I split: option.splits list.splits)[1] by (metis (no_types, lifting) Int_iff \ptr |\| object_ptr_kinds h\ assms(4) finite_set_in get_elements_by_class_name_is_strongly_dom_component_safe_step local.get_dom_component_impl local.get_dom_component_ptr select_result_I2) qed end interpretation i_get_dom_component_get_elements_by_class_name?: l_get_dom_component_get_elements_by_class_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs by(auto simp add: l_get_dom_component_get_elements_by_class_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_dom_component_get_elements_by_class_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_elements\_by\_tag\_name\ locale l_get_dom_component_get_elements_by_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_elements_by_tag_name_is_strongly_dom_component_safe_step: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_dom_component ptr \\<^sub>r c" assumes "h \ get_elements_by_tag_name ptr' idd \\<^sub>r results" assumes "result \ set results" shows "cast result \ set c \ ptr' \ set c" proof assume 1: "cast result \ set c" obtain root_ptr where root_ptr: "h \ get_root_node ptr \\<^sub>r root_ptr" by (metis assms(4) bind_is_OK_E get_dom_component_def is_OK_returns_result_I) then have "h \ to_tree_order root_ptr \\<^sub>r c" using \h \ get_dom_component ptr \\<^sub>r c\ by (simp add: get_dom_component_def bind_returns_result_E3) obtain to' where to': "h \ to_tree_order ptr' \\<^sub>r to'" using assms(5) apply(simp add: get_elements_by_tag_name_def first_in_tree_order_def) by (meson bind_is_OK_E is_OK_returns_result_I) then have "cast result \ set to'" using assms get_elements_by_tag_name_result_in_tree_order by blast have "h \ get_root_node (cast result) \\<^sub>r root_ptr" using "1" assms(1) assms(2) assms(3) assms(4) get_dom_component_root_node_same root_ptr by blast then have "h \ get_root_node ptr' \\<^sub>r root_ptr" using \cast result \ set to'\ \h \ to_tree_order ptr' \\<^sub>r to'\ using "1" assms(1) assms(2) assms(3) assms(4) get_dom_component_ptr get_dom_component_root_node_same get_dom_component_subset get_dom_component_to_tree_order by blast then have "h \ get_dom_component ptr' \\<^sub>r c" using \h \ to_tree_order root_ptr \\<^sub>r c\ using get_dom_component_def by auto then show "ptr' \ set c" using assms(1) assms(2) assms(3) get_dom_component_ptr by blast next assume "ptr' \ set c" moreover obtain to' where to': "h \ to_tree_order ptr' \\<^sub>r to'" by (meson assms(1) assms(2) assms(3) assms(4) calculation get_dom_component_ptr_in_heap get_dom_component_subset is_OK_returns_result_E is_OK_returns_result_I to_tree_order_ok) ultimately have "set to' \ set c" using assms(1) assms(2) assms(3) assms(4) get_dom_component_subset get_dom_component_to_tree_order_subset by blast moreover have "cast result \ set to'" using assms get_elements_by_tag_name_result_in_tree_order to' by blast ultimately show "cast result \ set c" by blast qed lemma get_elements_by_tag_name_is_strongly_dom_component_safe: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_elements_by_tag_name ptr name \\<^sub>r results" assumes "h \ get_elements_by_tag_name ptr name \\<^sub>h h'" shows "is_strongly_dom_component_safe {ptr} (cast ` set results) h h'" proof - have "h = h'" using assms(5) by (meson get_elements_by_tag_name_pure pure_returns_heap_eq) have "ptr |\| object_ptr_kinds h" using assms(4) apply(auto simp add: get_elements_by_tag_name_def)[1] by (metis (no_types, lifting) assms(1) assms(2) assms(3) bind_is_OK_E is_OK_returns_result_I local.first_in_tree_order_def local.to_tree_order_ptr_in_result local.to_tree_order_ptrs_in_heap) obtain to where to: "h \ to_tree_order ptr \\<^sub>r to" by (meson \ptr |\| object_ptr_kinds h\ assms(1) assms(2) assms(3) is_OK_returns_result_E local.to_tree_order_ok) then have "cast ` set results \ set to" using assms(4) local.get_elements_by_tag_name_result_in_tree_order by auto obtain c where c: "h \ a_get_dom_component ptr \\<^sub>r c" using \ptr |\| object_ptr_kinds h\ assms(1) assms(2) assms(3) local.get_dom_component_impl local.get_dom_component_ok by blast then show ?thesis using assms \h = h'\ apply(auto simp add: is_strongly_dom_component_safe_def Let_def preserved_def get_elements_by_class_name_def first_in_tree_order_def elim!: bind_returns_result_E2 intro!: map_filter_M_pure bind_pure_I split: option.splits list.splits)[1] by (metis (no_types, lifting) IntI \ptr |\| object_ptr_kinds h\ finite_set_in get_elements_by_tag_name_is_strongly_dom_component_safe_step local.get_dom_component_impl local.get_dom_component_ptr select_result_I2) qed end interpretation i_get_dom_component_get_elements_by_tag_name?: l_get_dom_component_get_elements_by_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs by(auto simp add: l_get_dom_component_get_elements_by_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_dom_component_get_elements_by_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \remove\_child\ lemma remove_child_unsafe: "\(\(h :: ('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap ) h' ptr child. heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ remove_child ptr child \\<^sub>h h' \ is_weakly_dom_component_safe {ptr, cast child} {} h h')" proof - obtain h document_ptr e1 e2 where h: "Inr ((document_ptr, e1, e2), h) = (Heap (fmempty) :: ('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap ) \ (do { document_ptr \ create_document; e1 \ create_element document_ptr ''div''; e2 \ create_element document_ptr ''div''; append_child (cast e1) (cast e2); return (document_ptr, e1, e2) })" by(code_simp, auto simp add: equal_eq List.member_def)+ then obtain h' where "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and h': "h \ remove_child (cast e1) (cast e2) \\<^sub>h h'" and "\(is_weakly_dom_component_safe {cast e1, cast e2} {} h h')" apply(code_simp) apply(clarify) by(code_simp, auto simp add: equal_eq List.member_def)+ then show ?thesis by auto qed subsubsection \adopt\_node\ lemma adopt_node_unsafe: "\(\(h :: ('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap ) h' document_ptr child. heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ adopt_node document_ptr child \\<^sub>h h' \ is_weakly_dom_component_safe {cast document_ptr, cast child} {} h h')" proof - obtain h document_ptr document_ptr2 e1 where h: "Inr ((document_ptr, document_ptr2, e1), h) = (Heap (fmempty) :: ('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap ) \ (do { document_ptr \ create_document; document_ptr2 \ create_document; e1 \ create_element document_ptr ''div''; adopt_node document_ptr2 (cast e1); return (document_ptr, document_ptr2, e1) })" by(code_simp, auto simp add: equal_eq List.member_def)+ then obtain h' where "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and h': "h \ adopt_node document_ptr (cast e1) \\<^sub>h h'" and "\(is_weakly_dom_component_safe {cast document_ptr, cast e1} {} h h')" apply(code_simp) apply(clarify) by(code_simp, auto simp add: equal_eq List.member_def)+ then show ?thesis by auto qed subsubsection \create\_element\ lemma create_element_not_strongly_dom_component_safe: obtains h :: "('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap" and h' and document_ptr and new_element_ptr and tag where "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and "h \ create_element document_ptr tag \\<^sub>r new_element_ptr \\<^sub>h h'" and "\ is_strongly_dom_component_safe {cast document_ptr} {cast new_element_ptr} h h'" proof - let ?h0 = "Heap fmempty ::('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap" let ?P = "create_document" let ?h1 = "|?h0 \ ?P|\<^sub>h" let ?document_ptr = "|?h0 \ ?P|\<^sub>r" show thesis apply(rule that[where h="?h1" and document_ptr="?document_ptr"]) by code_simp+ qed locale l_get_dom_component_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name + l_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs set_tag_name set_tag_name_locs type_wf create_element known_ptr + l_get_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_disconnected_nodes get_disconnected_nodes_locs + l_set_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_disconnected_nodes set_disconnected_nodes_locs + l_set_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_tag_name set_tag_name_locs + l_new_element_get_disconnected_nodes get_disconnected_nodes get_disconnected_nodes_locs + l_new_element_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs + l_set_tag_name_get_child_nodes type_wf set_tag_name set_tag_name_locs known_ptr get_child_nodes get_child_nodes_locs + l_create_element_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr known_ptrs type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed parent_child_rel set_tag_name set_tag_name_locs set_disconnected_nodes set_disconnected_nodes_locs create_element for known_ptr :: "(_::linorder) object_ptr \ bool" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and type_wf :: "(_) heap \ bool" and known_ptrs :: "(_) heap \ bool" and to_tree_order :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr list) prog" and get_parent :: "(_) node_ptr \ ((_) heap, exception, (_) object_ptr option) prog" and get_parent_locs :: "((_) heap \ (_) heap \ bool) set" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_dom_component :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr list) prog" and is_strongly_dom_component_safe :: "(_) object_ptr set \ (_) object_ptr set \ (_) heap \ (_) heap \ bool" and is_weakly_dom_component_safe :: "(_) object_ptr set \ (_) object_ptr set \ (_) heap \ (_) heap \ bool" and get_root_node :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr) prog" and get_root_node_locs :: "((_) heap \ (_) heap \ bool) set" and get_ancestors :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr list) prog" and get_ancestors_locs :: "((_) heap \ (_) heap \ bool) set" and get_element_by_id :: "(_) object_ptr \ char list \ ((_) heap, exception, (_) element_ptr option) prog" and get_elements_by_class_name :: "(_) object_ptr \ char list \ ((_) heap, exception, (_) element_ptr list) prog" and get_elements_by_tag_name :: "(_) object_ptr \ char list \ ((_) heap, exception, (_) element_ptr list) prog" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" and set_tag_name :: "(_) element_ptr \ char list \ ((_) heap, exception, unit) prog" and set_tag_name_locs :: "(_) element_ptr \ ((_) heap, exception, unit) prog set" and create_element :: "(_) document_ptr \ char list \ ((_) heap, exception, (_) element_ptr) prog" begin lemma create_element_is_weakly_dom_component_safe_step: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ create_element document_ptr tag \\<^sub>h h'" assumes "ptr \ set |h \ get_dom_component (cast document_ptr)|\<^sub>r" assumes "ptr \ cast |h \ create_element document_ptr tag|\<^sub>r" shows "preserved (get_M ptr getter) h h'" proof - obtain new_element_ptr h2 h3 disc_nodes where new_element_ptr: "h \ new_element \\<^sub>r new_element_ptr" and h2: "h \ new_element \\<^sub>h h2" and h3: "h2 \set_tag_name new_element_ptr tag \\<^sub>h h3" and disc_nodes: "h3 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes" and h': "h3 \ set_disconnected_nodes document_ptr (cast new_element_ptr # disc_nodes) \\<^sub>h h'" using assms(4) by(auto simp add: create_element_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated]) have "h \ create_element document_ptr tag \\<^sub>r new_element_ptr" using new_element_ptr h2 h3 disc_nodes h' apply(auto simp add: create_element_def intro!: bind_returns_result_I bind_pure_returns_result_I[OF get_disconnected_nodes_pure])[1] apply (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust) by (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust) have "preserved (get_M ptr getter) h h2" using h2 new_element_ptr apply(rule new_element_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t) using new_element_ptr assms(6) \h \ create_element document_ptr tag \\<^sub>r new_element_ptr\ by simp have "preserved (get_M ptr getter) h2 h3" using set_tag_name_writes h3 apply(rule reads_writes_preserved2) apply(auto simp add: set_tag_name_locs_impl a_set_tag_name_locs_def all_args_def)[1] by (metis (no_types, lifting) \h \ create_element document_ptr tag \\<^sub>r new_element_ptr\ assms(6) get_M_Element_preserved8 select_result_I2) have "document_ptr |\| document_ptr_kinds h" using create_element_document_in_heap using assms(4) by blast then have "ptr \ (cast document_ptr)" using assms(5) assms(1) assms(2) assms(3) local.get_dom_component_ok local.get_dom_component_ptr by auto have "preserved (get_M ptr getter) h3 h'" using set_disconnected_nodes_writes h' apply(rule reads_writes_preserved2) apply(auto simp add: set_disconnected_nodes_locs_def all_args_def)[1] by (metis \ptr \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr\ get_M_Mdocument_preserved3) show ?thesis using \preserved (get_M ptr getter) h h2\ \preserved (get_M ptr getter) h2 h3\ \preserved (get_M ptr getter) h3 h'\ by(auto simp add: preserved_def) qed lemma create_element_is_weakly_dom_component_safe: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ create_element document_ptr tag \\<^sub>r result" assumes "h \ create_element document_ptr tag \\<^sub>h h'" shows "is_weakly_dom_component_safe {cast document_ptr} {cast result} h h'" proof - obtain new_element_ptr h2 h3 disc_nodes_h3 where new_element_ptr: "h \ new_element \\<^sub>r new_element_ptr" and h2: "h \ new_element \\<^sub>h h2" and h3: "h2 \ set_tag_name new_element_ptr tag \\<^sub>h h3" and disc_nodes_h3: "h3 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" and h': "h3 \ set_disconnected_nodes document_ptr (cast new_element_ptr # disc_nodes_h3) \\<^sub>h h'" using assms(5) by(auto simp add: create_element_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] ) then have "h \ create_element document_ptr tag \\<^sub>r new_element_ptr" apply(auto simp add: create_element_def intro!: bind_returns_result_I)[1] apply (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust) apply (metis is_OK_returns_heap_E is_OK_returns_result_I local.get_disconnected_nodes_pure pure_returns_heap_eq) by (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust) have "new_element_ptr \ set |h \ element_ptr_kinds_M|\<^sub>r" using new_element_ptr ElementMonad.ptr_kinds_ptr_kinds_M h2 using new_element_ptr_not_in_heap by blast then have "cast new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r" by simp then have "cast new_element_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r" by simp have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\| {|cast new_element_ptr|}" using new_element_new_ptr h2 new_element_ptr by blast then have node_ptr_kinds_eq_h: "node_ptr_kinds h2 = node_ptr_kinds h |\| {|cast new_element_ptr|}" apply(simp add: node_ptr_kinds_def) by force then have element_ptr_kinds_eq_h: "element_ptr_kinds h2 = element_ptr_kinds h |\| {|new_element_ptr|}" apply(simp add: element_ptr_kinds_def) by force have character_data_ptr_kinds_eq_h: "character_data_ptr_kinds h2 = character_data_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: node_ptr_kinds_def character_data_ptr_kinds_def) have document_ptr_kinds_eq_h: "document_ptr_kinds h2 = document_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: document_ptr_kinds_def) have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_tag_name_writes h3]) using set_tag_name_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2" using object_ptr_kinds_eq_h2 by(auto simp add: node_ptr_kinds_def) have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_disconnected_nodes_writes h']) using set_disconnected_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3" using object_ptr_kinds_eq_h3 by(auto simp add: node_ptr_kinds_def) have "known_ptr (cast new_element_ptr)" using \h \ create_element document_ptr tag \\<^sub>r new_element_ptr\ local.create_element_known_ptr by blast then have "known_ptrs h2" using known_ptrs_new_ptr object_ptr_kinds_eq_h \known_ptrs h\ h2 by blast then have "known_ptrs h3" using known_ptrs_preserved object_ptr_kinds_eq_h2 by blast then have "known_ptrs h'" using known_ptrs_preserved object_ptr_kinds_eq_h3 by blast have "document_ptr |\| document_ptr_kinds h" using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2 get_disconnected_nodes_ptr_in_heap \type_wf h\ document_ptr_kinds_def by (metis is_OK_returns_result_I) have children_eq_h: "\(ptr'::(_) object_ptr) children. ptr' \ cast new_element_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads h2 get_child_nodes_new_element[rotated, OF new_element_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have children_eq2_h: "\ptr'. ptr' \ cast new_element_ptr \ |h \ get_child_nodes ptr'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have "h2 \ get_child_nodes (cast new_element_ptr) \\<^sub>r []" using new_element_ptr h2 new_element_ptr_in_heap[OF h2 new_element_ptr] new_element_is_element_ptr[OF new_element_ptr] new_element_no_child_nodes by blast have disconnected_nodes_eq_h: "\doc_ptr disc_nodes. h \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads h2 get_disconnected_nodes_new_element[OF new_element_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have disconnected_nodes_eq2_h: "\doc_ptr. |h \ get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. h2 \ get_child_nodes ptr' \\<^sub>r children = h3 \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads set_tag_name_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_tag_name_get_child_nodes) then have children_eq2_h2: "\ptr'. |h2 \ get_child_nodes ptr'|\<^sub>r = |h3 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_tag_name_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_tag_name_get_disconnected_nodes) then have disconnected_nodes_eq2_h2: "\doc_ptr. |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have "type_wf h2" using \type_wf h\ new_element_types_preserved h2 by blast then have "type_wf h3" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_tag_name_writes h3] using set_tag_name_types_preserved by(auto simp add: reflp_def transp_def) then have "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h'] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) have children_eq_h3: "\ptr' children. h3 \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_disconnected_nodes_get_child_nodes) then have children_eq2_h3: "\ptr'. |h3 \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h3: "\doc_ptr disc_nodes. document_ptr \ doc_ptr \ h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers) then have disconnected_nodes_eq2_h3: "\doc_ptr. document_ptr \ doc_ptr \ |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have disc_nodes_document_ptr_h2: "h2 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" using disconnected_nodes_eq_h2 disc_nodes_h3 by auto then have disc_nodes_document_ptr_h: "h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" using disconnected_nodes_eq_h by auto then have "cast new_element_ptr \ set disc_nodes_h3" using \heap_is_wellformed h\ using \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ a_all_ptrs_in_heap_def heap_is_wellformed_def using NodeMonad.ptr_kinds_ptr_kinds_M local.heap_is_wellformed_disc_nodes_in_heap by blast have "acyclic (parent_child_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def acyclic_heap_def) also have "parent_child_rel h = parent_child_rel h2" proof(auto simp add: parent_child_rel_def)[1] fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h2" by (simp add: object_ptr_kinds_eq_h) next fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "x \ set |h2 \ get_child_nodes a|\<^sub>r" by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2_h) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h" using object_ptr_kinds_eq_h \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ by(auto) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "x \ set |h \ get_child_nodes a|\<^sub>r" by (metis (no_types, lifting) \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ children_eq2_h empty_iff empty_set image_eqI select_result_I2) qed also have "\ = parent_child_rel h3" by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2) also have "\ = parent_child_rel h'" by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3) finally have "a_acyclic_heap h'" by (simp add: acyclic_heap_def) have "a_all_ptrs_in_heap h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) then have "a_all_ptrs_in_heap h2" apply(auto simp add: a_all_ptrs_in_heap_def)[1] apply (metis \known_ptrs h2\ \parent_child_rel h = parent_child_rel h2\ \type_wf h2\ assms funion_iff local.get_child_nodes_ok local.known_ptrs_known_ptr local.parent_child_rel_child_in_heap local.parent_child_rel_child_nodes2 node_ptr_kinds_commutes node_ptr_kinds_eq_h returns_result_select_result) by (metis assms disconnected_nodes_eq2_h document_ptr_kinds_eq_h funion_iff local.get_disconnected_nodes_ok local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq_h returns_result_select_result) then have "a_all_ptrs_in_heap h3" by (simp add: children_eq2_h2 disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 local.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 object_ptr_kinds_eq_h2) then have "a_all_ptrs_in_heap h'" by (smt \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ children_eq2_h3 disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 finite_set_in h' is_OK_returns_result_I l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes local.a_all_ptrs_in_heap_def local.get_child_nodes_ptr_in_heap local.l_set_disconnected_nodes_get_disconnected_nodes_axioms node_ptr_kinds_commutes object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 select_result_I2 set_ConsD subset_code(1)) have "\p. p |\| object_ptr_kinds h \ cast new_element_ptr \ set |h \ get_child_nodes p|\<^sub>r" using \heap_is_wellformed h\ \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ heap_is_wellformed_children_in_heap by (meson NodeMonad.ptr_kinds_ptr_kinds_M a_all_ptrs_in_heap_def assms fset_mp fset_of_list_elem get_child_nodes_ok known_ptrs_known_ptr returns_result_select_result) then have "\p. p |\| object_ptr_kinds h2 \ cast new_element_ptr \ set |h2 \ get_child_nodes p|\<^sub>r" using children_eq2_h apply(auto simp add: object_ptr_kinds_eq_h)[1] using \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ apply auto[1] by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\) then have "\p. p |\| object_ptr_kinds h3 \ cast new_element_ptr \ set |h3 \ get_child_nodes p|\<^sub>r" using object_ptr_kinds_eq_h2 children_eq2_h2 by auto then have new_element_ptr_not_in_any_children: "\p. p |\| object_ptr_kinds h' \ cast new_element_ptr \ set |h' \ get_child_nodes p|\<^sub>r" using object_ptr_kinds_eq_h3 children_eq2_h3 by auto have "a_distinct_lists h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) then have "a_distinct_lists h2" using \h2 \ get_child_nodes (cast new_element_ptr) \\<^sub>r []\ apply(auto simp add: a_distinct_lists_def object_ptr_kinds_eq_h document_ptr_kinds_eq_h disconnected_nodes_eq2_h intro!: distinct_concat_map_I)[1] - apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert) + apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert_remove) apply(case_tac "x=cast new_element_ptr") apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply (metis IntI assms empty_iff local.get_child_nodes_ok local.heap_is_wellformed_one_parent local.known_ptrs_known_ptr returns_result_select_result) apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] by (metis \local.a_distinct_lists h\ \type_wf h2\ disconnected_nodes_eq_h document_ptr_kinds_eq_h local.distinct_lists_no_parent local.get_disconnected_nodes_ok returns_result_select_result) then have "a_distinct_lists h3" by(auto simp add: a_distinct_lists_def disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 children_eq2_h2 object_ptr_kinds_eq_h2) then have "a_distinct_lists h'" proof(auto simp add: a_distinct_lists_def disconnected_nodes_eq2_h3 children_eq2_h3 object_ptr_kinds_eq_h3 document_ptr_kinds_eq_h3 intro!: distinct_concat_map_I)[1] fix x assume "distinct (concat (map (\document_ptr. |h3 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" and "x |\| document_ptr_kinds h3" then show "distinct |h' \ get_disconnected_nodes x|\<^sub>r" using document_ptr_kinds_eq_h3 disconnected_nodes_eq_h3 h' set_disconnected_nodes_get_disconnected_nodes by (metis (no_types, lifting) \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set disc_nodes_h3\ \a_distinct_lists h3\ \type_wf h'\ disc_nodes_h3 distinct.simps(2) distinct_lists_disconnected_nodes get_disconnected_nodes_ok returns_result_eq returns_result_select_result) next fix x y xa assume "distinct (concat (map (\document_ptr. |h3 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" and "x |\| document_ptr_kinds h3" and "y |\| document_ptr_kinds h3" and "x \ y" and "xa \ set |h' \ get_disconnected_nodes x|\<^sub>r" and "xa \ set |h' \ get_disconnected_nodes y|\<^sub>r" moreover have "set |h3 \ get_disconnected_nodes x|\<^sub>r \ set |h3 \ get_disconnected_nodes y|\<^sub>r = {}" using calculation by(auto dest: distinct_concat_map_E(1)) ultimately show "False" apply(-) apply(cases "x = document_ptr") apply (smt NodeMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ \local.a_all_ptrs_in_heap h\ disc_nodes_h3 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 finite_set_in h' l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms select_result_I2 set_ConsD subsetD) by (smt NodeMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ \local.a_all_ptrs_in_heap h\ disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 finite_set_in h' l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms select_result_I2 set_ConsD subsetD) next fix x xa xb assume 2: "(\x\fset (object_ptr_kinds h3). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h3). set |h3 \ get_disconnected_nodes x|\<^sub>r) = {}" and 3: "xa |\| object_ptr_kinds h3" and 4: "x \ set |h' \ get_child_nodes xa|\<^sub>r" and 5: "xb |\| document_ptr_kinds h3" and 6: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" show "False" using disc_nodes_document_ptr_h disconnected_nodes_eq2_h3 apply - apply(cases "xb = document_ptr") apply (metis (no_types, hide_lams) "3" "4" "6" \\p. p |\| object_ptr_kinds h3 \ cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h3 \ get_child_nodes p|\<^sub>r\ \a_distinct_lists h3\ children_eq2_h3 disc_nodes_h3 distinct_lists_no_parent h' select_result_I2 set_ConsD set_disconnected_nodes_get_disconnected_nodes) by (metis "3" "4" "5" "6" \a_distinct_lists h3\ \type_wf h3\ children_eq2_h3 distinct_lists_no_parent get_disconnected_nodes_ok returns_result_select_result) qed have "a_owner_document_valid h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) then have "a_owner_document_valid h'" using disc_nodes_h3 \document_ptr |\| document_ptr_kinds h\ apply(auto simp add: a_owner_document_valid_def)[1] apply(auto simp add: object_ptr_kinds_eq_h object_ptr_kinds_eq_h3 )[1] apply(auto simp add: object_ptr_kinds_eq_h2)[1] apply(auto simp add: document_ptr_kinds_eq_h document_ptr_kinds_eq_h3 )[1] apply(auto simp add: document_ptr_kinds_eq_h2)[1] apply(auto simp add: node_ptr_kinds_eq_h node_ptr_kinds_eq_h3 )[1] apply(auto simp add: node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h )[1] apply(auto simp add: children_eq2_h2[symmetric] children_eq2_h3[symmetric] disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3)[1] apply (metis (no_types, lifting) document_ptr_kinds_eq_h h' list.set_intros(1) local.set_disconnected_nodes_get_disconnected_nodes select_result_I2) apply(simp add: object_ptr_kinds_eq_h) by(metis (no_types, lifting) NodeMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ children_eq2_h children_eq2_h2 children_eq2_h3 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h finite_set_in h' l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes list.set_intros(2) local.l_set_disconnected_nodes_get_disconnected_nodes_axioms node_ptr_kinds_commutes select_result_I2) have "parent_child_rel h = parent_child_rel h'" proof - have "parent_child_rel h = parent_child_rel h2" proof(auto simp add: parent_child_rel_def)[1] fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h2" by (simp add: object_ptr_kinds_eq_h) next fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "x \ set |h2 \ get_child_nodes a|\<^sub>r" by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast new_element_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2_h) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h" using object_ptr_kinds_eq_h \h2 \ get_child_nodes (cast new_element_ptr) \\<^sub>r []\ by(auto) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "x \ set |h \ get_child_nodes a|\<^sub>r" by (metis (no_types, lifting) \h2 \ get_child_nodes (cast new_element_ptr) \\<^sub>r []\ children_eq2_h empty_iff empty_set image_eqI select_result_I2) qed also have "\ = parent_child_rel h3" by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2) also have "\ = parent_child_rel h'" by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3) finally show ?thesis by simp qed have root: "h \ get_root_node (cast document_ptr) \\<^sub>r cast document_ptr" by (simp add: \document_ptr |\| document_ptr_kinds h\ local.get_root_node_not_node_same) then have root': "h' \ get_root_node (cast document_ptr) \\<^sub>r cast document_ptr" by (simp add: \document_ptr |\| document_ptr_kinds h\ document_ptr_kinds_eq_h local.get_root_node_not_node_same object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3) have "heap_is_wellformed h'" and "known_ptrs h'" using create_element_preserves_wellformedness assms by blast+ have "cast result |\| object_ptr_kinds h" using \cast new_element_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ by (metis (full_types) ObjectMonad.ptr_kinds_ptr_kinds_M \h \ create_element document_ptr tag \\<^sub>r new_element_ptr\ assms(4) returns_result_eq) obtain to where to: "h \ to_tree_order (cast document_ptr) \\<^sub>r to" by (meson \document_ptr |\| document_ptr_kinds h\ assms(1) assms(2) assms(3) document_ptr_kinds_commutes is_OK_returns_result_E local.to_tree_order_ok) then have "h \ a_get_dom_component (cast document_ptr) \\<^sub>r to" using root by(auto simp add: a_get_dom_component_def) moreover obtain to' where to': "h' \ to_tree_order (cast document_ptr) \\<^sub>r to'" by (meson \heap_is_wellformed h'\ \known_ptrs h'\ \type_wf h'\ is_OK_returns_result_E local.get_root_node_root_in_heap local.to_tree_order_ok root') then have "h' \ a_get_dom_component (cast document_ptr) \\<^sub>r to'" using root' by(auto simp add: a_get_dom_component_def) moreover have "\child. child \ set to \ child \ set to'" by (metis \heap_is_wellformed h'\ \known_ptrs h'\ \parent_child_rel h = parent_child_rel h'\ \type_wf h'\ assms(1) assms(2) assms(3) local.to_tree_order_parent_child_rel to to') ultimately have "set |h \ local.a_get_dom_component (cast document_ptr)|\<^sub>r = set |h' \ local.a_get_dom_component (cast document_ptr)|\<^sub>r" by(auto simp add: a_get_dom_component_def) show ?thesis apply(auto simp add: is_weakly_dom_component_safe_def Let_def)[1] using \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ assms(2) assms(3) children_eq_h local.get_child_nodes_ok local.get_child_nodes_ptr_in_heap local.known_ptrs_known_ptr object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 returns_result_select_result apply (metis is_OK_returns_result_I) apply (metis \h \ create_element document_ptr tag \\<^sub>r new_element_ptr\ assms(4) element_ptr_kinds_commutes h2 new_element_ptr new_element_ptr_in_heap node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h3 returns_result_eq returns_result_heap_def) using \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r result |\| object_ptr_kinds h\ element_ptr_kinds_commutes node_ptr_kinds_commutes apply blast using assms(1) assms(2) assms(3) \h \ create_element document_ptr tag \\<^sub>h h'\ apply(rule create_element_is_weakly_dom_component_safe_step) apply (simp add: local.get_dom_component_impl) using \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ \h \ create_element document_ptr tag \\<^sub>r new_element_ptr\ by auto qed end interpretation i_get_dom_component_create_element?: l_get_dom_component_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr heap_is_wellformed parent_child_rel type_wf known_ptrs to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs set_tag_name set_tag_name_locs create_element by(auto simp add: l_get_dom_component_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_dom_component_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \create\_character\_data\ lemma create_character_data_not_strongly_dom_component_safe: obtains h :: "('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap" and h' and document_ptr and create_character_datanew_character_data_ptr and tag where "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and "h \ create_character_data document_ptr tag \\<^sub>r create_character_datanew_character_data_ptr \\<^sub>h h'" and "\ is_strongly_dom_component_safe {cast document_ptr} {cast create_character_datanew_character_data_ptr} h h'" proof - let ?h0 = "Heap fmempty ::('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap" let ?P = "create_document" let ?h1 = "|?h0 \ ?P|\<^sub>h" let ?document_ptr = "|?h0 \ ?P|\<^sub>r" show thesis apply(rule that[where h="?h1" and document_ptr="?document_ptr"]) by code_simp+ qed locale l_get_dom_component_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name + l_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs set_val set_val_locs type_wf create_character_data known_ptr + l_get_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_disconnected_nodes get_disconnected_nodes_locs + l_set_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_disconnected_nodes set_disconnected_nodes_locs + l_set_val\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_val set_val_locs + l_create_character_data_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed parent_child_rel set_val set_val_locs set_disconnected_nodes set_disconnected_nodes_locs create_character_data known_ptrs for known_ptr :: "(_::linorder) object_ptr \ bool" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and type_wf :: "(_) heap \ bool" and known_ptrs :: "(_) heap \ bool" and to_tree_order :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr list) prog" and get_parent :: "(_) node_ptr \ ((_) heap, exception, (_) object_ptr option) prog" and get_parent_locs :: "((_) heap \ (_) heap \ bool) set" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_dom_component :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr list) prog" and is_strongly_dom_component_safe :: "(_) object_ptr set \ (_) object_ptr set \ (_) heap \ (_) heap \ bool" and is_weakly_dom_component_safe :: "(_) object_ptr set \ (_) object_ptr set \ (_) heap \ (_) heap \ bool" and get_root_node :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr) prog" and get_root_node_locs :: "((_) heap \ (_) heap \ bool) set" and get_ancestors :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr list) prog" and get_ancestors_locs :: "((_) heap \ (_) heap \ bool) set" and get_element_by_id :: "(_) object_ptr \ char list \ ((_) heap, exception, (_) element_ptr option) prog" and get_elements_by_class_name :: "(_) object_ptr \ char list \ ((_) heap, exception, (_) element_ptr list) prog" and get_elements_by_tag_name :: "(_) object_ptr \ char list \ ((_) heap, exception, (_) element_ptr list) prog" and set_val :: "(_) character_data_ptr \ char list \ ((_) heap, exception, unit) prog" and set_val_locs :: "(_) character_data_ptr \ ((_) heap, exception, unit) prog set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" and create_character_data :: "(_) document_ptr \ char list \ ((_) heap, exception, (_) character_data_ptr) prog" begin lemma create_character_data_is_weakly_dom_component_safe_step: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ create_character_data document_ptr text \\<^sub>h h'" assumes "ptr \ set |h \ get_dom_component (cast document_ptr)|\<^sub>r" assumes "ptr \ cast |h \ create_character_data document_ptr text|\<^sub>r" shows "preserved (get_M ptr getter) h h'" proof - obtain new_character_data_ptr h2 h3 disc_nodes where new_character_data_ptr: "h \ new_character_data \\<^sub>r new_character_data_ptr" and h2: "h \ new_character_data \\<^sub>h h2" and h3: "h2 \ set_val new_character_data_ptr text \\<^sub>h h3" and disc_nodes: "h3 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes" and h': "h3 \ set_disconnected_nodes document_ptr (cast new_character_data_ptr # disc_nodes) \\<^sub>h h'" using assms(4) by(auto simp add: create_character_data_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated]) have "h \ create_character_data document_ptr text \\<^sub>r new_character_data_ptr" using new_character_data_ptr h2 h3 disc_nodes h' apply(auto simp add: create_character_data_def intro!: bind_returns_result_I bind_pure_returns_result_I[OF get_disconnected_nodes_pure])[1] apply (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust) by (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust) have "preserved (get_M ptr getter) h h2" using h2 new_character_data_ptr apply(rule new_character_data_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t) using new_character_data_ptr assms(6) \h \ create_character_data document_ptr text \\<^sub>r new_character_data_ptr\ by simp have "preserved (get_M ptr getter) h2 h3" using set_val_writes h3 apply(rule reads_writes_preserved2) apply(auto simp add: set_val_locs_impl a_set_val_locs_def all_args_def)[1] by (metis (mono_tags) CharacterData_simp11 \h \ create_character_data document_ptr text \\<^sub>r new_character_data_ptr\ assms(4) assms(6) is_OK_returns_heap_I is_OK_returns_result_E returns_result_eq select_result_I2) have "document_ptr |\| document_ptr_kinds h" using create_character_data_document_in_heap using assms(4) by blast then have "ptr \ (cast document_ptr)" using assms(5) assms(1) assms(2) assms(3) local.get_dom_component_ok local.get_dom_component_ptr by auto have "preserved (get_M ptr getter) h3 h'" using set_disconnected_nodes_writes h' apply(rule reads_writes_preserved2) apply(auto simp add: set_disconnected_nodes_locs_def all_args_def)[1] by (metis \ptr \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr\ get_M_Mdocument_preserved3) show ?thesis using \preserved (get_M ptr getter) h h2\ \preserved (get_M ptr getter) h2 h3\ \preserved (get_M ptr getter) h3 h'\ by(auto simp add: preserved_def) qed lemma create_character_data_is_weakly_dom_component_safe: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ create_character_data document_ptr text \\<^sub>r result" assumes "h \ create_character_data document_ptr text \\<^sub>h h'" shows "is_weakly_dom_component_safe {cast document_ptr} {cast result} h h'" proof - obtain new_character_data_ptr h2 h3 disc_nodes_h3 where new_character_data_ptr: "h \ new_character_data \\<^sub>r new_character_data_ptr" and h2: "h \ new_character_data \\<^sub>h h2" and h3: "h2 \ set_val new_character_data_ptr text \\<^sub>h h3" and disc_nodes_h3: "h3 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" and h': "h3 \ set_disconnected_nodes document_ptr (cast new_character_data_ptr # disc_nodes_h3) \\<^sub>h h'" using assms(5) by(auto simp add: create_character_data_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] ) then have "h \ create_character_data document_ptr text \\<^sub>r new_character_data_ptr" apply(auto simp add: create_character_data_def intro!: bind_returns_result_I)[1] apply (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust) apply (metis is_OK_returns_heap_E is_OK_returns_result_I local.get_disconnected_nodes_pure pure_returns_heap_eq) by (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust) have "new_character_data_ptr \ set |h \ character_data_ptr_kinds_M|\<^sub>r" using new_character_data_ptr CharacterDataMonad.ptr_kinds_ptr_kinds_M h2 using new_character_data_ptr_not_in_heap by blast then have "cast new_character_data_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r" by simp then have "cast new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r" by simp have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\| {|cast new_character_data_ptr|}" using new_character_data_new_ptr h2 new_character_data_ptr by blast then have node_ptr_kinds_eq_h: "node_ptr_kinds h2 = node_ptr_kinds h |\| {|cast new_character_data_ptr|}" apply(simp add: node_ptr_kinds_def) by force then have character_data_ptr_kinds_eq_h: "character_data_ptr_kinds h2 = character_data_ptr_kinds h |\| {|new_character_data_ptr|}" apply(simp add: character_data_ptr_kinds_def) by force have element_ptr_kinds_eq_h: "element_ptr_kinds h2 = element_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: node_ptr_kinds_def element_ptr_kinds_def) have document_ptr_kinds_eq_h: "document_ptr_kinds h2 = document_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: document_ptr_kinds_def) have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_val_writes h3]) using set_val_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2" using object_ptr_kinds_eq_h2 by(auto simp add: node_ptr_kinds_def) have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_disconnected_nodes_writes h']) using set_disconnected_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3" using object_ptr_kinds_eq_h3 by(auto simp add: node_ptr_kinds_def) have "document_ptr |\| document_ptr_kinds h" using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2 get_disconnected_nodes_ptr_in_heap \type_wf h\ document_ptr_kinds_def by (metis is_OK_returns_result_I) have children_eq_h: "\(ptr'::(_) object_ptr) children. ptr' \ cast new_character_data_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads h2 get_child_nodes_new_character_data[rotated, OF new_character_data_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have children_eq2_h: "\ptr'. ptr' \ cast new_character_data_ptr \ |h \ get_child_nodes ptr'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\| {|cast new_character_data_ptr|}" using new_character_data_new_ptr h2 new_character_data_ptr by blast then have node_ptr_kinds_eq_h: "node_ptr_kinds h2 = node_ptr_kinds h |\| {|cast new_character_data_ptr|}" apply(simp add: node_ptr_kinds_def) by force then have character_data_ptr_kinds_eq_h: "character_data_ptr_kinds h2 = character_data_ptr_kinds h |\| {|new_character_data_ptr|}" apply(simp add: character_data_ptr_kinds_def) by force have element_ptr_kinds_eq_h: "element_ptr_kinds h2 = element_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: node_ptr_kinds_def element_ptr_kinds_def) have document_ptr_kinds_eq_h: "document_ptr_kinds h2 = document_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: document_ptr_kinds_def) have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_val_writes h3]) using set_val_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2" using object_ptr_kinds_eq_h2 by(auto simp add: node_ptr_kinds_def) have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_disconnected_nodes_writes h']) using set_disconnected_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3" using object_ptr_kinds_eq_h3 by(auto simp add: node_ptr_kinds_def) have "document_ptr |\| document_ptr_kinds h" using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2 get_disconnected_nodes_ptr_in_heap \type_wf h\ document_ptr_kinds_def by (metis is_OK_returns_result_I) have children_eq_h: "\(ptr'::(_) object_ptr) children. ptr' \ cast new_character_data_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads h2 get_child_nodes_new_character_data[rotated, OF new_character_data_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have children_eq2_h: "\ptr'. ptr' \ cast new_character_data_ptr \ |h \ get_child_nodes ptr'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have "h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []" using new_character_data_ptr h2 new_character_data_ptr_in_heap[OF h2 new_character_data_ptr] new_character_data_is_character_data_ptr[OF new_character_data_ptr] new_character_data_no_child_nodes by blast have disconnected_nodes_eq_h: "\doc_ptr disc_nodes. h \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads h2 get_disconnected_nodes_new_character_data[OF new_character_data_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have disconnected_nodes_eq2_h: "\doc_ptr. |h \ get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. h2 \ get_child_nodes ptr' \\<^sub>r children = h3 \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads set_val_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_val_get_child_nodes) then have children_eq2_h2: "\ptr'. |h2 \ get_child_nodes ptr'|\<^sub>r = |h3 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_val_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_val_get_disconnected_nodes) then have disconnected_nodes_eq2_h2: "\doc_ptr. |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have "type_wf h2" using \type_wf h\ new_character_data_types_preserved h2 by blast then have "type_wf h3" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_val_writes h3] using set_val_types_preserved by(auto simp add: reflp_def transp_def) then have "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h'] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) have children_eq_h3: "\ptr' children. h3 \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_disconnected_nodes_get_child_nodes) then have children_eq2_h3: " \ptr'. |h3 \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h3: "\doc_ptr disc_nodes. document_ptr \ doc_ptr \ h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers) then have disconnected_nodes_eq2_h3: "\doc_ptr. document_ptr \ doc_ptr \ |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have disc_nodes_document_ptr_h2: "h2 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" using disconnected_nodes_eq_h2 disc_nodes_h3 by auto then have disc_nodes_document_ptr_h: "h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" using disconnected_nodes_eq_h by auto then have "cast new_character_data_ptr \ set disc_nodes_h3" using \heap_is_wellformed h\ using \cast new_character_data_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ a_all_ptrs_in_heap_def heap_is_wellformed_def using NodeMonad.ptr_kinds_ptr_kinds_M local.heap_is_wellformed_disc_nodes_in_heap by blast have "parent_child_rel h = parent_child_rel h'" proof - have "parent_child_rel h = parent_child_rel h2" proof(auto simp add: parent_child_rel_def)[1] fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h2" by (simp add: object_ptr_kinds_eq_h) next fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "x \ set |h2 \ get_child_nodes a|\<^sub>r" by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2_h) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h" using object_ptr_kinds_eq_h \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ by(auto) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "x \ set |h \ get_child_nodes a|\<^sub>r" by (metis (no_types, lifting) \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ children_eq2_h empty_iff empty_set image_eqI select_result_I2) qed also have "\ = parent_child_rel h3" by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2) also have "\ = parent_child_rel h'" by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3) finally show ?thesis by simp qed have root: "h \ get_root_node (cast document_ptr) \\<^sub>r cast document_ptr" by (simp add: \document_ptr |\| document_ptr_kinds h\ local.get_root_node_not_node_same) then have root': "h' \ get_root_node (cast document_ptr) \\<^sub>r cast document_ptr" by (simp add: \document_ptr |\| document_ptr_kinds h\ document_ptr_kinds_eq_h local.get_root_node_not_node_same object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3) have "heap_is_wellformed h'" and "known_ptrs h'" using create_character_data_preserves_wellformedness assms by blast+ have "cast result |\| object_ptr_kinds h" using \cast new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ by (metis (full_types) ObjectMonad.ptr_kinds_ptr_kinds_M \h \ create_character_data document_ptr text \\<^sub>r new_character_data_ptr\ assms(4) returns_result_eq) obtain to where to: "h \ to_tree_order (cast document_ptr) \\<^sub>r to" by (meson \document_ptr |\| document_ptr_kinds h\ assms(1) assms(2) assms(3) document_ptr_kinds_commutes is_OK_returns_result_E local.to_tree_order_ok) then have "h \ a_get_dom_component (cast document_ptr) \\<^sub>r to" using root by(auto simp add: a_get_dom_component_def) moreover obtain to' where to': "h' \ to_tree_order (cast document_ptr) \\<^sub>r to'" by (meson \heap_is_wellformed h'\ \known_ptrs h'\ \type_wf h'\ is_OK_returns_result_E local.get_root_node_root_in_heap local.to_tree_order_ok root') then have "h' \ a_get_dom_component (cast document_ptr) \\<^sub>r to'" using root' by(auto simp add: a_get_dom_component_def) moreover have "\child. child \ set to \ child \ set to'" by (metis \heap_is_wellformed h'\ \known_ptrs h'\ \parent_child_rel h = parent_child_rel h'\ \type_wf h'\ assms(1) assms(2) assms(3) local.to_tree_order_parent_child_rel to to') ultimately have "set |h \ local.a_get_dom_component (cast document_ptr)|\<^sub>r = set |h' \ local.a_get_dom_component (cast document_ptr)|\<^sub>r" by(auto simp add: a_get_dom_component_def) show ?thesis apply(auto simp add: is_weakly_dom_component_safe_def Let_def)[1] using assms(2) assms(3) children_eq_h local.get_child_nodes_ok local.get_child_nodes_ptr_in_heap local.known_ptrs_known_ptr object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 returns_result_select_result apply (metis \h2 \ get_child_nodes (cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_character_data_ptr) \\<^sub>r []\ is_OK_returns_result_I) apply (metis \h \ create_character_data document_ptr text \\<^sub>r new_character_data_ptr\ assms(4) character_data_ptr_kinds_commutes h2 new_character_data_ptr new_character_data_ptr_in_heap node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h3 returns_result_eq) using \h \ create_character_data document_ptr text \\<^sub>r new_character_data_ptr\ \new_character_data_ptr \ set |h \ character_data_ptr_kinds_M|\<^sub>r\ assms(4) returns_result_eq apply fastforce using assms(2) assms(3) children_eq_h local.get_child_nodes_ok local.get_child_nodes_ptr_in_heap local.known_ptrs_known_ptr object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 returns_result_select_result apply (smt ObjectMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ \h \ create_character_data document_ptr text \\<^sub>r new_character_data_ptr\ assms(1) assms(5) create_character_data_is_weakly_dom_component_safe_step local.get_dom_component_impl select_result_I2) done qed end interpretation i_get_dom_component_create_character_data?: l_get_dom_component_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr heap_is_wellformed parent_child_rel type_wf known_ptrs to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name set_val set_val_locs get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs create_character_data by(auto simp add: l_get_dom_component_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_dom_component_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \create\_document\ lemma create_document_unsafe: "\(\(h :: ('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap ) h' new_document_ptr. heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ create_document \\<^sub>r new_document_ptr \ h \ create_document \\<^sub>h h' \ is_strongly_dom_component_safe {} {cast new_document_ptr} h h')" proof - obtain h document_ptr where h: "Inr (document_ptr, h) = (Heap (fmempty) :: ('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap ) \ (do { document_ptr \ create_document; return (document_ptr) })" by(code_simp, auto simp add: equal_eq List.member_def)+ then obtain h' new_document_ptr where "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and h': "h \ create_document \\<^sub>r new_document_ptr" and h': "h \ create_document \\<^sub>h h'" and "\(is_strongly_dom_component_safe {} {cast new_document_ptr} h h')" by(code_simp, auto simp add: equal_eq List.member_def)+ then show ?thesis by blast qed locale l_get_dom_component_create_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name + l_create_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M create_document for known_ptr :: "(_::linorder) object_ptr \ bool" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and type_wf :: "(_) heap \ bool" and known_ptrs :: "(_) heap \ bool" and to_tree_order :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr list) prog" and get_parent :: "(_) node_ptr \ ((_) heap, exception, (_) object_ptr option) prog" and get_parent_locs :: "((_) heap \ (_) heap \ bool) set" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_dom_component :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr list) prog" and is_strongly_dom_component_safe :: "(_) object_ptr set \ (_) object_ptr set \ (_) heap \ (_) heap \ bool" and is_weakly_dom_component_safe :: "(_) object_ptr set \ (_) object_ptr set \ (_) heap \ (_) heap \ bool" and get_root_node :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr) prog" and get_root_node_locs :: "((_) heap \ (_) heap \ bool) set" and get_ancestors :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr list) prog" and get_ancestors_locs :: "((_) heap \ (_) heap \ bool) set" and get_element_by_id :: "(_) object_ptr \ char list \ ((_) heap, exception, (_) element_ptr option) prog" and get_elements_by_class_name :: "(_) object_ptr \ char list \ ((_) heap, exception, (_) element_ptr list) prog" and get_elements_by_tag_name :: "(_) object_ptr \ char list \ ((_) heap, exception, (_) element_ptr list) prog" and create_document :: "((_) heap, exception, (_) document_ptr) prog" begin lemma create_document_is_weakly_dom_component_safe_step: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ create_document \\<^sub>h h'" assumes "ptr \ cast |h \ create_document|\<^sub>r" shows "preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr getter) h h'" using assms apply(auto simp add: create_document_def)[1] by (metis assms(4) assms(5) is_OK_returns_heap_I local.create_document_def new_document_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t select_result_I) lemma create_document_is_weakly_dom_component_safe: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ create_document \\<^sub>r result" assumes "h \ create_document \\<^sub>h h'" shows "is_weakly_dom_component_safe {} {cast result} h h'" proof - have "object_ptr_kinds h' = object_ptr_kinds h |\| {|cast result|}" using assms(4) assms(5) local.create_document_def new_document_new_ptr by auto moreover have "result |\| document_ptr_kinds h" using assms(4) assms(5) local.create_document_def new_document_ptr_not_in_heap by auto ultimately show ?thesis using assms apply(auto simp add: is_weakly_dom_component_safe_def Let_def local.create_document_def new_document_ptr_not_in_heap)[1] by (metis \result |\| document_ptr_kinds h\ document_ptr_kinds_commutes new_document_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t) qed end interpretation i_get_dom_component_create_document?: l_get_dom_component_create_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr heap_is_wellformed parent_child_rel type_wf known_ptrs to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name create_document by(auto simp add: l_get_dom_component_create_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_dom_component_create_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \insert\_before\ lemma insert_before_unsafe: "\(\(h :: ('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap ) h' ptr child. heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ insert_before ptr child None \\<^sub>h h' \ is_weakly_dom_component_safe {ptr, cast child} {} h h')" proof - obtain h document_ptr e1 e2 where h: "Inr ((document_ptr, e1, e2), h) = (Heap (fmempty) :: ('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap ) \ (do { document_ptr \ create_document; e1 \ create_element document_ptr ''div''; e2 \ create_element document_ptr ''div''; return (document_ptr, e1, e2) })" by(code_simp, auto simp add: equal_eq List.member_def)+ then obtain h' where "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and h': "h \ insert_before (cast e1) (cast e2) None \\<^sub>h h'" and "\(is_weakly_dom_component_safe {cast e1, cast e2} {} h h')" by(code_simp, auto simp add: equal_eq List.member_def)+ then show ?thesis by auto qed lemma insert_before_unsafe2: "\(\(h :: ('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap ) h' ptr child ref. heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ insert_before ptr child (Some ref) \\<^sub>h h' \ is_weakly_dom_component_safe {ptr, cast child, cast ref} {} h h')" proof - obtain h document_ptr e1 e2 e3 where h: "Inr ((document_ptr, e1, e2, e3), h) = (Heap (fmempty) :: ('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap ) \ (do { document_ptr \ create_document; e1 \ create_element document_ptr ''div''; e2 \ create_element document_ptr ''div''; e3 \ create_element document_ptr ''div''; append_child (cast e1) (cast e2); return (document_ptr, e1, e2, e3) })" by(code_simp, auto simp add: equal_eq List.member_def)+ then obtain h' where "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and h': "h \ insert_before (cast e1) (cast e3) (Some (cast e2)) \\<^sub>h h'" and "\(is_weakly_dom_component_safe {cast e1, cast e3, cast e2} {} h h')" apply(code_simp) apply(clarify) by(code_simp, auto simp add: equal_eq List.member_def)+ then show ?thesis by fast qed lemma append_child_unsafe: obtains h :: "('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap" and h' and ptr and child where "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and "h \ append_child ptr child \\<^sub>h h'" and "\ is_weakly_dom_component_safe {ptr, cast child} {} h h'" proof - let ?h0 = "Heap fmempty ::('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap" let ?P = "do { document_ptr \ create_document; e1 \ create_element document_ptr ''div''; e2 \ create_element document_ptr ''div''; return (e1, e2) }" let ?h1 = "|?h0 \ ?P|\<^sub>h" let ?e1 = "fst |?h0 \ ?P|\<^sub>r" let ?e2 = "snd |?h0 \ ?P|\<^sub>r" show thesis apply(rule that[where h="?h1" and ptr="cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ?e1" and child="cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r ?e2"]) by code_simp+ qed subsubsection \get\_owner\_document\ lemma get_owner_document_unsafe: obtains h :: "('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap" and h' and ptr and owner_document where "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and "h \ get_owner_document ptr \\<^sub>r owner_document \\<^sub>h h'" and "\is_weakly_dom_component_safe {ptr} {cast owner_document} h h'" proof - let ?h0 = "Heap fmempty ::('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap" let ?P = "do { document_ptr \ create_document; e1 \ create_element document_ptr ''div''; return (document_ptr, e1) }" let ?h1 = "|?h0 \ ?P|\<^sub>h" let ?document_ptr = "fst |?h0 \ ?P|\<^sub>r" let ?e1 = "snd |?h0 \ ?P|\<^sub>r" show thesis apply(rule that[where h="?h1" and ptr="cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ?e1" and owner_document="?document_ptr"]) by code_simp+ qed end diff --git a/thys/Shadow_DOM/Shadow_DOM.thy b/thys/Shadow_DOM/Shadow_DOM.thy --- a/thys/Shadow_DOM/Shadow_DOM.thy +++ b/thys/Shadow_DOM/Shadow_DOM.thy @@ -1,10501 +1,10501 @@ (*********************************************************************************** * Copyright (c) 2016-2020 The University of Sheffield, UK * 2019-2020 University of Exeter, UK * * 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. * * 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 HOLDER 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. * * SPDX-License-Identifier: BSD-2-Clause ***********************************************************************************) section\The Shadow DOM\ theory Shadow_DOM imports "monads/ShadowRootMonad" Core_DOM.Core_DOM begin abbreviation "safe_shadow_root_element_types \ {''article'', ''aside'', ''blockquote'', ''body'', ''div'', ''footer'', ''h1'', ''h2'', ''h3'', ''h4'', ''h5'', ''h6'', ''header'', ''main'', ''nav'', ''p'', ''section'', ''span''}" subsection \Function Definitions\ subsubsection \get\_child\_nodes\ locale l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = CD: l_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs begin definition get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) shadow_root_ptr \ unit \ (_, (_) node_ptr list) dom_prog" where "get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr _ = get_M shadow_root_ptr RShadowRoot.child_nodes" definition a_get_child_nodes_tups :: "(((_) object_ptr \ bool) \ ((_) object_ptr \ unit \ (_, (_) node_ptr list) dom_prog)) list" where "a_get_child_nodes_tups \ [(is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r, get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast)]" definition a_get_child_nodes :: "(_) object_ptr \ (_, (_) node_ptr list) dom_prog" where "a_get_child_nodes ptr = invoke (CD.a_get_child_nodes_tups @ a_get_child_nodes_tups) ptr ()" definition a_get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" where "a_get_child_nodes_locs ptr \ (if is_shadow_root_ptr_kind ptr then {preserved (get_M (the (cast ptr)) RShadowRoot.child_nodes)} else {}) \ CD.a_get_child_nodes_locs ptr" definition first_child :: "(_) object_ptr \ (_, (_) node_ptr option) dom_prog" where "first_child ptr = do { children \ a_get_child_nodes ptr; return (case children of [] \ None | child#_ \ Some child)}" end global_interpretation l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs defines get_child_nodes = l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_child_nodes and get_child_nodes_locs = l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_child_nodes_locs . locale l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_type_wf type_wf + l_known_ptr known_ptr + l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_get_child_nodes_defs get_child_nodes get_child_nodes_locs + CD: l_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M for type_wf :: "(_) heap \ bool" and known_ptr :: "(_) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ bool" and get_child_nodes :: "(_) object_ptr \ (_, (_) node_ptr list) dom_prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ (_, (_) node_ptr list) dom_prog" and get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" + assumes known_ptr_impl: "known_ptr = ShadowRootClass.known_ptr" assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" assumes get_child_nodes_impl: "get_child_nodes = a_get_child_nodes" assumes get_child_nodes_locs_impl: "get_child_nodes_locs = a_get_child_nodes_locs" begin lemmas get_child_nodes_def = get_child_nodes_impl[unfolded a_get_child_nodes_def get_child_nodes_def] lemmas get_child_nodes_locs_def = get_child_nodes_locs_impl[unfolded a_get_child_nodes_locs_def get_child_nodes_locs_def, folded CD.get_child_nodes_locs_impl] lemma get_child_nodes_ok: assumes "known_ptr ptr" assumes "type_wf h" assumes "ptr |\| object_ptr_kinds h" shows "h \ ok (get_child_nodes ptr)" using assms[unfolded known_ptr_impl type_wf_impl] apply(auto simp add: get_child_nodes_def)[1] apply(split CD.get_child_nodes_splits, rule conjI)+ using ShadowRootClass.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t CD.get_child_nodes_ok CD.known_ptr_impl CD.type_wf_impl apply blast apply(auto simp add: CD.known_ptr_impl a_get_child_nodes_tups_def get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ok dest!: known_ptr_new_shadow_root_ptr intro!: bind_is_OK_I2)[1] by (metis is_shadow_root_ptr_kind_none l_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_lemmas.get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ok l_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_lemmas_axioms option.case_eq_if shadow_root_ptr_casts_commute3 shadow_root_ptr_kinds_commutes) lemma get_child_nodes_ptr_in_heap: assumes "h \ get_child_nodes ptr \\<^sub>r children" shows "ptr |\| object_ptr_kinds h" using assms by(auto simp add: get_child_nodes_def invoke_ptr_in_heap dest: is_OK_returns_result_I) lemma get_child_nodes_pure [simp]: "pure (get_child_nodes ptr) h" unfolding get_child_nodes_def a_get_child_nodes_tups_def proof(split CD.get_child_nodes_splits, rule conjI; clarify) assume "known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M ptr" then show "pure (get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M ptr) h" by simp next assume "\ known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M ptr" then show "pure (invoke [(is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r, get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r)] ptr ()) h" by(auto simp add: get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro: bind_pure_I split: invoke_splits) qed lemma get_child_nodes_reads: "reads (get_child_nodes_locs ptr) (get_child_nodes ptr) h h'" apply (simp add: get_child_nodes_def a_get_child_nodes_tups_def get_child_nodes_locs_def CD.get_child_nodes_locs_def) apply(split CD.get_child_nodes_splits, rule conjI)+ apply(auto intro!: reads_subset[OF CD.get_child_nodes_reads[unfolded CD.get_child_nodes_locs_def]] split: if_splits)[1] apply(split invoke_splits, rule conjI)+ apply(auto)[1] apply(auto simp add: get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro: reads_subset[OF reads_singleton] reads_subset[OF check_in_heap_reads] intro!: reads_bind_pure reads_subset[OF return_reads] split: option.splits)[1] done end interpretation i_get_child_nodes?: l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs by(simp add: l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_child_nodes_is_l_get_child_nodes [instances]: "l_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs" apply(auto simp add: l_get_child_nodes_def instances)[1] using get_child_nodes_reads get_child_nodes_ok get_child_nodes_ptr_in_heap get_child_nodes_pure by blast+ paragraph \new\_document\ locale l_new_document_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD: l_new_document_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M for type_wf :: "(_) heap \ bool" and known_ptr :: "(_) object_ptr \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ bool" and get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_child_nodes_new_document: "ptr' \ cast new_document_ptr \ h \ new_document \\<^sub>r new_document_ptr \ h \ new_document \\<^sub>h h' \ r \ get_child_nodes_locs ptr' \ r h h'" apply(auto simp add: get_child_nodes_locs_def)[1] using CD.get_child_nodes_new_document apply (metis document_ptr_casts_commute3 empty_iff is_document_ptr_kind_none new_document_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t option.case_eq_if shadow_root_ptr_casts_commute3 singletonD) by (simp add: CD.get_child_nodes_new_document) lemma new_document_no_child_nodes: "h \ new_document \\<^sub>r new_document_ptr \ h \ new_document \\<^sub>h h' \ h' \ get_child_nodes (cast new_document_ptr) \\<^sub>r []" apply(auto simp add: get_child_nodes_def)[1] apply(split CD.get_child_nodes_splits, rule conjI)+ using CD.new_document_no_child_nodes apply auto[1] by(auto simp add: DocumentClass.a_known_ptr_def CD.known_ptr_impl known_ptr_def dest!: new_document_is_document_ptr) end interpretation i_new_document_get_child_nodes?: l_new_document_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr get_child_nodes get_child_nodes_locs DocumentClass.type_wf DocumentClass.known_ptr Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs by(unfold_locales) declare l_new_document_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma new_document_get_child_nodes_is_l_new_document_get_child_nodes [instances]: "l_new_document_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs" using new_document_is_l_new_document get_child_nodes_is_l_get_child_nodes apply(simp add: l_new_document_get_child_nodes_def l_new_document_get_child_nodes_axioms_def) using get_child_nodes_new_document new_document_no_child_nodes by fast paragraph \new\_shadow\_root\ locale l_new_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M for type_wf :: "(_) heap \ bool" and known_ptr :: "(_) object_ptr \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ bool" and get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_child_nodes_new_shadow_root: "ptr' \ cast new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ r \ get_child_nodes_locs ptr' \ r h h'" apply(auto simp add: get_child_nodes_locs_def)[1] apply (metis document_ptr_casts_commute3 insert_absorb insert_not_empty is_document_ptr_kind_none new_shadow_root_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t option.case_eq_if shadow_root_ptr_casts_commute3 singletonD) apply(auto simp add: CD.get_child_nodes_locs_def)[1] using new_shadow_root_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t apply blast apply (smt insertCI new_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t singleton_iff) apply (metis document_ptr_casts_commute3 empty_iff new_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t singletonD) done lemma new_shadow_root_no_child_nodes: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ h' \ get_child_nodes (cast new_shadow_root_ptr) \\<^sub>r []" apply(auto simp add: get_child_nodes_def )[1] apply(split CD.get_child_nodes_splits, rule conjI)+ apply(auto simp add: CD.get_child_nodes_def CD.a_get_child_nodes_tups_def)[1] apply(split invoke_splits, rule conjI)+ using NodeClass.a_known_ptr_def known_ptr_not_character_data_ptr known_ptr_not_document_ptr known_ptr_not_element_ptr local.CD.known_ptr_impl apply blast apply(auto simp add: is_document_ptr_def split: option.splits document_ptr.splits)[1] apply(auto simp add: is_character_data_ptr_def split: option.splits document_ptr.splits)[1] apply(auto simp add: is_element_ptr_def split: option.splits document_ptr.splits)[1] apply(auto simp add: a_get_child_nodes_tups_def)[1] apply(split invoke_splits, rule conjI)+ apply(auto simp add: is_shadow_root_ptr_def split: shadow_root_ptr.splits dest!: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_is_shadow_root_ptr)[1] apply(auto intro!: bind_pure_returns_result_I)[1] apply(drule(1) new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ptr_in_heap) apply(auto simp add: shadow_root_ptr_kinds_def document_ptr_kinds_def)[1] apply (metis check_in_heap_ptr_in_heap is_OK_returns_result_E old.unit.exhaust) using new_shadow_root_children by (simp add: new_shadow_root_children get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def) end interpretation i_new_shadow_root_get_child_nodes?: l_new_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr get_child_nodes get_child_nodes_locs DocumentClass.type_wf DocumentClass.known_ptr Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs by(unfold_locales) declare l_new_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def[instances] locale l_new_shadow_root_get_child_nodes = l_get_child_nodes + assumes get_child_nodes_new_shadow_root: "ptr' \ cast new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ r \ get_child_nodes_locs ptr' \ r h h'" assumes new_shadow_root_no_child_nodes: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ h' \ get_child_nodes (cast new_shadow_root_ptr) \\<^sub>r []" lemma new_shadow_root_get_child_nodes_is_l_new_shadow_root_get_child_nodes [instances]: "l_new_shadow_root_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs" apply(simp add: l_new_shadow_root_get_child_nodes_def l_new_shadow_root_get_child_nodes_axioms_def instances) using get_child_nodes_new_shadow_root new_shadow_root_no_child_nodes by fast paragraph \new\_element\ locale l_new_element_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_new_element_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_child_nodes_new_element: "ptr' \ cast new_element_ptr \ h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ r \ get_child_nodes_locs ptr' \ r h h'" by (auto simp add: get_child_nodes_locs_def CD.get_child_nodes_locs_def new_element_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t new_element_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t new_element_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t new_element_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t split: prod.splits if_splits option.splits elim!: bind_returns_result_E bind_returns_heap_E intro: is_element_ptr_kind_obtains) lemma new_element_no_child_nodes: "h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ h' \ get_child_nodes (cast new_element_ptr) \\<^sub>r []" apply(auto simp add: get_child_nodes_def a_get_child_nodes_tups_def split: prod.splits elim!: bind_returns_result_E bind_returns_heap_E)[1] apply(split CD.get_child_nodes_splits, rule conjI)+ using local.new_element_no_child_nodes apply auto[1] apply(auto simp add: invoke_def)[1] apply(auto simp add: new_element_ptr_in_heap get_child_nodes\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def check_in_heap_def new_element_child_nodes intro!: bind_pure_returns_result_I intro: new_element_is_element_ptr elim!: new_element_ptr_in_heap)[1] proof - assume " h \ new_element \\<^sub>r new_element_ptr" assume "h \ new_element \\<^sub>h h'" assume "\ known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr)" moreover have "known_ptr (cast new_element_ptr)" using new_element_is_element_ptr \h \ new_element \\<^sub>r new_element_ptr\ by(auto simp add: known_ptr_impl ShadowRootClass.a_known_ptr_def DocumentClass.a_known_ptr_def CharacterDataClass.a_known_ptr_def ElementClass.a_known_ptr_def) ultimately show "False" by(simp add: known_ptr_impl CD.known_ptr_impl ShadowRootClass.a_known_ptr_def is_document_ptr_kind_none) qed end interpretation i_new_element_get_child_nodes?: l_new_element_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs by(unfold_locales) declare l_new_element_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma new_element_get_child_nodes_is_l_new_element_get_child_nodes [instances]: "l_new_element_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs" using new_element_is_l_new_element get_child_nodes_is_l_get_child_nodes apply(auto simp add: l_new_element_get_child_nodes_def l_new_element_get_child_nodes_axioms_def)[1] using get_child_nodes_new_element new_element_no_child_nodes by fast+ subsubsection \delete\_shadow\_root\ locale l_delete_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_child_nodes_delete_shadow_root: "ptr' \ cast shadow_root_ptr \ h \ delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M shadow_root_ptr \\<^sub>h h' \ r \ get_child_nodes_locs ptr' \ r h h'" by(auto simp add: get_child_nodes_locs_def CD.get_child_nodes_locs_def delete_shadow_root_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t delete_shadow_root_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t delete_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t delete_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t split: if_splits option.splits intro: is_shadow_root_ptr_kind_obtains) end locale l_delete_shadow_root_get_child_nodes = l_get_child_nodes_defs + assumes get_child_nodes_delete_shadow_root: "ptr' \ cast shadow_root_ptr \ h \ delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M shadow_root_ptr \\<^sub>h h' \ r \ get_child_nodes_locs ptr' \ r h h'" interpretation l_delete_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs by(auto simp add: l_delete_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) lemma l_delete_shadow_root_get_child_nodes_get_child_nodes_locs [instances]: "l_delete_shadow_root_get_child_nodes get_child_nodes_locs" apply(auto simp add: l_delete_shadow_root_get_child_nodes_def)[1] using get_child_nodes_delete_shadow_root apply fast done subsubsection \set\_child\_nodes\ locale l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = CD: l_set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs begin definition set_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) shadow_root_ptr \ (_) node_ptr list \ (_, unit) dom_prog" where "set_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr = put_M shadow_root_ptr RShadowRoot.child_nodes_update" definition a_set_child_nodes_tups :: "(((_) object_ptr \ bool) \ ((_) object_ptr \ (_) node_ptr list \ (_, unit) dom_prog)) list" where "a_set_child_nodes_tups \ [(is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r, set_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast)]" definition a_set_child_nodes :: "(_) object_ptr \ (_) node_ptr list \ (_, unit) dom_prog" where "a_set_child_nodes ptr children = invoke (CD.a_set_child_nodes_tups @ a_set_child_nodes_tups) ptr children" definition a_set_child_nodes_locs :: "(_) object_ptr \ (_, unit) dom_prog set" where "a_set_child_nodes_locs ptr \ (if is_shadow_root_ptr_kind ptr then all_args (put_M (the (cast ptr)) RShadowRoot.child_nodes_update) else {}) \ CD.a_set_child_nodes_locs ptr" end global_interpretation l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs defines set_child_nodes = l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_set_child_nodes and set_child_nodes_locs = l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_set_child_nodes_locs . locale l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_type_wf type_wf + l_known_ptr known_ptr + l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_set_child_nodes_defs set_child_nodes set_child_nodes_locs + CD: l_set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M for type_wf :: "(_) heap \ bool" and known_ptr :: "(_) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ bool" and set_child_nodes :: "(_) object_ptr \ (_) node_ptr list \ (_, unit) dom_prog" and set_child_nodes_locs :: "(_) object_ptr \ (_, unit) dom_prog set" and set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ (_) node_ptr list \ (_, unit) dom_prog" and set_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ (_, unit) dom_prog set" + assumes known_ptr_impl: "known_ptr = ShadowRootClass.known_ptr" assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" assumes set_child_nodes_impl: "set_child_nodes = a_set_child_nodes" assumes set_child_nodes_locs_impl: "set_child_nodes_locs = a_set_child_nodes_locs" begin lemmas set_child_nodes_def = set_child_nodes_impl[unfolded a_set_child_nodes_def set_child_nodes_def] lemmas set_child_nodes_locs_def = set_child_nodes_locs_impl[unfolded a_set_child_nodes_locs_def set_child_nodes_locs_def, folded CD.set_child_nodes_locs_impl] lemma set_child_nodes_writes: "writes (set_child_nodes_locs ptr) (set_child_nodes ptr children) h h'" apply (simp add: set_child_nodes_def a_set_child_nodes_tups_def set_child_nodes_locs_def) apply(split CD.set_child_nodes_splits, rule conjI)+ apply (simp add: CD.set_child_nodes_writes writes_union_right_I) apply(split invoke_splits, rule conjI)+ apply(auto simp add: a_set_child_nodes_def)[1] apply(auto simp add: set_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: writes_bind_pure intro: writes_union_right_I writes_union_left_I split: list.splits)[1] by (simp add: is_shadow_root_ptr_kind_none) lemma set_child_nodes_pointers_preserved: assumes "w \ set_child_nodes_locs object_ptr" assumes "h \ w \\<^sub>h h'" shows "object_ptr_kinds h = object_ptr_kinds h'" using assms(1) object_ptr_kinds_preserved[OF writes_singleton2 assms(2)] by(auto simp add: all_args_def set_child_nodes_locs_def CD.set_child_nodes_locs_def split: if_splits) lemma set_child_nodes_types_preserved: assumes "w \ set_child_nodes_locs object_ptr" assumes "h \ w \\<^sub>h h'" shows "type_wf h = type_wf h'" using assms(1) type_wf_preserved[OF writes_singleton2 assms(2)] by(auto simp add: all_args_def type_wf_impl a_set_child_nodes_tups_def set_child_nodes_locs_def CD.set_child_nodes_locs_def split: if_splits option.splits) end interpretation i_set_child_nodes?: l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr set_child_nodes set_child_nodes_locs Core_DOM_Functions.set_child_nodes Core_DOM_Functions.set_child_nodes_locs apply(unfold_locales) by (auto simp add: set_child_nodes_def set_child_nodes_locs_def) declare l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_child_nodes_is_l_set_child_nodes [instances]: "l_set_child_nodes type_wf set_child_nodes set_child_nodes_locs" using instances Shadow_DOM.i_set_child_nodes.set_child_nodes_pointers_preserved Shadow_DOM.i_set_child_nodes.set_child_nodes_writes i_set_child_nodes.set_child_nodes_types_preserved unfolding l_set_child_nodes_def by blast paragraph \get\_child\_nodes\ locale l_set_child_nodes_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_child_nodes set_child_nodes_locs set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + CD: l_set_child_nodes_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M for type_wf :: "(_) heap \ bool" and known_ptr :: "(_) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and set_child_nodes :: "(_) object_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_child_nodes_locs :: "(_) object_ptr \ ((_) heap, exception, unit) prog set" and set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap, exception, unit) prog set" begin lemma set_child_nodes_get_child_nodes: assumes "known_ptr ptr" assumes "type_wf h" assumes "h \ set_child_nodes ptr children \\<^sub>h h'" shows "h' \ get_child_nodes ptr \\<^sub>r children" proof - have "h \ check_in_heap ptr \\<^sub>r ()" using assms set_child_nodes_def invoke_ptr_in_heap by (metis (full_types) check_in_heap_ptr_in_heap is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust) then have ptr_in_h: "ptr |\| object_ptr_kinds h" by (simp add: check_in_heap_ptr_in_heap is_OK_returns_result_I) have "type_wf h'" apply(unfold type_wf_impl) apply(rule subst[where P=id, OF type_wf_preserved[OF set_child_nodes_writes assms(3), unfolded all_args_def], simplified]) by(auto simp add: all_args_def assms(2)[unfolded type_wf_impl] set_child_nodes_locs_def CD.set_child_nodes_locs_def split: if_splits) have "h' \ check_in_heap ptr \\<^sub>r ()" using check_in_heap_reads set_child_nodes_writes assms(3) \h \ check_in_heap ptr \\<^sub>r ()\ apply(rule reads_writes_separate_forwards) apply(auto simp add: all_args_def set_child_nodes_locs_def CD.set_child_nodes_locs_def)[1] done then have "ptr |\| object_ptr_kinds h'" using check_in_heap_ptr_in_heap by blast with assms ptr_in_h \type_wf h'\ show ?thesis apply(auto simp add: type_wf_impl known_ptr_impl get_child_nodes_def a_get_child_nodes_tups_def set_child_nodes_def a_set_child_nodes_tups_def del: bind_pure_returns_result_I2 intro!: bind_pure_returns_result_I2)[1] apply(split CD.get_child_nodes_splits, (rule conjI impI)+)+ apply(split CD.set_child_nodes_splits)+ apply(auto simp add: CD.set_child_nodes_get_child_nodes type_wf_impl CD.type_wf_impl dest: ShadowRootClass.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t)[1] apply(auto simp add: CD.set_child_nodes_get_child_nodes type_wf_impl CD.type_wf_impl dest: ShadowRootClass.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t)[1] apply(split CD.set_child_nodes_splits)+ by(auto simp add: known_ptr_impl CD.known_ptr_impl set_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def CD.type_wf_impl ShadowRootClass.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t dest: known_ptr_new_shadow_root_ptr)[2] qed lemma set_child_nodes_get_child_nodes_different_pointers: assumes "ptr \ ptr'" assumes "w \ set_child_nodes_locs ptr" assumes "h \ w \\<^sub>h h'" assumes "r \ get_child_nodes_locs ptr'" shows "r h h'" using assms apply(auto simp add: set_child_nodes_locs_def CD.set_child_nodes_locs_def get_child_nodes_locs_def CD.get_child_nodes_locs_def)[1] by(auto simp add: all_args_def elim!: is_document_ptr_kind_obtains is_shadow_root_ptr_kind_obtains is_element_ptr_kind_obtains split: if_splits option.splits) end interpretation i_set_child_nodes_get_child_nodes?: l_set_child_nodes_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs set_child_nodes set_child_nodes_locs Core_DOM_Functions.set_child_nodes Core_DOM_Functions.set_child_nodes_locs using instances by(auto simp add: l_set_child_nodes_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def ) declare l_set_child_nodes_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_child_nodes_get_child_nodes_is_l_set_child_nodes_get_child_nodes [instances]: "l_set_child_nodes_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs set_child_nodes set_child_nodes_locs" apply(auto simp add: instances l_set_child_nodes_get_child_nodes_def l_set_child_nodes_get_child_nodes_axioms_def)[1] using set_child_nodes_get_child_nodes apply fast using set_child_nodes_get_child_nodes_different_pointers apply fast done subsubsection \set\_tag\_type\ locale l_set_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD: l_set_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_tag_name set_tag_name_locs + l_type_wf type_wf for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and set_tag_name :: "(_) element_ptr \ tag_name \ (_, unit) dom_prog" and set_tag_name_locs :: "(_) element_ptr \ (_, unit) dom_prog set" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" begin lemmas set_tag_name_def = CD.set_tag_name_impl[unfolded CD.a_set_tag_name_def set_tag_name_def] lemmas set_tag_name_locs_def = CD.set_tag_name_locs_impl[unfolded CD.a_set_tag_name_locs_def set_tag_name_locs_def] lemma set_tag_name_ok: "type_wf h \ element_ptr |\| element_ptr_kinds h \ h \ ok (set_tag_name element_ptr tag)" apply(unfold type_wf_impl) unfolding set_tag_name_impl[unfolded a_set_tag_name_def] using get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ok put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ok using CD.set_tag_name_ok CD.type_wf_impl ShadowRootClass.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t by blast lemma set_tag_name_writes: "writes (set_tag_name_locs element_ptr) (set_tag_name element_ptr tag) h h'" using CD.set_tag_name_writes . lemma set_tag_name_pointers_preserved: assumes "w \ set_tag_name_locs element_ptr" assumes "h \ w \\<^sub>h h'" shows "object_ptr_kinds h = object_ptr_kinds h'" using assms by(simp add: CD.set_tag_name_pointers_preserved) lemma set_tag_name_typess_preserved: assumes "w \ set_tag_name_locs element_ptr" assumes "h \ w \\<^sub>h h'" shows "type_wf h = type_wf h'" apply(unfold type_wf_impl) apply(rule type_wf_preserved[OF writes_singleton2 assms(2)]) using assms(1) set_tag_name_locs_def by(auto simp add: all_args_def set_tag_name_locs_def split: if_splits) end interpretation i_set_tag_name?: l_set_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_tag_name set_tag_name_locs by(auto simp add: l_set_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_set_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_set_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma set_tag_name_is_l_set_tag_name [instances]: "l_set_tag_name type_wf set_tag_name set_tag_name_locs" apply(auto simp add: l_set_tag_name_def)[1] using set_tag_name_writes apply fast using set_tag_name_ok apply fast using set_tag_name_pointers_preserved apply (fast, fast) using set_tag_name_typess_preserved apply (fast, fast) done paragraph \get\_child\_nodes\ locale l_set_tag_name_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + CD: l_set_tag_name_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_tag_name set_tag_name_locs known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_tag_name_get_child_nodes: "\w \ set_tag_name_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_child_nodes_locs ptr'. r h h'))" apply(auto simp add: get_child_nodes_locs_def)[1] apply(auto simp add: set_tag_name_locs_def all_args_def)[1] using CD.set_tag_name_get_child_nodes apply(blast) using CD.set_tag_name_get_child_nodes apply(blast) done end interpretation i_set_tag_name_get_child_nodes?: l_set_tag_name_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_tag_name set_tag_name_locs known_ptr DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs by unfold_locales declare l_set_tag_name_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_tag_name_get_child_nodes_is_l_set_tag_name_get_child_nodes [instances]: "l_set_tag_name_get_child_nodes type_wf set_tag_name set_tag_name_locs known_ptr get_child_nodes get_child_nodes_locs" using set_tag_name_is_l_set_tag_name get_child_nodes_is_l_get_child_nodes apply(simp add: l_set_tag_name_get_child_nodes_def l_set_tag_name_get_child_nodes_axioms_def) using set_tag_name_get_child_nodes by fast subsubsection \get\_shadow\_root\ locale l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs begin definition a_get_shadow_root :: "(_) element_ptr \ (_, (_) shadow_root_ptr option) dom_prog" where "a_get_shadow_root element_ptr = get_M element_ptr shadow_root_opt" definition a_get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" where "a_get_shadow_root_locs element_ptr \ {preserved (get_M element_ptr shadow_root_opt)}" end global_interpretation l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs defines get_shadow_root = a_get_shadow_root and get_shadow_root_locs = a_get_shadow_root_locs . locale l_get_shadow_root_defs = fixes get_shadow_root :: "(_) element_ptr \ (_, (_) shadow_root_ptr option) dom_prog" fixes get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" locale l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_get_shadow_root_defs get_shadow_root get_shadow_root_locs + l_type_wf type_wf for type_wf :: "(_) heap \ bool" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" assumes get_shadow_root_impl: "get_shadow_root = a_get_shadow_root" assumes get_shadow_root_locs_impl: "get_shadow_root_locs = a_get_shadow_root_locs" begin lemmas get_shadow_root_def = get_shadow_root_impl[unfolded get_shadow_root_def a_get_shadow_root_def] lemmas get_shadow_root_locs_def = get_shadow_root_locs_impl[unfolded get_shadow_root_locs_def a_get_shadow_root_locs_def] lemma get_shadow_root_ok: "type_wf h \ element_ptr |\| element_ptr_kinds h \ h \ ok (get_shadow_root element_ptr)" unfolding get_shadow_root_def type_wf_impl using ShadowRootMonad.get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ok by blast lemma get_shadow_root_pure [simp]: "pure (get_shadow_root element_ptr) h" unfolding get_shadow_root_def by simp lemma get_shadow_root_ptr_in_heap: assumes "h \ get_shadow_root element_ptr \\<^sub>r children" shows "element_ptr |\| element_ptr_kinds h" using assms by(auto simp add: get_shadow_root_def get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap dest: is_OK_returns_result_I) lemma get_shadow_root_reads: "reads (get_shadow_root_locs element_ptr) (get_shadow_root element_ptr) h h'" by(simp add: get_shadow_root_def get_shadow_root_locs_def reads_bind_pure reads_insert_writes_set_right) end interpretation i_get_shadow_root?: l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs using instances by (auto simp add: l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_get_shadow_root = l_type_wf + l_get_shadow_root_defs + assumes get_shadow_root_reads: "reads (get_shadow_root_locs element_ptr) (get_shadow_root element_ptr) h h'" assumes get_shadow_root_ok: "type_wf h \ element_ptr |\| element_ptr_kinds h \ h \ ok (get_shadow_root element_ptr)" assumes get_shadow_root_ptr_in_heap: "h \ ok (get_shadow_root element_ptr) \ element_ptr |\| element_ptr_kinds h" assumes get_shadow_root_pure [simp]: "pure (get_shadow_root element_ptr) h" lemma get_shadow_root_is_l_get_shadow_root [instances]: "l_get_shadow_root type_wf get_shadow_root get_shadow_root_locs" using instances unfolding l_get_shadow_root_def by (metis (no_types, lifting) ElementClass.l_type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_axioms get_shadow_root_ok get_shadow_root_pure get_shadow_root_reads l_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas.get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap l_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas.intro l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_shadow_root_def) paragraph \set\_disconnected\_nodes\ locale l_set_disconnected_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_disconnected_nodes set_disconnected_nodes_locs + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma set_disconnected_nodes_get_shadow_root: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" by(auto simp add: set_disconnected_nodes_locs_def get_shadow_root_locs_def all_args_def) end locale l_set_disconnected_nodes_get_shadow_root = l_set_disconnected_nodes_defs + l_get_shadow_root_defs + assumes set_disconnected_nodes_get_shadow_root: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" interpretation i_set_disconnected_nodes_get_shadow_root?: l_set_disconnected_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_disconnected_nodes set_disconnected_nodes_locs get_shadow_root get_shadow_root_locs by(auto simp add: l_set_disconnected_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_disconnected_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_disconnected_nodes_get_shadow_root_is_l_set_disconnected_nodes_get_shadow_root [instances]: "l_set_disconnected_nodes_get_shadow_root set_disconnected_nodes_locs get_shadow_root_locs" apply(auto simp add: l_set_disconnected_nodes_get_shadow_root_def)[1] using set_disconnected_nodes_get_shadow_root apply fast done paragraph \set\_tag\_type\ locale l_set_tag_name_get_shadow_root\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_tag_name_get_shadow_root: "\w \ set_tag_name_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" by(auto simp add: set_tag_name_locs_def get_shadow_root_locs_def all_args_def intro: element_put_get_preserved[where setter=tag_name_update and getter=shadow_root_opt]) end locale l_set_tag_name_get_shadow_root = l_set_tag_name + l_get_shadow_root + assumes set_tag_name_get_shadow_root: "\w \ set_tag_name_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" interpretation i_set_tag_name_get_shadow_root?: l_set_tag_name_get_shadow_root\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_tag_name set_tag_name_locs get_shadow_root get_shadow_root_locs apply(auto simp add: l_set_tag_name_get_shadow_root\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)[1] using l_set_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms by unfold_locales declare l_set_tag_name_get_shadow_root\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_tag_name_get_shadow_root_is_l_set_tag_name_get_shadow_root [instances]: "l_set_tag_name_get_shadow_root type_wf set_tag_name set_tag_name_locs get_shadow_root get_shadow_root_locs" using set_tag_name_is_l_set_tag_name get_shadow_root_is_l_get_shadow_root apply(simp add: l_set_tag_name_get_shadow_root_def l_set_tag_name_get_shadow_root_axioms_def) using set_tag_name_get_shadow_root by fast paragraph \set\_child\_nodes\ locale l_set_child_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_child_nodes set_child_nodes_locs set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs for type_wf :: "(_) heap \ bool" and known_ptr :: "(_) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ bool" and set_child_nodes :: "(_) object_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_child_nodes_locs :: "(_) object_ptr \ ((_) heap, exception, unit) prog set" and set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap, exception, unit) prog set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma set_child_nodes_get_shadow_root: "\w \ set_child_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" apply(auto simp add: set_child_nodes_locs_def get_shadow_root_locs_def CD.set_child_nodes_locs_def all_args_def)[1] by(auto intro!: element_put_get_preserved[where getter=shadow_root_opt and setter=RElement.child_nodes_update]) end locale l_set_child_nodes_get_shadow_root = l_set_child_nodes_defs + l_get_shadow_root_defs + assumes set_child_nodes_get_shadow_root: "\w \ set_child_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" interpretation i_set_child_nodes_get_shadow_root?: l_set_child_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr set_child_nodes set_child_nodes_locs Core_DOM_Functions.set_child_nodes Core_DOM_Functions.set_child_nodes_locs get_shadow_root get_shadow_root_locs by(auto simp add: l_set_child_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_child_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_child_nodes_get_shadow_root_is_l_set_child_nodes_get_shadow_root [instances]: "l_set_child_nodes_get_shadow_root set_child_nodes_locs get_shadow_root_locs" apply(auto simp add: l_set_child_nodes_get_shadow_root_def)[1] using set_child_nodes_get_shadow_root apply fast done paragraph \delete\_shadow\_root\ locale l_delete_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_shadow_root_delete_shadow_root: "h \ delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M shadow_root_ptr \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" by(auto simp add: get_shadow_root_locs_def delete_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t) end locale l_delete_shadow_root_get_shadow_root = l_get_shadow_root_defs + assumes get_shadow_root_delete_shadow_root: "h \ delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M shadow_root_ptr \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" interpretation l_delete_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs by(auto simp add: l_delete_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) lemma l_delete_shadow_root_get_shadow_root_get_shadow_root_locs [instances]: "l_delete_shadow_root_get_shadow_root get_shadow_root_locs" apply(auto simp add: l_delete_shadow_root_get_shadow_root_def)[1] using get_shadow_root_delete_shadow_root apply fast done paragraph \new\_character\_data\ locale l_new_character_data_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs for type_wf :: "(_) heap \ bool" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_shadow_root_new_character_data: "h \ new_character_data \\<^sub>r new_character_data_ptr \ h \ new_character_data \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" by (auto simp add: get_shadow_root_locs_def new_character_data_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t new_character_data_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t split: prod.splits if_splits option.splits elim!: bind_returns_result_E bind_returns_heap_E intro: is_element_ptr_kind_obtains) end locale l_new_character_data_get_shadow_root = l_new_character_data + l_get_shadow_root + assumes get_shadow_root_new_character_data: "h \ new_character_data \\<^sub>r new_character_data_ptr \ h \ new_character_data \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" interpretation i_new_character_data_get_shadow_root?: l_new_character_data_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs by(unfold_locales) declare l_new_character_data_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma new_character_data_get_shadow_root_is_l_new_character_data_get_shadow_root [instances]: "l_new_character_data_get_shadow_root type_wf get_shadow_root get_shadow_root_locs" using new_character_data_is_l_new_character_data get_shadow_root_is_l_get_shadow_root apply(auto simp add: l_new_character_data_get_shadow_root_def l_new_character_data_get_shadow_root_axioms_def instances)[1] using get_shadow_root_new_character_data by fast paragraph \new\_document\ locale l_new_document_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs for type_wf :: "(_) heap \ bool" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_shadow_root_new_document: "h \ new_document \\<^sub>r new_document_ptr \ h \ new_document \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" by (auto simp add: get_shadow_root_locs_def new_document_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t new_document_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t split: prod.splits if_splits option.splits elim!: bind_returns_result_E bind_returns_heap_E intro: is_element_ptr_kind_obtains) end locale l_new_document_get_shadow_root = l_new_document + l_get_shadow_root + assumes get_shadow_root_new_document: "h \ new_document \\<^sub>r new_document_ptr \ h \ new_document \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" interpretation i_new_document_get_shadow_root?: l_new_document_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs by(unfold_locales) declare l_new_document_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma new_document_get_shadow_root_is_l_new_document_get_shadow_root [instances]: "l_new_document_get_shadow_root type_wf get_shadow_root get_shadow_root_locs" using new_document_is_l_new_document get_shadow_root_is_l_get_shadow_root apply(auto simp add: l_new_document_get_shadow_root_def l_new_document_get_shadow_root_axioms_def instances)[1] using get_shadow_root_new_document by fast paragraph \new\_element\ locale l_new_element_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs for type_wf :: "(_) heap \ bool" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_shadow_root_new_element: "ptr' \ new_element_ptr \ h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" by (auto simp add: get_shadow_root_locs_def new_element_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t new_element_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t new_element_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t split: prod.splits if_splits option.splits elim!: bind_returns_result_E bind_returns_heap_E intro: is_element_ptr_kind_obtains) lemma new_element_no_shadow_root: "h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ h' \ get_shadow_root new_element_ptr \\<^sub>r None" by(simp add: get_shadow_root_def new_element_shadow_root_opt) end locale l_new_element_get_shadow_root = l_new_element + l_get_shadow_root + assumes get_shadow_root_new_element: "ptr' \ new_element_ptr \ h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" assumes new_element_no_shadow_root: "h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ h' \ get_shadow_root new_element_ptr \\<^sub>r None" interpretation i_new_element_get_shadow_root?: l_new_element_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs by(unfold_locales) declare l_new_element_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma new_element_get_shadow_root_is_l_new_element_get_shadow_root [instances]: "l_new_element_get_shadow_root type_wf get_shadow_root get_shadow_root_locs" using new_element_is_l_new_element get_shadow_root_is_l_get_shadow_root apply(auto simp add: l_new_element_get_shadow_root_def l_new_element_get_shadow_root_axioms_def instances)[1] using get_shadow_root_new_element new_element_no_shadow_root by fast+ paragraph \new\_shadow\_root\ locale l_new_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs for type_wf :: "(_) heap \ bool" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_shadow_root_new_shadow_root: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" by (auto simp add: get_shadow_root_locs_def new_shadow_root_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t new_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t split: prod.splits if_splits option.splits elim!: bind_returns_result_E bind_returns_heap_E intro: is_element_ptr_kind_obtains) end locale l_new_shadow_root_get_shadow_root = l_get_shadow_root + assumes get_shadow_root_new_shadow_root: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" interpretation i_new_shadow_root_get_shadow_root?: l_new_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs by(unfold_locales) declare l_new_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma new_shadow_root_get_shadow_root_is_l_new_shadow_root_get_shadow_root [instances]: "l_new_shadow_root_get_shadow_root type_wf get_shadow_root get_shadow_root_locs" using get_shadow_root_is_l_get_shadow_root apply(auto simp add: l_new_shadow_root_get_shadow_root_def l_new_shadow_root_get_shadow_root_axioms_def instances)[1] using get_shadow_root_new_shadow_root by fast subsubsection \set\_shadow\_root\ locale l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs begin definition a_set_shadow_root :: "(_) element_ptr \ (_) shadow_root_ptr option \ (_, unit) dom_prog" where "a_set_shadow_root element_ptr = put_M element_ptr shadow_root_opt_update" definition a_set_shadow_root_locs :: "(_) element_ptr \ ((_, unit) dom_prog) set" where "a_set_shadow_root_locs element_ptr \ all_args (put_M element_ptr shadow_root_opt_update)" end global_interpretation l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs defines set_shadow_root = a_set_shadow_root and set_shadow_root_locs = a_set_shadow_root_locs . locale l_set_shadow_root_defs = fixes set_shadow_root :: "(_) element_ptr \ (_) shadow_root_ptr option \ (_, unit) dom_prog" fixes set_shadow_root_locs :: "(_) element_ptr \ (_, unit) dom_prog set" locale l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_type_wf type_wf + l_set_shadow_root_defs set_shadow_root set_shadow_root_locs + l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs for type_wf :: "(_) heap \ bool" and set_shadow_root :: "(_) element_ptr \ (_) shadow_root_ptr option \ (_, unit) dom_prog" and set_shadow_root_locs :: "(_) element_ptr \ (_, unit) dom_prog set" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" assumes set_shadow_root_impl: "set_shadow_root = a_set_shadow_root" assumes set_shadow_root_locs_impl: "set_shadow_root_locs = a_set_shadow_root_locs" begin lemmas set_shadow_root_def = set_shadow_root_impl[unfolded set_shadow_root_def a_set_shadow_root_def] lemmas set_shadow_root_locs_def = set_shadow_root_locs_impl[unfolded set_shadow_root_locs_def a_set_shadow_root_locs_def] lemma set_shadow_root_ok: "type_wf h \ element_ptr |\| element_ptr_kinds h \ h \ ok (set_shadow_root element_ptr tag)" apply(unfold type_wf_impl) unfolding set_shadow_root_def using get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ok put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ok by (simp add: ShadowRootMonad.put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ok) lemma set_shadow_root_ptr_in_heap: "h \ ok (set_shadow_root element_ptr shadow_root) \ element_ptr |\| element_ptr_kinds h" by(simp add: set_shadow_root_def ElementMonad.put_M_ptr_in_heap) lemma set_shadow_root_writes: "writes (set_shadow_root_locs element_ptr) (set_shadow_root element_ptr tag) h h'" by(auto simp add: set_shadow_root_def set_shadow_root_locs_def intro: writes_bind_pure) lemma set_shadow_root_pointers_preserved: assumes "w \ set_shadow_root_locs element_ptr" assumes "h \ w \\<^sub>h h'" shows "object_ptr_kinds h = object_ptr_kinds h'" using assms(1) object_ptr_kinds_preserved[OF writes_singleton2 assms(2)] by(auto simp add: all_args_def set_shadow_root_locs_def split: if_splits) lemma set_shadow_root_types_preserved: assumes "w \ set_shadow_root_locs element_ptr" assumes "h \ w \\<^sub>h h'" shows "type_wf h = type_wf h'" apply(unfold type_wf_impl) using assms(1) type_wf_preserved[OF writes_singleton2 assms(2)] by(auto simp add: all_args_def set_shadow_root_locs_def split: if_splits) end interpretation i_set_shadow_root?: l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_shadow_root set_shadow_root_locs by (auto simp add: l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_set_shadow_root = l_type_wf +l_set_shadow_root_defs + assumes set_shadow_root_writes: "writes (set_shadow_root_locs element_ptr) (set_shadow_root element_ptr disc_nodes) h h'" assumes set_shadow_root_ok: "type_wf h \ element_ptr |\| element_ptr_kinds h \ h \ ok (set_shadow_root element_ptr shadow_root)" assumes set_shadow_root_ptr_in_heap: "h \ ok (set_shadow_root element_ptr shadow_root) \ element_ptr |\| element_ptr_kinds h" assumes set_shadow_root_pointers_preserved: "w \ set_shadow_root_locs element_ptr \ h \ w \\<^sub>h h' \ object_ptr_kinds h = object_ptr_kinds h'" assumes set_shadow_root_types_preserved: "w \ set_shadow_root_locs element_ptr \ h \ w \\<^sub>h h' \ type_wf h = type_wf h'" lemma set_shadow_root_is_l_set_shadow_root [instances]: "l_set_shadow_root type_wf set_shadow_root set_shadow_root_locs" apply(auto simp add: l_set_shadow_root_def instances)[1] using set_shadow_root_writes apply blast using set_shadow_root_ok apply (blast) using set_shadow_root_ptr_in_heap apply blast using set_shadow_root_pointers_preserved apply(blast, blast) using set_shadow_root_types_preserved apply(blast, blast) done paragraph \get\_shadow\_root\ locale l_set_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_shadow_root_get_shadow_root: "type_wf h \ h \ set_shadow_root ptr shadow_root_ptr_opt \\<^sub>h h' \ h' \ get_shadow_root ptr \\<^sub>r shadow_root_ptr_opt" by(auto simp add: set_shadow_root_def get_shadow_root_def) lemma set_shadow_root_get_shadow_root_different_pointers: "ptr \ ptr' \ \w \ set_shadow_root_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" by(auto simp add: set_shadow_root_locs_def get_shadow_root_locs_def all_args_def) end interpretation i_set_shadow_root_get_shadow_root?: l_set_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_shadow_root set_shadow_root_locs get_shadow_root get_shadow_root_locs apply(auto simp add: l_set_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)[1] by(unfold_locales) declare l_set_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_set_shadow_root_get_shadow_root = l_type_wf + l_set_shadow_root_defs + l_get_shadow_root_defs + assumes set_shadow_root_get_shadow_root: "type_wf h \ h \ set_shadow_root ptr shadow_root_ptr_opt \\<^sub>h h' \ h' \ get_shadow_root ptr \\<^sub>r shadow_root_ptr_opt" assumes set_shadow_root_get_shadow_root_different_pointers: "ptr \ ptr' \ w \ set_shadow_root_locs ptr \ h \ w \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" lemma set_shadow_root_get_shadow_root_is_l_set_shadow_root_get_shadow_root [instances]: "l_set_shadow_root_get_shadow_root type_wf set_shadow_root set_shadow_root_locs get_shadow_root get_shadow_root_locs" apply(auto simp add: l_set_shadow_root_get_shadow_root_def instances)[1] using set_shadow_root_get_shadow_root apply fast using set_shadow_root_get_shadow_root_different_pointers apply fast done subsubsection \set\_mode\ locale l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs begin definition a_set_mode :: "(_) shadow_root_ptr \ shadow_root_mode \ (_, unit) dom_prog" where "a_set_mode shadow_root_ptr = put_M shadow_root_ptr mode_update" definition a_set_mode_locs :: "(_) shadow_root_ptr \ ((_, unit) dom_prog) set" where "a_set_mode_locs shadow_root_ptr \ all_args (put_M shadow_root_ptr mode_update)" end global_interpretation l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs defines set_mode = a_set_mode and set_mode_locs = a_set_mode_locs . locale l_set_mode_defs = fixes set_mode :: "(_) shadow_root_ptr \ shadow_root_mode \ (_, unit) dom_prog" fixes set_mode_locs :: "(_) shadow_root_ptr \ (_, unit) dom_prog set" locale l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_type_wf type_wf + l_set_mode_defs set_mode set_mode_locs + l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs for type_wf :: "(_) heap \ bool" and set_mode :: "(_) shadow_root_ptr \ shadow_root_mode \ (_, unit) dom_prog" and set_mode_locs :: "(_) shadow_root_ptr \ (_, unit) dom_prog set" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" assumes set_mode_impl: "set_mode = a_set_mode" assumes set_mode_locs_impl: "set_mode_locs = a_set_mode_locs" begin lemmas set_mode_def = set_mode_impl[unfolded set_mode_def a_set_mode_def] lemmas set_mode_locs_def = set_mode_locs_impl[unfolded set_mode_locs_def a_set_mode_locs_def] lemma set_mode_ok: "type_wf h \ shadow_root_ptr |\| shadow_root_ptr_kinds h \ h \ ok (set_mode shadow_root_ptr shadow_root_mode)" apply(unfold type_wf_impl) unfolding set_mode_def using get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ok put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ok by (simp add: ShadowRootMonad.put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ok) lemma set_mode_ptr_in_heap: "h \ ok (set_mode shadow_root_ptr shadow_root_mode) \ shadow_root_ptr |\| shadow_root_ptr_kinds h" by(simp add: set_mode_def put_M_ptr_in_heap) lemma set_mode_writes: "writes (set_mode_locs shadow_root_ptr) (set_mode shadow_root_ptr shadow_root_mode) h h'" by(auto simp add: set_mode_def set_mode_locs_def intro: writes_bind_pure) lemma set_mode_pointers_preserved: assumes "w \ set_mode_locs element_ptr" assumes "h \ w \\<^sub>h h'" shows "object_ptr_kinds h = object_ptr_kinds h'" using assms(1) object_ptr_kinds_preserved[OF writes_singleton2 assms(2)] by(auto simp add: all_args_def set_mode_locs_def split: if_splits) lemma set_mode_types_preserved: assumes "w \ set_mode_locs element_ptr" assumes "h \ w \\<^sub>h h'" shows "type_wf h = type_wf h'" apply(unfold type_wf_impl) using assms(1) type_wf_preserved[OF writes_singleton2 assms(2)] by(auto simp add: all_args_def set_mode_locs_def split: if_splits) end interpretation i_set_mode?: l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_mode set_mode_locs by (auto simp add: l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_set_mode = l_type_wf +l_set_mode_defs + assumes set_mode_writes: "writes (set_mode_locs shadow_root_ptr) (set_mode shadow_root_ptr shadow_root_mode) h h'" assumes set_mode_ok: "type_wf h \ shadow_root_ptr |\| shadow_root_ptr_kinds h \ h \ ok (set_mode shadow_root_ptr shadow_root_mode)" assumes set_mode_ptr_in_heap: "h \ ok (set_mode shadow_root_ptr shadow_root_mode) \ shadow_root_ptr |\| shadow_root_ptr_kinds h" assumes set_mode_pointers_preserved: "w \ set_mode_locs shadow_root_ptr \ h \ w \\<^sub>h h' \ object_ptr_kinds h = object_ptr_kinds h'" assumes set_mode_types_preserved: "w \ set_mode_locs shadow_root_ptr \ h \ w \\<^sub>h h' \ type_wf h = type_wf h'" lemma set_mode_is_l_set_mode [instances]: "l_set_mode type_wf set_mode set_mode_locs" apply(auto simp add: l_set_mode_def instances)[1] using set_mode_writes apply blast using set_mode_ok apply (blast) using set_mode_ptr_in_heap apply blast using set_mode_pointers_preserved apply(blast, blast) using set_mode_types_preserved apply(blast, blast) done paragraph \get\_child\_nodes\ locale l_set_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_shadow_root_get_child_nodes: "\w \ set_shadow_root_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_child_nodes_locs ptr'. r h h'))" by(auto simp add: get_child_nodes_locs_def set_shadow_root_locs_def CD.get_child_nodes_locs_def all_args_def intro: element_put_get_preserved[where setter=shadow_root_opt_update]) end interpretation i_set_shadow_root_get_child_nodes?: l_set_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs set_shadow_root set_shadow_root_locs by(unfold_locales) declare l_set_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_set_shadow_root_get_child_nodes = l_set_shadow_root + l_get_child_nodes + assumes set_shadow_root_get_child_nodes: "\w \ set_shadow_root_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_child_nodes_locs ptr'. r h h'))" lemma set_shadow_root_get_child_nodes_is_l_set_shadow_root_get_child_nodes [instances]: "l_set_shadow_root_get_child_nodes type_wf set_shadow_root set_shadow_root_locs known_ptr get_child_nodes get_child_nodes_locs" apply(auto simp add: l_set_shadow_root_get_child_nodes_def l_set_shadow_root_get_child_nodes_axioms_def instances)[1] using set_shadow_root_get_child_nodes apply blast done paragraph \get\_shadow\_root\ locale l_set_mode_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_mode_get_shadow_root: "\w \ set_mode_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" by(auto simp add: set_mode_locs_def get_shadow_root_locs_def all_args_def) end interpretation i_set_mode_get_shadow_root?: l_set_mode_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_mode set_mode_locs get_shadow_root get_shadow_root_locs by unfold_locales declare l_set_mode_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_set_mode_get_shadow_root = l_set_mode + l_get_shadow_root + assumes set_mode_get_shadow_root: "\w \ set_mode_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" lemma set_mode_get_shadow_root_is_l_set_mode_get_shadow_root [instances]: "l_set_mode_get_shadow_root type_wf set_mode set_mode_locs get_shadow_root get_shadow_root_locs" using set_mode_is_l_set_mode get_shadow_root_is_l_get_shadow_root apply(simp add: l_set_mode_get_shadow_root_def l_set_mode_get_shadow_root_axioms_def) using set_mode_get_shadow_root by fast paragraph \get\_child\_nodes\ locale l_set_mode_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_mode_get_child_nodes: "\w \ set_mode_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_child_nodes_locs ptr'. r h h'))" by(auto simp add: get_child_nodes_locs_def CD.get_child_nodes_locs_def set_mode_locs_def all_args_def)[1] end interpretation i_set_mode_get_child_nodes?: l_set_mode_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_mode set_mode_locs known_ptr DocumentClass.type_wf DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs by unfold_locales declare l_set_mode_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_set_mode_get_child_nodes = l_set_mode + l_get_child_nodes + assumes set_mode_get_child_nodes: "\w \ set_mode_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_child_nodes_locs ptr'. r h h'))" lemma set_mode_get_child_nodes_is_l_set_mode_get_child_nodes [instances]: "l_set_mode_get_child_nodes type_wf set_mode set_mode_locs known_ptr get_child_nodes get_child_nodes_locs" using set_mode_is_l_set_mode get_child_nodes_is_l_get_child_nodes apply(simp add: l_set_mode_get_child_nodes_def l_set_mode_get_child_nodes_axioms_def) using set_mode_get_child_nodes by fast subsubsection \get\_host\ locale l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_shadow_root_defs get_shadow_root get_shadow_root_locs for get_shadow_root :: "(_::linorder) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin definition a_get_host :: "(_) shadow_root_ptr \ (_, (_) element_ptr) dom_prog" where "a_get_host shadow_root_ptr = do { host_ptrs \ element_ptr_kinds_M \ filter_M (\element_ptr. do { shadow_root_opt \ get_shadow_root element_ptr; return (shadow_root_opt = Some shadow_root_ptr) }); (case host_ptrs of host_ptr#[] \ return host_ptr | _ \ error HierarchyRequestError) }" definition "a_get_host_locs \ (\element_ptr. (get_shadow_root_locs element_ptr)) \ (\ptr. {preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr RObject.nothing)})" end global_interpretation l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_shadow_root get_shadow_root_locs defines get_host = "a_get_host" and get_host_locs = "a_get_host_locs" . locale l_get_host_defs = fixes get_host :: "(_) shadow_root_ptr \ (_, (_) element_ptr) dom_prog" fixes get_host_locs :: "((_) heap \ (_) heap \ bool) set" locale l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_get_host_defs + l_get_shadow_root + assumes get_host_impl: "get_host = a_get_host" assumes get_host_locs_impl: "get_host_locs = a_get_host_locs" begin lemmas get_host_def = get_host_impl[unfolded a_get_host_def] lemmas get_host_locs_def = get_host_locs_impl[unfolded a_get_host_locs_def] lemma get_host_pure [simp]: "pure (get_host element_ptr) h" by(auto simp add: get_host_def intro!: bind_pure_I filter_M_pure_I split: list.splits) lemma get_host_reads: "reads get_host_locs (get_host element_ptr) h h'" using get_shadow_root_reads[unfolded reads_def] by(auto simp add: get_host_def get_host_locs_def intro!: reads_bind_pure reads_subset[OF check_in_heap_reads] reads_subset[OF error_reads] reads_subset[OF get_shadow_root_reads] reads_subset[OF return_reads] reads_subset[OF element_ptr_kinds_M_reads] filter_M_reads filter_M_pure_I bind_pure_I split: list.splits) end locale l_get_host = l_get_host_defs + assumes get_host_pure [simp]: "pure (get_host element_ptr) h" assumes get_host_reads: "reads get_host_locs (get_host node_ptr) h h'" interpretation i_get_host?: l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_shadow_root get_shadow_root_locs get_host get_host_locs type_wf using instances by (simp add: l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def get_host_def get_host_locs_def) declare l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_host_is_l_get_host [instances]: "l_get_host get_host get_host_locs" apply(auto simp add: l_get_host_def)[1] using get_host_reads apply fast done subsubsection \get\_mode\ locale l_get_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs begin definition a_get_mode :: "(_) shadow_root_ptr \ (_, shadow_root_mode) dom_prog" where "a_get_mode shadow_root_ptr = get_M shadow_root_ptr mode" definition a_get_mode_locs :: "(_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" where "a_get_mode_locs shadow_root_ptr \ {preserved (get_M shadow_root_ptr mode)}" end global_interpretation l_get_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs defines get_mode = a_get_mode and get_mode_locs = a_get_mode_locs . locale l_get_mode_defs = fixes get_mode :: "(_) shadow_root_ptr \ (_, shadow_root_mode) dom_prog" fixes get_mode_locs :: "(_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" locale l_get_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_get_mode_defs get_mode get_mode_locs + l_type_wf type_wf for get_mode :: "(_) shadow_root_ptr \ ((_) heap, exception, shadow_root_mode) prog" and get_mode_locs :: "(_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" and type_wf :: "(_) heap \ bool" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" assumes get_mode_impl: "get_mode = a_get_mode" assumes get_mode_locs_impl: "get_mode_locs = a_get_mode_locs" begin lemmas get_mode_def = get_mode_impl[unfolded get_mode_def a_get_mode_def] lemmas get_mode_locs_def = get_mode_locs_impl[unfolded get_mode_locs_def a_get_mode_locs_def] lemma get_mode_ok: "type_wf h \ shadow_root_ptr |\| shadow_root_ptr_kinds h \ h \ ok (get_mode shadow_root_ptr)" unfolding get_mode_def type_wf_impl using ShadowRootMonad.get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ok by blast lemma get_mode_pure [simp]: "pure (get_mode element_ptr) h" unfolding get_mode_def by simp lemma get_mode_ptr_in_heap: assumes "h \ get_mode shadow_root_ptr \\<^sub>r children" shows "shadow_root_ptr |\| shadow_root_ptr_kinds h" using assms by(auto simp add: get_mode_def get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ptr_in_heap dest: is_OK_returns_result_I) lemma get_mode_reads: "reads (get_mode_locs element_ptr) (get_mode element_ptr) h h'" by(simp add: get_mode_def get_mode_locs_def reads_bind_pure reads_insert_writes_set_right) end interpretation i_get_mode?: l_get_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_mode get_mode_locs type_wf using instances by (auto simp add: l_get_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_get_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_get_mode = l_type_wf + l_get_mode_defs + assumes get_mode_reads: "reads (get_mode_locs shadow_root_ptr) (get_mode shadow_root_ptr) h h'" assumes get_mode_ok: "type_wf h \ shadow_root_ptr |\| shadow_root_ptr_kinds h \ h \ ok (get_mode shadow_root_ptr)" assumes get_mode_ptr_in_heap: "h \ ok (get_mode shadow_root_ptr) \ shadow_root_ptr |\| shadow_root_ptr_kinds h" assumes get_mode_pure [simp]: "pure (get_mode shadow_root_ptr) h" lemma get_mode_is_l_get_mode [instances]: "l_get_mode type_wf get_mode get_mode_locs" apply(auto simp add: l_get_mode_def instances)[1] using get_mode_reads apply blast using get_mode_ok apply blast using get_mode_ptr_in_heap apply blast done subsubsection \get\_shadow\_root\_safe\ locale l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_shadow_root_defs get_shadow_root get_shadow_root_locs + l_get_mode_defs get_mode get_mode_locs for get_shadow_root :: "(_) element_ptr \ (_, (_) shadow_root_ptr option) dom_prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_mode :: "(_) shadow_root_ptr \ (_, shadow_root_mode) dom_prog" and get_mode_locs :: "(_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" begin definition a_get_shadow_root_safe :: "(_) element_ptr \ (_, (_) shadow_root_ptr option) dom_prog" where "a_get_shadow_root_safe element_ptr = do { shadow_root_ptr_opt \ get_shadow_root element_ptr; (case shadow_root_ptr_opt of Some shadow_root_ptr \ do { mode \ get_mode shadow_root_ptr; (if mode = Open then return (Some shadow_root_ptr) else return None ) } | None \ return None) }" definition a_get_shadow_root_safe_locs :: "(_) element_ptr \ (_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" where "a_get_shadow_root_safe_locs element_ptr shadow_root_ptr \ (get_shadow_root_locs element_ptr) \ (get_mode_locs shadow_root_ptr)" end global_interpretation l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_shadow_root get_shadow_root_locs get_mode get_mode_locs defines get_shadow_root_safe = a_get_shadow_root_safe and get_shadow_root_safe_locs = a_get_shadow_root_safe_locs . locale l_get_shadow_root_safe_defs = fixes get_shadow_root_safe :: "(_) element_ptr \ (_, (_) shadow_root_ptr option) dom_prog" fixes get_shadow_root_safe_locs :: "(_) element_ptr \ (_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" locale l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_shadow_root get_shadow_root_locs get_mode get_mode_locs + l_get_shadow_root_safe_defs get_shadow_root_safe get_shadow_root_safe_locs + l_get_shadow_root type_wf get_shadow_root get_shadow_root_locs + l_get_mode type_wf get_mode get_mode_locs + l_type_wf type_wf for type_wf :: "(_) heap \ bool" and get_shadow_root_safe :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_safe_locs :: "(_) element_ptr \ (_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ (_, (_) shadow_root_ptr option) dom_prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_mode :: "(_) shadow_root_ptr \ (_, shadow_root_mode) dom_prog" and get_mode_locs :: "(_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" assumes get_shadow_root_safe_impl: "get_shadow_root_safe = a_get_shadow_root_safe" assumes get_shadow_root_safe_locs_impl: "get_shadow_root_safe_locs = a_get_shadow_root_safe_locs" begin lemmas get_shadow_root_safe_def = get_shadow_root_safe_impl[unfolded get_shadow_root_safe_def a_get_shadow_root_safe_def] lemmas get_shadow_root_safe_locs_def = get_shadow_root_safe_locs_impl[unfolded get_shadow_root_safe_locs_def a_get_shadow_root_safe_locs_def] lemma get_shadow_root_safe_pure [simp]: "pure (get_shadow_root_safe element_ptr) h" apply(auto simp add: get_shadow_root_safe_def)[1] by (smt bind_returns_heap_E is_OK_returns_heap_E local.get_mode_pure local.get_shadow_root_pure option.case_eq_if pure_def pure_returns_heap_eq return_pure) end interpretation i_get_shadow_root_safe?: l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root_safe get_shadow_root_safe_locs get_shadow_root get_shadow_root_locs get_mode get_mode_locs using instances by (auto simp add: l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def get_shadow_root_safe_def get_shadow_root_safe_locs_def) declare l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_get_shadow_root_safe = l_get_shadow_root_safe_defs + assumes get_shadow_root_safe_pure [simp]: "pure (get_shadow_root_safe element_ptr) h" lemma get_shadow_root_safe_is_l_get_shadow_root_safe [instances]: "l_get_shadow_root_safe get_shadow_root_safe" using instances apply(auto simp add: l_get_shadow_root_safe_def)[1] done subsubsection \set\_disconnected\_nodes\ locale l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD: l_set_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_disconnected_nodes set_disconnected_nodes_locs + l_type_wf type_wf for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" begin lemma set_disconnected_nodes_ok: "type_wf h \ document_ptr |\| document_ptr_kinds h \ h \ ok (set_disconnected_nodes document_ptr node_ptrs)" using CD.set_disconnected_nodes_ok CD.type_wf_impl ShadowRootClass.type_wf_defs local.type_wf_impl by blast lemma set_disconnected_nodes_typess_preserved: assumes "w \ set_disconnected_nodes_locs object_ptr" assumes "h \ w \\<^sub>h h'" shows "type_wf h = type_wf h'" using assms(1) type_wf_preserved[OF writes_singleton2 assms(2)] apply(unfold type_wf_impl) by(auto simp add: all_args_def CD.set_disconnected_nodes_locs_def split: if_splits) end interpretation i_set_disconnected_nodes?: l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_disconnected_nodes set_disconnected_nodes_locs by(auto simp add: l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma set_disconnected_nodes_is_l_set_disconnected_nodes [instances]: "l_set_disconnected_nodes type_wf set_disconnected_nodes set_disconnected_nodes_locs" apply(auto simp add: l_set_disconnected_nodes_def)[1] apply (simp add: i_set_disconnected_nodes.set_disconnected_nodes_writes) using set_disconnected_nodes_ok apply blast apply (simp add: i_set_disconnected_nodes.set_disconnected_nodes_ptr_in_heap) using i_set_disconnected_nodes.set_disconnected_nodes_pointers_preserved apply (blast, blast) using set_disconnected_nodes_typess_preserved apply(blast, blast) done paragraph \get\_child\_nodes\ locale l_set_disconnected_nodes_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_disconnected_nodes set_disconnected_nodes_locs + l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M for type_wf :: "(_) heap \ bool" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" and known_ptr :: "(_) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma set_disconnected_nodes_get_child_nodes: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_child_nodes_locs ptr'. r h h'))" by(auto simp add: set_disconnected_nodes_locs_def get_child_nodes_locs_def CD.get_child_nodes_locs_def all_args_def) end interpretation i_set_disconnected_nodes_get_child_nodes?: l_set_disconnected_nodes_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_disconnected_nodes set_disconnected_nodes_locs known_ptr DocumentClass.type_wf DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs by(auto simp add: l_set_disconnected_nodes_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_disconnected_nodes_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_disconnected_nodes_get_child_nodes_is_l_set_disconnected_nodes_get_child_nodes [instances]: "l_set_disconnected_nodes_get_child_nodes set_disconnected_nodes_locs get_child_nodes_locs" apply(auto simp add: l_set_disconnected_nodes_get_child_nodes_def)[1] using set_disconnected_nodes_get_child_nodes apply fast done paragraph \get\_host\ locale l_set_disconnected_nodes_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_disconnected_nodes_get_host: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_host_locs. r h h'))" by(auto simp add: CD.set_disconnected_nodes_locs_def get_shadow_root_locs_def get_host_locs_def all_args_def) end interpretation i_set_disconnected_nodes_get_host?: l_set_disconnected_nodes_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_disconnected_nodes set_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_host get_host_locs by(auto simp add: l_set_disconnected_nodes_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_disconnected_nodes_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_set_disconnected_nodes_get_host = l_set_disconnected_nodes_defs + l_get_host_defs + assumes set_disconnected_nodes_get_host: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_host_locs. r h h'))" lemma set_disconnected_nodes_get_host_is_l_set_disconnected_nodes_get_host [instances]: "l_set_disconnected_nodes_get_host set_disconnected_nodes_locs get_host_locs" apply(auto simp add: l_set_disconnected_nodes_get_host_def instances)[1] using set_disconnected_nodes_get_host by fast subsubsection \get\_tag\_name\ locale l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD: l_get_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_tag_name get_tag_name_locs + l_type_wf type_wf for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_tag_name :: "(_) element_ptr \ (_, tag_name) dom_prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" begin lemma get_tag_name_ok: "type_wf h \ element_ptr |\| element_ptr_kinds h \ h \ ok (get_tag_name element_ptr)" apply(unfold type_wf_impl get_tag_name_impl[unfolded a_get_tag_name_def]) using CD.get_tag_name_ok CD.type_wf_impl ShadowRootClass.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t by blast end interpretation i_get_tag_name?: l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_tag_name get_tag_name_locs by(auto simp add: l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_tag_name_is_l_get_tag_name [instances]: "l_get_tag_name type_wf get_tag_name get_tag_name_locs" apply(auto simp add: l_get_tag_name_def)[1] using get_tag_name_reads apply fast using get_tag_name_ok apply fast done paragraph \set\_disconnected\_nodes\ locale l_set_disconnected_nodes_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_disconnected_nodes_get_tag_name: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_tag_name_locs ptr'. r h h'))" by(auto simp add: CD.set_disconnected_nodes_locs_def CD.get_tag_name_locs_def all_args_def) end interpretation i_set_disconnected_nodes_get_tag_name?: l_set_disconnected_nodes_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_disconnected_nodes set_disconnected_nodes_locs get_tag_name get_tag_name_locs by(auto simp add: l_set_disconnected_nodes_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_disconnected_nodes_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma set_disconnected_nodes_get_tag_name_is_l_set_disconnected_nodes_get_tag_name [instances]: "l_set_disconnected_nodes_get_tag_name type_wf set_disconnected_nodes set_disconnected_nodes_locs get_tag_name get_tag_name_locs" apply(auto simp add: l_set_disconnected_nodes_get_tag_name_def l_set_disconnected_nodes_get_tag_name_axioms_def instances)[1] using set_disconnected_nodes_get_tag_name by fast paragraph \set\_child\_nodes\ locale l_set_child_nodes_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_child_nodes_get_tag_name: "\w \ set_child_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_tag_name_locs ptr'. r h h'))" by(auto simp add: CD.set_child_nodes_locs_def set_child_nodes_locs_def CD.get_tag_name_locs_def all_args_def intro: element_put_get_preserved[where getter=tag_name and setter=RElement.child_nodes_update]) end interpretation i_set_child_nodes_get_tag_name?: l_set_child_nodes_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr set_child_nodes set_child_nodes_locs Core_DOM_Functions.set_child_nodes Core_DOM_Functions.set_child_nodes_locs get_tag_name get_tag_name_locs by(auto simp add: l_set_child_nodes_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_child_nodes_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma set_child_nodes_get_tag_name_is_l_set_child_nodes_get_tag_name [instances]: "l_set_child_nodes_get_tag_name type_wf set_child_nodes set_child_nodes_locs get_tag_name get_tag_name_locs" apply(auto simp add: l_set_child_nodes_get_tag_name_def l_set_child_nodes_get_tag_name_axioms_def instances)[1] using set_child_nodes_get_tag_name by fast paragraph \delete\_shadow\_root\ locale l_delete_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_tag_name_delete_shadow_root: "h \ delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M shadow_root_ptr \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" by(auto simp add: CD.get_tag_name_locs_def delete_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t) end locale l_delete_shadow_root_get_tag_name = l_get_tag_name_defs + assumes get_tag_name_delete_shadow_root: "h \ delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M shadow_root_ptr \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" interpretation l_delete_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_tag_name get_tag_name_locs by(auto simp add: l_delete_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) lemma l_delete_shadow_root_get_tag_name_get_tag_name_locs [instances]: "l_delete_shadow_root_get_tag_name get_tag_name_locs" apply(auto simp add: l_delete_shadow_root_get_tag_name_def)[1] using get_tag_name_delete_shadow_root apply fast done paragraph \set\_shadow\_root\ locale l_set_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_shadow_root_get_tag_name: "\w \ set_shadow_root_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_tag_name_locs ptr'. r h h'))" by(auto simp add: set_shadow_root_locs_def CD.get_tag_name_locs_def all_args_def element_put_get_preserved[where setter=shadow_root_opt_update]) end interpretation i_set_shadow_root_get_tag_name?: l_set_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_shadow_root set_shadow_root_locs DocumentClass.type_wf get_tag_name get_tag_name_locs apply(auto simp add: l_set_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)[1] by(unfold_locales) declare l_set_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_set_shadow_root_get_tag_name = l_set_shadow_root_defs + l_get_tag_name_defs + assumes set_shadow_root_get_tag_name: "\w \ set_shadow_root_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_tag_name_locs ptr'. r h h'))" lemma set_shadow_root_get_tag_name_is_l_set_shadow_root_get_tag_name [instances]: "l_set_shadow_root_get_tag_name set_shadow_root_locs get_tag_name_locs" using set_shadow_root_is_l_set_shadow_root get_tag_name_is_l_get_tag_name apply(simp add: l_set_shadow_root_get_tag_name_def ) using set_shadow_root_get_tag_name by fast paragraph \new\_element\ locale l_new_element_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_tag_name get_tag_name_locs for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_tag_name :: "(_) element_ptr \ (_, tag_name) dom_prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_tag_name_new_element: "ptr' \ new_element_ptr \ h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" by (auto simp add: CD.get_tag_name_locs_def new_element_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t new_element_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t new_element_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t split: prod.splits if_splits option.splits elim!: bind_returns_result_E bind_returns_heap_E intro: is_element_ptr_kind_obtains) lemma new_element_empty_tag_name: "h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ h' \ get_tag_name new_element_ptr \\<^sub>r ''''" by(simp add: CD.get_tag_name_def new_element_tag_name) end locale l_new_element_get_tag_name = l_new_element + l_get_tag_name + assumes get_tag_name_new_element: "ptr' \ new_element_ptr \ h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" assumes new_element_empty_tag_name: "h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ h' \ get_tag_name new_element_ptr \\<^sub>r ''''" interpretation i_new_element_get_tag_name?: l_new_element_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_tag_name get_tag_name_locs by(auto simp add: l_new_element_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_new_element_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma new_element_get_tag_name_is_l_new_element_get_tag_name [instances]: "l_new_element_get_tag_name type_wf get_tag_name get_tag_name_locs" using new_element_is_l_new_element get_tag_name_is_l_get_tag_name apply(auto simp add: l_new_element_get_tag_name_def l_new_element_get_tag_name_axioms_def instances)[1] using get_tag_name_new_element new_element_empty_tag_name by fast+ paragraph \get\_shadow\_root\ locale l_set_mode_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_mode_get_tag_name: "\w \ set_mode_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_tag_name_locs ptr'. r h h'))" by(auto simp add: set_mode_locs_def CD.get_tag_name_locs_def all_args_def) end interpretation i_set_mode_get_tag_name?: l_set_mode_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_mode set_mode_locs DocumentClass.type_wf get_tag_name get_tag_name_locs by unfold_locales declare l_set_mode_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_set_mode_get_tag_name = l_set_mode + l_get_tag_name + assumes set_mode_get_tag_name: "\w \ set_mode_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_tag_name_locs ptr'. r h h'))" lemma set_mode_get_tag_name_is_l_set_mode_get_tag_name [instances]: "l_set_mode_get_tag_name type_wf set_mode set_mode_locs get_tag_name get_tag_name_locs" using set_mode_is_l_set_mode get_tag_name_is_l_get_tag_name apply(simp add: l_set_mode_get_tag_name_def l_set_mode_get_tag_name_axioms_def) using set_mode_get_tag_name by fast paragraph \new\_document\ locale l_new_document_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_tag_name get_tag_name_locs for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, tag_name) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_tag_name_new_document: "h \ new_document \\<^sub>r new_document_ptr \ h \ new_document \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" by(auto simp add: CD.get_tag_name_locs_def new_document_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t) end locale l_new_document_get_tag_name = l_get_tag_name_defs + assumes get_tag_name_new_document: "h \ new_document \\<^sub>r new_document_ptr \ h \ new_document \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" interpretation i_new_document_get_tag_name?: l_new_document_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_tag_name get_tag_name_locs by unfold_locales declare l_new_document_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def[instances] lemma new_document_get_tag_name_is_l_new_document_get_tag_name [instances]: "l_new_document_get_tag_name get_tag_name_locs" unfolding l_new_document_get_tag_name_def unfolding get_tag_name_locs_def using new_document_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t by blast paragraph \new\_shadow\_root\ locale l_new_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_tag_name_new_shadow_root: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" by (auto simp add: CD.get_tag_name_locs_def new_shadow_root_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t new_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t split: prod.splits if_splits option.splits elim!: bind_returns_result_E bind_returns_heap_E intro: is_element_ptr_kind_obtains) end locale l_new_shadow_root_get_tag_name = l_get_tag_name + assumes get_tag_name_new_shadow_root: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" interpretation i_new_shadow_root_get_tag_name?: l_new_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_tag_name get_tag_name_locs by(unfold_locales) declare l_new_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma new_shadow_root_get_tag_name_is_l_new_shadow_root_get_tag_name [instances]: "l_new_shadow_root_get_tag_name type_wf get_tag_name get_tag_name_locs" using get_tag_name_is_l_get_tag_name apply(auto simp add: l_new_shadow_root_get_tag_name_def l_new_shadow_root_get_tag_name_axioms_def instances)[1] using get_tag_name_new_shadow_root by fast paragraph \new\_character\_data\ locale l_new_character_data_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_tag_name get_tag_name_locs for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, tag_name) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_tag_name_new_character_data: "h \ new_character_data \\<^sub>r new_character_data_ptr \ h \ new_character_data \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" by(auto simp add: CD.get_tag_name_locs_def new_character_data_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t) end locale l_new_character_data_get_tag_name = l_get_tag_name_defs + assumes get_tag_name_new_character_data: "h \ new_character_data \\<^sub>r new_character_data_ptr \ h \ new_character_data \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" interpretation i_new_character_data_get_tag_name?: l_new_character_data_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_tag_name get_tag_name_locs by unfold_locales declare l_new_character_data_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def[instances] lemma new_character_data_get_tag_name_is_l_new_character_data_get_tag_name [instances]: "l_new_character_data_get_tag_name get_tag_name_locs" unfolding l_new_character_data_get_tag_name_def unfolding get_tag_name_locs_def using new_character_data_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t by blast paragraph \get\_tag\_type\ locale l_set_tag_name_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_tag_name_get_tag_name: assumes "h \ CD.a_set_tag_name element_ptr tag \\<^sub>h h'" shows "h' \ CD.a_get_tag_name element_ptr \\<^sub>r tag" using assms by(auto simp add: CD.a_get_tag_name_def CD.a_set_tag_name_def) lemma set_tag_name_get_tag_name_different_pointers: assumes "ptr \ ptr'" assumes "w \ CD.a_set_tag_name_locs ptr" assumes "h \ w \\<^sub>h h'" assumes "r \ CD.a_get_tag_name_locs ptr'" shows "r h h'" using assms by(auto simp add: all_args_def CD.a_set_tag_name_locs_def CD.a_get_tag_name_locs_def split: if_splits option.splits ) end interpretation i_set_tag_name_get_tag_name?: l_set_tag_name_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_tag_name get_tag_name_locs set_tag_name set_tag_name_locs by unfold_locales declare l_set_tag_name_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_tag_name_get_tag_name_is_l_set_tag_name_get_tag_name [instances]: "l_set_tag_name_get_tag_name type_wf get_tag_name get_tag_name_locs set_tag_name set_tag_name_locs" using set_tag_name_is_l_set_tag_name get_tag_name_is_l_get_tag_name apply(simp add: l_set_tag_name_get_tag_name_def l_set_tag_name_get_tag_name_axioms_def) using set_tag_name_get_tag_name set_tag_name_get_tag_name_different_pointers by fast+ subsubsection \attach\_shadow\_root\ locale l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_set_shadow_root_defs set_shadow_root set_shadow_root_locs + l_set_mode_defs set_mode set_mode_locs + l_get_tag_name_defs get_tag_name get_tag_name_locs + l_get_shadow_root_defs get_shadow_root get_shadow_root_locs for set_shadow_root :: "(_) element_ptr \ (_) shadow_root_ptr option \ ((_) heap, exception, unit) prog" and set_shadow_root_locs :: "(_) element_ptr \ ((_) heap, exception, unit) prog set" and set_mode :: "(_) shadow_root_ptr \ shadow_root_mode \ ((_) heap, exception, unit) prog" and set_mode_locs :: "(_) shadow_root_ptr \ ((_) heap, exception, unit) prog set" and get_tag_name :: "(_) element_ptr \ (_, char list) dom_prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin definition a_attach_shadow_root :: "(_) element_ptr \ shadow_root_mode \ (_, (_) shadow_root_ptr) dom_prog" where "a_attach_shadow_root element_ptr shadow_root_mode = do { tag \ get_tag_name element_ptr; (if tag \ safe_shadow_root_element_types then error NotSupportedError else return ()); prev_shadow_root \ get_shadow_root element_ptr; (if prev_shadow_root \ None then error NotSupportedError else return ()); new_shadow_root_ptr \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M; set_mode new_shadow_root_ptr shadow_root_mode; set_shadow_root element_ptr (Some new_shadow_root_ptr); return new_shadow_root_ptr }" end locale l_attach_shadow_root_defs = fixes attach_shadow_root :: "(_) element_ptr \ shadow_root_mode \ (_, (_) shadow_root_ptr) dom_prog" global_interpretation l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs set_shadow_root set_shadow_root_locs set_mode set_mode_locs get_tag_name get_tag_name_locs get_shadow_root get_shadow_root_locs defines attach_shadow_root = a_attach_shadow_root . locale l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs set_shadow_root set_shadow_root_locs set_mode set_mode_locs get_tag_name get_tag_name_locs get_shadow_root get_shadow_root_locs + l_attach_shadow_root_defs attach_shadow_root + l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_shadow_root set_shadow_root_locs + l_set_mode type_wf set_mode set_mode_locs + l_get_tag_name type_wf get_tag_name get_tag_name_locs + l_get_shadow_root type_wf get_shadow_root get_shadow_root_locs + l_known_ptr known_ptr for known_ptr :: "(_) object_ptr \ bool" and set_shadow_root :: "(_) element_ptr \ (_) shadow_root_ptr option \ ((_) heap, exception, unit) prog" and set_shadow_root_locs :: "(_) element_ptr \ ((_) heap, exception, unit) prog set" and set_mode :: "(_) shadow_root_ptr \ shadow_root_mode \ ((_) heap, exception, unit) prog" and set_mode_locs :: "(_) shadow_root_ptr \ ((_) heap, exception, unit) prog set" and attach_shadow_root :: "(_) element_ptr \ shadow_root_mode \ ((_) heap, exception, (_) shadow_root_ptr) prog" and type_wf :: "(_) heap \ bool" and get_tag_name :: "(_) element_ptr \ (_, char list) dom_prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" + assumes known_ptr_impl: "known_ptr = a_known_ptr" assumes attach_shadow_root_impl: "attach_shadow_root = a_attach_shadow_root" begin lemmas attach_shadow_root_def = a_attach_shadow_root_def[folded attach_shadow_root_impl] lemma attach_shadow_root_element_ptr_in_heap: assumes "h \ ok (attach_shadow_root element_ptr shadow_root_mode)" shows "element_ptr |\| element_ptr_kinds h" proof - obtain h' where "h \ attach_shadow_root element_ptr shadow_root_mode \\<^sub>h h'" using assms by auto then obtain h2 h3 new_shadow_root_ptr where h2: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h2" and new_shadow_root_ptr: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr" and h3: "h2 \ set_mode new_shadow_root_ptr shadow_root_mode \\<^sub>h h3" and "h3 \ set_shadow_root element_ptr (Some new_shadow_root_ptr) \\<^sub>h h'" by(auto simp add: attach_shadow_root_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_tag_name_pure, rotated] bind_returns_heap_E2[rotated, OF get_shadow_root_pure, rotated] split: if_splits) then have "element_ptr |\| element_ptr_kinds h3" using set_shadow_root_ptr_in_heap by blast moreover have "object_ptr_kinds h2 = object_ptr_kinds h |\| {|cast new_shadow_root_ptr|}" using h2 new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_new_ptr new_shadow_root_ptr by auto moreover have "object_ptr_kinds h2 = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_mode_writes h3]) using set_mode_pointers_preserved apply blast by (auto simp add: reflp_def transp_def) ultimately show ?thesis by(auto simp add: element_ptr_kinds_def node_ptr_kinds_def) qed lemma create_shadow_root_known_ptr: assumes "h \ attach_shadow_root element_ptr shadow_root_mode \\<^sub>r new_shadow_root_ptr" shows "known_ptr (cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_shadow_root_ptr)" using assms by(auto simp add: attach_shadow_root_def known_ptr_impl ShadowRootClass.a_known_ptr_def new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def Let_def elim!: bind_returns_result_E) end locale l_attach_shadow_root = l_attach_shadow_root_defs interpretation i_attach_shadow_root?: l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr set_shadow_root set_shadow_root_locs set_mode set_mode_locs attach_shadow_root type_wf get_tag_name get_tag_name_locs get_shadow_root get_shadow_root_locs by(auto simp add: l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def attach_shadow_root_def instances) declare l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_parent\ global_interpretation l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs defines get_parent = "l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_parent get_child_nodes" and get_parent_locs = "l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_parent_locs get_child_nodes_locs" . interpretation i_get_parent?: l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs get_parent get_parent_locs by(simp add: l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def get_parent_def get_parent_locs_def instances) declare l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_parent_is_l_get_parent [instances]: "l_get_parent type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs" apply(simp add: l_get_parent_def l_get_parent_axioms_def instances) using get_parent_reads get_parent_ok get_parent_ptr_in_heap get_parent_pure get_parent_parent_in_heap get_parent_child_dual get_parent_reads_pointers by blast paragraph \set\_disconnected\_nodes\ locale l_set_disconnected_nodes_get_parent\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_disconnected_nodes_get_child_nodes + l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_disconnected_nodes_get_parent [simp]: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_parent_locs. r h h'))" by(auto simp add: get_parent_locs_def CD.set_disconnected_nodes_locs_def all_args_def) end interpretation i_set_disconnected_nodes_get_parent?: l_set_disconnected_nodes_get_parent\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_disconnected_nodes set_disconnected_nodes_locs get_child_nodes get_child_nodes_locs type_wf DocumentClass.type_wf known_ptr known_ptrs get_parent get_parent_locs by (simp add: l_set_disconnected_nodes_get_parent\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_disconnected_nodes_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_disconnected_nodes_get_parent_is_l_set_disconnected_nodes_get_parent [instances]: "l_set_disconnected_nodes_get_parent set_disconnected_nodes_locs get_parent_locs" by(simp add: l_set_disconnected_nodes_get_parent_def) subsubsection \get\_root\_node\ global_interpretation l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_parent get_parent_locs defines get_root_node = "l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_root_node get_parent" and get_root_node_locs = "l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_root_node_locs get_parent_locs" and get_ancestors = "l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_ancestors get_parent" and get_ancestors_locs = "l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_ancestors_locs get_parent_locs" . declare a_get_ancestors.simps [code] interpretation i_get_root_node?: l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_ancestors get_ancestors_locs get_root_node get_root_node_locs by(simp add: l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def get_root_node_def get_root_node_locs_def get_ancestors_def get_ancestors_locs_def instances) declare l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_ancestors_is_l_get_ancestors [instances]: "l_get_ancestors get_ancestors" apply(auto simp add: l_get_ancestors_def)[1] using get_ancestors_ptr_in_heap apply fast using get_ancestors_ptr apply fast done lemma get_root_node_is_l_get_root_node [instances]: "l_get_root_node get_root_node get_parent" by (simp add: l_get_root_node_def Shadow_DOM.i_get_root_node.get_root_node_no_parent) subsubsection \get\_root\_node\_si\ locale l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_parent_defs get_parent get_parent_locs + l_get_host_defs get_host get_host_locs for get_parent :: "(_) node_ptr \ ((_) heap, exception, (_::linorder) object_ptr option) prog" and get_parent_locs :: "((_) heap \ (_) heap \ bool) set" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" begin partial_function (dom_prog) a_get_ancestors_si :: "(_::linorder) object_ptr \ (_, (_) object_ptr list) dom_prog" where "a_get_ancestors_si ptr = do { check_in_heap ptr; ancestors \ (case cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr of Some node_ptr \ do { parent_ptr_opt \ get_parent node_ptr; (case parent_ptr_opt of Some parent_ptr \ a_get_ancestors_si parent_ptr | None \ return []) } | None \ (case cast ptr of Some shadow_root_ptr \ do { host \ get_host shadow_root_ptr; a_get_ancestors_si (cast host) } | None \ return [])); return (ptr # ancestors) }" definition "a_get_ancestors_si_locs = get_parent_locs \ get_host_locs" definition a_get_root_node_si :: "(_) object_ptr \ (_, (_) object_ptr) dom_prog" where "a_get_root_node_si ptr = do { ancestors \ a_get_ancestors_si ptr; return (last ancestors) }" definition "a_get_root_node_si_locs = a_get_ancestors_si_locs" end locale l_get_ancestors_si_defs = fixes get_ancestors_si :: "(_::linorder) object_ptr \ (_, (_) object_ptr list) dom_prog" fixes get_ancestors_si_locs :: "((_) heap \ (_) heap \ bool) set" locale l_get_root_node_si_defs = fixes get_root_node_si :: "(_) object_ptr \ (_, (_) object_ptr) dom_prog" fixes get_root_node_si_locs :: "((_) heap \ (_) heap \ bool) set" locale l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_parent + l_get_host + l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_get_ancestors_si_defs + l_get_root_node_si_defs + assumes get_ancestors_si_impl: "get_ancestors_si = a_get_ancestors_si" assumes get_ancestors_si_locs_impl: "get_ancestors_si_locs = a_get_ancestors_si_locs" assumes get_root_node_si_impl: "get_root_node_si = a_get_root_node_si" assumes get_root_node_si_locs_impl: "get_root_node_si_locs = a_get_root_node_si_locs" begin lemmas get_ancestors_si_def = a_get_ancestors_si.simps[folded get_ancestors_si_impl] lemmas get_ancestors_si_locs_def = a_get_ancestors_si_locs_def[folded get_ancestors_si_locs_impl] lemmas get_root_node_si_def = a_get_root_node_si_def[folded get_root_node_si_impl get_ancestors_si_impl] lemmas get_root_node_si_locs_def = a_get_root_node_si_locs_def[folded get_root_node_si_locs_impl get_ancestors_si_locs_impl] lemma get_ancestors_si_pure [simp]: "pure (get_ancestors_si ptr) h" proof - have "\ptr h h' x. h \ get_ancestors_si ptr \\<^sub>r x \ h \ get_ancestors_si ptr \\<^sub>h h' \ h = h'" proof (induct rule: a_get_ancestors_si.fixp_induct[folded get_ancestors_si_impl]) case 1 then show ?case by(rule admissible_dom_prog) next case 2 then show ?case by simp next case (3 f) then show ?case using get_parent_pure get_host_pure apply(auto simp add: pure_returns_heap_eq pure_def split: option.splits elim!: bind_returns_heap_E bind_returns_result_E dest!: pure_returns_heap_eq[rotated, OF check_in_heap_pure])[1] apply (meson option.simps(3) returns_result_eq) apply(metis get_parent_pure pure_returns_heap_eq) apply(metis get_host_pure pure_returns_heap_eq) done qed then show ?thesis by (meson pure_eq_iff) qed lemma get_root_node_si_pure [simp]: "pure (get_root_node_si ptr) h" by(auto simp add: get_root_node_si_def bind_pure_I) lemma get_ancestors_si_ptr_in_heap: assumes "h \ ok (get_ancestors_si ptr)" shows "ptr |\| object_ptr_kinds h" using assms by(auto simp add: get_ancestors_si_def check_in_heap_ptr_in_heap elim!: bind_is_OK_E dest: is_OK_returns_result_I) lemma get_ancestors_si_ptr: assumes "h \ get_ancestors_si ptr \\<^sub>r ancestors" shows "ptr \ set ancestors" using assms by(simp add: get_ancestors_si_def) (auto elim!: bind_returns_result_E2 split: option.splits intro!: bind_pure_I) lemma get_ancestors_si_never_empty: assumes "h \ get_ancestors_si child \\<^sub>r ancestors" shows "ancestors \ []" using assms apply(simp add: get_ancestors_si_def) by(auto elim!: bind_returns_result_E2 split: option.splits) lemma get_root_node_si_no_parent: "h \ get_parent node_ptr \\<^sub>r None \ h \ get_root_node_si (cast node_ptr) \\<^sub>r cast node_ptr" apply(auto simp add: check_in_heap_def get_root_node_si_def get_ancestors_si_def intro!: bind_pure_returns_result_I )[1] using get_parent_ptr_in_heap by blast lemma get_root_node_si_root_not_shadow_root: assumes "h \ get_root_node_si ptr \\<^sub>r root" shows "\ is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root" using assms proof(auto simp add: get_root_node_si_def elim!: bind_returns_result_E2) fix y assume "h \ get_ancestors_si ptr \\<^sub>r y" and "is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r (last y)" and "root = last y" then show False proof(induct y arbitrary: ptr) case Nil then show ?case using assms(1) get_ancestors_si_never_empty by blast next case (Cons a x) then show ?case apply(auto simp add: get_ancestors_si_def[of ptr] elim!: bind_returns_result_E2 split: option.splits if_splits)[1] using get_ancestors_si_never_empty apply blast using Cons.prems(2) apply auto[1] using \is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r (last y)\ \root = last y\ by auto qed qed end global_interpretation l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_parent get_parent_locs get_host get_host_locs defines get_root_node_si = a_get_root_node_si and get_root_node_si_locs = a_get_root_node_si_locs and get_ancestors_si = a_get_ancestors_si and get_ancestors_si_locs = a_get_ancestors_si_locs . declare a_get_ancestors_si.simps [code] interpretation i_get_root_node_si?: l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_host get_host_locs get_ancestors_si get_ancestors_si_locs get_root_node_si get_root_node_si_locs apply(auto simp add: l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances)[1] by(auto simp add: get_root_node_si_def get_root_node_si_locs_def get_ancestors_si_def get_ancestors_si_locs_def) declare l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma get_ancestors_si_is_l_get_ancestors [instances]: "l_get_ancestors get_ancestors_si" unfolding l_get_ancestors_def using get_ancestors_si_pure get_ancestors_si_ptr get_ancestors_si_ptr_in_heap by blast lemma get_root_node_si_is_l_get_root_node [instances]: "l_get_root_node get_root_node_si get_parent" apply(simp add: l_get_root_node_def) using get_root_node_si_no_parent by fast paragraph \set\_disconnected\_nodes\ locale l_set_disconnected_nodes_get_ancestors_si\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_disconnected_nodes_get_parent + l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_disconnected_nodes_get_host begin lemma set_disconnected_nodes_get_ancestors_si: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_ancestors_si_locs. r h h'))" by(auto simp add: get_parent_locs_def set_disconnected_nodes_locs_def set_disconnected_nodes_get_host get_ancestors_si_locs_def all_args_def) end locale l_set_disconnected_nodes_get_ancestors_si = l_set_disconnected_nodes_defs + l_get_ancestors_si_defs + assumes set_disconnected_nodes_get_ancestors_si: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_ancestors_si_locs. r h h'))" interpretation i_set_disconnected_nodes_get_ancestors_si?: l_set_disconnected_nodes_get_ancestors_si\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_disconnected_nodes set_disconnected_nodes_locs get_parent get_parent_locs type_wf known_ptr known_ptrs get_child_nodes get_child_nodes_locs get_host get_host_locs get_ancestors_si get_ancestors_si_locs get_root_node_si get_root_node_si_locs DocumentClass.type_wf by (auto simp add: l_set_disconnected_nodes_get_ancestors_si\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_disconnected_nodes_get_ancestors_si\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_disconnected_nodes_get_ancestors_si_is_l_set_disconnected_nodes_get_ancestors_si [instances]: "l_set_disconnected_nodes_get_ancestors_si set_disconnected_nodes_locs get_ancestors_si_locs" using instances apply(simp add: l_set_disconnected_nodes_get_ancestors_si_def) using set_disconnected_nodes_get_ancestors_si by fast subsubsection \get\_attribute\ lemma get_attribute_is_l_get_attribute [instances]: "l_get_attribute type_wf get_attribute get_attribute_locs" apply(auto simp add: l_get_attribute_def)[1] using i_get_attribute.get_attribute_reads apply fast using type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t i_get_attribute.get_attribute_ok apply blast using i_get_attribute.get_attribute_ptr_in_heap apply fast done subsubsection \to\_tree\_order\ global_interpretation l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs defines to_tree_order = "l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_to_tree_order get_child_nodes" . declare a_to_tree_order.simps [code] interpretation i_to_tree_order?: l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M ShadowRootClass.known_ptr ShadowRootClass.type_wf Shadow_DOM.get_child_nodes Shadow_DOM.get_child_nodes_locs to_tree_order by(auto simp add: l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def to_tree_order_def instances) declare l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \to\_tree\_order\_si\ locale l_to_tree_order_si\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_child_nodes_defs get_child_nodes get_child_nodes_locs + l_get_shadow_root_defs get_shadow_root get_shadow_root_locs for get_child_nodes :: "(_::linorder) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin partial_function (dom_prog) a_to_tree_order_si :: "(_) object_ptr \ (_, (_) object_ptr list) dom_prog" where "a_to_tree_order_si ptr = (do { children \ get_child_nodes ptr; shadow_root_part \ (case cast ptr of Some element_ptr \ do { shadow_root_opt \ get_shadow_root element_ptr; (case shadow_root_opt of Some shadow_root_ptr \ return [cast shadow_root_ptr] | None \ return []) } | None \ return []); treeorders \ map_M a_to_tree_order_si ((map (cast) children) @ shadow_root_part); return (ptr # concat treeorders) })" end locale l_to_tree_order_si_defs = fixes to_tree_order_si :: "(_) object_ptr \ (_, (_) object_ptr list) dom_prog" global_interpretation l_to_tree_order_si\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_shadow_root get_shadow_root_locs defines to_tree_order_si = "a_to_tree_order_si" . declare a_to_tree_order_si.simps [code] locale l_to_tree_order_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_to_tree_order_si_defs + l_to_tree_order_si\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_get_child_nodes + l_get_shadow_root + assumes to_tree_order_si_impl: "to_tree_order_si = a_to_tree_order_si" begin lemmas to_tree_order_si_def = a_to_tree_order_si.simps[folded to_tree_order_si_impl] lemma to_tree_order_si_pure [simp]: "pure (to_tree_order_si ptr) h" proof - have "\ptr h h' x. h \ to_tree_order_si ptr \\<^sub>r x \ h \ to_tree_order_si ptr \\<^sub>h h' \ h = h'" proof (induct rule: a_to_tree_order_si.fixp_induct[folded to_tree_order_si_impl]) case 1 then show ?case by (rule admissible_dom_prog) next case 2 then show ?case by simp next case (3 f) then have "\x h. pure (f x) h" by (metis is_OK_returns_heap_E is_OK_returns_result_E pure_def) then have "\xs h. pure (map_M f xs) h" by(rule map_M_pure_I) then show ?case by(auto elim!: bind_returns_heap_E2 split: option.splits) qed then show ?thesis unfolding pure_def by (metis is_OK_returns_heap_E is_OK_returns_result_E) qed end interpretation i_to_tree_order_si?: l_to_tree_order_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M to_tree_order_si get_child_nodes get_child_nodes_locs get_shadow_root get_shadow_root_locs type_wf known_ptr by(auto simp add: l_to_tree_order_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_to_tree_order_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def to_tree_order_si_def instances) declare l_to_tree_order_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \first\_in\_tree\_order\ global_interpretation l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs to_tree_order defines first_in_tree_order = "l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_first_in_tree_order to_tree_order" . interpretation i_first_in_tree_order?: l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M to_tree_order first_in_tree_order by(auto simp add: l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def first_in_tree_order_def) declare l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma to_tree_order_is_l_to_tree_order [instances]: "l_to_tree_order to_tree_order" by(auto simp add: l_to_tree_order_def) subsubsection \first\_in\_tree\_order\ global_interpretation l_dummy defines first_in_tree_order_si = "l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_first_in_tree_order to_tree_order_si" . subsubsection \get\_element\_by\ global_interpretation l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs to_tree_order first_in_tree_order get_attribute get_attribute_locs defines get_element_by_id = "l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_element_by_id first_in_tree_order get_attribute" and get_elements_by_class_name = "l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_elements_by_class_name to_tree_order get_attribute" and get_elements_by_tag_name = "l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_elements_by_tag_name to_tree_order" . interpretation i_get_element_by?: l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M to_tree_order first_in_tree_order get_attribute get_attribute_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name type_wf by(auto simp add: l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def get_element_by_id_def get_elements_by_class_name_def get_elements_by_tag_name_def instances) declare l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma get_element_by_is_l_get_element_by [instances]: "l_get_element_by get_element_by_id get_elements_by_tag_name to_tree_order" apply(auto simp add: l_get_element_by_def)[1] using get_element_by_id_result_in_tree_order apply fast done subsubsection \get\_element\_by\_si\ global_interpretation l_dummy defines get_element_by_id_si = "l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_element_by_id first_in_tree_order_si get_attribute" and get_elements_by_class_name_si = "l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_elements_by_class_name to_tree_order_si get_attribute" and get_elements_by_tag_name_si = "l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_elements_by_tag_name to_tree_order_si" . subsubsection \find\_slot\ locale l_find_slot\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_parent_defs get_parent get_parent_locs + l_get_shadow_root_defs get_shadow_root get_shadow_root_locs + l_get_mode_defs get_mode get_mode_locs + l_get_attribute_defs get_attribute get_attribute_locs + l_get_tag_name_defs get_tag_name get_tag_name_locs + l_first_in_tree_order_defs first_in_tree_order for get_parent :: "(_) node_ptr \ ((_) heap, exception, (_::linorder) object_ptr option) prog" and get_parent_locs :: "((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_mode :: "(_) shadow_root_ptr \ ((_) heap, exception, shadow_root_mode) prog" and get_mode_locs :: "(_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" and get_attribute :: "(_) element_ptr \ char list \ ((_) heap, exception, char list option) prog" and get_attribute_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and first_in_tree_order :: "(_) object_ptr \ ((_) object_ptr \ ((_) heap, exception, (_) element_ptr option) prog) \ ((_) heap, exception, (_) element_ptr option) prog" begin definition a_find_slot :: "bool \ (_) node_ptr \ (_, (_) element_ptr option) dom_prog" where "a_find_slot open_flag slotable = do { parent_opt \ get_parent slotable; (case parent_opt of Some parent \ if is_element_ptr_kind parent then do { shadow_root_ptr_opt \ get_shadow_root (the (cast parent)); (case shadow_root_ptr_opt of Some shadow_root_ptr \ do { shadow_root_mode \ get_mode shadow_root_ptr; if open_flag \ shadow_root_mode \ Open then return None else first_in_tree_order (cast shadow_root_ptr) (\ptr. if is_element_ptr_kind ptr then do { tag \ get_tag_name (the (cast ptr)); name_attr \ get_attribute (the (cast ptr)) ''name''; slotable_name_attr \ (if is_element_ptr_kind slotable then get_attribute (the (cast slotable)) ''slot'' else return None); (if (tag = ''slot'' \ (name_attr = slotable_name_attr \ (name_attr = None \ slotable_name_attr = Some '''') \ (name_attr = Some '''' \ slotable_name_attr = None))) then return (Some (the (cast ptr))) else return None)} else return None)} | None \ return None)} else return None | _ \ return None)}" definition a_assigned_slot :: "(_) node_ptr \ (_, (_) element_ptr option) dom_prog" where "a_assigned_slot = a_find_slot True" end global_interpretation l_find_slot\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_parent get_parent_locs get_shadow_root get_shadow_root_locs get_mode get_mode_locs get_attribute get_attribute_locs get_tag_name get_tag_name_locs first_in_tree_order defines find_slot = a_find_slot and assigned_slot = a_assigned_slot . locale l_find_slot_defs = fixes find_slot :: "bool \ (_) node_ptr \ (_, (_) element_ptr option) dom_prog" and assigned_slot :: "(_) node_ptr \ (_, (_) element_ptr option) dom_prog" locale l_find_slot\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_find_slot\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_find_slot_defs + l_get_parent + l_get_shadow_root + l_get_mode + l_get_attribute + l_get_tag_name + l_to_tree_order + l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + assumes find_slot_impl: "find_slot = a_find_slot" assumes assigned_slot_impl: "assigned_slot = a_assigned_slot" begin lemmas find_slot_def = find_slot_impl[unfolded a_find_slot_def] lemmas assigned_slot_def = assigned_slot_impl[unfolded a_assigned_slot_def] lemma find_slot_ptr_in_heap: assumes "h \ find_slot open_flag slotable \\<^sub>r slot_opt" shows "slotable |\| node_ptr_kinds h" using assms apply(auto simp add: find_slot_def elim!: bind_returns_result_E2)[1] using get_parent_ptr_in_heap by blast lemma find_slot_slot_in_heap: assumes "h \ find_slot open_flag slotable \\<^sub>r Some slot" shows "slot |\| element_ptr_kinds h" using assms apply(auto simp add: find_slot_def first_in_tree_order_def elim!: bind_returns_result_E2 map_filter_M_pure_E[where y=slot] split: option.splits if_splits list.splits intro!: map_filter_M_pure bind_pure_I)[1] using get_tag_name_ptr_in_heap by blast+ lemma find_slot_pure [simp]: "pure (find_slot open_flag slotable) h" by(auto simp add: find_slot_def first_in_tree_order_def intro!: bind_pure_I map_filter_M_pure split: option.splits list.splits) end interpretation i_find_slot?: l_find_slot\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_parent get_parent_locs get_shadow_root get_shadow_root_locs get_mode get_mode_locs get_attribute get_attribute_locs get_tag_name get_tag_name_locs first_in_tree_order find_slot assigned_slot type_wf known_ptr known_ptrs get_child_nodes get_child_nodes_locs to_tree_order by (auto simp add: find_slot_def assigned_slot_def l_find_slot\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_find_slot\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_find_slot\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_find_slot = l_find_slot_defs + assumes find_slot_ptr_in_heap: "h \ find_slot open_flag slotable \\<^sub>r slot_opt \ slotable |\| node_ptr_kinds h" assumes find_slot_slot_in_heap: "h \ find_slot open_flag slotable \\<^sub>r Some slot \ slot |\| element_ptr_kinds h" assumes find_slot_pure [simp]: "pure (find_slot open_flag slotable) h" lemma find_slot_is_l_find_slot [instances]: "l_find_slot find_slot" apply(auto simp add: l_find_slot_def)[1] using find_slot_ptr_in_heap apply fast using find_slot_slot_in_heap apply fast done subsubsection \get\_disconnected\_nodes\ locale l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD: l_get_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs + l_type_wf type_wf for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_disconnected_nodes :: "(_) document_ptr \ (_, (_) node_ptr list) dom_prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" begin lemma get_disconnected_nodes_ok: "type_wf h \ document_ptr |\| document_ptr_kinds h \ h \ ok (get_disconnected_nodes document_ptr)" apply(unfold type_wf_impl get_disconnected_nodes_impl[unfolded a_get_disconnected_nodes_def]) using CD.get_disconnected_nodes_ok CD.type_wf_impl ShadowRootClass.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t by blast end interpretation i_get_disconnected_nodes?: l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_disconnected_nodes get_disconnected_nodes_locs by(auto simp add: l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_disconnected_nodes_is_l_get_disconnected_nodes [instances]: "l_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs" apply(auto simp add: l_get_disconnected_nodes_def)[1] using i_get_disconnected_nodes.get_disconnected_nodes_reads apply fast using get_disconnected_nodes_ok apply fast using i_get_disconnected_nodes.get_disconnected_nodes_ptr_in_heap apply fast done paragraph \set\_child\_nodes\ locale l_set_child_nodes_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_child_nodes_get_disconnected_nodes: "\w \ set_child_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_disconnected_nodes_locs ptr'. r h h'))" by(auto simp add: set_child_nodes_locs_def CD.set_child_nodes_locs_def CD.get_disconnected_nodes_locs_def all_args_def) end interpretation i_set_child_nodes_get_disconnected_nodes?: l_set_child_nodes_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr set_child_nodes set_child_nodes_locs Core_DOM_Functions.set_child_nodes Core_DOM_Functions.set_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs apply(auto simp add: l_set_child_nodes_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)[1] by(unfold_locales) declare l_set_child_nodes_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_child_nodes_get_disconnected_nodes_is_l_set_child_nodes_get_disconnected_nodes [instances]: "l_set_child_nodes_get_disconnected_nodes type_wf set_child_nodes set_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs" using set_child_nodes_is_l_set_child_nodes get_disconnected_nodes_is_l_get_disconnected_nodes apply(simp add: l_set_child_nodes_get_disconnected_nodes_def l_set_child_nodes_get_disconnected_nodes_axioms_def) using set_child_nodes_get_disconnected_nodes by fast paragraph \set\_disconnected\_nodes\ lemma set_disconnected_nodes_get_disconnected_nodes_l_set_disconnected_nodes_get_disconnected_nodes [instances]: "l_set_disconnected_nodes_get_disconnected_nodes ShadowRootClass.type_wf get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs" apply(auto simp add: l_set_disconnected_nodes_get_disconnected_nodes_def l_set_disconnected_nodes_get_disconnected_nodes_axioms_def instances)[1] using i_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes apply fast using i_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes_different_pointers apply fast done paragraph \delete\_shadow\_root\ locale l_delete_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_disconnected_nodes_delete_shadow_root: "h \ delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M shadow_root_ptr \\<^sub>h h' \ r \ get_disconnected_nodes_locs ptr' \ r h h'" by(auto simp add: CD.get_disconnected_nodes_locs_def delete_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t) end locale l_delete_shadow_root_get_disconnected_nodes = l_get_disconnected_nodes_defs + assumes get_disconnected_nodes_delete_shadow_root: "h \ delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M shadow_root_ptr \\<^sub>h h' \ r \ get_disconnected_nodes_locs ptr' \ r h h'" interpretation l_delete_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_disconnected_nodes get_disconnected_nodes_locs by(auto simp add: l_delete_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) lemma l_delete_shadow_root_get_disconnected_nodes_get_disconnected_nodes_locs [instances]: "l_delete_shadow_root_get_disconnected_nodes get_disconnected_nodes_locs" apply(auto simp add: l_delete_shadow_root_get_disconnected_nodes_def)[1] using get_disconnected_nodes_delete_shadow_root apply fast done paragraph \set\_shadow\_root\ locale l_set_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_shadow_root_get_disconnected_nodes: "\w \ set_shadow_root_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_disconnected_nodes_locs ptr'. r h h'))" by(auto simp add: set_shadow_root_locs_def CD.get_disconnected_nodes_locs_def all_args_def) end interpretation i_set_shadow_root_get_disconnected_nodes?: l_set_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_shadow_root set_shadow_root_locs DocumentClass.type_wf get_disconnected_nodes get_disconnected_nodes_locs apply(auto simp add: l_set_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)[1] by(unfold_locales) declare l_set_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_set_shadow_root_get_disconnected_nodes = l_set_shadow_root_defs + l_get_disconnected_nodes_defs + assumes set_shadow_root_get_disconnected_nodes: "\w \ set_shadow_root_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_disconnected_nodes_locs ptr'. r h h'))" lemma set_shadow_root_get_disconnected_nodes_is_l_set_shadow_root_get_disconnected_nodes [instances]: "l_set_shadow_root_get_disconnected_nodes set_shadow_root_locs get_disconnected_nodes_locs" using set_shadow_root_is_l_set_shadow_root get_disconnected_nodes_is_l_get_disconnected_nodes apply(simp add: l_set_shadow_root_get_disconnected_nodes_def ) using set_shadow_root_get_disconnected_nodes by fast paragraph \set\_mode\ locale l_set_mode_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_mode_get_disconnected_nodes: "\w \ set_mode_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_disconnected_nodes_locs ptr'. r h h'))" by(auto simp add: set_mode_locs_def CD.get_disconnected_nodes_locs_impl[unfolded CD.a_get_disconnected_nodes_locs_def] all_args_def) end interpretation i_set_mode_get_disconnected_nodes?: l_set_mode_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_mode set_mode_locs DocumentClass.type_wf get_disconnected_nodes get_disconnected_nodes_locs by unfold_locales declare l_set_mode_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_set_mode_get_disconnected_nodes = l_set_mode + l_get_disconnected_nodes + assumes set_mode_get_disconnected_nodes: "\w \ set_mode_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_disconnected_nodes_locs ptr'. r h h'))" lemma set_mode_get_disconnected_nodes_is_l_set_mode_get_disconnected_nodes [instances]: "l_set_mode_get_disconnected_nodes type_wf set_mode set_mode_locs get_disconnected_nodes get_disconnected_nodes_locs" using set_mode_is_l_set_mode get_disconnected_nodes_is_l_get_disconnected_nodes apply(simp add: l_set_mode_get_disconnected_nodes_def l_set_mode_get_disconnected_nodes_axioms_def) using set_mode_get_disconnected_nodes by fast paragraph \new\_shadow\_root\ locale l_new_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_disconnected_nodes :: "(_) document_ptr \ (_, (_) node_ptr list) dom_prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_disconnected_nodes_new_shadow_root: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ r \ get_disconnected_nodes_locs ptr' \ r h h'" by(auto simp add: CD.get_disconnected_nodes_locs_def new_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t) end interpretation i_new_shadow_root_get_disconnected_nodes?: l_new_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_disconnected_nodes get_disconnected_nodes_locs by unfold_locales declare l_new_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_new_shadow_root_get_disconnected_nodes = l_get_disconnected_nodes_defs + assumes get_disconnected_nodes_new_shadow_root: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ r \ get_disconnected_nodes_locs ptr' \ r h h'" lemma new_shadow_root_get_disconnected_nodes_is_l_new_shadow_root_get_disconnected_nodes [instances]: "l_new_shadow_root_get_disconnected_nodes get_disconnected_nodes_locs" apply (auto simp add: l_new_shadow_root_get_disconnected_nodes_def)[1] using get_disconnected_nodes_new_shadow_root apply fast done subsubsection \remove\_shadow\_root\ locale l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_child_nodes_defs get_child_nodes get_child_nodes_locs + l_get_shadow_root_defs get_shadow_root get_shadow_root_locs + l_set_shadow_root_defs set_shadow_root set_shadow_root_locs + l_get_disconnected_nodes_defs get_disconnected_nodes get_disconnected_nodes_locs for get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and set_shadow_root :: "(_) element_ptr \ (_) shadow_root_ptr option \ ((_) heap, exception, unit) prog" and set_shadow_root_locs :: "(_) element_ptr \ ((_) heap, exception, unit) prog set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" begin definition a_remove_shadow_root :: "(_) element_ptr \ (_, unit) dom_prog" where "a_remove_shadow_root element_ptr = do { shadow_root_ptr_opt \ get_shadow_root element_ptr; (case shadow_root_ptr_opt of Some shadow_root_ptr \ do { children \ get_child_nodes (cast shadow_root_ptr); (if children = [] then do { set_shadow_root element_ptr None; delete_M shadow_root_ptr } else do { error HierarchyRequestError }) } | None \ error HierarchyRequestError) }" definition a_remove_shadow_root_locs :: "(_) element_ptr \ (_) shadow_root_ptr \ ((_, unit) dom_prog) set" where "a_remove_shadow_root_locs element_ptr shadow_root_ptr \ set_shadow_root_locs element_ptr \ {delete_M shadow_root_ptr}" end global_interpretation l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_shadow_root get_shadow_root_locs set_shadow_root set_shadow_root_locs get_disconnected_nodes get_disconnected_nodes_locs defines remove_shadow_root = "a_remove_shadow_root" and remove_shadow_root_locs = a_remove_shadow_root_locs . locale l_remove_shadow_root_defs = fixes remove_shadow_root :: "(_) element_ptr \ (_, unit) dom_prog" fixes remove_shadow_root_locs :: "(_) element_ptr \ (_) shadow_root_ptr \ ((_, unit) dom_prog) set" locale l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_remove_shadow_root_defs + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_child_nodes + l_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs + assumes remove_shadow_root_impl: "remove_shadow_root = a_remove_shadow_root" assumes remove_shadow_root_locs_impl: "remove_shadow_root_locs = a_remove_shadow_root_locs" begin lemmas remove_shadow_root_def = remove_shadow_root_impl[unfolded remove_shadow_root_def a_remove_shadow_root_def] lemmas remove_shadow_root_locs_def = remove_shadow_root_locs_impl[unfolded remove_shadow_root_locs_def a_remove_shadow_root_locs_def] lemma remove_shadow_root_writes: "writes (remove_shadow_root_locs element_ptr (the |h \ get_shadow_root element_ptr|\<^sub>r)) (remove_shadow_root element_ptr) h h'" apply(auto simp add: remove_shadow_root_locs_def remove_shadow_root_def all_args_def writes_union_right_I writes_union_left_I set_shadow_root_writes intro!: writes_bind writes_bind_pure[OF get_shadow_root_pure] writes_bind_pure[OF get_child_nodes_pure] intro: writes_subset[OF set_shadow_root_writes] writes_subset[OF writes_singleton2] split: option.splits)[1] using writes_union_left_I[OF set_shadow_root_writes] apply (metis inf_sup_aci(5) insert_is_Un) using writes_union_right_I[OF writes_singleton[of delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M]] by (smt insert_is_Un writes_singleton2 writes_union_left_I) end interpretation i_remove_shadow_root?: l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_shadow_root get_shadow_root_locs set_shadow_root set_shadow_root_locs get_disconnected_nodes get_disconnected_nodes_locs remove_shadow_root remove_shadow_root_locs type_wf known_ptr by(auto simp add: l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def remove_shadow_root_def remove_shadow_root_locs_def instances) declare l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] paragraph \get\_child\_nodes\ locale l_remove_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma remove_shadow_root_get_child_nodes_different_pointers: assumes "ptr \ cast shadow_root_ptr" assumes "w \ remove_shadow_root_locs element_ptr shadow_root_ptr" assumes "h \ w \\<^sub>h h'" assumes "r \ get_child_nodes_locs ptr" shows "r h h'" using assms apply(auto simp add: all_args_def get_child_nodes_locs_def CD.get_child_nodes_locs_def remove_shadow_root_locs_def set_shadow_root_locs_def delete_shadow_root_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t delete_shadow_root_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t delete_shadow_root_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t delete_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t delete_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t[rotated] element_put_get_preserved[where setter=shadow_root_opt_update] intro: is_shadow_root_ptr_kind_obtains elim: is_document_ptr_kind_obtains is_shadow_root_ptr_kind_obtains split: if_splits option.splits)[1] done end locale l_remove_shadow_root_get_child_nodes = l_get_child_nodes_defs + l_remove_shadow_root_defs + assumes remove_shadow_root_get_child_nodes_different_pointers: "ptr \ cast shadow_root_ptr \ w \ remove_shadow_root_locs element_ptr shadow_root_ptr \ h \ w \\<^sub>h h' \ r \ get_child_nodes_locs ptr \ r h h'" interpretation i_remove_shadow_root_get_child_nodes?: l_remove_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs get_shadow_root get_shadow_root_locs set_shadow_root set_shadow_root_locs get_disconnected_nodes get_disconnected_nodes_locs remove_shadow_root remove_shadow_root_locs by(auto simp add: l_remove_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_remove_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma remove_shadow_root_get_child_nodes_is_l_remove_shadow_root_get_child_nodes [instances]: "l_remove_shadow_root_get_child_nodes get_child_nodes_locs remove_shadow_root_locs" apply(auto simp add: l_remove_shadow_root_get_child_nodes_def instances )[1] using remove_shadow_root_get_child_nodes_different_pointers apply fast done paragraph \get\_tag\_name\ locale l_remove_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma remove_shadow_root_get_tag_name: assumes "w \ remove_shadow_root_locs element_ptr shadow_root_ptr" assumes "h \ w \\<^sub>h h'" assumes "r \ get_tag_name_locs ptr" shows "r h h'" using assms by(auto simp add: all_args_def remove_shadow_root_locs_def set_shadow_root_locs_def CD.get_tag_name_locs_def delete_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_put_get_preserved[where setter=shadow_root_opt_update] split: if_splits option.splits) end locale l_remove_shadow_root_get_tag_name = l_get_tag_name_defs + l_remove_shadow_root_defs + assumes remove_shadow_root_get_tag_name: "w \ remove_shadow_root_locs element_ptr shadow_root_ptr \ h \ w \\<^sub>h h' \ r \ get_tag_name_locs ptr \ r h h'" interpretation i_remove_shadow_root_get_tag_name?: l_remove_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_tag_name get_tag_name_locs get_child_nodes get_child_nodes_locs get_shadow_root get_shadow_root_locs set_shadow_root set_shadow_root_locs get_disconnected_nodes get_disconnected_nodes_locs remove_shadow_root remove_shadow_root_locs known_ptr by(auto simp add: l_remove_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_remove_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma remove_shadow_root_get_tag_name_is_l_remove_shadow_root_get_tag_name [instances]: "l_remove_shadow_root_get_tag_name get_tag_name_locs remove_shadow_root_locs" apply(auto simp add: l_remove_shadow_root_get_tag_name_def instances )[1] using remove_shadow_root_get_tag_name apply fast done subsubsection \get\_owner\_document\ locale l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_host_defs get_host get_host_locs + CD: l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_root_node get_root_node_locs get_disconnected_nodes get_disconnected_nodes_locs for get_root_node :: "(_::linorder) object_ptr \ ((_) heap, exception, (_) object_ptr) prog" and get_root_node_locs :: "((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" begin definition a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) shadow_root_ptr \ unit \ (_, (_) document_ptr) dom_prog" where "a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr _ = do { host \ get_host shadow_root_ptr; CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r (cast host) () }" definition a_get_owner_document_tups :: "(((_) object_ptr \ bool) \ ((_) object_ptr \ unit \ (_, (_) document_ptr) dom_prog)) list" where "a_get_owner_document_tups = [(is_shadow_root_ptr, a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast)]" definition a_get_owner_document :: "(_::linorder) object_ptr \ (_, (_) document_ptr) dom_prog" where "a_get_owner_document ptr = invoke (CD.a_get_owner_document_tups @ a_get_owner_document_tups) ptr ()" end global_interpretation l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_root_node_si get_root_node_si_locs get_disconnected_nodes get_disconnected_nodes_locs get_host get_host_locs defines get_owner_document_tups = a_get_owner_document_tups and get_owner_document = a_get_owner_document and get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r = a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r and get_owner_document_tups\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = "l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_owner_document_tups get_root_node_si get_disconnected_nodes" and get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r = "l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r get_root_node_si get_disconnected_nodes" . locale l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_known_ptr known_ptr + l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_root_node_si get_root_node_si_locs get_disconnected_nodes get_disconnected_nodes_locs get_host get_host_locs + l_get_owner_document_defs get_owner_document + l_get_host get_host get_host_locs + CD: l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_parent get_parent_locs known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_disconnected_nodes get_disconnected_nodes_locs get_root_node_si get_root_node_si_locs get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M for known_ptr :: "(_::linorder) object_ptr \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_::linorder) object_ptr \ bool" and get_parent :: "(_) node_ptr \ ((_) heap, exception, (_) object_ptr option) prog" and get_parent_locs :: "((_) heap \ (_) heap \ bool) set" and type_wf :: "(_) heap \ bool" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and get_root_node_si :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr) prog" and get_root_node_si_locs :: "((_) heap \ (_) heap \ bool) set" and get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap, exception, (_) document_ptr) prog" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_owner_document :: "(_) object_ptr \ ((_) heap, exception, (_) document_ptr) prog" + assumes known_ptr_impl: "known_ptr = a_known_ptr" assumes get_owner_document_impl: "get_owner_document = a_get_owner_document" begin lemmas known_ptr_def = known_ptr_impl[unfolded a_known_ptr_def] lemmas get_owner_document_def = a_get_owner_document_def[folded get_owner_document_impl] lemma get_owner_document_pure [simp]: "pure (get_owner_document ptr) h" proof - have "\shadow_root_ptr. pure (a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr ()) h" apply(auto simp add: a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_pure_I filter_M_pure_I split: option.splits)[1] by(auto simp add: CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_pure_I filter_M_pure_I split: option.splits) then show ?thesis apply(auto simp add: get_owner_document_def)[1] apply(split CD.get_owner_document_splits, rule conjI)+ apply(simp) apply(auto simp add: a_get_owner_document_tups_def)[1] apply(split invoke_splits, rule conjI)+ apply simp by(auto intro!: bind_pure_I) qed lemma get_owner_document_ptr_in_heap: assumes "h \ ok (get_owner_document ptr)" shows "ptr |\| object_ptr_kinds h" using assms by(auto simp add: get_owner_document_def invoke_ptr_in_heap dest: is_OK_returns_heap_I) end interpretation i_get_owner_document?: l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr DocumentClass.known_ptr get_parent get_parent_locs type_wf get_disconnected_nodes get_disconnected_nodes_locs get_root_node_si get_root_node_si_locs CD.a_get_owner_document get_host get_host_locs get_owner_document by(auto simp add: instances l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def get_owner_document_def Core_DOM_Functions.get_owner_document_def) declare l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_owner_document_is_l_get_owner_document [instances]: "l_get_owner_document get_owner_document" apply(auto simp add: l_get_owner_document_def)[1] using get_owner_document_ptr_in_heap apply fast done subsubsection \remove\_child\ global_interpretation l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs set_child_nodes set_child_nodes_locs get_parent get_parent_locs get_owner_document get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs defines remove = "l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_remove get_child_nodes set_child_nodes get_parent get_owner_document get_disconnected_nodes set_disconnected_nodes" and remove_child = "l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_remove_child get_child_nodes set_child_nodes get_owner_document get_disconnected_nodes set_disconnected_nodes" and remove_child_locs = "l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_remove_child_locs set_child_nodes_locs set_disconnected_nodes_locs" . interpretation i_remove_child?: l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M Shadow_DOM.get_child_nodes Shadow_DOM.get_child_nodes_locs Shadow_DOM.set_child_nodes Shadow_DOM.set_child_nodes_locs Shadow_DOM.get_parent Shadow_DOM.get_parent_locs Shadow_DOM.get_owner_document get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs remove_child remove_child_locs remove ShadowRootClass.type_wf ShadowRootClass.known_ptr ShadowRootClass.known_ptrs by(auto simp add: l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def remove_child_def remove_child_locs_def remove_def instances) declare l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_disconnected\_document\ locale l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_disconnected_nodes_defs get_disconnected_nodes get_disconnected_nodes_locs for get_disconnected_nodes :: "(_::linorder) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" begin definition a_get_disconnected_document :: "(_) node_ptr \ (_, (_) document_ptr) dom_prog" where "a_get_disconnected_document node_ptr = do { check_in_heap (cast node_ptr); ptrs \ document_ptr_kinds_M; candidates \ filter_M (\document_ptr. do { disconnected_nodes \ get_disconnected_nodes document_ptr; return (node_ptr \ set disconnected_nodes) }) ptrs; (case candidates of Cons document_ptr [] \ return document_ptr | _ \ error HierarchyRequestError ) }" definition "a_get_disconnected_document_locs = (\document_ptr. get_disconnected_nodes_locs document_ptr) \ (\ptr. {preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr RObject.nothing)})" end locale l_get_disconnected_document_defs = fixes get_disconnected_document :: "(_) node_ptr \ (_, (_::linorder) document_ptr) dom_prog" fixes get_disconnected_document_locs :: "((_) heap \ (_) heap \ bool) set" locale l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_get_disconnected_document_defs + l_get_disconnected_nodes + assumes get_disconnected_document_impl: "get_disconnected_document = a_get_disconnected_document" assumes get_disconnected_document_locs_impl: "get_disconnected_document_locs = a_get_disconnected_document_locs" begin lemmas get_disconnected_document_def = get_disconnected_document_impl[unfolded a_get_disconnected_document_def] lemmas get_disconnected_document_locs_def = get_disconnected_document_locs_impl[unfolded a_get_disconnected_document_locs_def] lemma get_disconnected_document_pure [simp]: "pure (get_disconnected_document ptr) h" using get_disconnected_nodes_pure by(auto simp add: get_disconnected_document_def intro!: bind_pure_I filter_M_pure_I split: list.splits) lemma get_disconnected_document_ptr_in_heap [simp]: "h \ ok (get_disconnected_document node_ptr) \ node_ptr |\| node_ptr_kinds h" using get_disconnected_document_def is_OK_returns_result_I check_in_heap_ptr_in_heap by (metis (no_types, lifting) bind_returns_heap_E get_disconnected_document_pure node_ptr_kinds_commutes pure_pure) lemma get_disconnected_document_disconnected_document_in_heap: assumes "h \ get_disconnected_document child_node \\<^sub>r disconnected_document" shows "disconnected_document |\| document_ptr_kinds h" using assms get_disconnected_nodes_pure by(auto simp add: get_disconnected_document_def elim!: bind_returns_result_E2 dest!: filter_M_not_more_elements[where x=disconnected_document] intro!: filter_M_pure_I bind_pure_I split: if_splits list.splits) lemma get_disconnected_document_reads: "reads get_disconnected_document_locs (get_disconnected_document node_ptr) h h'" using get_disconnected_nodes_reads[unfolded reads_def] by(auto simp add: get_disconnected_document_def get_disconnected_document_locs_def intro!: reads_bind_pure reads_subset[OF check_in_heap_reads] reads_subset[OF error_reads] reads_subset[OF get_disconnected_nodes_reads] reads_subset[OF return_reads] reads_subset[OF document_ptr_kinds_M_reads] filter_M_reads filter_M_pure_I bind_pure_I split: list.splits) end locale l_get_disconnected_document = l_get_disconnected_document_defs + assumes get_disconnected_document_reads: "reads get_disconnected_document_locs (get_disconnected_document node_ptr) h h'" assumes get_disconnected_document_ptr_in_heap: "h \ ok (get_disconnected_document node_ptr) \ node_ptr |\| node_ptr_kinds h" assumes get_disconnected_document_pure [simp]: "pure (get_disconnected_document node_ptr) h" assumes get_disconnected_document_disconnected_document_in_heap: "h \ get_disconnected_document child_node \\<^sub>r disconnected_document \ disconnected_document |\| document_ptr_kinds h" global_interpretation l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_disconnected_nodes get_disconnected_nodes_locs defines get_disconnected_document = a_get_disconnected_document and get_disconnected_document_locs = a_get_disconnected_document_locs . interpretation i_get_disconnected_document?: l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs get_disconnected_document get_disconnected_document_locs type_wf by(auto simp add: l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def get_disconnected_document_def get_disconnected_document_locs_def instances) declare l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_disconnected_document_is_l_get_disconnected_document [instances]: "l_get_disconnected_document get_disconnected_document get_disconnected_document_locs" apply(auto simp add: l_get_disconnected_document_def instances)[1] using get_disconnected_document_ptr_in_heap get_disconnected_document_pure get_disconnected_document_disconnected_document_in_heap get_disconnected_document_reads by blast+ paragraph \get\_disconnected\_nodes\ locale l_set_tag_name_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_tag_name_get_disconnected_nodes: "\w \ set_tag_name_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_disconnected_nodes_locs ptr'. r h h'))" by(auto simp add: CD.set_tag_name_locs_impl[unfolded CD.a_set_tag_name_locs_def] CD.get_disconnected_nodes_locs_impl[unfolded CD.a_get_disconnected_nodes_locs_def] all_args_def) end interpretation i_set_tag_name_get_disconnected_nodes?: l_set_tag_name_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_tag_name set_tag_name_locs get_disconnected_nodes get_disconnected_nodes_locs by unfold_locales declare l_set_tag_name_get_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_tag_name_get_disconnected_nodes_is_l_set_tag_name_get_disconnected_nodes [instances]: "l_set_tag_name_get_disconnected_nodes type_wf set_tag_name set_tag_name_locs get_disconnected_nodes get_disconnected_nodes_locs" using set_tag_name_is_l_set_tag_name get_disconnected_nodes_is_l_get_disconnected_nodes apply(simp add: l_set_tag_name_get_disconnected_nodes_def l_set_tag_name_get_disconnected_nodes_axioms_def) using set_tag_name_get_disconnected_nodes by fast subsubsection \adopt\_node\ global_interpretation l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_owner_document get_parent get_parent_locs remove_child remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs defines adopt_node = a_adopt_node and adopt_node_locs = a_adopt_node_locs . interpretation i_adopt_node?: l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_owner_document get_parent get_parent_locs remove_child remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs adopt_node adopt_node_locs known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs set_child_nodes set_child_nodes_locs remove by(auto simp add: l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def adopt_node_def adopt_node_locs_def instances) declare l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma adopt_node_is_l_adopt_node [instances]: "l_adopt_node type_wf known_ptr known_ptrs get_parent adopt_node adopt_node_locs get_child_nodes get_owner_document" apply(auto simp add: l_adopt_node_def l_adopt_node_axioms_def instances)[1] using adopt_node_writes apply fast using adopt_node_pointers_preserved apply (fast, fast) using adopt_node_types_preserved apply (fast, fast) using adopt_node_child_in_heap apply fast using adopt_node_children_subset apply fast done paragraph \get\_shadow\_root\ locale l_adopt_node_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_child_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_disconnected_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma adopt_node_get_shadow_root: "\w \ adopt_node_locs parent owner_document document_ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" by(auto simp add: adopt_node_locs_def remove_child_locs_def all_args_def set_disconnected_nodes_get_shadow_root set_child_nodes_get_shadow_root) end locale l_adopt_node_get_shadow_root = l_adopt_node_defs + l_get_shadow_root_defs + assumes adopt_node_get_shadow_root: "\w \ adopt_node_locs parent owner_document document_ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" interpretation i_adopt_node_get_shadow_root?: l_adopt_node_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr set_child_nodes set_child_nodes_locs Core_DOM_Functions.set_child_nodes Core_DOM_Functions.set_child_nodes_locs get_shadow_root get_shadow_root_locs set_disconnected_nodes set_disconnected_nodes_locs get_owner_document get_parent get_parent_locs remove_child remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs adopt_node adopt_node_locs get_child_nodes get_child_nodes_locs known_ptrs remove by(auto simp add: l_adopt_node_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_adopt_node_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma adopt_node_get_shadow_root_is_l_adopt_node_get_shadow_root [instances]: "l_adopt_node_get_shadow_root adopt_node_locs get_shadow_root_locs" apply(auto simp add: l_adopt_node_get_shadow_root_def)[1] using adopt_node_get_shadow_root apply fast done subsubsection \insert\_before\ global_interpretation l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_parent get_parent_locs get_child_nodes get_child_nodes_locs set_child_nodes set_child_nodes_locs get_ancestors_si get_ancestors_si_locs adopt_node adopt_node_locs set_disconnected_nodes set_disconnected_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_owner_document defines next_sibling = a_next_sibling and insert_node = a_insert_node and ensure_pre_insertion_validity = a_ensure_pre_insertion_validity and insert_before = a_insert_before and insert_before_locs = a_insert_before_locs . global_interpretation l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs insert_before defines append_child = "l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_append_child insert_before" . interpretation i_insert_before?: l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_parent get_parent_locs get_child_nodes get_child_nodes_locs set_child_nodes set_child_nodes_locs get_ancestors_si get_ancestors_si_locs adopt_node adopt_node_locs set_disconnected_nodes set_disconnected_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_owner_document insert_before insert_before_locs append_child type_wf known_ptr known_ptrs by(auto simp add: l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def insert_before_def insert_before_locs_def instances) declare l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] interpretation i_append_child?: l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M append_child insert_before insert_before_locs by(simp add: l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances append_child_def) declare l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] subsubsection \get\_assigned\_nodes\ fun map_filter_M2 :: "('x \ ('heap, 'e, 'y option) prog) \ 'x list \ ('heap, 'e, 'y list) prog" where "map_filter_M2 f [] = return []" | "map_filter_M2 f (x # xs) = do { res \ f x; remainder \ map_filter_M2 f xs; return ((case res of Some r \ [r] | None \ []) @ remainder) }" lemma map_filter_M2_pure [simp]: assumes "\x. x \ set xs \ pure (f x) h" shows "pure (map_filter_M2 f xs) h" using assms apply(induct xs arbitrary: h) by(auto elim!: bind_returns_result_E2 intro!: bind_pure_I) lemma map_filter_pure_no_monad: assumes "\x. x \ set xs \ pure (f x) h" assumes "h \ map_filter_M2 f xs \\<^sub>r ys" shows "ys = map the (filter (\x. x \ None) (map (\x. |h \ f x|\<^sub>r) xs))" and "\x. x \ set xs \ h \ ok (f x)" using assms apply(induct xs arbitrary: h ys) by(auto elim!: bind_returns_result_E2) lemma map_filter_pure_foo: assumes "\x. x \ set xs \ pure (f x) h" assumes "h \ map_filter_M2 f xs \\<^sub>r ys" assumes "y \ set ys" obtains x where "h \ f x \\<^sub>r Some y" and "x \ set xs" using assms apply(induct xs arbitrary: ys) by(auto elim!: bind_returns_result_E2) lemma map_filter_M2_in_result: assumes "h \ map_filter_M2 P xs \\<^sub>r ys" assumes "a \ set xs" assumes "\x. x \ set xs \ pure (P x) h" assumes "h \ P a \\<^sub>r Some b" shows "b \ set ys" using assms apply(induct xs arbitrary: h ys) by(auto elim!: bind_returns_result_E2 ) locale l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_tag_name_defs get_tag_name get_tag_name_locs + l_get_root_node_defs get_root_node get_root_node_locs + l_get_host_defs get_host get_host_locs + l_get_child_nodes_defs get_child_nodes get_child_nodes_locs + l_find_slot_defs find_slot assigned_slot + l_remove_defs remove + l_insert_before_defs insert_before insert_before_locs + l_append_child_defs append_child + l_remove_shadow_root_defs remove_shadow_root remove_shadow_root_locs for get_child_nodes :: "(_::linorder) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_root_node :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr) prog" and get_root_node_locs :: "((_) heap \ (_) heap \ bool) set" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and find_slot :: "bool \ (_) node_ptr \ ((_) heap, exception, (_) element_ptr option) prog" and assigned_slot :: "(_) node_ptr \ ((_) heap, exception, (_) element_ptr option) prog" and remove :: "(_) node_ptr \ ((_) heap, exception, unit) prog" and insert_before :: "(_) object_ptr \ (_) node_ptr \ (_) node_ptr option \ ((_) heap, exception, unit) prog" and insert_before_locs :: "(_) object_ptr \ (_) object_ptr option \ (_) document_ptr \ (_) document_ptr \ (_, unit) dom_prog set" and append_child :: "(_) object_ptr \ (_) node_ptr \ ((_) heap, exception, unit) prog" and remove_shadow_root :: "(_) element_ptr \ ((_) heap, exception, unit) prog" and remove_shadow_root_locs :: "(_) element_ptr \ (_) shadow_root_ptr \ ((_) heap, exception, unit) prog set" begin definition a_assigned_nodes :: "(_) element_ptr \ (_, (_) node_ptr list) dom_prog" where "a_assigned_nodes slot = do { tag \ get_tag_name slot; (if tag \ ''slot'' then error HierarchyRequestError else return ()); root \ get_root_node (cast slot); if is_shadow_root_ptr_kind root then do { host \ get_host (the (cast root)); children \ get_child_nodes (cast host); filter_M (\slotable. do { found_slot \ find_slot False slotable; return (found_slot = Some slot)}) children} else return []}" partial_function (dom_prog) a_assigned_nodes_flatten :: "(_) element_ptr \ (_, (_) node_ptr list) dom_prog" where "a_assigned_nodes_flatten slot = do { tag \ get_tag_name slot; (if tag \ ''slot'' then error HierarchyRequestError else return ()); root \ get_root_node (cast slot); (if is_shadow_root_ptr_kind root then do { slotables \ a_assigned_nodes slot; slotables_or_child_nodes \ (if slotables = [] then do { get_child_nodes (cast slot) } else do { return slotables }); list_of_lists \ map_M (\node_ptr. do { (case cast node_ptr of Some element_ptr \ do { tag \ get_tag_name element_ptr; (if tag = ''slot'' then do { root \ get_root_node (cast element_ptr); (if is_shadow_root_ptr_kind root then do { a_assigned_nodes_flatten element_ptr } else do { return [node_ptr] }) } else do { return [node_ptr] }) } | None \ return [node_ptr]) }) slotables_or_child_nodes; return (concat list_of_lists) } else return []) }" definition a_flatten_dom :: "(_, unit) dom_prog" where "a_flatten_dom = do { tups \ element_ptr_kinds_M \ map_filter_M2 (\element_ptr. do { tag \ get_tag_name element_ptr; assigned_nodes \ a_assigned_nodes element_ptr; (if tag = ''slot'' \ assigned_nodes \ [] then return (Some (element_ptr, assigned_nodes)) else return None)}); forall_M (\(slot, assigned_nodes). do { get_child_nodes (cast slot) \ forall_M remove; forall_M (append_child (cast slot)) assigned_nodes }) tups; shadow_root_ptr_kinds_M \ forall_M (\shadow_root_ptr. do { host \ get_host shadow_root_ptr; get_child_nodes (cast host) \ forall_M remove; get_child_nodes (cast shadow_root_ptr) \ forall_M (append_child (cast host)); remove_shadow_root host }); return () }" end global_interpretation l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_tag_name get_tag_name_locs get_root_node get_root_node_locs get_host get_host_locs find_slot assigned_slot remove insert_before insert_before_locs append_child remove_shadow_root remove_shadow_root_locs defines assigned_nodes = a_assigned_nodes and assigned_nodes_flatten = a_assigned_nodes_flatten and flatten_dom = a_flatten_dom . declare a_assigned_nodes_flatten.simps [code] locale l_assigned_nodes_defs = fixes assigned_nodes :: "(_) element_ptr \ (_, (_) node_ptr list) dom_prog" fixes assigned_nodes_flatten :: "(_) element_ptr \ (_, (_) node_ptr list) dom_prog" fixes flatten_dom :: "(_, unit) dom_prog" locale l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_assigned_nodes_defs assigned_nodes assigned_nodes_flatten flatten_dom + l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_tag_name get_tag_name_locs get_root_node get_root_node_locs get_host get_host_locs find_slot assigned_slot remove insert_before insert_before_locs append_child remove_shadow_root remove_shadow_root_locs + l_get_shadow_root type_wf get_shadow_root get_shadow_root_locs + l_set_shadow_root type_wf set_shadow_root set_shadow_root_locs + l_remove + l_insert_before insert_before insert_before_locs + l_find_slot find_slot assigned_slot + l_get_tag_name type_wf get_tag_name get_tag_name_locs + l_get_root_node get_root_node get_root_node_locs get_parent get_parent_locs + l_get_host get_host get_host_locs + l_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs + l_to_tree_order to_tree_order for known_ptr :: "(_::linorder) object_ptr \ bool" and assigned_nodes :: "(_) element_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and assigned_nodes_flatten :: "(_) element_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and flatten_dom :: "((_) heap, exception, unit) prog" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_root_node :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr) prog" and get_root_node_locs :: "((_) heap \ (_) heap \ bool) set" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and find_slot :: "bool \ (_) node_ptr \ ((_) heap, exception, (_) element_ptr option) prog" and assigned_slot :: "(_) node_ptr \ ((_) heap, exception, (_) element_ptr option) prog" and remove :: "(_) node_ptr \ ((_) heap, exception, unit) prog" and insert_before :: "(_) object_ptr \ (_) node_ptr \ (_) node_ptr option \ ((_) heap, exception, unit) prog" and insert_before_locs :: "(_) object_ptr \ (_) object_ptr option \ (_) document_ptr \ (_) document_ptr \ (_, unit) dom_prog set" and append_child :: "(_) object_ptr \ (_) node_ptr \ ((_) heap, exception, unit) prog" and remove_shadow_root :: "(_) element_ptr \ ((_) heap, exception, unit) prog" and remove_shadow_root_locs :: "(_) element_ptr \ (_) shadow_root_ptr \ ((_) heap, exception, unit) prog set" and type_wf :: "(_) heap \ bool" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and set_shadow_root :: "(_) element_ptr \ (_) shadow_root_ptr option \ ((_) heap, exception, unit) prog" and set_shadow_root_locs :: "(_) element_ptr \ ((_) heap, exception, unit) prog set" and get_parent :: "(_) node_ptr \ ((_) heap, exception, (_) object_ptr option) prog" and get_parent_locs :: "((_) heap \ (_) heap \ bool) set" and to_tree_order :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr list) prog" + assumes assigned_nodes_impl: "assigned_nodes = a_assigned_nodes" assumes flatten_dom_impl: "flatten_dom = a_flatten_dom" begin lemmas assigned_nodes_def = assigned_nodes_impl[unfolded a_assigned_nodes_def] lemmas flatten_dom_def = flatten_dom_impl[unfolded a_flatten_dom_def, folded assigned_nodes_impl] lemma assigned_nodes_pure [simp]: "pure (assigned_nodes slot) h" by(auto simp add: assigned_nodes_def intro!: bind_pure_I filter_M_pure_I) lemma assigned_nodes_ptr_in_heap: assumes "h \ ok (assigned_nodes slot)" shows "slot |\| element_ptr_kinds h" using assms apply(auto simp add: assigned_nodes_def)[1] by (meson bind_is_OK_E is_OK_returns_result_I local.get_tag_name_ptr_in_heap) lemma assigned_nodes_slot_is_slot: assumes "h \ ok (assigned_nodes slot)" shows "h \ get_tag_name slot \\<^sub>r ''slot''" using assms by(auto simp add: assigned_nodes_def elim!: bind_is_OK_E split: if_splits) lemma assigned_nodes_different_ptr: assumes "h \ assigned_nodes slot \\<^sub>r nodes" assumes "h \ assigned_nodes slot' \\<^sub>r nodes'" assumes "slot \ slot'" shows "set nodes \ set nodes' = {}" proof (rule ccontr) assume "set nodes \ set nodes' \ {} " then obtain common_ptr where "common_ptr \ set nodes" and "common_ptr \ set nodes'" by auto have "h \ find_slot False common_ptr \\<^sub>r Some slot" using \common_ptr \ set nodes\ using assms(1) by(auto simp add: assigned_nodes_def elim!: bind_returns_result_E2 split: if_splits dest!: filter_M_holds_for_result[where x=common_ptr] intro!: bind_pure_I) moreover have "h \ find_slot False common_ptr \\<^sub>r Some slot'" using \common_ptr \ set nodes'\ using assms(2) by(auto simp add: assigned_nodes_def elim!: bind_returns_result_E2 split: if_splits dest!: filter_M_holds_for_result[where x=common_ptr] intro!: bind_pure_I) ultimately show False using assms(3) by (meson option.inject returns_result_eq) qed end interpretation i_assigned_nodes?: l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr assigned_nodes assigned_nodes_flatten flatten_dom get_child_nodes get_child_nodes_locs get_tag_name get_tag_name_locs get_root_node get_root_node_locs get_host get_host_locs find_slot assigned_slot remove insert_before insert_before_locs append_child remove_shadow_root remove_shadow_root_locs type_wf get_shadow_root get_shadow_root_locs set_shadow_root set_shadow_root_locs get_parent get_parent_locs to_tree_order by(auto simp add: instances l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def assigned_nodes_def flatten_dom_def) declare l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_assigned_nodes = l_assigned_nodes_defs + assumes assigned_nodes_pure [simp]: "pure (assigned_nodes slot) h" assumes assigned_nodes_ptr_in_heap: "h \ ok (assigned_nodes slot) \ slot |\| element_ptr_kinds h" assumes assigned_nodes_slot_is_slot: "h \ ok (assigned_nodes slot) \ h \ get_tag_name slot \\<^sub>r ''slot''" assumes assigned_nodes_different_ptr: "h \ assigned_nodes slot \\<^sub>r nodes \ h \ assigned_nodes slot' \\<^sub>r nodes' \ slot \ slot' \ set nodes \ set nodes' = {}" lemma assigned_nodes_is_l_assigned_nodes [instances]: "l_assigned_nodes assigned_nodes" apply(auto simp add: l_assigned_nodes_def)[1] using assigned_nodes_ptr_in_heap apply fast using assigned_nodes_slot_is_slot apply fast using assigned_nodes_different_ptr apply fast done subsubsection \set\_val\ locale l_set_val\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD: l_set_val\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_val set_val_locs + l_type_wf type_wf for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and set_val :: "(_) character_data_ptr \ char list \ (_, unit) dom_prog" and set_val_locs :: "(_) character_data_ptr \ (_, unit) dom_prog set" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" begin lemma set_val_ok: "type_wf h \ character_data_ptr |\| character_data_ptr_kinds h \ h \ ok (set_val character_data_ptr tag)" using CD.set_val_ok CD.type_wf_impl ShadowRootClass.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t local.type_wf_impl by blast lemma set_val_writes: "writes (set_val_locs character_data_ptr) (set_val character_data_ptr tag) h h'" using CD.set_val_writes . lemma set_val_pointers_preserved: assumes "w \ set_val_locs character_data_ptr" assumes "h \ w \\<^sub>h h'" shows "object_ptr_kinds h = object_ptr_kinds h'" using assms CD.set_val_pointers_preserved by simp lemma set_val_typess_preserved: assumes "w \ set_val_locs character_data_ptr" assumes "h \ w \\<^sub>h h'" shows "type_wf h = type_wf h'" apply(unfold type_wf_impl) using assms(1) type_wf_preserved[OF writes_singleton2 assms(2)] by(auto simp add: all_args_def CD.set_val_locs_impl[unfolded CD.a_set_val_locs_def] split: if_splits) end interpretation i_set_val?: l_set_val\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_val set_val_locs apply(unfold_locales) by (auto simp add: set_val_def set_val_locs_def) declare l_set_val\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_val_is_l_set_val [instances]: "l_set_val type_wf set_val set_val_locs" apply(simp add: l_set_val_def) using set_val_ok set_val_writes set_val_pointers_preserved set_val_typess_preserved by blast paragraph \get\_shadow\_root\ locale l_set_val_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_val\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_val_get_shadow_root: "\w \ set_val_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" by(auto simp add: CD.set_val_locs_impl[unfolded CD.a_set_val_locs_def] get_shadow_root_locs_def all_args_def) end locale l_set_val_get_shadow_root = l_set_val + l_get_shadow_root + assumes set_val_get_shadow_root: "\w \ set_val_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" interpretation i_set_val_get_shadow_root?: l_set_val_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_val set_val_locs get_shadow_root get_shadow_root_locs apply(auto simp add: l_set_val_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)[1] using l_set_val\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms by unfold_locales declare l_set_val_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_val_get_shadow_root_is_l_set_val_get_shadow_root [instances]: "l_set_val_get_shadow_root type_wf set_val set_val_locs get_shadow_root get_shadow_root_locs" using set_val_is_l_set_val get_shadow_root_is_l_get_shadow_root apply(simp add: l_set_val_get_shadow_root_def l_set_val_get_shadow_root_axioms_def) using set_val_get_shadow_root by fast paragraph \get\_tag\_type\ locale l_set_val_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_val\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_val_get_tag_name: "\w \ set_val_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_tag_name_locs ptr'. r h h'))" by(auto simp add: CD.set_val_locs_impl[unfolded CD.a_set_val_locs_def] CD.get_tag_name_locs_impl[unfolded CD.a_get_tag_name_locs_def] all_args_def) end locale l_set_val_get_tag_name = l_set_val + l_get_tag_name + assumes set_val_get_tag_name: "\w \ set_val_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_tag_name_locs ptr'. r h h'))" interpretation i_set_val_get_tag_name?: l_set_val_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_val set_val_locs get_tag_name get_tag_name_locs by unfold_locales declare l_set_val_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_val_get_tag_name_is_l_set_val_get_tag_name [instances]: "l_set_val_get_tag_name type_wf set_val set_val_locs get_tag_name get_tag_name_locs" using set_val_is_l_set_val get_tag_name_is_l_get_tag_name apply(simp add: l_set_val_get_tag_name_def l_set_val_get_tag_name_axioms_def) using set_val_get_tag_name by fast subsubsection \create\_character\_data\ locale l_create_character_data\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD: l_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M _ _ _ _ _ _ type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M _ known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_known_ptr known_ptr for known_ptr :: "(_) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ bool" + assumes known_ptr_impl: "known_ptr = a_known_ptr" begin lemma create_character_data_document_in_heap: assumes "h \ ok (create_character_data document_ptr text)" shows "document_ptr |\| document_ptr_kinds h" using assms CD.create_character_data_document_in_heap by simp lemma create_character_data_known_ptr: assumes "h \ create_character_data document_ptr text \\<^sub>r new_character_data_ptr" shows "known_ptr (cast new_character_data_ptr)" using assms CD.create_character_data_known_ptr by(simp add: known_ptr_impl CD.known_ptr_impl ShadowRootClass.a_known_ptr_def) end locale l_create_character_data = l_create_character_data_defs interpretation i_create_character_data?: l_create_character_data\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs set_val set_val_locs create_character_data known_ptr DocumentClass.type_wf DocumentClass.known_ptr by(auto simp add: l_create_character_data\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_create_character_data\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \create\_element\ locale l_create_element\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD: l_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs set_tag_name set_tag_name_locs type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M create_element known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_known_ptr known_ptr for get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" and set_tag_name :: "(_) element_ptr \ char list \ ((_) heap, exception, unit) prog" and set_tag_name_locs :: "(_) element_ptr \ ((_) heap, exception, unit) prog set" and type_wf :: "(_) heap \ bool" and create_element :: "(_) document_ptr \ char list \ ((_) heap, exception, (_) element_ptr) prog" and known_ptr :: "(_) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ bool" + assumes known_ptr_impl: "known_ptr = a_known_ptr" begin lemmas create_element_def = CD.create_element_def lemma create_element_document_in_heap: assumes "h \ ok (create_element document_ptr tag)" shows "document_ptr |\| document_ptr_kinds h" using CD.create_element_document_in_heap assms . lemma create_element_known_ptr: assumes "h \ create_element document_ptr tag \\<^sub>r new_element_ptr" shows "known_ptr (cast new_element_ptr)" proof - have "is_element_ptr new_element_ptr" using assms apply(auto simp add: create_element_def elim!: bind_returns_result_E)[1] using new_element_is_element_ptr by blast then show ?thesis by(auto simp add: known_ptr_impl known_ptr_defs DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs) qed end interpretation i_create_element?: l_create_element\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs set_tag_name set_tag_name_locs type_wf create_element known_ptr DocumentClass.type_wf DocumentClass.known_ptr by(auto simp add: l_create_element\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_create_element\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_create_element\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] subsection \A wellformed heap (Core DOM)\ subsubsection \wellformed\_heap\ locale l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = CD: l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs + l_get_shadow_root_defs get_shadow_root get_shadow_root_locs + l_get_tag_name_defs get_tag_name get_tag_name_locs for get_child_nodes :: "(_::linorder) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin definition a_host_shadow_root_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" where "a_host_shadow_root_rel h = (\(x, y). (cast x, cast y)) ` {(host, shadow_root). host |\| element_ptr_kinds h \ |h \ get_shadow_root host|\<^sub>r = Some shadow_root}" lemma a_host_shadow_root_rel_code [code]: "a_host_shadow_root_rel h = set (concat (map (\host. (case |h \ get_shadow_root host|\<^sub>r of Some shadow_root \ [(cast host, cast shadow_root)] | None \ [])) (sorted_list_of_fset (element_ptr_kinds h))) )" by(auto simp add: a_host_shadow_root_rel_def) definition a_all_ptrs_in_heap :: "(_) heap \ bool" where "a_all_ptrs_in_heap h = ((\host shadow_root_ptr. (h \ get_shadow_root host \\<^sub>r Some shadow_root_ptr) \ shadow_root_ptr |\| shadow_root_ptr_kinds h))" definition a_distinct_lists :: "(_) heap \ bool" where "a_distinct_lists h = distinct (concat ( map (\element_ptr. (case |h \ get_shadow_root element_ptr|\<^sub>r of Some shadow_root_ptr \ [shadow_root_ptr] | None \ [])) |h \ element_ptr_kinds_M|\<^sub>r ))" definition a_shadow_root_valid :: "(_) heap \ bool" where "a_shadow_root_valid h = (\shadow_root_ptr \ fset (shadow_root_ptr_kinds h). (\host \ fset(element_ptr_kinds h). |h \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types \ |h \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr))" definition a_heap_is_wellformed :: "(_) heap \ bool" where "a_heap_is_wellformed h \ CD.a_heap_is_wellformed h \ acyclic (CD.a_parent_child_rel h \ a_host_shadow_root_rel h) \ a_all_ptrs_in_heap h \ a_distinct_lists h \ a_shadow_root_valid h" end global_interpretation l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs defines heap_is_wellformed = a_heap_is_wellformed and parent_child_rel = CD.a_parent_child_rel and host_shadow_root_rel = a_host_shadow_root_rel and all_ptrs_in_heap = a_all_ptrs_in_heap and distinct_lists = a_distinct_lists and shadow_root_valid = a_shadow_root_valid and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD.a_heap_is_wellformed and parent_child_rel\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD.a_parent_child_rel and acyclic_heap\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD.a_acyclic_heap and all_ptrs_in_heap\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD.a_all_ptrs_in_heap and distinct_lists\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD.a_distinct_lists and owner_document_valid\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD.a_owner_document_valid . interpretation i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M: l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel by (auto simp add: l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def parent_child_rel_def instances) declare i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_is_l_heap_is_wellformed [instances]: "l_heap_is_wellformed type_wf known_ptr heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel get_child_nodes get_disconnected_nodes" apply(auto simp add: l_heap_is_wellformed_def)[1] using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_children_in_heap apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_disc_nodes_in_heap apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_one_parent apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_one_disc_parent apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_children_disc_nodes_different apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_disconnected_nodes_distinct apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_children_distinct apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_children_disc_nodes apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_child apply (blast, blast) using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_finite apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_acyclic apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_node_ptr apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_parent_in_heap apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_child_in_heap apply blast done locale l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs + CD: l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel + l_heap_is_wellformed_defs heap_is_wellformed parent_child_rel + l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_shadow_root get_shadow_root_locs get_host get_host_locs type_wf + l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs get_disconnected_document get_disconnected_document_locs type_wf + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs for get_child_nodes :: "(_::linorder) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and known_ptr :: "(_) object_ptr \ bool" and type_wf :: "(_) heap \ bool" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_disconnected_document :: "(_) node_ptr \ ((_) heap, exception, (_) document_ptr) prog" and get_disconnected_document_locs :: "((_) heap \ (_) heap \ bool) set" + assumes heap_is_wellformed_impl: "heap_is_wellformed = a_heap_is_wellformed" begin lemmas heap_is_wellformed_def = heap_is_wellformed_impl[unfolded a_heap_is_wellformed_def, folded CD.heap_is_wellformed_impl CD.parent_child_rel_impl] lemma a_distinct_lists_code [code]: "a_all_ptrs_in_heap h = ((\host \ fset (element_ptr_kinds h). h \ ok (get_shadow_root host) \ (case |h \ get_shadow_root host|\<^sub>r of Some shadow_root_ptr \ shadow_root_ptr |\| shadow_root_ptr_kinds h | None \ True)))" apply(auto simp add: a_all_ptrs_in_heap_def split: option.splits)[1] by (meson is_OK_returns_result_I local.get_shadow_root_ptr_in_heap notin_fset select_result_I2) lemma get_shadow_root_shadow_root_ptr_in_heap: assumes "heap_is_wellformed h" assumes "h \ get_shadow_root host \\<^sub>r Some shadow_root_ptr" shows "shadow_root_ptr |\| shadow_root_ptr_kinds h" using assms by(auto simp add: heap_is_wellformed_def a_all_ptrs_in_heap_def) lemma get_host_ptr_in_heap: assumes "heap_is_wellformed h" assumes "h \ get_host shadow_root_ptr \\<^sub>r host" shows "shadow_root_ptr |\| shadow_root_ptr_kinds h" using assms get_shadow_root_shadow_root_ptr_in_heap by(auto simp add: get_host_def elim!: bind_returns_result_E2 dest!: filter_M_holds_for_result intro!: bind_pure_I split: list.splits) lemma shadow_root_same_host: assumes "heap_is_wellformed h" and "type_wf h" assumes "h \ get_shadow_root host \\<^sub>r Some shadow_root_ptr" assumes "h \ get_shadow_root host' \\<^sub>r Some shadow_root_ptr" shows "host = host'" proof (rule ccontr) assume " host \ host'" have "host |\| element_ptr_kinds h" using assms(3) by (meson is_OK_returns_result_I local.get_shadow_root_ptr_in_heap) moreover have "host' |\| element_ptr_kinds h" using assms(4) by (meson is_OK_returns_result_I local.get_shadow_root_ptr_in_heap) ultimately show False using assms apply(auto simp add: heap_is_wellformed_def a_distinct_lists_def)[1] apply(drule distinct_concat_map_E(1)[where x=host and y=host']) apply(simp) apply(simp) using \host \ host'\ apply(simp) apply(auto)[1] done qed lemma shadow_root_host_dual: assumes "h \ get_host shadow_root_ptr \\<^sub>r host" shows "h \ get_shadow_root host \\<^sub>r Some shadow_root_ptr" using assms by(auto simp add: get_host_def dest: filter_M_holds_for_result elim!: bind_returns_result_E2 intro!: bind_pure_I split: list.splits) lemma disc_doc_disc_node_dual: assumes "h \ get_disconnected_document disc_node \\<^sub>r disc_doc" obtains disc_nodes where "h \ get_disconnected_nodes disc_doc \\<^sub>r disc_nodes" and "disc_node \ set disc_nodes" using assms get_disconnected_nodes_pure by(auto simp add: get_disconnected_document_def bind_pure_I dest!: filter_M_holds_for_result elim!: bind_returns_result_E2 intro!: filter_M_pure_I split: if_splits list.splits) lemma get_host_valid_tag_name: assumes "heap_is_wellformed h" and "type_wf h" assumes "h \ get_host shadow_root_ptr \\<^sub>r host" assumes "h \ get_tag_name host \\<^sub>r tag" shows "tag \ safe_shadow_root_element_types" proof - obtain host' where "host' |\| element_ptr_kinds h" and "|h \ get_tag_name host'|\<^sub>r \ safe_shadow_root_element_types" and "h \ get_shadow_root host' \\<^sub>r Some shadow_root_ptr" using assms apply(auto simp add: heap_is_wellformed_def a_shadow_root_valid_def)[1] by (smt assms(1) finite_set_in get_host_ptr_in_heap local.get_shadow_root_ok returns_result_select_result) then have "host = host'" by (meson assms(1) assms(2) assms(3) shadow_root_host_dual shadow_root_same_host) then show ?thesis by (smt\\thesis. (\host'. \host' |\| element_ptr_kinds h; |h \ get_tag_name host'|\<^sub>r \ safe_shadow_root_element_types; h \ get_shadow_root host' \\<^sub>r Some shadow_root_ptr\ \ thesis) \ thesis\ \h \ get_shadow_root host' \\<^sub>r Some shadow_root_ptr\ assms(1) assms(2) assms(4) select_result_I2 shadow_root_same_host) qed lemma a_host_shadow_root_rel_finite: "finite (a_host_shadow_root_rel h)" proof - have "a_host_shadow_root_rel h = (\host \ fset (element_ptr_kinds h). (case |h \ get_shadow_root host|\<^sub>r of Some shadow_root \ {(cast host, cast shadow_root)} | None \ {}))" by(auto simp add: a_host_shadow_root_rel_def split: option.splits) moreover have "finite (\host \ fset (element_ptr_kinds h). (case |h \ get_shadow_root host|\<^sub>r of Some shadow_root \ {(cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r host, cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root)} | None \ {}))" by(auto split: option.splits) ultimately show ?thesis by auto qed lemma heap_is_wellformed_children_in_heap: "heap_is_wellformed h \ h \ get_child_nodes ptr \\<^sub>r children \ child \ set children \ child |\| node_ptr_kinds h" using CD.heap_is_wellformed_children_in_heap local.heap_is_wellformed_def by blast lemma heap_is_wellformed_disc_nodes_in_heap: "heap_is_wellformed h \ h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes \ node \ set disc_nodes \ node |\| node_ptr_kinds h" using CD.heap_is_wellformed_disc_nodes_in_heap local.heap_is_wellformed_def by blast lemma heap_is_wellformed_one_parent: "heap_is_wellformed h \ h \ get_child_nodes ptr \\<^sub>r children \ h \ get_child_nodes ptr' \\<^sub>r children' \ set children \ set children' \ {} \ ptr = ptr'" using CD.heap_is_wellformed_one_parent local.heap_is_wellformed_def by blast lemma heap_is_wellformed_one_disc_parent: "heap_is_wellformed h \ h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes \ h \ get_disconnected_nodes document_ptr' \\<^sub>r disc_nodes' \ set disc_nodes \ set disc_nodes' \ {} \ document_ptr = document_ptr'" using CD.heap_is_wellformed_one_disc_parent local.heap_is_wellformed_def by blast lemma heap_is_wellformed_children_disc_nodes_different: "heap_is_wellformed h \ h \ get_child_nodes ptr \\<^sub>r children \ h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes \ set children \ set disc_nodes = {}" using CD.heap_is_wellformed_children_disc_nodes_different local.heap_is_wellformed_def by blast lemma heap_is_wellformed_disconnected_nodes_distinct: "heap_is_wellformed h \ h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes \ distinct disc_nodes" using CD.heap_is_wellformed_disconnected_nodes_distinct local.heap_is_wellformed_def by blast lemma heap_is_wellformed_children_distinct: "heap_is_wellformed h \ h \ get_child_nodes ptr \\<^sub>r children \ distinct children" using CD.heap_is_wellformed_children_distinct local.heap_is_wellformed_def by blast lemma heap_is_wellformed_children_disc_nodes: "heap_is_wellformed h \ node_ptr |\| node_ptr_kinds h \ \(\parent \ fset (object_ptr_kinds h). node_ptr \ set |h \ get_child_nodes parent|\<^sub>r) \ (\document_ptr \ fset (document_ptr_kinds h). node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r)" using CD.heap_is_wellformed_children_disc_nodes local.heap_is_wellformed_def by blast lemma parent_child_rel_finite: "heap_is_wellformed h \ finite (parent_child_rel h)" using CD.parent_child_rel_finite by blast lemma parent_child_rel_acyclic: "heap_is_wellformed h \ acyclic (parent_child_rel h)" using CD.parent_child_rel_acyclic heap_is_wellformed_def by blast lemma parent_child_rel_child_in_heap: "heap_is_wellformed h \ type_wf h \ known_ptr parent \ (parent, child_ptr) \ parent_child_rel h \ child_ptr |\| object_ptr_kinds h" using CD.parent_child_rel_child_in_heap local.heap_is_wellformed_def by blast end interpretation i_heap_is_wellformed?: l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs by(auto simp add: l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def parent_child_rel_def heap_is_wellformed_def instances) declare l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma heap_is_wellformed_is_l_heap_is_wellformed [instances]: "l_heap_is_wellformed ShadowRootClass.type_wf ShadowRootClass.known_ptr Shadow_DOM.heap_is_wellformed Shadow_DOM.parent_child_rel Shadow_DOM.get_child_nodes get_disconnected_nodes" apply(auto simp add: l_heap_is_wellformed_def instances)[1] using heap_is_wellformed_children_in_heap apply metis using heap_is_wellformed_disc_nodes_in_heap apply metis using heap_is_wellformed_one_parent apply blast using heap_is_wellformed_one_disc_parent apply blast using heap_is_wellformed_children_disc_nodes_different apply blast using heap_is_wellformed_disconnected_nodes_distinct apply metis using heap_is_wellformed_children_distinct apply metis using heap_is_wellformed_children_disc_nodes apply metis using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_child apply(blast, blast) using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_finite apply blast using parent_child_rel_acyclic apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_node_ptr apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_parent_in_heap apply blast using parent_child_rel_child_in_heap apply metis done subsubsection \get\_parent\ interpretation i_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M: l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs get_parent get_parent_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel get_disconnected_nodes by(simp add: l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare i_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] interpretation i_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M: l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs get_parent get_parent_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel get_disconnected_nodes get_disconnected_nodes_locs by(auto simp add: l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare i_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_is_l_get_parent_wf [instances]: "l_get_parent_wf type_wf known_ptr known_ptrs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel get_child_nodes get_parent" apply(auto simp add: l_get_parent_wf_def l_get_parent_wf_axioms_def instances)[1] using i_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.child_parent_dual apply fast using i_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_wellformed_induct apply metis using i_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_wellformed_induct_rev apply metis using i_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_parent apply fast done subsubsection \get\_disconnected\_nodes\ subsubsection \set\_disconnected\_nodes\ paragraph \get\_disconnected\_nodes\ interpretation i_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M: l_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel get_child_nodes by (simp add: l_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare i_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.l_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_is_l_set_disconnected_nodes_get_disconnected_nodes_wf [instances]: "l_set_disconnected_nodes_get_disconnected_nodes_wf type_wf known_ptr heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel get_child_nodes get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs" apply(auto simp add: l_set_disconnected_nodes_get_disconnected_nodes_wf_def l_set_disconnected_nodes_get_disconnected_nodes_wf_axioms_def instances)[1] using i_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.remove_from_disconnected_nodes_removes apply fast done paragraph \get\_root\_node\ interpretation i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M:l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf known_ptrs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_parent get_parent_locs get_ancestors get_ancestors_locs get_root_node get_root_node_locs by(simp add: l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma get_ancestors_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_is_l_get_ancestors_wf [instances]: "l_get_ancestors_wf heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel known_ptr known_ptrs type_wf get_ancestors get_ancestors_locs get_child_nodes get_parent" apply(auto simp add: l_get_ancestors_wf_def l_get_ancestors_wf_axioms_def instances)[1] using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_never_empty apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_ok apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_reads apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_ptrs_in_heap apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_remains_not_in_ancestors apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_also_parent apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_obtains_children apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_parent_child_rel apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_parent_child_rel apply blast done lemma get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_is_l_get_root_node_wf [instances]: "l_get_root_node_wf heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_root_node type_wf known_ptr known_ptrs get_ancestors get_parent" apply(auto simp add: l_get_root_node_wf_def l_get_root_node_wf_axioms_def instances)[1] using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_root_node_ok apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_root_node_ptr_in_heap apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_root_node_root_in_heap apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_same_root_node apply(blast, blast) using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_root_node_same_no_parent apply blast (* using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_root_node_not_node_same apply blast *) using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_root_node_parent_same apply (blast, blast) done subsubsection \to\_tree\_order\ interpretation i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M: l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs to_tree_order known_ptrs get_parent get_parent_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel get_disconnected_nodes get_disconnected_nodes_locs apply(simp add: l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) done declare i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_is_l_to_tree_order_wf [instances]: "l_to_tree_order_wf heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent get_child_nodes" apply(auto simp add: l_to_tree_order_wf_def l_to_tree_order_wf_axioms_def instances)[1] using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_ok apply blast using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_ptrs_in_heap apply blast using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_parent_child_rel apply(blast, blast) using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_child2 apply blast using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_node_ptrs apply blast using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_child apply blast using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_ptr_in_result apply blast using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_parent apply blast using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_subset apply blast done paragraph \get\_root\_node\ interpretation i_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M: l_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf known_ptrs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_parent get_parent_locs get_ancestors get_ancestors_locs get_root_node get_root_node_locs to_tree_order by(auto simp add: l_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare i_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.l_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_is_l_to_tree_order_wf_get_root_node_wf [instances]: "l_to_tree_order_wf_get_root_node_wf type_wf known_ptr known_ptrs to_tree_order get_root_node heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M" apply(auto simp add: l_to_tree_order_wf_get_root_node_wf_def l_to_tree_order_wf_get_root_node_wf_axioms_def instances)[1] using i_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_get_root_node apply blast using i_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_same_root apply blast done subsubsection \remove\_child\ interpretation i_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M: l_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs set_child_nodes set_child_nodes_locs get_parent get_parent_locs get_owner_document get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs remove_child remove_child_locs remove type_wf known_ptr known_ptrs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel by unfold_locales declare i_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.l_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_is_l_remove_child_wf2 [instances]: "l_remove_child_wf2 type_wf known_ptr known_ptrs remove_child heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes remove" apply(auto simp add: l_remove_child_wf2_def l_remove_child_wf2_axioms_def instances)[1] using i_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.remove_child_heap_is_wellformed_preserved apply(fast, fast, fast) using i_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.remove_heap_is_wellformed_preserved apply(fast, fast, fast) using i_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.remove_child_removes_child apply fast using i_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.remove_child_removes_first_child apply fast using i_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.remove_removes_child apply fast using i_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.remove_for_all_empty_children apply fast done subsection \A wellformed heap\ subsubsection \get\_parent\ interpretation i_get_parent_wf?: l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs get_parent get_parent_locs heap_is_wellformed parent_child_rel get_disconnected_nodes using instances by(simp add: l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma get_parent_wf_is_l_get_parent_wf [instances]: "l_get_parent_wf ShadowRootClass.type_wf ShadowRootClass.known_ptr ShadowRootClass.known_ptrs heap_is_wellformed parent_child_rel Shadow_DOM.get_child_nodes Shadow_DOM.get_parent" apply(auto simp add: l_get_parent_wf_def l_get_parent_wf_axioms_def instances)[1] using child_parent_dual apply blast using heap_wellformed_induct apply metis using heap_wellformed_induct_rev apply metis using parent_child_rel_parent apply metis done subsubsection \remove\_shadow\_root\ locale l_remove_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_tag_name + l_get_disconnected_nodes + l_set_shadow_root_get_tag_name + l_get_child_nodes + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_delete_shadow_root_get_disconnected_nodes + l_delete_shadow_root_get_child_nodes + l_set_shadow_root_get_disconnected_nodes + l_set_shadow_root_get_child_nodes + l_delete_shadow_root_get_tag_name + l_set_shadow_root_get_shadow_root + l_delete_shadow_root_get_shadow_root + l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma remove_shadow_root_preserves: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ remove_shadow_root ptr \\<^sub>h h'" shows "known_ptrs h'" and "type_wf h'" "heap_is_wellformed h'" proof - obtain shadow_root_ptr h2 where "h \ get_shadow_root ptr \\<^sub>r Some shadow_root_ptr" and "h \ get_child_nodes (cast shadow_root_ptr) \\<^sub>r []" and h2: "h \ set_shadow_root ptr None \\<^sub>h h2" and h': "h2 \ delete_M shadow_root_ptr \\<^sub>h h'" using assms(4) by(auto simp add: remove_shadow_root_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_shadow_root_pure, rotated] bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated] split: option.splits if_splits) have "type_wf h2" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_shadow_root_writes h2] using \type_wf h\ set_shadow_root_types_preserved by(auto simp add: reflp_def transp_def) then show "type_wf h'" using h' delete_shadow_root_type_wf_preserved local.type_wf_impl by blast have object_ptr_kinds_eq_h: "object_ptr_kinds h = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_shadow_root_writes h2]) using set_shadow_root_pointers_preserved apply blast by (auto simp add: reflp_def transp_def) have node_ptr_kinds_eq_h: "node_ptr_kinds h = node_ptr_kinds h2" using object_ptr_kinds_eq_h by (simp add: node_ptr_kinds_def) have element_ptr_kinds_eq_h: "element_ptr_kinds h = element_ptr_kinds h2" using node_ptr_kinds_eq_h by (simp add: element_ptr_kinds_def) have document_ptr_kinds_eq_h: "document_ptr_kinds h = document_ptr_kinds h2" using object_ptr_kinds_eq_h by (simp add: document_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h: "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h2" using object_ptr_kinds_eq_h by (simp add: shadow_root_ptr_kinds_def) have "known_ptrs h2" using \known_ptrs h\ object_ptr_kinds_eq_h known_ptrs_subset by blast have object_ptr_kinds_eq_h2: "object_ptr_kinds h' |\| object_ptr_kinds h2" using h' delete_shadow_root_pointers by auto have object_ptr_kinds_eq2_h2: "object_ptr_kinds h2 = object_ptr_kinds h' |\| {|cast shadow_root_ptr|}" using h' delete_shadow_root_pointers by auto have node_ptr_kinds_eq_h2: "node_ptr_kinds h2 = node_ptr_kinds h'" using object_ptr_kinds_eq_h2 by(auto simp add: node_ptr_kinds_def delete_shadow_root_pointers[OF h']) have element_ptr_kinds_eq_h2: "element_ptr_kinds h2 = element_ptr_kinds h'" using node_ptr_kinds_eq_h2 by (simp add: element_ptr_kinds_def) have document_ptr_kinds_eq_h2: "document_ptr_kinds h2 = document_ptr_kinds h'" using object_ptr_kinds_eq_h2 by(auto simp add: document_ptr_kinds_def delete_shadow_root_pointers[OF h']) have shadow_root_ptr_kinds_eq_h2: "shadow_root_ptr_kinds h' |\| shadow_root_ptr_kinds h2" using object_ptr_kinds_eq_h2 by (auto simp add: shadow_root_ptr_kinds_def) have shadow_root_ptr_kinds_eq2_h2: "shadow_root_ptr_kinds h2 = shadow_root_ptr_kinds h' |\| {|shadow_root_ptr|}" using object_ptr_kinds_eq2_h2 apply (auto simp add: shadow_root_ptr_kinds_def)[1] by (metis \h \ get_shadow_root ptr \\<^sub>r Some shadow_root_ptr\ assms(1) fset.map_comp local.get_shadow_root_shadow_root_ptr_in_heap object_ptr_kinds_eq_h shadow_root_ptr_kinds_def) show "known_ptrs h'" using object_ptr_kinds_eq_h2 \known_ptrs h2\ known_ptrs_subset by blast have disconnected_nodes_eq_h: "\doc_ptr disc_nodes. h \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_shadow_root_writes h2 set_shadow_root_get_disconnected_nodes by(rule reads_writes_preserved) then have disconnected_nodes_eq2_h: "\doc_ptr. |h \ get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads get_disconnected_nodes_delete_shadow_root[OF h'] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have disconnected_nodes_eq2_h2: "\doc_ptr. |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have tag_name_eq_h: "\doc_ptr disc_nodes. h \ get_tag_name doc_ptr \\<^sub>r disc_nodes = h2 \ get_tag_name doc_ptr \\<^sub>r disc_nodes" using get_tag_name_reads set_shadow_root_writes h2 set_shadow_root_get_tag_name by(rule reads_writes_preserved) then have tag_name_eq2_h: "\doc_ptr. |h \ get_tag_name doc_ptr|\<^sub>r = |h2 \ get_tag_name doc_ptr|\<^sub>r" using select_result_eq by force have tag_name_eq_h2: "\doc_ptr disc_nodes. h2 \ get_tag_name doc_ptr \\<^sub>r disc_nodes = h' \ get_tag_name doc_ptr \\<^sub>r disc_nodes" using get_tag_name_reads get_tag_name_delete_shadow_root[OF h'] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have tag_name_eq2_h2: "\doc_ptr. |h2 \ get_tag_name doc_ptr|\<^sub>r = |h' \ get_tag_name doc_ptr|\<^sub>r" using select_result_eq by force have children_eq_h: "\ptr' children. h \ get_child_nodes ptr' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads set_shadow_root_writes h2 set_shadow_root_get_child_nodes by(rule reads_writes_preserved) then have children_eq2_h: "\ptr'. |h \ get_child_nodes ptr'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. ptr' \ cast shadow_root_ptr \ h2 \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads h' get_child_nodes_delete_shadow_root apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have children_eq2_h2: "\ptr'. ptr' \ cast shadow_root_ptr \ |h2 \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have "cast shadow_root_ptr |\| object_ptr_kinds h'" using h' delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ptr_not_in_heap by auto have get_shadow_root_eq_h: "\shadow_root_opt ptr'. ptr \ ptr' \ h \ get_shadow_root ptr' \\<^sub>r shadow_root_opt = h2 \ get_shadow_root ptr' \\<^sub>r shadow_root_opt" using get_shadow_root_reads set_shadow_root_writes h2 apply(rule reads_writes_preserved) using set_shadow_root_get_shadow_root_different_pointers by fast have get_shadow_root_eq_h2: "\shadow_root_opt ptr'. h2 \ get_shadow_root ptr' \\<^sub>r shadow_root_opt = h' \ get_shadow_root ptr' \\<^sub>r shadow_root_opt" using get_shadow_root_reads get_shadow_root_delete_shadow_root[OF h'] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have get_shadow_root_eq2_h2: "\ptr'. |h2 \ get_shadow_root ptr'|\<^sub>r = |h' \ get_shadow_root ptr'|\<^sub>r" using select_result_eq by force have "acyclic (parent_child_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def) moreover have "parent_child_rel h = parent_child_rel h2" by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq_h children_eq2_h) moreover have "parent_child_rel h' \ parent_child_rel h2" using object_ptr_kinds_eq_h2 apply(auto simp add: CD.parent_child_rel_def)[1] by (metis \cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr |\| object_ptr_kinds h'\ children_eq2_h2) ultimately have "CD.a_acyclic_heap h'" using acyclic_subset by (auto simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def) moreover have "CD.a_all_ptrs_in_heap h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_all_ptrs_in_heap h2" by(auto simp add: children_eq2_h disconnected_nodes_eq2_h document_ptr_kinds_eq_h CD.a_all_ptrs_in_heap_def object_ptr_kinds_eq_h node_ptr_kinds_def children_eq_h disconnected_nodes_eq_h) then have "CD.a_all_ptrs_in_heap h'" apply(auto simp add: CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 children_eq_h2 disconnected_nodes_eq_h2)[1] apply(case_tac "ptr = cast shadow_root_ptr") using object_ptr_kinds_eq_h2 children_eq_h2 apply (meson \cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr |\| object_ptr_kinds h'\ is_OK_returns_result_I local.get_child_nodes_ptr_in_heap) apply (metis (no_types, lifting) children_eq2_h2 fin_mono finite_set_in object_ptr_kinds_eq_h2 subsetD) by (metis (full_types) assms(1) assms(2) disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 local.get_disconnected_nodes_ok local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq_h node_ptr_kinds_eq_h2 returns_result_select_result) moreover have "CD.a_distinct_lists h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_distinct_lists h2" by(auto simp add: CD.a_distinct_lists_def object_ptr_kinds_eq_h document_ptr_kinds_eq_h children_eq2_h disconnected_nodes_eq2_h) then have "CD.a_distinct_lists h'" apply(auto simp add: CD.a_distinct_lists_def document_ptr_kinds_eq_h2 disconnected_nodes_eq2_h2)[1] apply(auto simp add: intro!: distinct_concat_map_I)[1] apply(case_tac "x = cast shadow_root_ptr") using \cast shadow_root_ptr |\| object_ptr_kinds h'\ apply simp using children_eq_h2 concat_map_all_distinct[of "(\ptr. |h2 \ get_child_nodes ptr|\<^sub>r)"] apply (metis (no_types, lifting) children_eq2_h2 finite_fset fmember.rep_eq fset_mp object_ptr_kinds_eq_h2 set_sorted_list_of_set) apply(case_tac "x = cast shadow_root_ptr") using \cast shadow_root_ptr |\| object_ptr_kinds h'\ apply simp apply(case_tac "y = cast shadow_root_ptr") using \cast shadow_root_ptr |\| object_ptr_kinds h'\ apply simp using children_eq_h2 distinct_concat_map_E(1)[of "(\ptr. |h2 \ get_child_nodes ptr|\<^sub>r)"] apply (smt IntI children_eq2_h2 empty_iff finite_fset fmember.rep_eq fset_mp object_ptr_kinds_eq_h2 set_sorted_list_of_set) apply(case_tac "xa = cast shadow_root_ptr") using \cast shadow_root_ptr |\| object_ptr_kinds h'\ apply simp by (smt \local.CD.a_distinct_lists h2\ \type_wf h'\ children_eq2_h2 disconnected_nodes_eq_h2 fset_mp is_OK_returns_result_E local.CD.distinct_lists_no_parent local.get_disconnected_nodes_ok object_ptr_kinds_eq_h2 select_result_I2) moreover have "CD.a_owner_document_valid h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_owner_document_valid h2" by(auto simp add: CD.a_owner_document_valid_def object_ptr_kinds_eq_h document_ptr_kinds_eq_h node_ptr_kinds_eq_h children_eq2_h disconnected_nodes_eq2_h) then have "CD.a_owner_document_valid h'" apply(auto simp add: CD.a_owner_document_valid_def document_ptr_kinds_eq_h2 node_ptr_kinds_eq_h2 disconnected_nodes_eq2_h2)[1] proof - fix node_ptr assume 0: "\node_ptr\fset (node_ptr_kinds h'). (\document_ptr. document_ptr |\| document_ptr_kinds h' \ node_ptr \ set |h' \ get_disconnected_nodes document_ptr|\<^sub>r) \ (\parent_ptr. parent_ptr |\| object_ptr_kinds h2 \ node_ptr \ set |h2 \ get_child_nodes parent_ptr|\<^sub>r)" and 1: "node_ptr |\| node_ptr_kinds h'" and 2: "\parent_ptr. parent_ptr |\| object_ptr_kinds h' \ node_ptr \ set |h' \ get_child_nodes parent_ptr|\<^sub>r" then have "\parent_ptr. parent_ptr |\| object_ptr_kinds h2 \ node_ptr \ set |h2 \ get_child_nodes parent_ptr|\<^sub>r" apply(auto)[1] apply(case_tac "parent_ptr = cast shadow_root_ptr") using \h \ get_child_nodes (cast shadow_root_ptr) \\<^sub>r []\ children_eq_h apply auto[1] using children_eq2_h2 object_ptr_kinds_eq_h2 h' delete_shadow_root_pointers by (metis fempty_iff finsert_iff funionE) then show "\document_ptr. document_ptr |\| document_ptr_kinds h' \ node_ptr \ set |h' \ get_disconnected_nodes document_ptr|\<^sub>r" using 0 1 by auto qed ultimately have "heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'" by(simp add: CD.heap_is_wellformed_def) moreover have "acyclic (parent_child_rel h \ a_host_shadow_root_rel h)" using \heap_is_wellformed h\ by(simp add: heap_is_wellformed_def) then have "acyclic (parent_child_rel h2 \ a_host_shadow_root_rel h2)" proof - have "a_host_shadow_root_rel h2 \ a_host_shadow_root_rel h" apply(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h)[1] apply(case_tac "aa = ptr") apply(simp) apply (metis (no_types, lifting) \type_wf h2\ assms(2) h2 local.get_shadow_root_ok local.type_wf_impl option.distinct(1) returns_result_eq returns_result_select_result set_shadow_root_get_shadow_root) using get_shadow_root_eq_h by (metis (mono_tags, lifting) \type_wf h2\ image_eqI is_OK_returns_result_E local.get_shadow_root_ok mem_Collect_eq prod.simps(2) select_result_I2) then show ?thesis using \parent_child_rel h = parent_child_rel h2\ by (metis (no_types, hide_lams) \acyclic (parent_child_rel h \ a_host_shadow_root_rel h)\ acyclic_subset order_refl sup_mono) qed then have "acyclic (parent_child_rel h' \ a_host_shadow_root_rel h')" proof - have "a_host_shadow_root_rel h' \ a_host_shadow_root_rel h2" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h2 get_shadow_root_eq2_h2) then show ?thesis using \parent_child_rel h' \ parent_child_rel h2\ \acyclic (parent_child_rel h2 \ a_host_shadow_root_rel h2)\ using acyclic_subset sup_mono by (metis (no_types, hide_lams)) qed moreover have "a_all_ptrs_in_heap h" using \heap_is_wellformed h\ by(simp add: heap_is_wellformed_def) then have "a_all_ptrs_in_heap h2" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h)[1] apply(case_tac "host = ptr") apply(simp) apply (metis assms(2) h2 local.type_wf_impl option.distinct(1) returns_result_eq set_shadow_root_get_shadow_root) using get_shadow_root_eq_h by fastforce then have "a_all_ptrs_in_heap h'" apply(auto simp add: a_all_ptrs_in_heap_def get_shadow_root_eq_h2)[1] apply(auto simp add: shadow_root_ptr_kinds_eq2_h2)[1] by (metis (no_types, lifting) \h \ get_shadow_root ptr \\<^sub>r Some shadow_root_ptr\ assms(1) assms(2) get_shadow_root_eq_h get_shadow_root_eq_h2 h2 local.shadow_root_same_host local.type_wf_impl option.distinct(1) select_result_I2 set_shadow_root_get_shadow_root) moreover have "a_distinct_lists h" using \heap_is_wellformed h\ by(simp add: heap_is_wellformed_def) then have "a_distinct_lists h2" apply(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h)[1] apply(auto intro!: distinct_concat_map_I split: option.splits)[1] apply(case_tac "x = ptr") apply(simp) apply (metis (no_types, hide_lams) assms(2) h2 is_OK_returns_result_I l_set_shadow_root_get_shadow_root.set_shadow_root_get_shadow_root l_set_shadow_root_get_shadow_root_axioms local.type_wf_impl option.discI returns_result_eq returns_result_select_result) apply(case_tac "y = ptr") apply(simp) apply (metis (no_types, hide_lams) assms(2) h2 is_OK_returns_result_I l_set_shadow_root_get_shadow_root.set_shadow_root_get_shadow_root l_set_shadow_root_get_shadow_root_axioms local.type_wf_impl option.discI returns_result_eq returns_result_select_result) by (metis \type_wf h2\ assms(1) assms(2) get_shadow_root_eq_h local.get_shadow_root_ok local.shadow_root_same_host returns_result_select_result) then have "a_distinct_lists h'" by(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h2 get_shadow_root_eq2_h2) moreover have "a_shadow_root_valid h" using \heap_is_wellformed h\ by(simp add: heap_is_wellformed_def) then have "a_shadow_root_valid h'" apply(auto simp add: a_shadow_root_valid_def shadow_root_ptr_kinds_eq_h element_ptr_kinds_eq_h tag_name_eq2_h)[1] apply(simp add: shadow_root_ptr_kinds_eq2_h2 element_ptr_kinds_eq_h2 tag_name_eq2_h2) using get_shadow_root_eq_h get_shadow_root_eq_h2 by (smt \cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr |\| object_ptr_kinds h'\ \h \ get_shadow_root ptr \\<^sub>r Some shadow_root_ptr\ assms(2) element_ptr_kinds_eq_h element_ptr_kinds_eq_h2 finite_set_in local.get_shadow_root_ok option.inject returns_result_select_result select_result_I2 shadow_root_ptr_kinds_commutes) ultimately show "heap_is_wellformed h'" by(simp add: heap_is_wellformed_def) qed end interpretation i_remove_shadow_root_wf?: l_remove_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_tag_name get_tag_name_locs get_disconnected_nodes get_disconnected_nodes_locs set_shadow_root set_shadow_root_locs known_ptr get_child_nodes get_child_nodes_locs get_shadow_root get_shadow_root_locs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs remove_shadow_root remove_shadow_root_locs known_ptrs get_parent get_parent_locs by(auto simp add: l_remove_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_remove_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_root\_node\ interpretation i_get_root_node_wf?: l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf known_ptrs heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_parent get_parent_locs get_ancestors get_ancestors_locs get_root_node get_root_node_locs by(simp add: l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma get_ancestors_wf_is_l_get_ancestors_wf [instances]: "l_get_ancestors_wf heap_is_wellformed parent_child_rel known_ptr known_ptrs type_wf get_ancestors get_ancestors_locs get_child_nodes get_parent" apply(auto simp add: l_get_ancestors_wf_def l_get_ancestors_wf_axioms_def instances)[1] using get_ancestors_never_empty apply blast using get_ancestors_ok apply blast using get_ancestors_reads apply blast using get_ancestors_ptrs_in_heap apply blast using get_ancestors_remains_not_in_ancestors apply blast using get_ancestors_also_parent apply blast using get_ancestors_obtains_children apply blast using get_ancestors_parent_child_rel apply blast using get_ancestors_parent_child_rel apply blast done lemma get_root_node_wf_is_l_get_root_node_wf [instances]: "l_get_root_node_wf heap_is_wellformed get_root_node type_wf known_ptr known_ptrs get_ancestors get_parent" using known_ptrs_is_l_known_ptrs apply(auto simp add: l_get_root_node_wf_def l_get_root_node_wf_axioms_def)[1] using get_root_node_ok apply blast using get_root_node_ptr_in_heap apply blast using get_root_node_root_in_heap apply blast using get_ancestors_same_root_node apply(blast, blast) using get_root_node_same_no_parent apply blast (* using get_root_node_not_node_same apply blast *) using get_root_node_parent_same apply (blast, blast) done subsubsection \get\_parent\_get\_host\ locale l_get_parent_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent_wf + l_get_shadow_root + l_get_host + l_get_child_nodes begin lemma host_shadow_root_rel_finite: "finite (a_host_shadow_root_rel h)" proof - have "a_host_shadow_root_rel h = (\host \ fset (element_ptr_kinds h). (case |h \ get_shadow_root host|\<^sub>r of Some shadow_root \ {(cast host, cast shadow_root)} | None \ {}))" by(auto simp add: a_host_shadow_root_rel_def split: option.splits) moreover have "finite (\host \ fset (element_ptr_kinds h). (case |h \ get_shadow_root host|\<^sub>r of Some shadow_root \ {(cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r host, cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root)} | None \ {}))" by(auto split: option.splits) ultimately show ?thesis by auto qed lemma host_shadow_root_rel_shadow_root: "h \ get_shadow_root host \\<^sub>r shadow_root_option \ shadow_root_option = Some shadow_root \ ((cast host, cast shadow_root) \ a_host_shadow_root_rel h)" apply(auto simp add: a_host_shadow_root_rel_def)[1] by(metis (mono_tags, lifting) case_prodI is_OK_returns_result_I l_get_shadow_root.get_shadow_root_ptr_in_heap local.l_get_shadow_root_axioms mem_Collect_eq pair_imageI select_result_I2) lemma host_shadow_root_rel_host: "heap_is_wellformed h \ h \ get_host shadow_root \\<^sub>r host \ (cast host, cast shadow_root) \ a_host_shadow_root_rel h" apply(auto simp add: a_host_shadow_root_rel_def)[1] using shadow_root_host_dual by (metis (no_types, lifting) Collect_cong host_shadow_root_rel_shadow_root local.a_host_shadow_root_rel_def split_cong) lemma heap_wellformed_induct_si [consumes 1, case_names step]: assumes "heap_is_wellformed h" assumes "\parent. (\children child. h \ get_child_nodes parent \\<^sub>r children \ child \ set children \ P (cast child)) \ (\shadow_root host. parent = cast host \ h \ get_shadow_root host \\<^sub>r Some shadow_root \ P (cast shadow_root)) \ P parent" shows "P ptr" proof - fix ptr have "finite (parent_child_rel h \ a_host_shadow_root_rel h)" using host_shadow_root_rel_finite using local.CD.parent_child_rel_finite local.CD.parent_child_rel_impl by auto then have "wf ((parent_child_rel h \ a_host_shadow_root_rel h)\)" using assms(1) apply(simp add: heap_is_wellformed_def) by (simp add: finite_acyclic_wf_converse local.CD.parent_child_rel_impl) then show "?thesis" proof (induct rule: wf_induct_rule) case (less parent) then show ?case apply(auto)[1] using assms host_shadow_root_rel_shadow_root local.CD.parent_child_rel_child by blast qed qed lemma heap_wellformed_induct_rev_si [consumes 1, case_names step]: assumes "heap_is_wellformed h" assumes "\child. (\parent child_node. child = cast child_node \ h \ get_parent child_node \\<^sub>r Some parent \ P parent) \ (\host shadow_root. child = cast shadow_root \ h \ get_host shadow_root \\<^sub>r host \ P (cast host)) \ P child" shows "P ptr" proof - fix ptr have "finite (parent_child_rel h \ a_host_shadow_root_rel h)" using host_shadow_root_rel_finite using local.CD.parent_child_rel_finite local.CD.parent_child_rel_impl by auto then have "wf (parent_child_rel h \ a_host_shadow_root_rel h)" using assms(1) apply(simp add: heap_is_wellformed_def) by (simp add: finite_acyclic_wf) then show "?thesis" proof (induct rule: wf_induct_rule) case (less parent) then show ?case apply(auto)[1] using parent_child_rel_parent host_shadow_root_rel_host using assms(1) assms(2) by auto qed qed end interpretation i_get_parent_get_host_wf?: l_get_parent_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs known_ptrs get_parent get_parent_locs by(auto simp add: l_get_parent_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_parent_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_get_parent_get_host_wf = l_heap_is_wellformed_defs + l_get_parent_defs + l_get_shadow_root_defs + l_get_host_defs + l_get_child_nodes_defs + assumes heap_wellformed_induct_si [consumes 1, case_names step]: "heap_is_wellformed h \ (\parent. (\children child. h \ get_child_nodes parent \\<^sub>r children \ child \ set children \ P (cast child)) \ (\shadow_root host. parent = cast host \ h \ get_shadow_root host \\<^sub>r Some shadow_root \ P (cast shadow_root)) \ P parent) \ P ptr" assumes heap_wellformed_induct_rev_si [consumes 1, case_names step]: "heap_is_wellformed h \ (\child. (\parent child_node. child = cast child_node \ h \ get_parent child_node \\<^sub>r Some parent \ P parent) \ (\host shadow_root. child = cast shadow_root \ h \ get_host shadow_root \\<^sub>r host \ P (cast host)) \ P child) \ P ptr" lemma l_get_parent_get_host_wf_is_get_parent_get_host_wf [instances]: "l_get_parent_get_host_wf heap_is_wellformed get_parent get_shadow_root get_host get_child_nodes" apply(auto simp add: l_get_parent_get_host_wf_def instances)[1] using heap_wellformed_induct_si apply metis using heap_wellformed_induct_rev_si apply blast done subsubsection \get\_host\ locale l_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs + l_type_wf type_wf + l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_shadow_root get_shadow_root_locs get_host get_host_locs type_wf + l_get_shadow_root type_wf get_shadow_root get_shadow_root_locs for known_ptr :: "(_::linorder) object_ptr \ bool" and known_ptrs :: "(_) heap \ bool" and type_wf :: "(_) heap \ bool" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_child_nodes :: "(_::linorder) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" begin lemma get_host_ok [simp]: assumes "heap_is_wellformed h" assumes "type_wf h" assumes "known_ptrs h" assumes "shadow_root_ptr |\| shadow_root_ptr_kinds h" shows "h \ ok (get_host shadow_root_ptr)" proof - obtain host where host: "host |\| element_ptr_kinds h" and "|h \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types" and shadow_root: "h \ get_shadow_root host \\<^sub>r Some shadow_root_ptr" using assms(1) assms(4) get_shadow_root_ok assms(2) apply(auto simp add: heap_is_wellformed_def a_shadow_root_valid_def)[1] by (smt finite_set_in returns_result_select_result) obtain host_candidates where host_candidates: "h \ filter_M (\element_ptr. Heap_Error_Monad.bind (get_shadow_root element_ptr) (\shadow_root_opt. return (shadow_root_opt = Some shadow_root_ptr))) (sorted_list_of_set (fset (element_ptr_kinds h))) \\<^sub>r host_candidates" apply(drule is_OK_returns_result_E[rotated]) using get_shadow_root_ok assms(2) by(auto intro!: filter_M_is_OK_I bind_pure_I bind_is_OK_I2) then have "host_candidates = [host]" apply(rule filter_M_ex1) apply (simp add: host) apply (smt assms(1) assms(2) bind_pure_returns_result_I2 bind_returns_result_E finite_set_in host local.get_shadow_root_ok local.get_shadow_root_pure local.shadow_root_same_host return_returns_result returns_result_eq shadow_root sorted_list_of_fset.rep_eq sorted_list_of_fset_simps(1)) by (simp_all add: assms(2) bind_pure_I bind_pure_returns_result_I2 host local.get_shadow_root_ok returns_result_eq shadow_root) then show ?thesis using host_candidates host assms(1) get_shadow_root_ok apply(auto simp add: get_host_def known_ptrs_known_ptr intro!: bind_is_OK_pure_I filter_M_pure_I filter_M_is_OK_I bind_pure_I split: list.splits)[1] using assms(2) apply blast apply (meson list.distinct(1) returns_result_eq) by (meson list.distinct(1) list.inject returns_result_eq) qed end interpretation i_get_host_wf?: l_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_document get_disconnected_document_locs known_ptr known_ptrs type_wf get_host get_host_locs get_shadow_root get_shadow_root_locs get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M by(auto simp add: l_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_get_host_wf = l_heap_is_wellformed_defs + l_known_ptrs + l_type_wf + l_get_host_defs + assumes get_host_ok: "heap_is_wellformed h \ known_ptrs h \ type_wf h \ shadow_root_ptr |\| shadow_root_ptr_kinds h \ h \ ok (get_host shadow_root_ptr)" lemma get_host_wf_is_l_get_host_wf [instances]: "l_get_host_wf heap_is_wellformed known_ptr known_ptrs type_wf get_host" by(auto simp add: l_get_host_wf_def l_get_host_wf_axioms_def instances) subsubsection \get\_root\_node\_si\ locale l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent_wf + l_get_parent_get_host_wf + l_get_host_wf begin lemma get_root_node_si_ptr_in_heap: assumes "h \ ok (get_root_node_si ptr)" shows "ptr |\| object_ptr_kinds h" using assms unfolding get_root_node_si_def using get_ancestors_si_ptr_in_heap by auto lemma get_ancestors_si_ok: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" and "ptr |\| object_ptr_kinds h" shows "h \ ok (get_ancestors_si ptr)" proof (insert assms(1) assms(4), induct rule: heap_wellformed_induct_rev_si) case (step child) then show ?case using assms(2) assms(3) apply(auto simp add: get_ancestors_si_def[of child] assms(1) get_parent_parent_in_heap intro!: bind_is_OK_pure_I split: option.splits)[1] using local.get_parent_ok apply blast using get_host_ok assms(1) apply blast by (meson assms(1) is_OK_returns_result_I local.get_shadow_root_ptr_in_heap local.shadow_root_host_dual) qed lemma get_ancestors_si_remains_not_in_ancestors: assumes "heap_is_wellformed h" and "heap_is_wellformed h'" and "h \ get_ancestors_si ptr \\<^sub>r ancestors" and "h' \ get_ancestors_si ptr \\<^sub>r ancestors'" and "\p children children'. h \ get_child_nodes p \\<^sub>r children \ h' \ get_child_nodes p \\<^sub>r children' \ set children' \ set children" and "\p shadow_root_option shadow_root_option'. h \ get_shadow_root p \\<^sub>r shadow_root_option \ h' \ get_shadow_root p \\<^sub>r shadow_root_option' \ (if shadow_root_option = None then shadow_root_option' = None else shadow_root_option' = None \ shadow_root_option' = shadow_root_option)" and "node \ set ancestors" and object_ptr_kinds_eq3: "object_ptr_kinds h = object_ptr_kinds h'" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" and type_wf': "type_wf h'" shows "node \ set ancestors'" proof - have object_ptr_kinds_M_eq: "\ptrs. h \ object_ptr_kinds_M \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" using object_ptr_kinds_eq3 by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_eq: "|h \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" by(simp) show ?thesis proof (insert assms(1) assms(3) assms(4) assms(7), induct ptr arbitrary: ancestors ancestors' rule: heap_wellformed_induct_rev_si) case (step child) obtain ancestors_remains where ancestors_remains: "ancestors = child # ancestors_remains" using \h \ get_ancestors_si child \\<^sub>r ancestors\ get_ancestors_si_never_empty by(auto simp add: get_ancestors_si_def[of child] elim!: bind_returns_result_E2 split: option.splits) obtain ancestors_remains' where ancestors_remains': "ancestors' = child # ancestors_remains'" using \h' \ get_ancestors_si child \\<^sub>r ancestors'\ get_ancestors_si_never_empty by(auto simp add: get_ancestors_si_def[of child] elim!: bind_returns_result_E2 split: option.splits) have "child |\| object_ptr_kinds h" using local.get_ancestors_si_ptr_in_heap object_ptr_kinds_eq3 step.prems(2) by fastforce have "node \ child" using ancestors_remains step.prems(3) by auto have 1: "\p parent. h' \ get_parent p \\<^sub>r Some parent \ h \ get_parent p \\<^sub>r Some parent" proof - fix p parent assume "h' \ get_parent p \\<^sub>r Some parent" then obtain children' where children': "h' \ get_child_nodes parent \\<^sub>r children'" and p_in_children': "p \ set children'" using get_parent_child_dual by blast obtain children where children: "h \ get_child_nodes parent \\<^sub>r children" using get_child_nodes_ok assms(1) get_child_nodes_ptr_in_heap object_ptr_kinds_eq children' known_ptrs using type_wf type_wf' by (metis \h' \ get_parent p \\<^sub>r Some parent\ get_parent_parent_in_heap is_OK_returns_result_E local.known_ptrs_known_ptr object_ptr_kinds_eq3) have "p \ set children" using assms(5) children children' p_in_children' by blast then show "h \ get_parent p \\<^sub>r Some parent" using child_parent_dual assms(1) children known_ptrs type_wf by blast qed have 2: "\p host. h' \ get_host p \\<^sub>r host \ h \ get_host p \\<^sub>r host" proof - fix p host assume "h' \ get_host p \\<^sub>r host" then have "h' \ get_shadow_root host \\<^sub>r Some p" using local.shadow_root_host_dual by blast then have "h \ get_shadow_root host \\<^sub>r Some p" by (metis assms(6) element_ptr_kinds_commutes is_OK_returns_result_I local.get_shadow_root_ok local.get_shadow_root_ptr_in_heap node_ptr_kinds_commutes object_ptr_kinds_eq3 option.distinct(1) returns_result_select_result type_wf) then show "h \ get_host p \\<^sub>r host" by (metis assms(1) is_OK_returns_result_E known_ptrs local.get_host_ok local.get_shadow_root_shadow_root_ptr_in_heap local.shadow_root_host_dual local.shadow_root_same_host type_wf) qed show ?case proof (cases "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child") case None then show ?thesis using step(3) step(4) \node \ child\ apply(auto simp add: get_ancestors_si_def[of child] elim!: bind_returns_result_E2 split: option.splits)[1] by (metis "2" assms(1) l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M.shadow_root_same_host list.set_intros(2) local.l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms local.shadow_root_host_dual step.hyps(2) step.prems(3) type_wf) next case (Some node_child) then show ?thesis using step(3) step(4) \node \ child\ apply(auto simp add: get_ancestors_si_def[of child] elim!: bind_returns_result_E2 split: option.splits)[1] apply (meson "1" option.distinct(1) returns_result_eq) by (metis "1" list.set_intros(2) option.inject returns_result_eq step.hyps(1) step.prems(3)) qed qed qed lemma get_ancestors_si_ptrs_in_heap: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_ancestors_si ptr \\<^sub>r ancestors" assumes "ptr' \ set ancestors" shows "ptr' |\| object_ptr_kinds h" proof (insert assms(4) assms(5), induct ancestors arbitrary: ptr) case Nil then show ?case by(auto) next case (Cons a ancestors) then obtain x where x: "h \ get_ancestors_si x \\<^sub>r a # ancestors" by(auto simp add: get_ancestors_si_def[of a] elim!: bind_returns_result_E2 split: option.splits) then have "x = a" by(auto simp add: get_ancestors_si_def[of x] elim!: bind_returns_result_E2 split: option.splits) then show ?case proof (cases "ptr' = a") case True then show ?thesis using Cons.hyps Cons.prems(2) get_ancestors_si_ptr_in_heap x using \x = a\ by blast next case False obtain ptr'' where ptr'': "h \ get_ancestors_si ptr'' \\<^sub>r ancestors" using \ h \ get_ancestors_si x \\<^sub>r a # ancestors\ Cons.prems(2) False by(auto simp add: get_ancestors_si_def[of x] elim!: bind_returns_result_E2 split: option.splits) then show ?thesis using Cons.hyps Cons.prems(2) False by auto qed qed lemma get_ancestors_si_reads: assumes "heap_is_wellformed h" shows "reads get_ancestors_si_locs (get_ancestors_si node_ptr) h h'" proof (insert assms(1), induct rule: heap_wellformed_induct_rev_si) case (step child) then show ?case using [[simproc del: Product_Type.unit_eq]] get_parent_reads[unfolded reads_def] get_host_reads[unfolded reads_def] apply(simp (no_asm) add: get_ancestors_si_def) by(auto simp add: get_ancestors_si_locs_def get_parent_reads_pointers intro!: reads_bind_pure reads_subset[OF check_in_heap_reads] reads_subset[OF return_reads] reads_subset[OF get_parent_reads] reads_subset[OF get_child_nodes_reads] reads_subset[OF get_host_reads] split: option.splits) qed lemma get_ancestors_si_subset: assumes "heap_is_wellformed h" and "h \ get_ancestors_si ptr \\<^sub>r ancestors" and "ancestor \ set ancestors" and "h \ get_ancestors_si ancestor \\<^sub>r ancestor_ancestors" and type_wf: "type_wf h" and known_ptrs: "known_ptrs h" shows "set ancestor_ancestors \ set ancestors" proof (insert assms(1) assms(2) assms(3), induct ptr arbitrary: ancestors rule: heap_wellformed_induct_rev_si) case (step child) have "child |\| object_ptr_kinds h" using get_ancestors_si_ptr_in_heap step(3) by auto (* then have "h \ check_in_heap child \\<^sub>r ()" using returns_result_select_result by force *) obtain tl_ancestors where tl_ancestors: "ancestors = child # tl_ancestors" using step(3) by(auto simp add: get_ancestors_si_def[of child] intro!: bind_pure_I elim!: bind_returns_result_E2 split: option.splits) show ?case proof (induct "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child") case None show ?case proof (induct "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child") case None then show ?case using step(3) \None = cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child\ apply(auto simp add: get_ancestors_si_def[of child] elim!: bind_returns_result_E2)[1] by (metis (no_types, lifting) assms(4) empty_iff empty_set select_result_I2 set_ConsD step.prems(1) step.prems(2)) next case (Some shadow_root_child) then have "shadow_root_child |\| shadow_root_ptr_kinds h" using \child |\| object_ptr_kinds h\ by (metis (no_types, lifting) shadow_root_ptr_casts_commute shadow_root_ptr_kinds_commutes) obtain host where host: "h \ get_host shadow_root_child \\<^sub>r host" using get_host_ok assms by (meson \shadow_root_child |\| shadow_root_ptr_kinds h\ is_OK_returns_result_E) then have "h \ get_ancestors_si (cast host) \\<^sub>r tl_ancestors" using Some step(3) tl_ancestors None by(auto simp add: get_ancestors_si_def[of child] intro!: bind_pure_returns_result_I elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq) then show ?case using step(2) Some host step(4) tl_ancestors by (metis (no_types, lifting) assms(4) dual_order.trans eq_iff returns_result_eq set_ConsD set_subset_Cons shadow_root_ptr_casts_commute step.prems(1)) qed next case (Some child_node) note s1 = Some obtain parent_opt where parent_opt: "h \ get_parent child_node \\<^sub>r parent_opt" using \child |\| object_ptr_kinds h\ assms(1) Some[symmetric] get_parent_ok[OF type_wf known_ptrs] by (metis (no_types, lifting) is_OK_returns_result_E known_ptrs get_parent_ok l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms node_ptr_casts_commute node_ptr_kinds_commutes) then show ?case proof (induct parent_opt) case None then have "ancestors = [child]" using step(3) s1 apply(simp add: get_ancestors_si_def) by(auto elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq) show ?case using step(3) step(4) apply(auto simp add: \ancestors = [child]\)[1] using assms(4) returns_result_eq by fastforce next case (Some parent) then have "h \ get_ancestors_si parent \\<^sub>r tl_ancestors" using s1 tl_ancestors step(3) by(auto simp add: get_ancestors_si_def[of child] elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq) show ?case by (metis (no_types, lifting) Some.prems \h \ get_ancestors_si parent \\<^sub>r tl_ancestors\ assms(4) eq_iff node_ptr_casts_commute order_trans s1 select_result_I2 set_ConsD set_subset_Cons step.hyps(1) step.prems(1) step.prems(2) tl_ancestors) qed qed qed lemma get_ancestors_si_also_parent: assumes "heap_is_wellformed h" and "h \ get_ancestors_si some_ptr \\<^sub>r ancestors" and "cast child \ set ancestors" and "h \ get_parent child \\<^sub>r Some parent" and type_wf: "type_wf h" and known_ptrs: "known_ptrs h" shows "parent \ set ancestors" proof - obtain child_ancestors where child_ancestors: "h \ get_ancestors_si (cast child) \\<^sub>r child_ancestors" by (meson assms(1) assms(4) get_ancestors_si_ok is_OK_returns_result_I known_ptrs local.get_parent_ptr_in_heap node_ptr_kinds_commutes returns_result_select_result type_wf) then have "parent \ set child_ancestors" apply(simp add: get_ancestors_si_def) by(auto elim!: bind_returns_result_E2 split: option.splits dest!: returns_result_eq[OF assms(4)] get_ancestors_si_ptr) then show ?thesis using assms child_ancestors get_ancestors_si_subset by blast qed lemma get_ancestors_si_also_host: assumes "heap_is_wellformed h" and "h \ get_ancestors_si some_ptr \\<^sub>r ancestors" and "cast shadow_root \ set ancestors" and "h \ get_host shadow_root \\<^sub>r host" and type_wf: "type_wf h" and known_ptrs: "known_ptrs h" shows "cast host \ set ancestors" proof - obtain child_ancestors where child_ancestors: "h \ get_ancestors_si (cast shadow_root) \\<^sub>r child_ancestors" by (meson assms(1) assms(2) assms(3) get_ancestors_si_ok get_ancestors_si_ptrs_in_heap is_OK_returns_result_E known_ptrs type_wf) then have "cast host \ set child_ancestors" apply(simp add: get_ancestors_si_def) by(auto elim!: bind_returns_result_E2 split: option.splits dest!: returns_result_eq[OF assms(4)] get_ancestors_si_ptr) then show ?thesis using assms child_ancestors get_ancestors_si_subset by blast qed lemma get_ancestors_si_parent_child_rel: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "h \ get_ancestors_si child \\<^sub>r ancestors" assumes "((ptr, child) \ (parent_child_rel h)\<^sup>*)" shows "ptr \ set ancestors" proof (insert assms(5), induct ptr rule: heap_wellformed_induct_si[OF assms(1)]) case (1 ptr) then show ?case proof (cases "ptr = child") case True then show ?thesis using assms(4) local.get_ancestors_si_ptr by blast next case False obtain ptr_child where ptr_child: "(ptr, ptr_child) \ (parent_child_rel h) \ (ptr_child, child) \ (parent_child_rel h)\<^sup>*" using converse_rtranclE[OF 1(3)] \ptr \ child\ by metis then obtain ptr_child_node where ptr_child_ptr_child_node: "ptr_child = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node" using ptr_child node_ptr_casts_commute3 CD.parent_child_rel_node_ptr by (metis ) then obtain children where children: "h \ get_child_nodes ptr \\<^sub>r children" and ptr_child_node: "ptr_child_node \ set children" proof - assume a1: "\children. \h \ get_child_nodes ptr \\<^sub>r children; ptr_child_node \ set children\ \ thesis" have "ptr |\| object_ptr_kinds h" using CD.parent_child_rel_parent_in_heap ptr_child by blast moreover have "ptr_child_node \ set |h \ get_child_nodes ptr|\<^sub>r" by (metis calculation \known_ptrs h\ local.get_child_nodes_ok local.known_ptrs_known_ptr CD.parent_child_rel_child ptr_child ptr_child_ptr_child_node returns_result_select_result \type_wf h\) ultimately show ?thesis using a1 get_child_nodes_ok \type_wf h\ \known_ptrs h\ by (meson local.known_ptrs_known_ptr returns_result_select_result) qed moreover have "(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node, child) \ (parent_child_rel h)\<^sup>*" using ptr_child ptr_child_ptr_child_node by auto ultimately have "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node \ set ancestors" using 1 by auto moreover have "h \ get_parent ptr_child_node \\<^sub>r Some ptr" using assms(1) children ptr_child_node child_parent_dual using \known_ptrs h\ \type_wf h\ by blast ultimately show ?thesis using get_ancestors_si_also_parent assms \type_wf h\ by blast qed qed lemma get_ancestors_si_parent_child_host_shadow_root_rel: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "h \ get_ancestors_si child \\<^sub>r ancestors" assumes "((ptr, child) \ (parent_child_rel h \ a_host_shadow_root_rel h)\<^sup>*)" shows "ptr \ set ancestors" proof (insert assms(5), induct ptr rule: heap_wellformed_induct_si[OF assms(1)]) case (1 ptr) then show ?case proof (cases "ptr = child") case True then show ?thesis using assms(4) local.get_ancestors_si_ptr by blast next case False obtain ptr_child where ptr_child: "(ptr, ptr_child) \ (parent_child_rel h \ local.a_host_shadow_root_rel h) \ (ptr_child, child) \ (parent_child_rel h \ local.a_host_shadow_root_rel h)\<^sup>*" using converse_rtranclE[OF 1(3)] \ptr \ child\ by metis then show ?thesis proof(cases "(ptr, ptr_child) \ parent_child_rel h") case True then obtain ptr_child_node where ptr_child_ptr_child_node: "ptr_child = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node" using ptr_child node_ptr_casts_commute3 CD.parent_child_rel_node_ptr by (metis) then obtain children where children: "h \ get_child_nodes ptr \\<^sub>r children" and ptr_child_node: "ptr_child_node \ set children" proof - assume a1: "\children. \h \ get_child_nodes ptr \\<^sub>r children; ptr_child_node \ set children\ \ thesis" have "ptr |\| object_ptr_kinds h" using CD.parent_child_rel_parent_in_heap True by blast moreover have "ptr_child_node \ set |h \ get_child_nodes ptr|\<^sub>r" by (metis True assms(2) assms(3) calculation local.CD.parent_child_rel_child local.get_child_nodes_ok local.known_ptrs_known_ptr ptr_child_ptr_child_node returns_result_select_result) ultimately show ?thesis using a1 get_child_nodes_ok \type_wf h\ \known_ptrs h\ by (meson local.known_ptrs_known_ptr returns_result_select_result) qed moreover have "(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node, child) \ (parent_child_rel h \ local.a_host_shadow_root_rel h)\<^sup>*" using ptr_child True ptr_child_ptr_child_node by auto ultimately have "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node \ set ancestors" using 1 by auto moreover have "h \ get_parent ptr_child_node \\<^sub>r Some ptr" using assms(1) children ptr_child_node child_parent_dual using \known_ptrs h\ \type_wf h\ by blast ultimately show ?thesis using get_ancestors_si_also_parent assms \type_wf h\ by blast next case False then obtain host where host: "ptr = cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r host" using ptr_child by(auto simp add: a_host_shadow_root_rel_def) then obtain shadow_root where shadow_root: "h \ get_shadow_root host \\<^sub>r Some shadow_root" and ptr_child_shadow_root: "ptr_child = cast shadow_root" using ptr_child False apply(auto simp add: a_host_shadow_root_rel_def)[1] by (metis (no_types, lifting) assms(3) local.get_shadow_root_ok select_result_I) moreover have "(cast shadow_root, child) \ (parent_child_rel h \ local.a_host_shadow_root_rel h)\<^sup>*" using ptr_child ptr_child_shadow_root by blast ultimately have "cast shadow_root \ set ancestors" using "1.hyps"(2) host by blast moreover have "h \ get_host shadow_root \\<^sub>r host" by (metis assms(1) assms(2) assms(3) is_OK_returns_result_E local.get_host_ok local.get_shadow_root_shadow_root_ptr_in_heap local.shadow_root_host_dual local.shadow_root_same_host shadow_root) ultimately show ?thesis using get_ancestors_si_also_host assms(1) assms(2) assms(3) assms(4) host by blast qed qed qed lemma get_root_node_si_ok: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" and "ptr |\| object_ptr_kinds h" shows "h \ ok (get_root_node_si ptr)" using assms get_ancestors_si_ok by(auto simp add: get_root_node_si_def) lemma get_root_node_si_root_in_heap: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_root_node_si ptr \\<^sub>r root" shows "root |\| object_ptr_kinds h" using assms apply(auto simp add: get_root_node_si_def elim!: bind_returns_result_E2)[1] by (simp add: get_ancestors_si_never_empty get_ancestors_si_ptrs_in_heap) lemma get_root_node_si_same_no_parent: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_root_node_si ptr \\<^sub>r cast child" shows "h \ get_parent child \\<^sub>r None" proof (insert assms(1) assms(4), induct ptr rule: heap_wellformed_induct_rev_si) case (step c) then show ?case proof (cases "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r c") case None then show ?thesis using step(3) by(auto simp add: get_root_node_si_def get_ancestors_si_def[of c] elim!: bind_returns_result_E2 split: if_splits option.splits intro!: step(2) bind_pure_returns_result_I) next case (Some child_node) note s = this then obtain parent_opt where parent_opt: "h \ get_parent child_node \\<^sub>r parent_opt" using step(3) apply(auto simp add: get_root_node_si_def get_ancestors_si_def intro!: bind_pure_I elim!: bind_returns_result_E2)[1] by(auto split: option.splits) then show ?thesis proof(induct parent_opt) case None then show ?case using Some get_root_node_si_no_parent returns_result_eq step.prems by fastforce next case (Some parent) then show ?case using step(3) s apply(auto simp add: get_root_node_si_def get_ancestors_si_def[of c] elim!: bind_returns_result_E2 split: option.splits list.splits if_splits)[1] using assms(1) get_ancestors_si_never_empty apply blast by(auto simp add: get_root_node_si_def dest: returns_result_eq intro!: step(1) bind_pure_returns_result_I) qed qed qed end interpretation i_get_root_node_si_wf?: l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_host get_host_locs get_ancestors_si get_ancestors_si_locs get_root_node_si get_root_node_si_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_document get_disconnected_document_locs by(auto simp add: instances l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_disconnected\_document\ locale l_get_disconnected_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent_wf + l_get_parent begin lemma get_disconnected_document_ok: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_parent node_ptr \\<^sub>r None" shows "h \ ok (get_disconnected_document node_ptr)" proof - have "node_ptr |\| node_ptr_kinds h" by (meson assms(4) is_OK_returns_result_I local.get_parent_ptr_in_heap) have "\(\parent \ fset (object_ptr_kinds h). node_ptr \ set |h \ get_child_nodes parent|\<^sub>r)" apply(auto)[1] using assms(4) child_parent_dual[OF assms(1)] assms(1) assms(2) assms(3) known_ptrs_known_ptr option.simps(3) returns_result_eq returns_result_select_result by (metis (no_types, lifting) CD.get_child_nodes_ok) then have "(\document_ptr \ fset (document_ptr_kinds h). node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r)" using heap_is_wellformed_children_disc_nodes using \node_ptr |\| node_ptr_kinds h\ assms(1) by blast then obtain some_owner_document where "some_owner_document \ set (sorted_list_of_set (fset (document_ptr_kinds h)))" and "node_ptr \ set |h \ get_disconnected_nodes some_owner_document|\<^sub>r" by auto have h5: "\!x. x \ set (sorted_list_of_set (fset (document_ptr_kinds h))) \ h \ Heap_Error_Monad.bind (get_disconnected_nodes x) (\children. return (node_ptr \ set children)) \\<^sub>r True" apply(auto intro!: bind_pure_returns_result_I)[1] apply (smt CD.get_disconnected_nodes_ok CD.get_disconnected_nodes_pure \\document_ptr\fset (document_ptr_kinds h). node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r\ assms(2) bind_pure_returns_result_I2 notin_fset return_returns_result select_result_I2) apply(auto elim!: bind_returns_result_E2 intro!: bind_pure_returns_result_I)[1] using heap_is_wellformed_one_disc_parent assms(1) by blast let ?filter_M = "filter_M (\document_ptr. Heap_Error_Monad.bind (get_disconnected_nodes document_ptr) (\disconnected_nodes. return (node_ptr \ set disconnected_nodes))) (sorted_list_of_set (fset (document_ptr_kinds h)))" have "h \ ok (?filter_M)" using CD.get_disconnected_nodes_ok by (smt CD.get_disconnected_nodes_pure DocumentMonad.ptr_kinds_M_ptr_kinds DocumentMonad.ptr_kinds_ptr_kinds_M assms(2) bind_is_OK_pure_I bind_pure_I document_ptr_kinds_M_def filter_M_is_OK_I l_ptr_kinds_M.ptr_kinds_M_ok return_ok return_pure returns_result_select_result) then obtain candidates where candidates: "h \ filter_M (\document_ptr. Heap_Error_Monad.bind (get_disconnected_nodes document_ptr) (\disconnected_nodes. return (node_ptr \ set disconnected_nodes))) (sorted_list_of_set (fset (document_ptr_kinds h))) \\<^sub>r candidates" by auto have "candidates = [some_owner_document]" apply(rule filter_M_ex1[OF candidates \some_owner_document \ set (sorted_list_of_set (fset (document_ptr_kinds h)))\ h5]) using \node_ptr \ set |h \ get_disconnected_nodes some_owner_document|\<^sub>r\ \some_owner_document \ set (sorted_list_of_set (fset (document_ptr_kinds h)))\ by(auto simp add: CD.get_disconnected_nodes_ok assms(2) intro!: bind_pure_I intro!: bind_pure_returns_result_I) then show ?thesis using candidates \node_ptr |\| node_ptr_kinds h\ apply(auto simp add: get_disconnected_document_def intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I split: list.splits)[1] apply (meson not_Cons_self2 returns_result_eq) by (meson list.distinct(1) list.inject returns_result_eq) qed end interpretation i_get_disconnected_document_wf?: l_get_disconnected_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs known_ptrs get_parent get_parent_locs by(auto simp add: l_get_disconnected_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_disconnected_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_owner\_document\ locale l_get_owner_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_disconnected_nodes + l_get_child_nodes + l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent_wf + l_known_ptrs + l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + assumes known_ptr_impl: "known_ptr = ShadowRootClass.known_ptr" begin lemma get_owner_document_disconnected_nodes: assumes "heap_is_wellformed h" assumes "h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes" assumes "node_ptr \ set disc_nodes" assumes known_ptrs: "known_ptrs h" assumes type_wf: "type_wf h" shows "h \ get_owner_document (cast node_ptr) \\<^sub>r document_ptr" proof - have 2: "node_ptr |\| node_ptr_kinds h" using assms apply(auto simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.a_all_ptrs_in_heap_def)[1] using assms(1) local.heap_is_wellformed_disc_nodes_in_heap by blast have 3: "document_ptr |\| document_ptr_kinds h" using assms(2) get_disconnected_nodes_ptr_in_heap by blast then have 4: "\(\parent_ptr. parent_ptr |\| object_ptr_kinds h \ node_ptr \ set |h \ get_child_nodes parent_ptr|\<^sub>r)" using CD.distinct_lists_no_parent assms unfolding heap_is_wellformed_def CD.heap_is_wellformed_def by simp moreover have "(\document_ptr. document_ptr |\| document_ptr_kinds h \ node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r) \ (\parent_ptr. parent_ptr |\| object_ptr_kinds h \ node_ptr \ set |h \ get_child_nodes parent_ptr|\<^sub>r)" using assms(1) 2 "3" assms(2) assms(3) by auto ultimately have 0: "\!document_ptr\set |h \ document_ptr_kinds_M|\<^sub>r. node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r" using concat_map_distinct assms(1) known_ptrs_implies by (smt CD.heap_is_wellformed_one_disc_parent DocumentMonad.ptr_kinds_ptr_kinds_M disjoint_iff_not_equal local.get_disconnected_nodes_ok local.heap_is_wellformed_def returns_result_select_result type_wf) have "h \ get_parent node_ptr \\<^sub>r None" using 4 2 apply(auto simp add: get_parent_def intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I )[1] apply(auto intro!: filter_M_empty_I bind_pure_I bind_pure_returns_result_I)[1] using get_child_nodes_ok assms(4) type_wf by (metis get_child_nodes_ok known_ptrs_known_ptr returns_result_select_result) then have 4: "h \ get_root_node_si (cast node_ptr) \\<^sub>r cast node_ptr" using get_root_node_si_no_parent by simp obtain document_ptrs where document_ptrs: "h \ document_ptr_kinds_M \\<^sub>r document_ptrs" by simp then have "h \ ok (filter_M (\document_ptr. do { disconnected_nodes \ get_disconnected_nodes document_ptr; return (((cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)) \ cast ` set disconnected_nodes) }) document_ptrs)" using assms(1) get_disconnected_nodes_ok type_wf by(auto intro!: bind_is_OK_I2 filter_M_is_OK_I bind_pure_I) then obtain candidates where candidates: "h \ filter_M (\document_ptr. do { disconnected_nodes \ get_disconnected_nodes document_ptr; return (((cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)) \ cast ` set disconnected_nodes) }) document_ptrs \\<^sub>r candidates" by auto have filter: "filter (\document_ptr. |h \ do { disconnected_nodes \ get_disconnected_nodes document_ptr; return (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr \ cast ` set disconnected_nodes) }|\<^sub>r) document_ptrs = [document_ptr]" apply(rule filter_ex1) using 0 document_ptrs apply(simp)[1] apply (smt "0" "3" assms bind_is_OK_pure_I bind_pure_returns_result_I bind_pure_returns_result_I2 bind_returns_result_E2 bind_returns_result_E3 document_ptr_kinds_M_def get_disconnected_nodes_ok get_disconnected_nodes_pure image_eqI is_OK_returns_result_E l_ptr_kinds_M.ptr_kinds_ptr_kinds_M return_ok return_returns_result returns_result_eq select_result_E select_result_I select_result_I2 select_result_I2) using assms(2) assms(3) apply (metis (no_types, lifting) bind_pure_returns_result_I2 is_OK_returns_result_I local.get_disconnected_nodes_pure node_ptr_inclusion return_returns_result select_result_I2) using document_ptrs 3 apply(simp) using document_ptrs by simp have "h \ filter_M (\document_ptr. do { disconnected_nodes \ get_disconnected_nodes document_ptr; return (((cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)) \ cast ` set disconnected_nodes) }) document_ptrs \\<^sub>r [document_ptr]" apply(rule filter_M_filter2) using get_disconnected_nodes_ok document_ptrs 3 assms(1) type_wf filter by(auto intro: bind_pure_I bind_is_OK_I2) with 4 document_ptrs have "h \ CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r document_ptr" by(auto simp add: CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I split: option.splits) moreover have "known_ptr (cast node_ptr)" using known_ptrs_known_ptr[OF known_ptrs, where ptr="cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr"] 2 known_ptrs_implies by(simp) ultimately show ?thesis using 2 apply(auto simp add: CD.a_get_owner_document_tups_def get_owner_document_def a_get_owner_document_tups_def known_ptr_impl)[1] apply(split invoke_splits, (rule conjI | rule impI)+)+ apply(drule(1) known_ptr_not_shadow_root_ptr) apply(drule(1) known_ptr_not_document_ptr) apply(drule(1) known_ptr_not_character_data_ptr) apply(drule(1) known_ptr_not_element_ptr) apply(simp add: NodeClass.known_ptr_defs) by(auto split: option.splits intro!: bind_pure_returns_result_I) qed lemma in_disconnected_nodes_no_parent: assumes "heap_is_wellformed h" assumes "h \ get_parent node_ptr \\<^sub>r None" assumes "h \ get_owner_document (cast node_ptr) \\<^sub>r owner_document" assumes "h \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes" assumes "known_ptrs h" assumes "type_wf h" shows "node_ptr \ set disc_nodes" proof - have "\parent. parent |\| object_ptr_kinds h \ node_ptr \ set |h \ get_child_nodes parent|\<^sub>r" using assms(2) by (meson get_child_nodes_ok assms(1) assms(5) assms(6) local.child_parent_dual local.known_ptrs_known_ptr option.distinct(1) returns_result_eq returns_result_select_result) then show ?thesis by (smt assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) finite_set_in is_OK_returns_result_I local.get_disconnected_nodes_ok local.get_owner_document_disconnected_nodes local.get_parent_ptr_in_heap local.heap_is_wellformed_children_disc_nodes returns_result_select_result select_result_I2) qed lemma get_owner_document_owner_document_in_heap_node: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r owner_document" shows "owner_document |\| document_ptr_kinds h" proof - obtain root where root: "h \ get_root_node_si (cast node_ptr) \\<^sub>r root" using assms(4) by(auto simp add: CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 split: option.splits) then show ?thesis proof (cases "is_document_ptr root") case True then show ?thesis using assms(4) root apply(auto simp add: CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 intro!: filter_M_pure_I bind_pure_I split: option.splits)[1] apply(drule(1) returns_result_eq) apply(auto)[1] using assms document_ptr_kinds_commutes get_root_node_si_root_in_heap by blast next case False have "known_ptr root" using assms local.get_root_node_si_root_in_heap local.known_ptrs_known_ptr root by blast have "root |\| object_ptr_kinds h" using root using assms local.get_root_node_si_root_in_heap by blast have "\is_shadow_root_ptr root" using root using local.get_root_node_si_root_not_shadow_root by blast then have "is_node_ptr_kind root" using False \known_ptr root\ \root |\| object_ptr_kinds h\ apply(simp add: known_ptr_impl known_ptr_defs DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs) using is_node_ptr_kind_none by force then have "(\document_ptr \ fset (document_ptr_kinds h). root \ cast ` set |h \ get_disconnected_nodes document_ptr|\<^sub>r)" using local.child_parent_dual local.get_child_nodes_ok local.get_root_node_si_same_no_parent local.heap_is_wellformed_children_disc_nodes local.known_ptrs_known_ptr node_ptr_casts_commute3 node_ptr_inclusion node_ptr_kinds_commutes notin_fset option.distinct(1) returns_result_eq returns_result_select_result root by (metis (no_types, lifting) assms \root |\| object_ptr_kinds h\) then obtain some_owner_document where "some_owner_document |\| document_ptr_kinds h" and "root \ cast ` set |h \ get_disconnected_nodes some_owner_document|\<^sub>r" by auto then obtain candidates where candidates: "h \ filter_M (\document_ptr. Heap_Error_Monad.bind (get_disconnected_nodes document_ptr) (\disconnected_nodes. return (root \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ` set disconnected_nodes))) (sorted_list_of_set (fset (document_ptr_kinds h))) \\<^sub>r candidates" by (metis (no_types, lifting) assms bind_is_OK_I2 bind_pure_I filter_M_is_OK_I finite_fset is_OK_returns_result_E local.get_disconnected_nodes_ok local.get_disconnected_nodes_pure notin_fset return_ok return_pure sorted_list_of_set(1)) then have "some_owner_document \ set candidates" apply(rule filter_M_in_result_if_ok) using \some_owner_document |\| document_ptr_kinds h\ \root \ cast ` set |h \ get_disconnected_nodes some_owner_document|\<^sub>r\ apply(auto intro!: bind_pure_I bind_pure_returns_result_I)[1] using \some_owner_document |\| document_ptr_kinds h\ \root \ cast ` set |h \ get_disconnected_nodes some_owner_document|\<^sub>r\ apply(auto intro!: bind_pure_I bind_pure_returns_result_I)[1] using \some_owner_document |\| document_ptr_kinds h\ \root \ cast ` set |h \ get_disconnected_nodes some_owner_document|\<^sub>r\ apply(auto simp add: assms local.get_disconnected_nodes_ok intro!: bind_pure_I bind_pure_returns_result_I)[1] done then have "candidates \ []" by auto then have "owner_document \ set candidates" using assms(4) root apply(auto simp add: CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 intro!: filter_M_pure_I bind_pure_I split: option.splits)[1] apply (metis candidates list.set_sel(1) returns_result_eq) by (metis \is_node_ptr_kind root\ node_ptr_no_document_ptr_cast returns_result_eq) then show ?thesis using candidates by (meson bind_pure_I bind_returns_result_E2 filter_M_holds_for_result is_OK_returns_result_I local.get_disconnected_nodes_ptr_in_heap local.get_disconnected_nodes_pure return_pure) qed qed lemma get_owner_document_owner_document_in_heap: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_owner_document ptr \\<^sub>r owner_document" shows "owner_document |\| document_ptr_kinds h" using assms apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def CD.a_get_owner_document_tups_def)[1] apply(split invoke_split_asm)+ proof - assume "h \ invoke [] ptr () \\<^sub>r owner_document" then show "owner_document |\| document_ptr_kinds h" by (meson invoke_empty is_OK_returns_result_I) next assume "h \ Heap_Error_Monad.bind (check_in_heap ptr) (\_. (CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r) ptr ()) \\<^sub>r owner_document" then show "owner_document |\| document_ptr_kinds h" by(auto simp add: CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 split: if_splits) next assume 0: "heap_is_wellformed h" and 1: "type_wf h" and 2: "known_ptrs h" and 3: "\ is_element_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 4: "is_character_data_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 5: "h \ Heap_Error_Monad.bind (check_in_heap ptr) (\_. (CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r) ptr ()) \\<^sub>r owner_document" then show ?thesis by (metis bind_returns_result_E2 check_in_heap_pure comp_apply get_owner_document_owner_document_in_heap_node) next assume 0: "heap_is_wellformed h" and 1: "type_wf h" and 2: "known_ptrs h" and 3: "is_element_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 4: "h \ Heap_Error_Monad.bind (check_in_heap ptr) (\_. (CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r) ptr ()) \\<^sub>r owner_document" then show ?thesis by (metis bind_returns_result_E2 check_in_heap_pure comp_apply get_owner_document_owner_document_in_heap_node) next assume 0: "heap_is_wellformed h" and 1: "type_wf h" and 2: "known_ptrs h" and 3: "\ is_element_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 4: "\ is_character_data_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 5: "\ is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 6: "is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 7: "h \ Heap_Error_Monad.bind (check_in_heap ptr) (\_. (local.a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r) ptr ()) \\<^sub>r owner_document" then show "owner_document |\| document_ptr_kinds h" apply(auto simp add: CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: filter_M_pure_I bind_pure_I elim!: bind_returns_result_E2 split: if_splits option.splits)[1] using get_owner_document_owner_document_in_heap_node by blast qed lemma get_owner_document_ok: assumes "heap_is_wellformed h" "known_ptrs h" "type_wf h" assumes "ptr |\| object_ptr_kinds h" shows "h \ ok (get_owner_document ptr)" proof - have "known_ptr ptr" using assms(2) assms(4) local.known_ptrs_known_ptr by blast then show ?thesis apply(simp add: get_owner_document_def a_get_owner_document_tups_def CD.a_get_owner_document_tups_def) apply(split invoke_splits, (rule conjI | rule impI)+)+ proof - assume 0: "known_ptr ptr" and 1: "\ is_element_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 2: "\ is_character_data_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 3: "\ is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 4: "\ is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" then show "h \ ok invoke [] ptr ()" using NodeClass.a_known_ptr_def known_ptr_not_character_data_ptr known_ptr_not_document_ptr known_ptr_not_shadow_root_ptr known_ptr_not_element_ptr known_ptr_impl by blast next assume 0: "known_ptr ptr" and 1: "\ is_element_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 2: "\ is_character_data_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 3: "\ is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" then show "is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \ h \ ok Heap_Error_Monad.bind (check_in_heap ptr) (\_. (local.a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r) ptr ())" using assms(1) assms(2) assms(3) assms(4) by(auto simp add: local.get_host_ok get_root_node_def CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I filter_M_is_OK_I get_root_node_si_ok get_disconnected_nodes_ok intro!: local.get_shadow_root_ptr_in_heap local.shadow_root_host_dual split: option.splits) next show "is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \ h \ ok Heap_Error_Monad.bind (check_in_heap ptr) (\_. (local.CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r) ptr ())" using assms(4) by(auto simp add: CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def split: option.splits) next show "is_character_data_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \ h \ ok Heap_Error_Monad.bind (check_in_heap ptr) (\_. (local.CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r) ptr ())" using assms(1) assms(2) assms(3) assms(4) by(auto simp add: local.get_host_ok get_root_node_def CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I filter_M_is_OK_I get_root_node_si_ok get_disconnected_nodes_ok intro!: local.get_shadow_root_ptr_in_heap local.shadow_root_host_dual split: option.splits) next show "is_element_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \ h \ ok Heap_Error_Monad.bind (check_in_heap ptr) (\_. (local.CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r) ptr ())" using assms(1) assms(2) assms(3) assms(4) by(auto simp add: local.get_host_ok get_root_node_def CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I filter_M_is_OK_I get_root_node_si_ok get_disconnected_nodes_ok intro!: local.get_shadow_root_ptr_in_heap local.shadow_root_host_dual split: option.splits) qed qed end interpretation i_get_owner_document_wf?: l_get_owner_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_disconnected_nodes get_disconnected_nodes_locs known_ptr get_child_nodes get_child_nodes_locs DocumentClass.known_ptr get_parent get_parent_locs get_root_node_si get_root_node_si_locs CD.a_get_owner_document get_host get_host_locs get_owner_document get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_document get_disconnected_document_locs known_ptrs get_ancestors_si get_ancestors_si_locs by(auto simp add: l_get_owner_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_owner_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_get_owner_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_owner_document_wf_is_l_get_owner_document_wf [instances]: "l_get_owner_document_wf heap_is_wellformed type_wf known_ptr known_ptrs get_disconnected_nodes get_owner_document get_parent" apply(auto simp add: l_get_owner_document_wf_def l_get_owner_document_wf_axioms_def instances)[1] using get_owner_document_disconnected_nodes apply fast using in_disconnected_nodes_no_parent apply fast using get_owner_document_owner_document_in_heap apply fast using get_owner_document_ok apply fast done subsubsection \remove\_child\ locale l_remove_child_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_disconnected_nodes_get_disconnected_nodes + l_get_child_nodes + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_owner_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_child_nodes_get_shadow_root + l_set_disconnected_nodes_get_shadow_root + l_set_child_nodes_get_tag_name + l_set_disconnected_nodes_get_tag_name + CD: l_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma remove_child_preserves_type_wf: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ remove_child ptr child \\<^sub>h h'" shows "type_wf h'" using CD.remove_child_heap_is_wellformed_preserved(1) assms unfolding heap_is_wellformed_def by auto lemma remove_child_preserves_known_ptrs: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ remove_child ptr child \\<^sub>h h'" shows "known_ptrs h'" using CD.remove_child_heap_is_wellformed_preserved(2) assms unfolding heap_is_wellformed_def by auto lemma remove_child_heap_is_wellformed_preserved: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ remove_child ptr child \\<^sub>h h'" shows "heap_is_wellformed h'" proof - have "heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'" using CD.remove_child_heap_is_wellformed_preserved(3) assms unfolding heap_is_wellformed_def by auto have shadow_root_eq: "\ptr' shadow_root_ptr_opt. h \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt = h' \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt" using get_shadow_root_reads remove_child_writes assms(4) apply(rule reads_writes_preserved) by(auto simp add: remove_child_locs_def set_child_nodes_get_shadow_root set_disconnected_nodes_get_shadow_root) then have shadow_root_eq2: "\ptr'. |h \ get_shadow_root ptr'|\<^sub>r = |h' \ get_shadow_root ptr'|\<^sub>r" by (meson select_result_eq) have tag_name_eq: "\ptr' tag. h \ get_tag_name ptr' \\<^sub>r tag = h' \ get_tag_name ptr' \\<^sub>r tag" using get_tag_name_reads remove_child_writes assms(4) apply(rule reads_writes_preserved) by(auto simp add: remove_child_locs_def set_child_nodes_get_tag_name set_disconnected_nodes_get_tag_name) then have tag_name_eq2: "\ptr'. |h \ get_tag_name ptr'|\<^sub>r = |h' \ get_tag_name ptr'|\<^sub>r" by (meson select_result_eq) have object_ptr_kinds_eq: "object_ptr_kinds h = object_ptr_kinds h'" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF remove_child_writes assms(4)]) unfolding remove_child_locs_def using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) have shadow_root_ptr_kinds_eq: "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'" using object_ptr_kinds_eq by(auto simp add: shadow_root_ptr_kinds_def) have element_ptr_kinds_eq: "element_ptr_kinds h = element_ptr_kinds h'" using object_ptr_kinds_eq by(auto simp add: element_ptr_kinds_def node_ptr_kinds_def) have "parent_child_rel h' \ parent_child_rel h" using \heap_is_wellformed h\ heap_is_wellformed_def using CD.remove_child_parent_child_rel_subset using \known_ptrs h\ \type_wf h\ assms(4) by simp show ?thesis using \heap_is_wellformed h\ using \heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'\ \parent_child_rel h' \ parent_child_rel h\ apply(auto simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def a_host_shadow_root_rel_def a_all_ptrs_in_heap_def a_distinct_lists_def a_shadow_root_valid_def object_ptr_kinds_eq element_ptr_kinds_eq shadow_root_ptr_kinds_eq shadow_root_eq shadow_root_eq2 tag_name_eq tag_name_eq2)[1] by (meson acyclic_subset order_refl sup_mono) qed lemma remove_preserves_type_wf: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ remove child \\<^sub>h h'" shows "type_wf h'" using CD.remove_heap_is_wellformed_preserved(1) assms unfolding heap_is_wellformed_def by auto lemma remove_preserves_known_ptrs: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ remove child \\<^sub>h h'" shows "known_ptrs h'" using CD.remove_heap_is_wellformed_preserved(2) assms unfolding heap_is_wellformed_def by auto lemma remove_heap_is_wellformed_preserved: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ remove child \\<^sub>h h'" shows "heap_is_wellformed h'" using assms by(auto simp add: remove_def elim!: bind_returns_heap_E2 intro: remove_child_heap_is_wellformed_preserved split: option.splits) lemma remove_child_removes_child: "heap_is_wellformed h \ h \ remove_child ptr' child \\<^sub>h h' \ h' \ get_child_nodes ptr \\<^sub>r children \ known_ptrs h \ type_wf h \ child \ set children" using CD.remove_child_removes_child local.heap_is_wellformed_def by blast lemma remove_child_removes_first_child: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_child_nodes ptr \\<^sub>r node_ptr # children \ h \ remove_child ptr node_ptr \\<^sub>h h' \ h' \ get_child_nodes ptr \\<^sub>r children" using CD.remove_child_removes_first_child local.heap_is_wellformed_def by blast lemma remove_removes_child: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_child_nodes ptr \\<^sub>r node_ptr # children \ h \ remove node_ptr \\<^sub>h h' \ h' \ get_child_nodes ptr \\<^sub>r children" using CD.remove_removes_child local.heap_is_wellformed_def by blast lemma remove_for_all_empty_children: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_child_nodes ptr \\<^sub>r children \ h \ forall_M remove children \\<^sub>h h' \ h' \ get_child_nodes ptr \\<^sub>r []" using CD.remove_for_all_empty_children local.heap_is_wellformed_def by blast end interpretation i_remove_child_wf2?: l_remove_child_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs known_ptr get_child_nodes get_child_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs DocumentClass.known_ptr get_parent get_parent_locs get_root_node_si get_root_node_si_locs CD.a_get_owner_document get_owner_document known_ptrs get_ancestors_si get_ancestors_si_locs set_child_nodes set_child_nodes_locs remove_child remove_child_locs remove by(auto simp add: l_remove_child_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_remove_child_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma remove_child_wf2_is_l_remove_child_wf2 [instances]: "l_remove_child_wf2 type_wf known_ptr known_ptrs remove_child heap_is_wellformed get_child_nodes remove" apply(auto simp add: l_remove_child_wf2_def l_remove_child_wf2_axioms_def instances)[1] using remove_child_preserves_type_wf apply fast using remove_child_preserves_known_ptrs apply fast using remove_child_heap_is_wellformed_preserved apply (fast) using remove_preserves_type_wf apply fast using remove_preserves_known_ptrs apply fast using remove_heap_is_wellformed_preserved apply (fast) using remove_child_removes_child apply fast using remove_child_removes_first_child apply fast using remove_removes_child apply fast using remove_for_all_empty_children apply fast done subsubsection \adopt\_node\ locale l_adopt_node_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes + l_get_disconnected_nodes + l_set_child_nodes_get_shadow_root + l_set_disconnected_nodes_get_shadow_root + l_set_child_nodes_get_tag_name + l_set_disconnected_nodes_get_tag_name + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_root_node + l_set_disconnected_nodes_get_child_nodes + l_get_owner_document_wf + l_remove_child_wf2 + l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma adopt_node_removes_child: assumes wellformed: "heap_is_wellformed h" and adopt_node: "h \ adopt_node owner_document node_ptr \\<^sub>h h2" and children: "h2 \ get_child_nodes ptr \\<^sub>r children" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "node_ptr \ set children" proof - obtain old_document parent_opt h' where old_document: "h \ get_owner_document (cast node_ptr) \\<^sub>r old_document" and parent_opt: "h \ get_parent node_ptr \\<^sub>r parent_opt" and h': "h \ (case parent_opt of Some parent \ remove_child parent node_ptr | None \ return () ) \\<^sub>h h'" using adopt_node by(auto simp add: adopt_node_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated] bind_returns_heap_E2[rotated, OF get_parent_pure, rotated] bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] split: if_splits) then have "h' \ get_child_nodes ptr \\<^sub>r children" using adopt_node apply(auto simp add: adopt_node_def dest!: bind_returns_heap_E3[rotated, OF old_document, rotated] bind_returns_heap_E3[rotated, OF parent_opt, rotated] elim!: bind_returns_heap_E4[rotated, OF h', rotated])[1] apply(auto split: if_splits elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated])[1] apply (simp add: set_disconnected_nodes_get_child_nodes children reads_writes_preserved[OF get_child_nodes_reads set_disconnected_nodes_writes]) using children by blast show ?thesis proof(insert parent_opt h', induct parent_opt) case None then show ?case using child_parent_dual wellformed known_ptrs type_wf \h' \ get_child_nodes ptr \\<^sub>r children\ returns_result_eq by fastforce next case (Some option) then show ?case using remove_child_removes_child \h' \ get_child_nodes ptr \\<^sub>r children\ known_ptrs type_wf wellformed by auto qed qed lemma adopt_node_preserves_wellformedness: assumes "heap_is_wellformed h" and "h \ adopt_node document_ptr child \\<^sub>h h'" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "heap_is_wellformed h'" and "known_ptrs h'" and "type_wf h'" proof - obtain old_document parent_opt h2 where old_document: "h \ get_owner_document (cast child) \\<^sub>r old_document" and parent_opt: "h \ get_parent child \\<^sub>r parent_opt" and h2: "h \ (case parent_opt of Some parent \ remove_child parent child | None \ return ()) \\<^sub>h h2" and h': "h2 \ (if document_ptr \ old_document then do { old_disc_nodes \ get_disconnected_nodes old_document; set_disconnected_nodes old_document (remove1 child old_disc_nodes); disc_nodes \ get_disconnected_nodes document_ptr; set_disconnected_nodes document_ptr (child # disc_nodes) } else do { return () }) \\<^sub>h h'" using assms(2) by(auto simp add: adopt_node_def elim!: bind_returns_heap_E dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure] pure_returns_heap_eq[rotated, OF get_parent_pure]) have object_ptr_kinds_h_eq3: "object_ptr_kinds h = object_ptr_kinds h2" using h2 apply(simp split: option.splits) apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF remove_child_writes]) using remove_child_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h: "\ptrs. h \ object_ptr_kinds_M \\<^sub>r ptrs = h2 \ object_ptr_kinds_M \\<^sub>r ptrs" unfolding object_ptr_kinds_M_defs by simp then have object_ptr_kinds_eq_h: "|h \ object_ptr_kinds_M|\<^sub>r = |h2 \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq_h: "|h \ node_ptr_kinds_M|\<^sub>r = |h2 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have wellformed_h2: "heap_is_wellformed h2" using h2 remove_child_heap_is_wellformed_preserved known_ptrs type_wf by (metis (no_types, lifting) assms(1) option.case_eq_if pure_returns_heap_eq return_pure) have "type_wf h2" using h2 remove_child_preserves_type_wf assms by(auto split: option.splits) have "known_ptrs h2" using h2 remove_child_preserves_known_ptrs assms by(auto split: option.splits) then have "heap_is_wellformed h' \ known_ptrs h' \ type_wf h'" proof(cases "document_ptr = old_document") case True then show "heap_is_wellformed h' \ known_ptrs h' \ type_wf h'" using h' wellformed_h2 \known_ptrs h2\ \type_wf h2\ by auto next case False then obtain h3 old_disc_nodes disc_nodes_document_ptr_h3 where docs_neq: "document_ptr \ old_document" and old_disc_nodes: "h2 \ get_disconnected_nodes old_document \\<^sub>r old_disc_nodes" and h3: "h2 \ set_disconnected_nodes old_document (remove1 child old_disc_nodes) \\<^sub>h h3" and disc_nodes_document_ptr_h3: "h3 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_document_ptr_h3" and h': "h3 \ set_disconnected_nodes document_ptr (child # disc_nodes_document_ptr_h3) \\<^sub>h h'" using h' by(auto elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] ) have object_ptr_kinds_h2_eq3: "object_ptr_kinds h2 = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes h3]) using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h2: "\ptrs. h2 \ object_ptr_kinds_M \\<^sub>r ptrs = h3 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_eq_h2: "|h2 \ object_ptr_kinds_M|\<^sub>r = |h3 \ object_ptr_kinds_M|\<^sub>r" by(simp) then have node_ptr_kinds_eq_h2: "|h2 \ node_ptr_kinds_M|\<^sub>r = |h3 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast then have node_ptr_kinds_eq3_h2: "node_ptr_kinds h2 = node_ptr_kinds h3" by auto have document_ptr_kinds_eq2_h2: "|h2 \ document_ptr_kinds_M|\<^sub>r = |h3 \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_eq_h2 document_ptr_kinds_M_eq by auto then have document_ptr_kinds_eq3_h2: "document_ptr_kinds h2 = document_ptr_kinds h3" using object_ptr_kinds_eq_h2 document_ptr_kinds_M_eq by auto have children_eq_h2: "\ptr children. h2 \ get_child_nodes ptr \\<^sub>r children = h3 \ get_child_nodes ptr \\<^sub>r children" using get_child_nodes_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by (simp add: set_disconnected_nodes_get_child_nodes) then have children_eq2_h2: "\ptr. |h2 \ get_child_nodes ptr|\<^sub>r = |h3 \ get_child_nodes ptr|\<^sub>r" using select_result_eq by force have object_ptr_kinds_h3_eq3: "object_ptr_kinds h3 = object_ptr_kinds h'" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes h']) using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h3: "\ptrs. h3 \ object_ptr_kinds_M \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_eq_h3: "|h3 \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" by(simp) then have node_ptr_kinds_eq_h3: "|h3 \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast then have node_ptr_kinds_eq3_h3: "node_ptr_kinds h3 = node_ptr_kinds h'" by auto have document_ptr_kinds_eq2_h3: "|h3 \ document_ptr_kinds_M|\<^sub>r = |h' \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_eq_h3 document_ptr_kinds_M_eq by auto then have document_ptr_kinds_eq3_h3: "document_ptr_kinds h3 = document_ptr_kinds h'" using object_ptr_kinds_eq_h3 document_ptr_kinds_M_eq by auto have children_eq_h3: "\ptr children. h3 \ get_child_nodes ptr \\<^sub>r children = h' \ get_child_nodes ptr \\<^sub>r children" using get_child_nodes_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by (simp add: set_disconnected_nodes_get_child_nodes) then have children_eq2_h3: "\ptr. |h3 \ get_child_nodes ptr|\<^sub>r = |h' \ get_child_nodes ptr|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. old_document \ doc_ptr \ h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by (simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers) then have disconnected_nodes_eq2_h2: "\doc_ptr. old_document \ doc_ptr \ |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force obtain disc_nodes_old_document_h2 where disc_nodes_old_document_h2: "h2 \ get_disconnected_nodes old_document \\<^sub>r disc_nodes_old_document_h2" using old_disc_nodes by blast then have disc_nodes_old_document_h3: "h3 \ get_disconnected_nodes old_document \\<^sub>r remove1 child disc_nodes_old_document_h2" using h3 old_disc_nodes returns_result_eq set_disconnected_nodes_get_disconnected_nodes by fastforce have "distinct disc_nodes_old_document_h2" using disc_nodes_old_document_h2 local.heap_is_wellformed_disconnected_nodes_distinct wellformed_h2 by blast have "type_wf h2" proof (insert h2, induct parent_opt) case None then show ?case using type_wf by simp next case (Some option) then show ?case using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF remove_child_writes] type_wf remove_child_types_preserved by (simp add: reflp_def transp_def) qed then have "type_wf h3" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h3] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) then have "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h'] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) have disconnected_nodes_eq_h3: "\doc_ptr disc_nodes. document_ptr \ doc_ptr \ h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by (simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers) then have disconnected_nodes_eq2_h3: "\doc_ptr. document_ptr \ doc_ptr \ |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have disc_nodes_document_ptr_h2: "h2 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_document_ptr_h3" using disconnected_nodes_eq_h2 docs_neq disc_nodes_document_ptr_h3 by auto have disc_nodes_document_ptr_h': "h' \ get_disconnected_nodes document_ptr \\<^sub>r child # disc_nodes_document_ptr_h3" using h' disc_nodes_document_ptr_h3 using set_disconnected_nodes_get_disconnected_nodes by blast have document_ptr_in_heap: "document_ptr |\| document_ptr_kinds h2" using disc_nodes_document_ptr_h3 document_ptr_kinds_eq2_h2 get_disconnected_nodes_ok assms(1) unfolding heap_is_wellformed_def using disc_nodes_document_ptr_h2 get_disconnected_nodes_ptr_in_heap by blast have old_document_in_heap: "old_document |\| document_ptr_kinds h2" using disc_nodes_old_document_h3 document_ptr_kinds_eq2_h2 get_disconnected_nodes_ok assms(1) unfolding heap_is_wellformed_def using get_disconnected_nodes_ptr_in_heap old_disc_nodes by blast have "child \ set disc_nodes_old_document_h2" proof (insert parent_opt h2, induct parent_opt) case None then have "h = h2" by(auto) moreover have "CD.a_owner_document_valid h" using assms(1) by(simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) ultimately show ?case using old_document disc_nodes_old_document_h2 None(1) child_parent_dual[OF assms(1)] in_disconnected_nodes_no_parent assms(1) known_ptrs type_wf by blast next case (Some option) then show ?case apply(simp split: option.splits) using assms(1) disc_nodes_old_document_h2 old_document remove_child_in_disconnected_nodes known_ptrs by blast qed have "child \ set (remove1 child disc_nodes_old_document_h2)" using disc_nodes_old_document_h3 h3 known_ptrs wellformed_h2 \distinct disc_nodes_old_document_h2\ by auto have "child \ set disc_nodes_document_ptr_h3" proof - have "CD.a_distinct_lists h2" using heap_is_wellformed_def CD.heap_is_wellformed_def wellformed_h2 by blast then have 0: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>r) |h2 \ document_ptr_kinds_M|\<^sub>r))" by(simp add: CD.a_distinct_lists_def) show ?thesis using distinct_concat_map_E(1)[OF 0] \child \ set disc_nodes_old_document_h2\ disc_nodes_old_document_h2 disc_nodes_document_ptr_h2 by (meson \type_wf h2\ docs_neq known_ptrs local.get_owner_document_disconnected_nodes local.known_ptrs_preserved object_ptr_kinds_h_eq3 returns_result_eq wellformed_h2) qed have child_in_heap: "child |\| node_ptr_kinds h" using get_owner_document_ptr_in_heap[OF is_OK_returns_result_I[OF old_document]] node_ptr_kinds_commutes by blast have "CD.a_acyclic_heap h2" using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) have "parent_child_rel h' \ parent_child_rel h2" proof fix x assume "x \ parent_child_rel h'" then show "x \ parent_child_rel h2" using object_ptr_kinds_h2_eq3 object_ptr_kinds_h3_eq3 children_eq2_h2 children_eq2_h3 mem_Collect_eq object_ptr_kinds_M_eq_h3 select_result_eq split_cong unfolding CD.parent_child_rel_def by(simp) qed then have " CD.a_acyclic_heap h'" using \ CD.a_acyclic_heap h2\ CD.acyclic_heap_def acyclic_subset by blast moreover have " CD.a_all_ptrs_in_heap h2" using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_all_ptrs_in_heap h3" apply(auto simp add: CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq3_h2 children_eq_h2)[1] apply (metis \type_wf h'\ children_eq2_h3 children_eq_h2 children_eq_h3 known_ptrs l_heap_is_wellformed.heap_is_wellformed_children_in_heap local.get_child_nodes_ok local.known_ptrs_known_ptr local.l_heap_is_wellformed_axioms node_ptr_kinds_eq3_h2 object_ptr_kinds_h2_eq3 object_ptr_kinds_h3_eq3 object_ptr_kinds_h_eq3 returns_result_select_result wellformed_h2) by (metis (no_types, hide_lams) disc_nodes_old_document_h2 disc_nodes_old_document_h3 disconnected_nodes_eq2_h2 document_ptr_kinds_eq3_h2 finite_set_in select_result_I2 set_remove1_subset subsetD) then have "CD.a_all_ptrs_in_heap h'" apply(auto simp add: CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq3_h3 children_eq_h3)[1] apply (metis (no_types, hide_lams) children_eq2_h3 finite_set_in object_ptr_kinds_h3_eq3 subsetD) by (metis (no_types, hide_lams) \child \ set disc_nodes_old_document_h2\ disc_nodes_document_ptr_h' disc_nodes_document_ptr_h2 disc_nodes_old_document_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq3_h3 finite_set_in local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq3_h2 node_ptr_kinds_eq3_h3 select_result_I2 set_ConsD subsetD wellformed_h2) moreover have "CD.a_owner_document_valid h2" using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_owner_document_valid h'" apply(simp add: CD.a_owner_document_valid_def node_ptr_kinds_eq_h2 node_ptr_kinds_eq3_h3 object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 children_eq2_h2 children_eq2_h3 ) by (metis (no_types) disc_nodes_document_ptr_h' disc_nodes_document_ptr_h2 disc_nodes_old_document_h2 disc_nodes_old_document_h3 disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_in_heap document_ptr_kinds_eq3_h2 document_ptr_kinds_eq3_h3 in_set_remove1 list.set_intros(1) list.set_intros(2) node_ptr_kinds_eq3_h2 node_ptr_kinds_eq3_h3 object_ptr_kinds_h2_eq3 object_ptr_kinds_h3_eq3 select_result_I2) have a_distinct_lists_h2: "CD.a_distinct_lists h2" using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_distinct_lists h'" apply(auto simp add: CD.a_distinct_lists_def object_ptr_kinds_eq_h3 object_ptr_kinds_eq_h2 children_eq2_h2 children_eq2_h3)[1] proof - assume 1: "distinct (concat (map (\ptr. |h' \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h')))))" and 2: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h2)))))" and 3: "(\x\fset (object_ptr_kinds h'). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h2). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" show "distinct (concat (map (\document_ptr. |h' \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h')))))" proof(rule distinct_concat_map_I) show "distinct (sorted_list_of_set (fset (document_ptr_kinds h')))" by(auto simp add: document_ptr_kinds_M_def ) next fix x assume a1: "x \ set (sorted_list_of_set (fset (document_ptr_kinds h')))" have 4: "distinct |h2 \ get_disconnected_nodes x|\<^sub>r" using a_distinct_lists_h2 "2" a1 concat_map_all_distinct document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 by fastforce then show "distinct |h' \ get_disconnected_nodes x|\<^sub>r" proof (cases "old_document \ x") case True then show ?thesis proof (cases "document_ptr \ x") case True then show ?thesis using disconnected_nodes_eq2_h2[OF \old_document \ x\] disconnected_nodes_eq2_h3[OF \document_ptr \ x\] 4 by(auto) next case False then show ?thesis using disc_nodes_document_ptr_h3 disc_nodes_document_ptr_h' 4 \child \ set disc_nodes_document_ptr_h3\ by(auto simp add: disconnected_nodes_eq2_h2[OF \old_document \ x\] ) qed next case False then show ?thesis by (metis (no_types, hide_lams) \distinct disc_nodes_old_document_h2\ disc_nodes_old_document_h3 disconnected_nodes_eq2_h3 distinct_remove1 docs_neq select_result_I2) qed next fix x y assume a0: "x \ set (sorted_list_of_set (fset (document_ptr_kinds h')))" and a1: "y \ set (sorted_list_of_set (fset (document_ptr_kinds h')))" and a2: "x \ y" moreover have 5: "set |h2 \ get_disconnected_nodes x|\<^sub>r \ set |h2 \ get_disconnected_nodes y|\<^sub>r = {}" using 2 calculation by (auto simp add: document_ptr_kinds_eq3_h2 document_ptr_kinds_eq3_h3 dest: distinct_concat_map_E(1)) ultimately show "set |h' \ get_disconnected_nodes x|\<^sub>r \ set |h' \ get_disconnected_nodes y|\<^sub>r = {}" proof(cases "old_document = x") case True have "old_document \ y" using \x \ y\ \old_document = x\ by simp have "document_ptr \ x" using docs_neq \old_document = x\ by auto show ?thesis proof(cases "document_ptr = y") case True then show ?thesis using 5 True select_result_I2[OF disc_nodes_document_ptr_h'] select_result_I2[OF disc_nodes_document_ptr_h2] select_result_I2[OF disc_nodes_old_document_h2] select_result_I2[OF disc_nodes_old_document_h3] \old_document = x\ by (metis (no_types, lifting) \child \ set (remove1 child disc_nodes_old_document_h2)\ \document_ptr \ x\ disconnected_nodes_eq2_h3 disjoint_iff_not_equal notin_set_remove1 set_ConsD) next case False then show ?thesis using 5 select_result_I2[OF disc_nodes_document_ptr_h'] select_result_I2[OF disc_nodes_document_ptr_h2] select_result_I2[OF disc_nodes_old_document_h2] select_result_I2[OF disc_nodes_old_document_h3] disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 \old_document = x\ docs_neq \old_document \ y\ by (metis (no_types, lifting) disjoint_iff_not_equal notin_set_remove1) qed next case False then show ?thesis proof(cases "old_document = y") case True then show ?thesis proof(cases "document_ptr = x") case True show ?thesis using 5 select_result_I2[OF disc_nodes_document_ptr_h'] select_result_I2[OF disc_nodes_document_ptr_h2] select_result_I2[OF disc_nodes_old_document_h2] select_result_I2[OF disc_nodes_old_document_h3] \old_document \ x\ \old_document = y\ \document_ptr = x\ apply(simp) by (metis (no_types, lifting) \child \ set (remove1 child disc_nodes_old_document_h2)\ disconnected_nodes_eq2_h3 disjoint_iff_not_equal notin_set_remove1) next case False then show ?thesis using 5 select_result_I2[OF disc_nodes_document_ptr_h'] select_result_I2[OF disc_nodes_document_ptr_h2] select_result_I2[OF disc_nodes_old_document_h2] select_result_I2[OF disc_nodes_old_document_h3] \old_document \ x\ \old_document = y\ \document_ptr \ x\ by (metis (no_types, lifting) disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disjoint_iff_not_equal docs_neq notin_set_remove1) qed next case False have "set |h2 \ get_disconnected_nodes y|\<^sub>r \ set disc_nodes_old_document_h2 = {}" by (metis DocumentMonad.ptr_kinds_M_ok DocumentMonad.ptr_kinds_M_ptr_kinds False \type_wf h2\ a1 disc_nodes_old_document_h2 document_ptr_kinds_M_def document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 l_ptr_kinds_M.ptr_kinds_ptr_kinds_M local.get_disconnected_nodes_ok local.heap_is_wellformed_one_disc_parent returns_result_select_result wellformed_h2) then show ?thesis proof(cases "document_ptr = x") case True then have "document_ptr \ y" using \x \ y\ by auto have "set |h2 \ get_disconnected_nodes y|\<^sub>r \ set disc_nodes_old_document_h2 = {}" using \set |h2 \ get_disconnected_nodes y|\<^sub>r \ set disc_nodes_old_document_h2 = {}\ by blast then show ?thesis using 5 select_result_I2[OF disc_nodes_document_ptr_h'] select_result_I2[OF disc_nodes_document_ptr_h2] select_result_I2[OF disc_nodes_old_document_h2] select_result_I2[OF disc_nodes_old_document_h3] \old_document \ x\ \old_document \ y\ \document_ptr = x\ \document_ptr \ y\ \child \ set disc_nodes_old_document_h2\ disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 \set |h2 \ get_disconnected_nodes y|\<^sub>r \ set disc_nodes_old_document_h2 = {}\ by(auto) next case False then show ?thesis proof(cases "document_ptr = y") case True have f1: "set |h2 \ get_disconnected_nodes x|\<^sub>r \ set disc_nodes_document_ptr_h3 = {}" using 2 a1 document_ptr_in_heap document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 \document_ptr \ x\ select_result_I2[OF disc_nodes_document_ptr_h3, symmetric] disconnected_nodes_eq2_h2[OF docs_neq[symmetric], symmetric] by (simp add: "5" True) moreover have f1: "set |h2 \ get_disconnected_nodes x|\<^sub>r \ set |h2 \ get_disconnected_nodes old_document|\<^sub>r = {}" using 2 a1 old_document_in_heap document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 \old_document \ x\ by (metis (no_types, lifting) a0 distinct_concat_map_E(1) document_ptr_kinds_eq3_h2 document_ptr_kinds_eq3_h3 finite_fset fmember.rep_eq set_sorted_list_of_set) ultimately show ?thesis using 5 select_result_I2[OF disc_nodes_document_ptr_h'] select_result_I2[OF disc_nodes_old_document_h2] \old_document \ x\ \document_ptr \ x\ \document_ptr = y\ \child \ set disc_nodes_old_document_h2\ disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 by auto next case False then show ?thesis using 5 select_result_I2[OF disc_nodes_old_document_h2] \old_document \ x\ \document_ptr \ x\ \document_ptr \ y\ \child \ set disc_nodes_old_document_h2\ disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 by (metis \set |h2 \ get_disconnected_nodes y|\<^sub>r \ set disc_nodes_old_document_h2 = {}\ empty_iff inf.idem) qed qed qed qed qed next fix x xa xb assume 0: "distinct (concat (map (\ptr. |h' \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h')))))" and 1: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h2)))))" and 2: "(\x\fset (object_ptr_kinds h'). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h2). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" and 3: "xa |\| object_ptr_kinds h'" and 4: "x \ set |h' \ get_child_nodes xa|\<^sub>r" and 5: "xb |\| document_ptr_kinds h'" and 6: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" then show False using \child \ set disc_nodes_old_document_h2\ disc_nodes_document_ptr_h' disc_nodes_document_ptr_h2 disc_nodes_old_document_h2 disc_nodes_old_document_h3 disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 old_document_in_heap apply(auto)[1] apply(cases "xb = old_document") proof - assume a1: "xb = old_document" assume a2: "h2 \ get_disconnected_nodes old_document \\<^sub>r disc_nodes_old_document_h2" assume a3: "h3 \ get_disconnected_nodes old_document \\<^sub>r remove1 child disc_nodes_old_document_h2" assume a4: "x \ set |h' \ get_child_nodes xa|\<^sub>r" assume "document_ptr_kinds h2 = document_ptr_kinds h'" assume a5: "(\x\fset (object_ptr_kinds h'). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h'). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" have f6: "old_document |\| document_ptr_kinds h'" using a1 \xb |\| document_ptr_kinds h'\ by blast have f7: "|h2 \ get_disconnected_nodes old_document|\<^sub>r = disc_nodes_old_document_h2" using a2 by simp have "x \ set disc_nodes_old_document_h2" using f6 a3 a1 by (metis (no_types) \type_wf h'\ \x \ set |h' \ get_disconnected_nodes xb|\<^sub>r\ disconnected_nodes_eq_h3 docs_neq get_disconnected_nodes_ok returns_result_eq returns_result_select_result set_remove1_subset subsetCE) then have "set |h' \ get_child_nodes xa|\<^sub>r \ set |h2 \ get_disconnected_nodes xb|\<^sub>r = {}" using f7 f6 a5 a4 \xa |\| object_ptr_kinds h'\ by fastforce then show ?thesis using \x \ set disc_nodes_old_document_h2\ a1 a4 f7 by blast next assume a1: "xb \ old_document" assume a2: "h2 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_document_ptr_h3" assume a3: "h2 \ get_disconnected_nodes old_document \\<^sub>r disc_nodes_old_document_h2" assume a4: "xa |\| object_ptr_kinds h'" assume a5: "h' \ get_disconnected_nodes document_ptr \\<^sub>r child # disc_nodes_document_ptr_h3" assume a6: "old_document |\| document_ptr_kinds h'" assume a7: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" assume a8: "x \ set |h' \ get_child_nodes xa|\<^sub>r" assume a9: "document_ptr_kinds h2 = document_ptr_kinds h'" assume a10: "\doc_ptr. old_document \ doc_ptr \ |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" assume a11: "\doc_ptr. document_ptr \ doc_ptr \ |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" assume a12: "(\x\fset (object_ptr_kinds h'). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h'). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" have f13: "\d. d \ set |h' \ document_ptr_kinds_M|\<^sub>r \ h2 \ ok get_disconnected_nodes d" using a9 \type_wf h2\ get_disconnected_nodes_ok by simp then have f14: "|h2 \ get_disconnected_nodes old_document|\<^sub>r = disc_nodes_old_document_h2" using a6 a3 by simp have "x \ set |h2 \ get_disconnected_nodes xb|\<^sub>r" using a12 a8 a4 \xb |\| document_ptr_kinds h'\ by (meson UN_I disjoint_iff_not_equal fmember.rep_eq) then have "x = child" using f13 a11 a10 a7 a5 a2 a1 by (metis (no_types, lifting) select_result_I2 set_ConsD) then have "child \ set disc_nodes_old_document_h2" using f14 a12 a8 a6 a4 by (metis \type_wf h'\ adopt_node_removes_child assms(1) assms(2) type_wf get_child_nodes_ok known_ptrs local.known_ptrs_known_ptr object_ptr_kinds_h2_eq3 object_ptr_kinds_h3_eq3 object_ptr_kinds_h_eq3 returns_result_select_result) then show ?thesis using \child \ set disc_nodes_old_document_h2\ by fastforce qed qed ultimately have "heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'" using \CD.a_owner_document_valid h'\ CD.heap_is_wellformed_def by simp have shadow_root_eq_h2: "\ptr' shadow_root_ptr_opt. h2 \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt = h3 \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt" using get_shadow_root_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by(auto simp add: adopt_node_locs_def remove_child_locs_def set_child_nodes_get_shadow_root set_disconnected_nodes_get_shadow_root) then have shadow_root_eq2_h2: "\ptr'. |h2 \ get_shadow_root ptr'|\<^sub>r = |h3 \ get_shadow_root ptr'|\<^sub>r" by (meson select_result_eq) have shadow_root_eq_h3: "\ptr' shadow_root_ptr_opt. h3 \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt = h' \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt" using get_shadow_root_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: adopt_node_locs_def remove_child_locs_def set_child_nodes_get_shadow_root set_disconnected_nodes_get_shadow_root) then have shadow_root_eq2_h3: "\ptr'. |h3 \ get_shadow_root ptr'|\<^sub>r = |h' \ get_shadow_root ptr'|\<^sub>r" by (meson select_result_eq) have tag_name_eq_h2: "\ptr' tag. h2 \ get_tag_name ptr' \\<^sub>r tag = h3 \ get_tag_name ptr' \\<^sub>r tag" using get_tag_name_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by(auto simp add: adopt_node_locs_def remove_child_locs_def set_child_nodes_get_tag_name set_disconnected_nodes_get_tag_name) then have tag_name_eq2_h2: "\ptr'. |h2 \ get_tag_name ptr'|\<^sub>r = |h3 \ get_tag_name ptr'|\<^sub>r" by (meson select_result_eq) have tag_name_eq_h3: "\ptr' tag. h3 \ get_tag_name ptr' \\<^sub>r tag = h' \ get_tag_name ptr' \\<^sub>r tag" using get_tag_name_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: adopt_node_locs_def remove_child_locs_def set_child_nodes_get_tag_name set_disconnected_nodes_get_tag_name) then have tag_name_eq2_h3: "\ptr'. |h3 \ get_tag_name ptr'|\<^sub>r = |h' \ get_tag_name ptr'|\<^sub>r" by (meson select_result_eq) have object_ptr_kinds_eq_h2: "object_ptr_kinds h2 = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes h3]) unfolding adopt_node_locs_def remove_child_locs_def using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def split: if_splits) have object_ptr_kinds_eq_h3: "object_ptr_kinds h3 = object_ptr_kinds h'" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes h']) unfolding adopt_node_locs_def remove_child_locs_def using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def split: if_splits) have shadow_root_ptr_kinds_eq_h2: "shadow_root_ptr_kinds h2 = shadow_root_ptr_kinds h3" using object_ptr_kinds_eq_h2 by(auto simp add: shadow_root_ptr_kinds_def) have element_ptr_kinds_eq_h2: "element_ptr_kinds h2 = element_ptr_kinds h3" using object_ptr_kinds_eq_h2 by(auto simp add: element_ptr_kinds_def node_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h3: "shadow_root_ptr_kinds h3 = shadow_root_ptr_kinds h'" using object_ptr_kinds_eq_h3 by(auto simp add: shadow_root_ptr_kinds_def) have element_ptr_kinds_eq_h3: "element_ptr_kinds h3 = element_ptr_kinds h'" using object_ptr_kinds_eq_h3 by(auto simp add: element_ptr_kinds_def node_ptr_kinds_def) have "known_ptrs h3" using known_ptrs local.known_ptrs_preserved object_ptr_kinds_h2_eq3 object_ptr_kinds_h_eq3 by blast then have "known_ptrs h'" using local.known_ptrs_preserved object_ptr_kinds_h3_eq3 by blast show "heap_is_wellformed h' \ known_ptrs h' \ type_wf h'" using \heap_is_wellformed h2\ using \heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'\ \known_ptrs h'\ \type_wf h'\ using \parent_child_rel h' \ parent_child_rel h2\ apply(auto simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def a_host_shadow_root_rel_def a_all_ptrs_in_heap_def a_distinct_lists_def a_shadow_root_valid_def object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 element_ptr_kinds_eq_h2 element_ptr_kinds_eq_h3 shadow_root_ptr_kinds_eq_h2 shadow_root_ptr_kinds_eq_h3 shadow_root_eq_h2 shadow_root_eq_h3 shadow_root_eq2_h2 shadow_root_eq2_h3 tag_name_eq_h2 tag_name_eq_h3 tag_name_eq2_h2 tag_name_eq2_h3 CD.parent_child_rel_def children_eq2_h2 children_eq2_h3 object_ptr_kinds_h2_eq3 object_ptr_kinds_h3_eq3)[1] done qed then show "heap_is_wellformed h'" and "known_ptrs h'" and "type_wf h'" by auto qed lemma adopt_node_node_in_disconnected_nodes: assumes wellformed: "heap_is_wellformed h" and adopt_node: "h \ adopt_node owner_document node_ptr \\<^sub>h h'" and "h' \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "node_ptr \ set disc_nodes" proof - obtain old_document parent_opt h2 where old_document: "h \ get_owner_document (cast node_ptr) \\<^sub>r old_document" and parent_opt: "h \ get_parent node_ptr \\<^sub>r parent_opt" and h2: "h \ (case parent_opt of Some parent \ remove_child parent node_ptr | None \ return ()) \\<^sub>h h2" and h': "h2 \ (if owner_document \ old_document then do { old_disc_nodes \ get_disconnected_nodes old_document; set_disconnected_nodes old_document (remove1 node_ptr old_disc_nodes); disc_nodes \ get_disconnected_nodes owner_document; set_disconnected_nodes owner_document (node_ptr # disc_nodes) } else do { return () }) \\<^sub>h h'" using assms(2) by(auto simp add: adopt_node_def elim!: bind_returns_heap_E dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure] pure_returns_heap_eq[rotated, OF get_parent_pure]) show ?thesis proof (cases "owner_document = old_document") case True then show ?thesis proof (insert parent_opt h2, induct parent_opt) case None then have "h = h'" using h2 h' by(auto) then show ?case using in_disconnected_nodes_no_parent assms None old_document by blast next case (Some parent) then show ?case using remove_child_in_disconnected_nodes known_ptrs True h' assms(3) old_document by auto qed next case False then show ?thesis using assms(3) h' list.set_intros(1) select_result_I2 set_disconnected_nodes_get_disconnected_nodes apply(auto elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated])[1] proof - fix x and h'a and xb assume a1: "h' \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes" assume a2: "\h document_ptr disc_nodes h'. h \ set_disconnected_nodes document_ptr disc_nodes \\<^sub>h h' \ h' \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes" assume "h'a \ set_disconnected_nodes owner_document (node_ptr # xb) \\<^sub>h h'" then have "node_ptr # xb = disc_nodes" using a2 a1 by (meson returns_result_eq) then show ?thesis by (meson list.set_intros(1)) qed qed qed end interpretation l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M Shadow_DOM.get_owner_document Shadow_DOM.get_parent Shadow_DOM.get_parent_locs Shadow_DOM.remove_child Shadow_DOM.remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs Shadow_DOM.adopt_node Shadow_DOM.adopt_node_locs ShadowRootClass.known_ptr ShadowRootClass.type_wf Shadow_DOM.get_child_nodes Shadow_DOM.get_child_nodes_locs ShadowRootClass.known_ptrs Shadow_DOM.set_child_nodes Shadow_DOM.set_child_nodes_locs Shadow_DOM.remove Shadow_DOM.heap_is_wellformed Shadow_DOM.parent_child_rel by(auto simp add: l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] interpretation i_adopt_node_wf2?: l_adopt_node_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs set_child_nodes set_child_nodes_locs get_shadow_root get_shadow_root_locs set_disconnected_nodes set_disconnected_nodes_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs get_root_node get_root_node_locs get_parent get_parent_locs known_ptrs get_owner_document remove_child remove_child_locs remove adopt_node adopt_node_locs by(auto simp add: l_adopt_node_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_adopt_node_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma adopt_node_wf_is_l_adopt_node_wf [instances]: "l_adopt_node_wf type_wf known_ptr heap_is_wellformed parent_child_rel get_child_nodes get_disconnected_nodes known_ptrs adopt_node" apply(auto simp add: l_adopt_node_wf_def l_adopt_node_wf_axioms_def instances)[1] using adopt_node_preserves_wellformedness apply blast using adopt_node_removes_child apply blast using adopt_node_node_in_disconnected_nodes apply blast using adopt_node_removes_first_child apply blast using adopt_node_document_in_heap apply blast using adopt_node_preserves_wellformedness apply blast using adopt_node_preserves_wellformedness apply blast done subsubsection \insert\_before\ locale l_insert_before_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes + l_get_disconnected_nodes + l_set_child_nodes_get_shadow_root + l_set_disconnected_nodes_get_shadow_root + l_set_child_nodes_get_tag_name + l_set_disconnected_nodes_get_tag_name + l_set_disconnected_nodes_get_disconnected_nodes + l_set_child_nodes_get_disconnected_nodes + l_set_disconnected_nodes_get_disconnected_nodes_wf + l_set_disconnected_nodes_get_ancestors_si + l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M _ _ _ _ _ _ get_ancestors_si get_ancestors_si_locs + l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_owner_document + l_adopt_node + l_adopt_node_wf + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_adopt_node_get_shadow_root begin lemma insert_before_child_preserves: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ insert_before ptr node child \\<^sub>h h'" shows "type_wf h'" and "known_ptrs h'" and "heap_is_wellformed h'" proof - obtain ancestors reference_child owner_document h2 h3 disconnected_nodes_h2 where ancestors: "h \ get_ancestors_si ptr \\<^sub>r ancestors" and node_not_in_ancestors: "cast node \ set ancestors" and reference_child: "h \ (if Some node = child then a_next_sibling node else return child) \\<^sub>r reference_child" and owner_document: "h \ get_owner_document ptr \\<^sub>r owner_document" and h2: "h \ adopt_node owner_document node \\<^sub>h h2" and disconnected_nodes_h2: "h2 \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h2" and h3: "h2 \ set_disconnected_nodes owner_document (remove1 node disconnected_nodes_h2) \\<^sub>h h3" and h': "h3 \ a_insert_node ptr node reference_child \\<^sub>h h'" (* children: "h3 \ get_child_nodes ptr \\<^sub>r children" and *) (* h': "h3 \ set_child_nodes ptr (insert_before_list node reference_child children) \\<^sub>h h'" *) using assms(4) by(auto simp add: insert_before_def a_ensure_pre_insertion_validity_def elim!: bind_returns_heap_E bind_returns_result_E bind_returns_heap_E2[rotated, OF get_parent_pure, rotated] bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated] bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] bind_returns_heap_E2[rotated, OF get_ancestors_pure, rotated] bind_returns_heap_E2[rotated, OF next_sibling_pure, rotated] bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated] split: if_splits option.splits) have "type_wf h2" using \type_wf h\ using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF adopt_node_writes h2] using adopt_node_types_preserved by(auto simp add: reflp_def transp_def) then have "type_wf h3" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h3] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) then show "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF insert_node_writes h'] using set_child_nodes_types_preserved by(auto simp add: reflp_def transp_def) have "object_ptr_kinds h = object_ptr_kinds h2" using adopt_node_writes h2 apply(rule writes_small_big) using adopt_node_pointers_preserved by(auto simp add: reflp_def transp_def) moreover have "\ = object_ptr_kinds h3" using set_disconnected_nodes_writes h3 apply(rule writes_small_big) using set_disconnected_nodes_pointers_preserved by(auto simp add: reflp_def transp_def) moreover have "\ = object_ptr_kinds h'" using insert_node_writes h' apply(rule writes_small_big) using set_child_nodes_pointers_preserved by(auto simp add: reflp_def transp_def) ultimately show "known_ptrs h'" using \known_ptrs h\ known_ptrs_preserved by blast have "known_ptrs h2" using \known_ptrs h\ known_ptrs_preserved \object_ptr_kinds h = object_ptr_kinds h2\ by blast then have "known_ptrs h3" using known_ptrs_preserved \object_ptr_kinds h2 = object_ptr_kinds h3\ by blast have "known_ptr ptr" by (meson get_owner_document_ptr_in_heap is_OK_returns_result_I \known_ptrs h\ l_known_ptrs.known_ptrs_known_ptr l_known_ptrs_axioms owner_document) have object_ptr_kinds_M_eq3_h: "object_ptr_kinds h = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF adopt_node_writes h2]) using adopt_node_pointers_preserved apply blast by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h: "\ptrs. h \ object_ptr_kinds_M \\<^sub>r ptrs = h2 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs ) then have object_ptr_kinds_M_eq2_h: "|h \ object_ptr_kinds_M|\<^sub>r = |h2 \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h: "|h \ node_ptr_kinds_M|\<^sub>r = |h2 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have wellformed_h2: "heap_is_wellformed h2" using adopt_node_preserves_wellformedness[OF \heap_is_wellformed h\ h2] \known_ptrs h\ \type_wf h\ . have object_ptr_kinds_M_eq3_h2: "object_ptr_kinds h2 = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes h3]) unfolding a_remove_child_locs_def using set_disconnected_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h2: "\ptrs. h2 \ object_ptr_kinds_M \\<^sub>r ptrs = h3 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_M_eq2_h2: "|h2 \ object_ptr_kinds_M|\<^sub>r = |h3 \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h2: "|h2 \ node_ptr_kinds_M|\<^sub>r = |h3 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have document_ptr_kinds_eq2_h2: "|h2 \ document_ptr_kinds_M|\<^sub>r = |h3 \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_M_eq2_h2 document_ptr_kinds_M_eq by auto have object_ptr_kinds_M_eq3_h': "object_ptr_kinds h3 = object_ptr_kinds h'" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF insert_node_writes h']) unfolding a_remove_child_locs_def using set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h3: "\ptrs. h3 \ object_ptr_kinds_M \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_M_eq2_h3: "|h3 \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h3: "|h3 \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have document_ptr_kinds_eq2_h3: "|h3 \ document_ptr_kinds_M|\<^sub>r = |h' \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_M_eq2_h3 document_ptr_kinds_M_eq by auto have shadow_root_eq_h2: "\ptr' shadow_root. h \ get_shadow_root ptr' \\<^sub>r shadow_root = h2 \ get_shadow_root ptr' \\<^sub>r shadow_root" using get_shadow_root_reads adopt_node_writes h2 apply(rule reads_writes_preserved) using local.adopt_node_get_shadow_root by blast have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. owner_document \ doc_ptr \ h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by (auto simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers) then have disconnected_nodes_eq2_h2: "\doc_ptr. doc_ptr \ owner_document \ |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have disconnected_nodes_h3: "h3 \ get_disconnected_nodes owner_document \\<^sub>r remove1 node disconnected_nodes_h2" using h3 set_disconnected_nodes_get_disconnected_nodes by blast have disconnected_nodes_eq_h3: "\doc_ptr disc_nodes. h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads insert_node_writes h' apply(rule reads_writes_preserved) using set_child_nodes_get_disconnected_nodes by fast then have disconnected_nodes_eq2_h3: "\doc_ptr. |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. h2 \ get_child_nodes ptr' \\<^sub>r children = h3 \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by (auto simp add: set_disconnected_nodes_get_child_nodes) then have children_eq2_h2: "\ptr'. |h2 \ get_child_nodes ptr'|\<^sub>r = |h3 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have children_eq_h3: "\ptr' children. ptr \ ptr' \ h3 \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads insert_node_writes h' apply(rule reads_writes_preserved) by (auto simp add: set_child_nodes_get_child_nodes_different_pointers) then have children_eq2_h3: "\ptr'. ptr \ ptr' \ |h3 \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force obtain children_h3 where children_h3: "h3 \ get_child_nodes ptr \\<^sub>r children_h3" using h' a_insert_node_def by auto have children_h': "h' \ get_child_nodes ptr \\<^sub>r insert_before_list node reference_child children_h3" using h' \type_wf h3\ \known_ptr ptr\ by(auto simp add: a_insert_node_def elim!: bind_returns_heap_E2 dest!: set_child_nodes_get_child_nodes returns_result_eq[OF children_h3]) have ptr_in_heap: "ptr |\| object_ptr_kinds h3" using children_h3 get_child_nodes_ptr_in_heap by blast have node_in_heap: "node |\| node_ptr_kinds h" using h2 adopt_node_child_in_heap by fast have child_not_in_any_children: "\p children. h2 \ get_child_nodes p \\<^sub>r children \ node \ set children" using \heap_is_wellformed h\ h2 adopt_node_removes_child \type_wf h\ \known_ptrs h\ by auto have "node \ set disconnected_nodes_h2" using disconnected_nodes_h2 h2 adopt_node_node_in_disconnected_nodes assms(1) \type_wf h\ \known_ptrs h\ by blast have node_not_in_disconnected_nodes: "\d. d |\| document_ptr_kinds h3 \ node \ set |h3 \ get_disconnected_nodes d|\<^sub>r" proof - fix d assume "d |\| document_ptr_kinds h3" show "node \ set |h3 \ get_disconnected_nodes d|\<^sub>r" proof (cases "d = owner_document") case True then show ?thesis using disconnected_nodes_h2 wellformed_h2 h3 remove_from_disconnected_nodes_removes wellformed_h2 \d |\| document_ptr_kinds h3\ disconnected_nodes_h3 by fastforce next case False then have "set |h2 \ get_disconnected_nodes d|\<^sub>r \ set |h2 \ get_disconnected_nodes owner_document|\<^sub>r = {}" using distinct_concat_map_E(1) wellformed_h2 by (metis (no_types, lifting) \d |\| document_ptr_kinds h3\ \type_wf h2\ disconnected_nodes_h2 document_ptr_kinds_M_def document_ptr_kinds_eq2_h2 l_ptr_kinds_M.ptr_kinds_ptr_kinds_M local.get_disconnected_nodes_ok local.heap_is_wellformed_one_disc_parent returns_result_select_result select_result_I2) then show ?thesis using disconnected_nodes_eq2_h2[OF False] \node \ set disconnected_nodes_h2\ disconnected_nodes_h2 by fastforce qed qed have "cast node \ ptr" using ancestors node_not_in_ancestors get_ancestors_ptr by fast obtain ancestors_h2 where ancestors_h2: "h2 \ get_ancestors_si ptr \\<^sub>r ancestors_h2" using get_ancestors_si_ok by (metis \known_ptrs h2\ \type_wf h2\ is_OK_returns_result_E object_ptr_kinds_M_eq3_h2 ptr_in_heap wellformed_h2) have ancestors_h3: "h3 \ get_ancestors_si ptr \\<^sub>r ancestors_h2" using get_ancestors_si_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_separate_forwards) using \heap_is_wellformed h2\ apply simp using ancestors_h2 apply simp apply(auto simp add: get_ancestors_si_locs_def get_parent_locs_def)[1] apply (simp add: local.get_ancestors_si_locs_def local.get_parent_reads_pointers local.set_disconnected_nodes_get_ancestors_si) using local.get_ancestors_si_locs_def local.set_disconnected_nodes_get_ancestors_si by blast have node_not_in_ancestors_h2: "cast node \ set ancestors_h2" using \heap_is_wellformed h\ \heap_is_wellformed h2\ ancestors ancestors_h2 apply(rule get_ancestors_si_remains_not_in_ancestors) using assms(2) assms(3) h2 local.adopt_node_children_subset apply blast using shadow_root_eq_h2 node_not_in_ancestors object_ptr_kinds_M_eq2_h assms(2) assms(3) \type_wf h2\ by(auto dest: returns_result_eq) moreover have "parent_child_rel h2 = parent_child_rel h3" by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_M_eq3_h2 children_eq2_h2) have "parent_child_rel h' = insert (ptr, cast node) ((parent_child_rel h3))" using children_h3 children_h' ptr_in_heap apply(auto simp add: CD.parent_child_rel_def object_ptr_kinds_M_eq3_h' children_eq2_h3 insert_before_list_node_in_set)[1] apply (metis (no_types, lifting) children_eq2_h3 insert_before_list_in_set select_result_I2) by (metis (no_types, lifting) children_eq2_h3 imageI insert_before_list_in_set select_result_I2) have "CD.a_acyclic_heap h'" proof - have "acyclic (parent_child_rel h2)" using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def) then have "acyclic (parent_child_rel h3)" by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_M_eq3_h2 children_eq2_h2) moreover have "cast node \ {x. (x, ptr) \ (parent_child_rel h2)\<^sup>*}" using get_ancestors_si_parent_child_rel using \known_ptrs h2\ \type_wf h2\ ancestors_h2 node_not_in_ancestors_h2 wellformed_h2 by blast then have "cast node \ {x. (x, ptr) \ (parent_child_rel h3)\<^sup>*}" by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_M_eq3_h2 children_eq2_h2) ultimately show ?thesis using \parent_child_rel h' = insert (ptr, cast node) ((parent_child_rel h3))\ by(auto simp add: CD.acyclic_heap_def) qed moreover have "CD.a_all_ptrs_in_heap h2" using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) have "CD.a_all_ptrs_in_heap h'" proof - have "CD.a_all_ptrs_in_heap h3" using \CD.a_all_ptrs_in_heap h2\ apply(auto simp add: CD.a_all_ptrs_in_heap_def object_ptr_kinds_M_eq2_h2 node_ptr_kinds_eq2_h2 children_eq_h2)[1] using disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3 using node_ptr_kinds_eq2_h2 apply auto[1] apply (metis (no_types, lifting) children_eq2_h2 in_mono notin_fset object_ptr_kinds_M_eq3_h2) by (metis (no_types, hide_lams) NodeMonad.ptr_kinds_ptr_kinds_M disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3 document_ptr_kinds_commutes finite_set_in node_ptr_kinds_eq2_h2 object_ptr_kinds_M_eq3_h2 select_result_I2 set_remove1_subset subsetD) have "set children_h3 \ set |h' \ node_ptr_kinds_M|\<^sub>r" using children_h3 \CD.a_all_ptrs_in_heap h3\ apply(auto simp add: CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq2_h3)[1] using CD.parent_child_rel_child_nodes2 \known_ptr ptr\ \parent_child_rel h2 = parent_child_rel h3\ \type_wf h2\ local.parent_child_rel_child_in_heap node_ptr_kinds_commutes object_ptr_kinds_M_eq3_h' object_ptr_kinds_M_eq3_h2 wellformed_h2 by blast then have "set (insert_before_list node reference_child children_h3) \ set |h' \ node_ptr_kinds_M|\<^sub>r" using node_in_heap apply(auto simp add: node_ptr_kinds_eq2_h node_ptr_kinds_eq2_h2 node_ptr_kinds_eq2_h3)[1] by (metis (no_types, hide_lams) contra_subsetD finite_set_in insert_before_list_in_set node_ptr_kinds_commutes object_ptr_kinds_M_eq3_h object_ptr_kinds_M_eq3_h' object_ptr_kinds_M_eq3_h2) then show ?thesis using \CD.a_all_ptrs_in_heap h3\ apply(auto simp add: object_ptr_kinds_M_eq3_h' CD.a_all_ptrs_in_heap_def node_ptr_kinds_def node_ptr_kinds_eq2_h3 disconnected_nodes_eq_h3)[1] using children_eq_h3 children_h' apply (metis (no_types, lifting) children_eq2_h3 finite_set_in select_result_I2 subsetD) by (metis (no_types, lifting) DocumentMonad.ptr_kinds_ptr_kinds_M disconnected_nodes_eq2_h3 document_ptr_kinds_eq2_h3 finite_set_in subsetD) qed moreover have "CD.a_distinct_lists h2" using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def ) then have "CD.a_distinct_lists h3" proof(auto simp add: CD.a_distinct_lists_def object_ptr_kinds_M_eq2_h2 document_ptr_kinds_eq2_h2 children_eq2_h2 intro!: distinct_concat_map_I) fix x assume 1: "x |\| document_ptr_kinds h3" and 2: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" show "distinct |h3 \ get_disconnected_nodes x|\<^sub>r" using distinct_concat_map_E(2)[OF 2] select_result_I2[OF disconnected_nodes_h3] disconnected_nodes_eq2_h2 select_result_I2[OF disconnected_nodes_h2] 1 by (metis (full_types) distinct_remove1 finite_fset fmember.rep_eq set_sorted_list_of_set) next fix x y xa assume 1: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" and 2: "x |\| document_ptr_kinds h3" and 3: "y |\| document_ptr_kinds h3" and 4: "x \ y" and 5: "xa \ set |h3 \ get_disconnected_nodes x|\<^sub>r" and 6: "xa \ set |h3 \ get_disconnected_nodes y|\<^sub>r" show False proof (cases "x = owner_document") case True then have "y \ owner_document" using 4 by simp show ?thesis using distinct_concat_map_E(1)[OF 1] using 2 3 4 5 6 select_result_I2[OF disconnected_nodes_h3] select_result_I2[OF disconnected_nodes_h2] apply(auto simp add: True disconnected_nodes_eq2_h2[OF \y \ owner_document\])[1] by (metis (no_types, hide_lams) disconnected_nodes_eq2_h2 disjoint_iff_not_equal notin_set_remove1) next case False then show ?thesis proof (cases "y = owner_document") case True then show ?thesis using distinct_concat_map_E(1)[OF 1] using 2 3 4 5 6 select_result_I2[OF disconnected_nodes_h3] select_result_I2[OF disconnected_nodes_h2] apply(auto simp add: True disconnected_nodes_eq2_h2[OF \x \ owner_document\])[1] by (metis (no_types, hide_lams) disconnected_nodes_eq2_h2 disjoint_iff_not_equal notin_set_remove1) next case False then show ?thesis using distinct_concat_map_E(1)[OF 1, simplified, OF 2 3 4] 5 6 using disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3 disjoint_iff_not_equal finite_fset fmember.rep_eq notin_set_remove1 select_result_I2 set_sorted_list_of_set by (metis (no_types, lifting)) qed qed next fix x xa xb assume 1: "(\x\fset (object_ptr_kinds h3). set |h3 \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h3). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" and 2: "xa |\| object_ptr_kinds h3" and 3: "x \ set |h3 \ get_child_nodes xa|\<^sub>r" and 4: "xb |\| document_ptr_kinds h3" and 5: "x \ set |h3 \ get_disconnected_nodes xb|\<^sub>r" have 6: "set |h3 \ get_child_nodes xa|\<^sub>r \ set |h2 \ get_disconnected_nodes xb|\<^sub>r = {}" using 1 2 4 by (metis \type_wf h2\ children_eq2_h2 document_ptr_kinds_commutes \known_ptrs h\ local.get_child_nodes_ok local.get_disconnected_nodes_ok local.heap_is_wellformed_children_disc_nodes_different local.known_ptrs_known_ptr object_ptr_kinds_M_eq3_h object_ptr_kinds_M_eq3_h2 returns_result_select_result wellformed_h2) show False proof (cases "xb = owner_document") case True then show ?thesis using select_result_I2[OF disconnected_nodes_h3,folded select_result_I2[OF disconnected_nodes_h2]] by (metis (no_types, lifting) "3" "5" "6" disjoint_iff_not_equal notin_set_remove1) next case False show ?thesis using 2 3 4 5 6 unfolding disconnected_nodes_eq2_h2[OF False] by auto qed qed then have "CD.a_distinct_lists h'" proof(auto simp add: CD.a_distinct_lists_def document_ptr_kinds_eq2_h3 object_ptr_kinds_M_eq2_h3 disconnected_nodes_eq2_h3 intro!: distinct_concat_map_I) fix x assume 1: "distinct (concat (map (\ptr. |h3 \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h')))))" and 2: "x |\| object_ptr_kinds h'" have 3: "\p. p |\| object_ptr_kinds h' \ distinct |h3 \ get_child_nodes p|\<^sub>r" using 1 by (auto elim: distinct_concat_map_E) show "distinct |h' \ get_child_nodes x|\<^sub>r" proof(cases "ptr = x") case True show ?thesis using 3[OF 2] children_h3 children_h' by(auto simp add: True insert_before_list_distinct dest: child_not_in_any_children[unfolded children_eq_h2]) next case False show ?thesis using children_eq2_h3[OF False] 3[OF 2] by auto qed next fix x y xa assume 1:"distinct (concat (map (\ptr. |h3 \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h')))))" and 2: "x |\| object_ptr_kinds h'" and 3: "y |\| object_ptr_kinds h'" and 4: "x \ y" and 5: "xa \ set |h' \ get_child_nodes x|\<^sub>r" and 6: "xa \ set |h' \ get_child_nodes y|\<^sub>r" have 7:"set |h3 \ get_child_nodes x|\<^sub>r \ set |h3 \ get_child_nodes y|\<^sub>r = {}" using distinct_concat_map_E(1)[OF 1] 2 3 4 by auto show False proof (cases "ptr = x") case True then have "ptr \ y" using 4 by simp then show ?thesis using children_h3 children_h' child_not_in_any_children[unfolded children_eq_h2] 5 6 apply(auto simp add: True children_eq2_h3[OF \ptr \ y\])[1] by (metis (no_types, hide_lams) "3" "7" \type_wf h3\ children_eq2_h3 disjoint_iff_not_equal get_child_nodes_ok insert_before_list_in_set \known_ptrs h\ local.known_ptrs_known_ptr object_ptr_kinds_M_eq3_h object_ptr_kinds_M_eq3_h' object_ptr_kinds_M_eq3_h2 returns_result_select_result select_result_I2) next case False then show ?thesis proof (cases "ptr = y") case True then show ?thesis using children_h3 children_h' child_not_in_any_children[unfolded children_eq_h2] 5 6 apply(auto simp add: True children_eq2_h3[OF \ptr \ x\])[1] by (metis (no_types, hide_lams) "2" "4" "7" IntI \known_ptrs h3\ \type_wf h'\ children_eq_h3 empty_iff insert_before_list_in_set local.get_child_nodes_ok local.known_ptrs_known_ptr object_ptr_kinds_M_eq3_h' returns_result_select_result select_result_I2) next case False then show ?thesis using children_eq2_h3[OF \ptr \ x\] children_eq2_h3[OF \ptr \ y\] 5 6 7 by auto qed qed next fix x xa xb assume 1: " (\x\fset (object_ptr_kinds h'). set |h3 \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h'). set |h' \ get_disconnected_nodes x|\<^sub>r) = {} " and 2: "xa |\| object_ptr_kinds h'" and 3: "x \ set |h' \ get_child_nodes xa|\<^sub>r" and 4: "xb |\| document_ptr_kinds h'" and 5: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" have 6: "set |h3 \ get_child_nodes xa|\<^sub>r \ set |h' \ get_disconnected_nodes xb|\<^sub>r = {}" using 1 2 3 4 5 proof - have "\h d. \ type_wf h \ d |\| document_ptr_kinds h \ h \ ok get_disconnected_nodes d" using local.get_disconnected_nodes_ok by satx then have "h' \ ok get_disconnected_nodes xb" using "4" \type_wf h'\ by fastforce then have f1: "h3 \ get_disconnected_nodes xb \\<^sub>r |h' \ get_disconnected_nodes xb|\<^sub>r" by (simp add: disconnected_nodes_eq_h3) have "xa |\| object_ptr_kinds h3" using "2" object_ptr_kinds_M_eq3_h' by blast then show ?thesis using f1 \local.CD.a_distinct_lists h3\ CD.distinct_lists_no_parent by fastforce qed show False proof (cases "ptr = xa") case True show ?thesis using 6 node_not_in_disconnected_nodes 3 4 5 select_result_I2[OF children_h'] select_result_I2[OF children_h3] True disconnected_nodes_eq2_h3 by (metis (no_types, lifting) "2" DocumentMonad.ptr_kinds_ptr_kinds_M \CD.a_distinct_lists h3\ \type_wf h'\ disconnected_nodes_eq_h3 CD.distinct_lists_no_parent document_ptr_kinds_eq2_h3 get_disconnected_nodes_ok insert_before_list_in_set object_ptr_kinds_M_eq3_h' returns_result_select_result) next case False then show ?thesis using 1 2 3 4 5 children_eq2_h3[OF False] by fastforce qed qed moreover have "CD.a_owner_document_valid h2" using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_owner_document_valid h'" apply(auto simp add: CD.a_owner_document_valid_def object_ptr_kinds_M_eq2_h2 object_ptr_kinds_M_eq2_h3 node_ptr_kinds_eq2_h2 node_ptr_kinds_eq2_h3 document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 children_eq2_h2 )[1] thm children_eq2_h3 apply(auto simp add: document_ptr_kinds_eq2_h2[simplified] document_ptr_kinds_eq2_h3[simplified] object_ptr_kinds_M_eq2_h2[simplified] object_ptr_kinds_M_eq2_h3[simplified] node_ptr_kinds_eq2_h2[simplified] node_ptr_kinds_eq2_h3[simplified])[1] apply(auto simp add: disconnected_nodes_eq2_h3[symmetric])[1] by (metis (no_types, lifting) Core_DOM_Functions.i_insert_before.insert_before_list_in_set children_eq2_h3 children_h' children_h3 disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3 finite_set_in in_set_remove1 is_OK_returns_result_I object_ptr_kinds_M_eq3_h' ptr_in_heap returns_result_eq returns_result_select_result) ultimately have "heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'" by (simp add: CD.heap_is_wellformed_def) have shadow_root_eq_h2: "\ptr' shadow_root_ptr_opt. h2 \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt = h3 \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt" using get_shadow_root_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by(auto simp add: adopt_node_locs_def remove_child_locs_def set_child_nodes_get_shadow_root set_disconnected_nodes_get_shadow_root) then have shadow_root_eq2_h2: "\ptr'. |h2 \ get_shadow_root ptr'|\<^sub>r = |h3 \ get_shadow_root ptr'|\<^sub>r" by (meson select_result_eq) have shadow_root_eq_h3: "\ptr' shadow_root_ptr_opt. h3 \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt = h' \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt" using get_shadow_root_reads insert_node_writes h' apply(rule reads_writes_preserved) by(auto simp add: adopt_node_locs_def remove_child_locs_def set_child_nodes_get_shadow_root set_disconnected_nodes_get_shadow_root) then have shadow_root_eq2_h3: "\ptr'. |h3 \ get_shadow_root ptr'|\<^sub>r = |h' \ get_shadow_root ptr'|\<^sub>r" by (meson select_result_eq) have tag_name_eq_h2: "\ptr' tag. h2 \ get_tag_name ptr' \\<^sub>r tag = h3 \ get_tag_name ptr' \\<^sub>r tag" using get_tag_name_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by(auto simp add: adopt_node_locs_def remove_child_locs_def set_child_nodes_get_tag_name set_disconnected_nodes_get_tag_name) then have tag_name_eq2_h2: "\ptr'. |h2 \ get_tag_name ptr'|\<^sub>r = |h3 \ get_tag_name ptr'|\<^sub>r" by (meson select_result_eq) have tag_name_eq_h3: "\ptr' tag. h3 \ get_tag_name ptr' \\<^sub>r tag = h' \ get_tag_name ptr' \\<^sub>r tag" using get_tag_name_reads insert_node_writes h' apply(rule reads_writes_preserved) by(auto simp add: adopt_node_locs_def remove_child_locs_def set_child_nodes_get_tag_name set_disconnected_nodes_get_tag_name) then have tag_name_eq2_h3: "\ptr'. |h3 \ get_tag_name ptr'|\<^sub>r = |h' \ get_tag_name ptr'|\<^sub>r" by (meson select_result_eq) have object_ptr_kinds_eq_h2: "object_ptr_kinds h2 = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes h3]) unfolding adopt_node_locs_def remove_child_locs_def using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def split: if_splits) have object_ptr_kinds_eq_h3: "object_ptr_kinds h3 = object_ptr_kinds h'" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF insert_node_writes h']) using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def split: if_splits) have shadow_root_ptr_kinds_eq_h2: "shadow_root_ptr_kinds h2 = shadow_root_ptr_kinds h3" using object_ptr_kinds_eq_h2 by(auto simp add: shadow_root_ptr_kinds_def) have element_ptr_kinds_eq_h2: "element_ptr_kinds h2 = element_ptr_kinds h3" using object_ptr_kinds_eq_h2 by(auto simp add: element_ptr_kinds_def node_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h3: "shadow_root_ptr_kinds h3 = shadow_root_ptr_kinds h'" using object_ptr_kinds_eq_h3 by(auto simp add: shadow_root_ptr_kinds_def) have element_ptr_kinds_eq_h3: "element_ptr_kinds h3 = element_ptr_kinds h'" using object_ptr_kinds_eq_h3 by(auto simp add: element_ptr_kinds_def node_ptr_kinds_def) have "a_host_shadow_root_rel h2 = a_host_shadow_root_rel h3" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h2 shadow_root_eq2_h2) have "a_host_shadow_root_rel h3 = a_host_shadow_root_rel h'" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h3 shadow_root_eq2_h3) have "cast node \ {x. (x, ptr) \ (parent_child_rel h3 \ a_host_shadow_root_rel h3)\<^sup>*}" using get_ancestors_si_parent_child_host_shadow_root_rel using \known_ptrs h2\ \local.a_host_shadow_root_rel h2 = local.a_host_shadow_root_rel h3\ \parent_child_rel h2 = parent_child_rel h3\ \type_wf h2\ ancestors_h2 node_not_in_ancestors_h2 wellformed_h2 by auto have "acyclic (parent_child_rel h3 \ a_host_shadow_root_rel h3)" using \heap_is_wellformed h2\ by(auto simp add: heap_is_wellformed_def \parent_child_rel h2 = parent_child_rel h3\ \a_host_shadow_root_rel h2 = a_host_shadow_root_rel h3\) then have "acyclic (parent_child_rel h' \ a_host_shadow_root_rel h')" apply(auto simp add: \a_host_shadow_root_rel h3 = a_host_shadow_root_rel h'\ \parent_child_rel h' = insert (ptr, cast node) ((parent_child_rel h3))\)[1] using \cast node \ {x. (x, ptr) \ (parent_child_rel h3 \ a_host_shadow_root_rel h3)\<^sup>*}\ by (simp add: \local.a_host_shadow_root_rel h3 = local.a_host_shadow_root_rel h'\) then show "heap_is_wellformed h'" using \heap_is_wellformed h2\ using \heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'\ apply(auto simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def a_all_ptrs_in_heap_def a_distinct_lists_def a_shadow_root_valid_def)[1] by(auto simp add: object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 element_ptr_kinds_eq_h2 element_ptr_kinds_eq_h3 shadow_root_ptr_kinds_eq_h2 shadow_root_ptr_kinds_eq_h3 shadow_root_eq_h2 shadow_root_eq_h3 shadow_root_eq2_h2 shadow_root_eq2_h3 tag_name_eq_h2 tag_name_eq_h3 tag_name_eq2_h2 tag_name_eq2_h3) qed end interpretation i_insert_before_wf?: l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_parent get_parent_locs get_child_nodes get_child_nodes_locs set_child_nodes set_child_nodes_locs get_ancestors_si get_ancestors_si_locs adopt_node adopt_node_locs set_disconnected_nodes set_disconnected_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_owner_document insert_before insert_before_locs append_child type_wf known_ptr known_ptrs heap_is_wellformed parent_child_rel by(simp add: l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma insert_before_wf_is_l_insert_before_wf [instances]: "l_insert_before_wf Shadow_DOM.heap_is_wellformed ShadowRootClass.type_wf ShadowRootClass.known_ptr ShadowRootClass.known_ptrs Shadow_DOM.insert_before Shadow_DOM.get_child_nodes" apply(auto simp add: l_insert_before_wf_def l_insert_before_wf_axioms_def instances)[1] using insert_before_removes_child apply fast done lemma l_set_disconnected_nodes_get_disconnected_nodes_wf [instances]: "l_set_disconnected_nodes_get_disconnected_nodes_wf ShadowRootClass.type_wf ShadowRootClass.known_ptr Shadow_DOM.heap_is_wellformed Shadow_DOM.parent_child_rel Shadow_DOM.get_child_nodes get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs" apply(auto simp add: l_set_disconnected_nodes_get_disconnected_nodes_wf_def l_set_disconnected_nodes_get_disconnected_nodes_wf_axioms_def instances)[1] by (metis Diff_iff Shadow_DOM.i_heap_is_wellformed.heap_is_wellformed_disconnected_nodes_distinct Shadow_DOM.i_remove_child.set_disconnected_nodes_get_disconnected_nodes insert_iff returns_result_eq set_remove1_eq) interpretation l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_host get_host_locs get_ancestors_si get_ancestors_si_locs get_root_node_si get_root_node_si_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_document get_disconnected_document_locs by(auto simp add: l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] interpretation i_insert_before_wf2?: l_insert_before_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs set_child_nodes set_child_nodes_locs get_shadow_root get_shadow_root_locs set_disconnected_nodes set_disconnected_nodes_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel get_ancestors_si get_ancestors_si_locs get_parent get_parent_locs adopt_node adopt_node_locs get_owner_document insert_before insert_before_locs append_child known_ptrs get_host get_host_locs get_root_node_si get_root_node_si_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_document get_disconnected_document_locs by(auto simp add: l_insert_before_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_insert_before_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma insert_before_wf2_is_l_insert_before_wf2 [instances]: "l_insert_before_wf2 ShadowRootClass.type_wf ShadowRootClass.known_ptr ShadowRootClass.known_ptrs Shadow_DOM.insert_before Shadow_DOM.heap_is_wellformed" apply(auto simp add: l_insert_before_wf2_def l_insert_before_wf2_axioms_def instances)[1] using insert_before_child_preserves apply(fast, fast, fast) done subsubsection \append\_child\ interpretation i_append_child_wf?: l_append_child_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_owner_document get_parent get_parent_locs remove_child remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs adopt_node adopt_node_locs known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs set_child_nodes set_child_nodes_locs remove get_ancestors_si get_ancestors_si_locs insert_before insert_before_locs append_child heap_is_wellformed parent_child_rel by(auto simp add: l_append_child_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_append_child_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma append_child_wf_is_l_append_child_wf [instances]: "l_append_child_wf type_wf known_ptr known_ptrs append_child heap_is_wellformed" apply(auto simp add: l_append_child_wf_def l_append_child_wf_axioms_def instances)[1] using append_child_heap_is_wellformed_preserved by fast+ subsubsection \to\_tree\_order\ interpretation i_to_tree_order_wf?: l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs to_tree_order known_ptrs get_parent get_parent_locs heap_is_wellformed parent_child_rel get_disconnected_nodes get_disconnected_nodes_locs apply(auto simp add: l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)[1] done declare l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma to_tree_order_wf_is_l_to_tree_order_wf [instances]: "l_to_tree_order_wf heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent get_child_nodes" apply(auto simp add: l_to_tree_order_wf_def l_to_tree_order_wf_axioms_def instances)[1] using to_tree_order_ok apply fast using to_tree_order_ptrs_in_heap apply fast using to_tree_order_parent_child_rel apply(fast, fast) using to_tree_order_child2 apply blast using to_tree_order_node_ptrs apply fast using to_tree_order_child apply fast using to_tree_order_ptr_in_result apply fast using to_tree_order_parent apply fast using to_tree_order_subset apply fast done paragraph \get\_root\_node\ interpretation i_to_tree_order_wf_get_root_node_wf?: l_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf known_ptrs heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_parent get_parent_locs get_ancestors get_ancestors_locs get_root_node get_root_node_locs to_tree_order by(auto simp add: l_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma to_tree_order_wf_get_root_node_wf_is_l_to_tree_order_wf_get_root_node_wf [instances]: "l_to_tree_order_wf_get_root_node_wf ShadowRootClass.type_wf ShadowRootClass.known_ptr ShadowRootClass.known_ptrs to_tree_order Shadow_DOM.get_root_node Shadow_DOM.heap_is_wellformed" apply(auto simp add: l_to_tree_order_wf_get_root_node_wf_def l_to_tree_order_wf_get_root_node_wf_axioms_def instances)[1] using to_tree_order_get_root_node apply fast using to_tree_order_same_root apply fast done subsubsection \to\_tree\_order\_si\ locale l_to_tree_order_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes + l_get_parent_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_to_tree_order_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma to_tree_order_si_ok: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" and "ptr |\| object_ptr_kinds h" shows "h \ ok (to_tree_order_si ptr)" proof(insert assms(1) assms(4), induct rule: heap_wellformed_induct_si) case (step parent) have "known_ptr parent" using assms(2) local.known_ptrs_known_ptr step.prems by blast then show ?case using step using assms(1) assms(2) assms(3) using local.heap_is_wellformed_children_in_heap local.get_shadow_root_shadow_root_ptr_in_heap by(auto simp add: to_tree_order_si_def[of parent] intro: get_child_nodes_ok get_shadow_root_ok intro!: bind_is_OK_pure_I map_M_pure_I bind_pure_I map_M_ok_I split: option.splits) qed end interpretation i_to_tree_order_si_wf?: l_to_tree_order_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs known_ptrs get_parent get_parent_locs to_tree_order_si by(auto simp add: l_to_tree_order_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_to_tree_order_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_assigned\_nodes\ lemma forall_M_small_big: "h \ forall_M f xs \\<^sub>h h' \ P h \ (\h h' x. x \ set xs \ h \ f x \\<^sub>h h' \ P h \ P h') \ P h'" by(induct xs arbitrary: h) (auto elim!: bind_returns_heap_E) locale l_assigned_nodes_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_heap_is_wellformed + l_remove_child_wf2 + l_append_child_wf + l_remove_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma assigned_nodes_distinct: assumes "heap_is_wellformed h" assumes "h \ assigned_nodes slot \\<^sub>r nodes" shows "distinct nodes" proof - have "\ptr children. h \ get_child_nodes ptr \\<^sub>r children \ distinct children" using assms(1) local.heap_is_wellformed_children_distinct by blast then show ?thesis using assms apply(auto simp add: assigned_nodes_def elim!: bind_returns_result_E2 split: if_splits)[1] by (simp add: filter_M_distinct) qed lemma flatten_dom_preserves: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "h \ flatten_dom \\<^sub>h h'" shows "heap_is_wellformed h'" and "known_ptrs h'" and "type_wf h'" proof - obtain tups h2 element_ptrs shadow_root_ptrs where "h \ element_ptr_kinds_M \\<^sub>r element_ptrs" and tups: "h \ map_filter_M2 (\element_ptr. do { tag \ get_tag_name element_ptr; assigned_nodes \ assigned_nodes element_ptr; (if tag = ''slot'' \ assigned_nodes \ [] then return (Some (element_ptr, assigned_nodes)) else return None)}) element_ptrs \\<^sub>r tups" (is "h \ map_filter_M2 ?f element_ptrs \\<^sub>r tups") and h2: "h \ forall_M (\(slot, assigned_nodes). do { get_child_nodes (cast slot) \ forall_M remove; forall_M (append_child (cast slot)) assigned_nodes }) tups \\<^sub>h h2" and "h2 \ shadow_root_ptr_kinds_M \\<^sub>r shadow_root_ptrs" and h': "h2 \ forall_M (\shadow_root_ptr. do { host \ get_host shadow_root_ptr; get_child_nodes (cast host) \ forall_M remove; get_child_nodes (cast shadow_root_ptr) \ forall_M (append_child (cast host)); remove_shadow_root host }) shadow_root_ptrs \\<^sub>h h'" using \h \ flatten_dom \\<^sub>h h'\ apply(auto simp add: flatten_dom_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF ElementMonad.ptr_kinds_M_pure, rotated] bind_returns_heap_E2[rotated, OF ShadowRootMonad.ptr_kinds_M_pure, rotated])[1] apply(drule pure_returns_heap_eq) by(auto intro!: map_filter_M2_pure bind_pure_I) have "heap_is_wellformed h2 \ known_ptrs h2 \ type_wf h2" using h2 \heap_is_wellformed h\ \known_ptrs h\ \type_wf h\ by(auto elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated] elim!: forall_M_small_big[where P = "\h. heap_is_wellformed h \ known_ptrs h \ type_wf h", simplified] intro: remove_preserves_known_ptrs remove_heap_is_wellformed_preserved remove_preserves_type_wf append_child_preserves_known_ptrs append_child_heap_is_wellformed_preserved append_child_preserves_type_wf) then show "heap_is_wellformed h'" and "known_ptrs h'" and "type_wf h'" using h' by(auto elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_host_pure, rotated] bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated] dest!: forall_M_small_big[where P = "\h. heap_is_wellformed h \ known_ptrs h \ type_wf h", simplified] intro: remove_preserves_known_ptrs remove_heap_is_wellformed_preserved remove_preserves_type_wf append_child_preserves_known_ptrs append_child_heap_is_wellformed_preserved append_child_preserves_type_wf remove_shadow_root_preserves ) qed end interpretation i_assigned_nodes_wf?: l_assigned_nodes_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr assigned_nodes assigned_nodes_flatten flatten_dom get_child_nodes get_child_nodes_locs get_tag_name get_tag_name_locs get_root_node get_root_node_locs get_host get_host_locs find_slot assigned_slot remove insert_before insert_before_locs append_child remove_shadow_root remove_shadow_root_locs type_wf get_shadow_root get_shadow_root_locs set_shadow_root set_shadow_root_locs get_parent get_parent_locs to_tree_order heap_is_wellformed parent_child_rel get_disconnected_nodes get_disconnected_nodes_locs known_ptrs remove_child remove_child_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_document get_disconnected_document_locs by(auto simp add: l_assigned_nodes_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_assigned_nodes_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_shadow\_root\_safe\ locale l_get_shadow_root_safe_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs + l_type_wf type_wf + l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root_safe get_shadow_root_safe_locs get_shadow_root get_shadow_root_locs get_mode get_mode_locs for known_ptr :: "(_::linorder) object_ptr \ bool" and known_ptrs :: "(_) heap \ bool" and type_wf :: "(_) heap \ bool" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_shadow_root_safe :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_safe_locs :: "(_) element_ptr \ (_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" and get_child_nodes :: "(_::linorder) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_mode :: "(_) shadow_root_ptr \ ((_) heap, exception, shadow_root_mode) prog" and get_mode_locs :: "(_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" begin end subsubsection \create\_element\ locale l_create_element_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs + l_set_tag_name type_wf set_tag_name set_tag_name_locs + l_create_element_defs create_element + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs + l_new_element_get_disconnected_nodes get_disconnected_nodes get_disconnected_nodes_locs + l_set_tag_name_get_disconnected_nodes type_wf set_tag_name set_tag_name_locs get_disconnected_nodes get_disconnected_nodes_locs + l_create_element\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs set_tag_name set_tag_name_locs type_wf create_element known_ptr type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_new_element_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs + l_set_tag_name_get_child_nodes type_wf set_tag_name set_tag_name_locs known_ptr get_child_nodes get_child_nodes_locs + l_set_tag_name_get_tag_name type_wf get_tag_name get_tag_name_locs set_tag_name set_tag_name_locs + l_new_element_get_tag_name type_wf get_tag_name get_tag_name_locs + l_set_disconnected_nodes_get_child_nodes set_disconnected_nodes set_disconnected_nodes_locs get_child_nodes get_child_nodes_locs + l_set_disconnected_nodes_get_shadow_root set_disconnected_nodes set_disconnected_nodes_locs get_shadow_root get_shadow_root_locs + l_set_disconnected_nodes_get_tag_name type_wf set_disconnected_nodes set_disconnected_nodes_locs get_tag_name get_tag_name_locs + l_set_disconnected_nodes type_wf set_disconnected_nodes set_disconnected_nodes_locs + l_set_disconnected_nodes_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs + l_new_element_get_shadow_root type_wf get_shadow_root get_shadow_root_locs + l_set_tag_name_get_shadow_root type_wf set_tag_name set_tag_name_locs get_shadow_root get_shadow_root_locs + l_new_element type_wf + l_known_ptrs known_ptr known_ptrs for known_ptr :: "(_::linorder) object_ptr \ bool" and known_ptrs :: "(_) heap \ bool" and type_wf :: "(_) heap \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and set_tag_name :: "(_) element_ptr \ char list \ ((_) heap, exception, unit) prog" and set_tag_name_locs :: "(_) element_ptr \ ((_) heap, exception, unit) prog set" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" and create_element :: "(_) document_ptr \ char list \ ((_) heap, exception, (_) element_ptr) prog" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_disconnected_document :: "(_) node_ptr \ ((_) heap, exception, (_) document_ptr) prog" and get_disconnected_document_locs :: "((_) heap \ (_) heap \ bool) set" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_::linorder) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" begin lemma create_element_preserves_wellformedness: assumes "heap_is_wellformed h" and "h \ create_element document_ptr tag \\<^sub>h h'" and "type_wf h" and "known_ptrs h" shows "heap_is_wellformed h'" and "type_wf h'" and "known_ptrs h'" proof - obtain new_element_ptr h2 h3 disc_nodes_h3 where new_element_ptr: "h \ new_element \\<^sub>r new_element_ptr" and h2: "h \ new_element \\<^sub>h h2" and h3: "h2 \ set_tag_name new_element_ptr tag \\<^sub>h h3" and disc_nodes_h3: "h3 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" and h': "h3 \ set_disconnected_nodes document_ptr (cast new_element_ptr # disc_nodes_h3) \\<^sub>h h'" using assms(2) by(auto simp add: create_element_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF CD.get_disconnected_nodes_pure, rotated] ) then have "h \ create_element document_ptr tag \\<^sub>r new_element_ptr" apply(auto simp add: create_element_def intro!: bind_returns_result_I)[1] apply (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust) apply (metis is_OK_returns_heap_E is_OK_returns_result_I CD.get_disconnected_nodes_pure pure_returns_heap_eq) by (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust) have "new_element_ptr \ set |h \ element_ptr_kinds_M|\<^sub>r" using new_element_ptr ElementMonad.ptr_kinds_ptr_kinds_M h2 using new_element_ptr_not_in_heap by blast then have "cast new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r" by simp then have "cast new_element_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r" by simp have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\| {|cast new_element_ptr|}" using new_element_new_ptr h2 new_element_ptr by blast then have node_ptr_kinds_eq_h: "node_ptr_kinds h2 = node_ptr_kinds h |\| {|cast new_element_ptr|}" apply(simp add: node_ptr_kinds_def) by force then have element_ptr_kinds_eq_h: "element_ptr_kinds h2 = element_ptr_kinds h |\| {|new_element_ptr|}" apply(simp add: element_ptr_kinds_def) by force have character_data_ptr_kinds_eq_h: "character_data_ptr_kinds h2 = character_data_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: node_ptr_kinds_def character_data_ptr_kinds_def) have document_ptr_kinds_eq_h: "document_ptr_kinds h2 = document_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: document_ptr_kinds_def) have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_tag_name_writes h3]) using set_tag_name_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2" using object_ptr_kinds_eq_h2 by(auto simp add: node_ptr_kinds_def) then have element_ptr_kinds_eq_h2: "element_ptr_kinds h3 = element_ptr_kinds h2" by(simp add: element_ptr_kinds_def) have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_disconnected_nodes_writes h']) using set_disconnected_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3" using object_ptr_kinds_eq_h3 by(auto simp add: node_ptr_kinds_def) then have element_ptr_kinds_eq_h3: "element_ptr_kinds h' = element_ptr_kinds h3" by(simp add: element_ptr_kinds_def) have "known_ptr (cast new_element_ptr)" using \h \ create_element document_ptr tag \\<^sub>r new_element_ptr\ local.create_element_known_ptr by blast then have "known_ptrs h2" using known_ptrs_new_ptr object_ptr_kinds_eq_h \known_ptrs h\ h2 by blast then have "known_ptrs h3" using known_ptrs_preserved object_ptr_kinds_eq_h2 by blast then show "known_ptrs h'" using known_ptrs_preserved object_ptr_kinds_eq_h3 by blast have "document_ptr |\| document_ptr_kinds h" using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2 CD.get_disconnected_nodes_ptr_in_heap \type_wf h\ document_ptr_kinds_def by (metis is_OK_returns_result_I) have children_eq_h: "\(ptr'::(_) object_ptr) children. ptr' \ cast new_element_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads h2 get_child_nodes_new_element[rotated, OF new_element_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have children_eq2_h: "\ptr'. ptr' \ cast new_element_ptr \ |h \ get_child_nodes ptr'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have "h2 \ get_child_nodes (cast new_element_ptr) \\<^sub>r []" using new_element_ptr h2 new_element_ptr_in_heap[OF h2 new_element_ptr] new_element_is_element_ptr[OF new_element_ptr] new_element_no_child_nodes by blast have tag_name_eq_h: "\ptr' disc_nodes. ptr' \ new_element_ptr \ h \ get_tag_name ptr' \\<^sub>r disc_nodes = h2 \ get_tag_name ptr' \\<^sub>r disc_nodes" using get_tag_name_reads h2 get_tag_name_new_element[rotated, OF new_element_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by(blast)+ then have tag_name_eq2_h: "\ptr'. ptr' \ new_element_ptr \ |h \ get_tag_name ptr'|\<^sub>r = |h2 \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have "h2 \ get_tag_name new_element_ptr \\<^sub>r ''''" using new_element_ptr h2 new_element_ptr_in_heap[OF h2 new_element_ptr] new_element_is_element_ptr[OF new_element_ptr] new_element_empty_tag_name by blast have disconnected_nodes_eq_h: "\doc_ptr disc_nodes. h \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using CD.get_disconnected_nodes_reads h2 get_disconnected_nodes_new_element[OF new_element_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have disconnected_nodes_eq2_h: "\doc_ptr. |h \ get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. h2 \ get_child_nodes ptr' \\<^sub>r children = h3 \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads set_tag_name_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_tag_name_get_child_nodes) then have children_eq2_h2: "\ptr'. |h2 \ get_child_nodes ptr'|\<^sub>r = |h3 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using CD.get_disconnected_nodes_reads set_tag_name_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_tag_name_get_disconnected_nodes) then have disconnected_nodes_eq2_h2: "\doc_ptr. |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have tag_name_eq_h2: "\ptr' disc_nodes. ptr' \ new_element_ptr \ h2 \ get_tag_name ptr' \\<^sub>r disc_nodes = h3 \ get_tag_name ptr' \\<^sub>r disc_nodes" apply(rule reads_writes_preserved[OF get_tag_name_reads set_tag_name_writes h3]) by (metis local.set_tag_name_get_tag_name_different_pointers) then have tag_name_eq2_h2: "\ptr'. ptr' \ new_element_ptr \ |h2 \ get_tag_name ptr'|\<^sub>r = |h3 \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have "h2 \ get_tag_name new_element_ptr \\<^sub>r ''''" using new_element_ptr h2 new_element_ptr_in_heap[OF h2 new_element_ptr] new_element_is_element_ptr[OF new_element_ptr] new_element_empty_tag_name by blast have "type_wf h2" using \type_wf h\ new_element_types_preserved h2 by blast then have "type_wf h3" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_tag_name_writes h3] using set_tag_name_types_preserved by(auto simp add: reflp_def transp_def) then show "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h'] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) have children_eq_h3: "\ptr' children. h3 \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_disconnected_nodes_get_child_nodes) then have children_eq2_h3: "\ptr'. |h3 \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h3: "\doc_ptr disc_nodes. document_ptr \ doc_ptr \ h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using CD.get_disconnected_nodes_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers) then have disconnected_nodes_eq2_h3: "\doc_ptr. document_ptr \ doc_ptr \ |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have tag_name_eq_h2: "\ptr' disc_nodes. ptr' \ new_element_ptr \ h2 \ get_tag_name ptr' \\<^sub>r disc_nodes = h3 \ get_tag_name ptr' \\<^sub>r disc_nodes" apply(rule reads_writes_preserved[OF get_tag_name_reads set_tag_name_writes h3]) by (metis local.set_tag_name_get_tag_name_different_pointers) then have tag_name_eq2_h2: "\ptr'. ptr' \ new_element_ptr \ |h2 \ get_tag_name ptr'|\<^sub>r = |h3 \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have disc_nodes_document_ptr_h2: "h2 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" using disconnected_nodes_eq_h2 disc_nodes_h3 by auto then have disc_nodes_document_ptr_h: "h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" using disconnected_nodes_eq_h by auto then have "cast new_element_ptr \ set disc_nodes_h3" using \heap_is_wellformed h\ using \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ a_all_ptrs_in_heap_def heap_is_wellformed_def using NodeMonad.ptr_kinds_ptr_kinds_M local.heap_is_wellformed_disc_nodes_in_heap by blast have tag_name_eq_h3: "\ptr' disc_nodes. h3 \ get_tag_name ptr' \\<^sub>r disc_nodes = h' \ get_tag_name ptr' \\<^sub>r disc_nodes" apply(rule reads_writes_preserved[OF get_tag_name_reads set_disconnected_nodes_writes h']) using set_disconnected_nodes_get_tag_name by blast then have tag_name_eq2_h3: "\ptr'. |h3 \ get_tag_name ptr'|\<^sub>r = |h' \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have "acyclic (parent_child_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def) also have "parent_child_rel h = parent_child_rel h2" proof(auto simp add: CD.parent_child_rel_def)[1] fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h2" by (simp add: object_ptr_kinds_eq_h) next fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "x \ set |h2 \ get_child_nodes a|\<^sub>r" by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2_h) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h" using object_ptr_kinds_eq_h \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ by(auto) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "x \ set |h \ get_child_nodes a|\<^sub>r" by (metis (no_types, lifting) \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ children_eq2_h empty_iff empty_set image_eqI select_result_I2) qed also have "\ = parent_child_rel h3" by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2) also have "\ = parent_child_rel h'" by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3) finally have "CD.a_acyclic_heap h'" by (simp add: CD.acyclic_heap_def) have "CD.a_all_ptrs_in_heap h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_all_ptrs_in_heap h2" apply(auto simp add: CD.a_all_ptrs_in_heap_def)[1] apply (metis \known_ptrs h2\ \parent_child_rel h = parent_child_rel h2\ \type_wf h2\ assms(1) assms(3) funion_iff CD.get_child_nodes_ok local.known_ptrs_known_ptr local.parent_child_rel_child_in_heap CD.parent_child_rel_child_nodes2 node_ptr_kinds_commutes node_ptr_kinds_eq_h returns_result_select_result) by (metis (no_types, lifting) CD.get_child_nodes_ok CD.get_child_nodes_ptr_in_heap \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ assms(3) assms(4) children_eq_h disconnected_nodes_eq2_h document_ptr_kinds_eq_h finite_set_in is_OK_returns_result_I local.known_ptrs_known_ptr node_ptr_kinds_commutes returns_result_select_result subsetD) then have "CD.a_all_ptrs_in_heap h3" by (simp add: children_eq2_h2 disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 object_ptr_kinds_eq_h2) then have "CD.a_all_ptrs_in_heap h'" by (smt children_eq2_h3 disc_nodes_h3 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 element_ptr_kinds_commutes h' h2 local.CD.a_all_ptrs_in_heap_def local.set_disconnected_nodes_get_disconnected_nodes new_element_ptr new_element_ptr_in_heap node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h3 notin_fset object_ptr_kinds_eq_h3 select_result_I2 set_ConsD subset_code(1)) have "\p. p |\| object_ptr_kinds h \ cast new_element_ptr \ set |h \ get_child_nodes p|\<^sub>r" using \heap_is_wellformed h\ \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ heap_is_wellformed_children_in_heap by (meson NodeMonad.ptr_kinds_ptr_kinds_M CD.a_all_ptrs_in_heap_def assms(3) assms(4) fset_mp fset_of_list_elem CD.get_child_nodes_ok known_ptrs_known_ptr returns_result_select_result) then have "\p. p |\| object_ptr_kinds h2 \ cast new_element_ptr \ set |h2 \ get_child_nodes p|\<^sub>r" using children_eq2_h apply(auto simp add: object_ptr_kinds_eq_h)[1] using \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ apply auto[1] by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\) then have "\p. p |\| object_ptr_kinds h3 \ cast new_element_ptr \ set |h3 \ get_child_nodes p|\<^sub>r" using object_ptr_kinds_eq_h2 children_eq2_h2 by auto then have new_element_ptr_not_in_any_children: "\p. p |\| object_ptr_kinds h' \ cast new_element_ptr \ set |h' \ get_child_nodes p|\<^sub>r" using object_ptr_kinds_eq_h3 children_eq2_h3 by auto have "CD.a_distinct_lists h" using \heap_is_wellformed h\ by (simp add: CD.heap_is_wellformed_def heap_is_wellformed_def) then have "CD.a_distinct_lists h2" using \h2 \ get_child_nodes (cast new_element_ptr) \\<^sub>r []\ apply(auto simp add: CD.a_distinct_lists_def object_ptr_kinds_eq_h document_ptr_kinds_eq_h disconnected_nodes_eq2_h intro!: distinct_concat_map_I)[1] - apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert) + apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert_remove) apply(case_tac "x=cast new_element_ptr") apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply (metis IntI assms(1) assms(3) assms(4) empty_iff CD.get_child_nodes_ok local.heap_is_wellformed_one_parent local.known_ptrs_known_ptr returns_result_select_result) apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] by (metis \ CD.a_distinct_lists h\ \type_wf h2\ disconnected_nodes_eq_h document_ptr_kinds_eq_h CD.distinct_lists_no_parent get_disconnected_nodes_ok returns_result_select_result) then have " CD.a_distinct_lists h3" by(auto simp add: CD.a_distinct_lists_def disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 children_eq2_h2 object_ptr_kinds_eq_h2) then have " CD.a_distinct_lists h'" proof(auto simp add: CD.a_distinct_lists_def disconnected_nodes_eq2_h3 children_eq2_h3 object_ptr_kinds_eq_h3 document_ptr_kinds_eq_h3 intro!: distinct_concat_map_I)[1] fix x assume "distinct (concat (map (\document_ptr. |h3 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" and "x |\| document_ptr_kinds h3" then show "distinct |h' \ get_disconnected_nodes x|\<^sub>r" using document_ptr_kinds_eq_h3 disconnected_nodes_eq_h3 h' set_disconnected_nodes_get_disconnected_nodes by (metis (no_types, lifting) \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set disc_nodes_h3\ \ CD.a_distinct_lists h3\ \type_wf h'\ disc_nodes_h3 distinct.simps(2) CD.distinct_lists_disconnected_nodes get_disconnected_nodes_ok returns_result_eq returns_result_select_result) next fix x y xa assume "distinct (concat (map (\document_ptr. |h3 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" and "x |\| document_ptr_kinds h3" and "y |\| document_ptr_kinds h3" and "x \ y" and "xa \ set |h' \ get_disconnected_nodes x|\<^sub>r" and "xa \ set |h' \ get_disconnected_nodes y|\<^sub>r" moreover have "set |h3 \ get_disconnected_nodes x|\<^sub>r \ set |h3 \ get_disconnected_nodes y|\<^sub>r = {}" using calculation by(auto dest: distinct_concat_map_E(1)) ultimately show "False" apply(-) apply(cases "x = document_ptr") apply(smt NodeMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ \CD.a_all_ptrs_in_heap h\ disc_nodes_h3 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 finite_set_in h' set_disconnected_nodes_get_disconnected_nodes CD.a_all_ptrs_in_heap_def select_result_I2 set_ConsD subsetD) by (smt NodeMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ \CD.a_all_ptrs_in_heap h\ disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 finite_set_in h' l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes CD.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms select_result_I2 set_ConsD subsetD) next fix x xa xb assume 2: "(\x\fset (object_ptr_kinds h3). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h3). set |h3 \ get_disconnected_nodes x|\<^sub>r) = {}" and 3: "xa |\| object_ptr_kinds h3" and 4: "x \ set |h' \ get_child_nodes xa|\<^sub>r" and 5: "xb |\| document_ptr_kinds h3" and 6: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" show "False" using disc_nodes_document_ptr_h disconnected_nodes_eq2_h3 apply - apply(cases "xb = document_ptr") apply (metis (no_types, hide_lams) "3" "4" "6" \\p. p |\| object_ptr_kinds h3 \ cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h3 \ get_child_nodes p|\<^sub>r\ \ CD.a_distinct_lists h3\ children_eq2_h3 disc_nodes_h3 CD.distinct_lists_no_parent h' select_result_I2 set_ConsD set_disconnected_nodes_get_disconnected_nodes) by (metis "3" "4" "5" "6" \ CD.a_distinct_lists h3\ \type_wf h3\ children_eq2_h3 CD.distinct_lists_no_parent get_disconnected_nodes_ok returns_result_select_result) qed have "CD.a_owner_document_valid h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_owner_document_valid h'" using disc_nodes_h3 \document_ptr |\| document_ptr_kinds h\ apply(auto simp add: CD.a_owner_document_valid_def)[1] apply(auto simp add: object_ptr_kinds_eq_h object_ptr_kinds_eq_h3 )[1] apply(auto simp add: object_ptr_kinds_eq_h2)[1] apply(auto simp add: document_ptr_kinds_eq_h document_ptr_kinds_eq_h3 )[1] apply(auto simp add: document_ptr_kinds_eq_h2)[1] apply(auto simp add: node_ptr_kinds_eq_h node_ptr_kinds_eq_h3 )[1] apply(auto simp add: node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h )[1] apply(auto simp add: children_eq2_h2[symmetric] children_eq2_h3[symmetric] disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3)[1] apply (metis (no_types, lifting) document_ptr_kinds_eq_h h' list.set_intros(1) local.set_disconnected_nodes_get_disconnected_nodes select_result_I2) apply(simp add: object_ptr_kinds_eq_h) by(metis (no_types, lifting) NodeMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ children_eq2_h children_eq2_h2 children_eq2_h3 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h finite_set_in h' l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes list.set_intros(2) local.l_set_disconnected_nodes_get_disconnected_nodes_axioms node_ptr_kinds_commutes select_result_I2) have "CD.a_heap_is_wellformed h'" using \CD.a_acyclic_heap h'\ \CD.a_all_ptrs_in_heap h'\ \CD.a_distinct_lists h'\ \CD.a_owner_document_valid h'\ by(simp add: CD.a_heap_is_wellformed_def) have shadow_root_ptr_kinds_eq_h: "shadow_root_ptr_kinds h2 = shadow_root_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: shadow_root_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h2: "shadow_root_ptr_kinds h3 = shadow_root_ptr_kinds h2" using object_ptr_kinds_eq_h2 by(auto simp add: shadow_root_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h3: "shadow_root_ptr_kinds h' = shadow_root_ptr_kinds h3" using object_ptr_kinds_eq_h3 by(auto simp add: shadow_root_ptr_kinds_def) have shadow_root_eq_h: "\element_ptr shadow_root_opt. element_ptr \ new_element_ptr \ h \ get_shadow_root element_ptr \\<^sub>r shadow_root_opt = h2 \ get_shadow_root element_ptr \\<^sub>r shadow_root_opt" proof - fix element_ptr shadow_root_opt assume "element_ptr \ new_element_ptr " have "\P \ get_shadow_root_locs element_ptr. P h h2" using get_shadow_root_new_element new_element_ptr h2 using \element_ptr \ new_element_ptr\ by blast then have "preserved (get_shadow_root element_ptr) h h2" using get_shadow_root_new_element[rotated, OF new_element_ptr h2] using get_shadow_root_reads by(simp add: reads_def) then show "h \ get_shadow_root element_ptr \\<^sub>r shadow_root_opt = h2 \ get_shadow_root element_ptr \\<^sub>r shadow_root_opt" by (simp add: preserved_def) qed have shadow_root_none: "h2 \ get_shadow_root (new_element_ptr) \\<^sub>r None" using new_element_ptr h2 new_element_ptr_in_heap[OF h2 new_element_ptr] new_element_is_element_ptr[OF new_element_ptr] new_element_no_shadow_root by blast have shadow_root_eq_h2: "\ptr' children. h2 \ get_shadow_root ptr' \\<^sub>r children = h3 \ get_shadow_root ptr' \\<^sub>r children" using get_shadow_root_reads set_tag_name_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_tag_name_get_shadow_root) have shadow_root_eq_h3: "\ptr' children. h3 \ get_shadow_root ptr' \\<^sub>r children = h' \ get_shadow_root ptr' \\<^sub>r children" using get_shadow_root_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) using set_disconnected_nodes_get_shadow_root by(auto simp add: set_disconnected_nodes_get_shadow_root) have "a_all_ptrs_in_heap h" by (simp add: assms(1) local.a_all_ptrs_in_heap_def local.get_shadow_root_shadow_root_ptr_in_heap) then have "a_all_ptrs_in_heap h2" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h)[1] using returns_result_eq shadow_root_eq_h shadow_root_none by fastforce then have "a_all_ptrs_in_heap h3" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h2)[1] using shadow_root_eq_h2 by blast then have "a_all_ptrs_in_heap h'" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h3)[1] by (simp add: shadow_root_eq_h3) have "a_distinct_lists h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_distinct_lists h2" apply(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h)[1] apply(auto simp add: distinct_insort intro!: distinct_concat_map_I split: option.splits)[1] apply(case_tac "x = new_element_ptr") using shadow_root_none apply auto[1] using shadow_root_eq_h by (smt Diff_empty Diff_insert0 ElementMonad.ptr_kinds_M_ptr_kinds ElementMonad.ptr_kinds_ptr_kinds_M assms(1) assms(3) finite_set_in h2 insort_split local.get_shadow_root_ok local.shadow_root_same_host new_element_ptr new_element_ptr_not_in_heap option.distinct(1) returns_result_select_result select_result_I2 shadow_root_none) then have "a_distinct_lists h3" by(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h2 select_result_eq[OF shadow_root_eq_h2]) then have "a_distinct_lists h'" by(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h3 select_result_eq[OF shadow_root_eq_h3]) have "a_shadow_root_valid h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_shadow_root_valid h2" proof (unfold a_shadow_root_valid_def; safe) fix shadow_root_ptr assume "\shadow_root_ptr\fset (shadow_root_ptr_kinds h). \host\fset (element_ptr_kinds h). |h \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types \ |h \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr" assume "shadow_root_ptr \ fset (shadow_root_ptr_kinds h2)" obtain previous_host where "previous_host \ fset (element_ptr_kinds h)" and "|h \ get_tag_name previous_host|\<^sub>r \ safe_shadow_root_element_types" and "|h \ get_shadow_root previous_host|\<^sub>r = Some shadow_root_ptr" by (metis \local.a_shadow_root_valid h\ \shadow_root_ptr \ fset (shadow_root_ptr_kinds h2)\ local.a_shadow_root_valid_def shadow_root_ptr_kinds_eq_h) moreover have "previous_host \ new_element_ptr" using calculation(1) h2 new_element_ptr new_element_ptr_not_in_heap by auto ultimately have "|h2 \ get_tag_name previous_host|\<^sub>r \ safe_shadow_root_element_types" and "|h2 \ get_shadow_root previous_host|\<^sub>r = Some shadow_root_ptr" using shadow_root_eq_h apply (simp add: tag_name_eq2_h) by (metis \previous_host \ new_element_ptr\ \|h \ get_shadow_root previous_host|\<^sub>r = Some shadow_root_ptr\ select_result_eq shadow_root_eq_h) then show "\host\fset (element_ptr_kinds h2). |h2 \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types \ |h2 \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr" by (meson \previous_host \ fset (element_ptr_kinds h)\ \previous_host \ new_element_ptr\ assms(3) local.get_shadow_root_ok local.get_shadow_root_ptr_in_heap notin_fset returns_result_select_result shadow_root_eq_h) qed then have "a_shadow_root_valid h3" proof (unfold a_shadow_root_valid_def; safe) fix shadow_root_ptr assume "\shadow_root_ptr\fset (shadow_root_ptr_kinds h2). \host\fset (element_ptr_kinds h2). |h2 \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types \ |h2 \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr" assume "shadow_root_ptr \ fset (shadow_root_ptr_kinds h3)" obtain previous_host where "previous_host \ fset (element_ptr_kinds h2)" and "|h2 \ get_tag_name previous_host|\<^sub>r \ safe_shadow_root_element_types" and "|h2 \ get_shadow_root previous_host|\<^sub>r = Some shadow_root_ptr" by (metis \local.a_shadow_root_valid h2\ \shadow_root_ptr \ fset (shadow_root_ptr_kinds h3)\ local.a_shadow_root_valid_def shadow_root_ptr_kinds_eq_h2) moreover have "previous_host \ new_element_ptr" using calculation(1) h3 new_element_ptr new_element_ptr_not_in_heap using calculation(3) shadow_root_none by auto ultimately have "|h2 \ get_tag_name previous_host|\<^sub>r \ safe_shadow_root_element_types" and "|h2 \ get_shadow_root previous_host|\<^sub>r = Some shadow_root_ptr" using shadow_root_eq_h2 apply (simp add: tag_name_eq2_h2) by (metis \previous_host \ new_element_ptr\ \|h2 \ get_shadow_root previous_host|\<^sub>r = Some shadow_root_ptr\ select_result_eq shadow_root_eq_h) then show "\host\fset (element_ptr_kinds h3). |h3 \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types \ |h3 \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr" by (smt \previous_host \ fset (element_ptr_kinds h2)\ \previous_host \ new_element_ptr\ \type_wf h2\ \type_wf h3\ element_ptr_kinds_eq_h2 finite_set_in local.get_shadow_root_ok returns_result_eq returns_result_select_result shadow_root_eq_h2 tag_name_eq2_h2) qed then have "a_shadow_root_valid h'" apply(auto simp add: a_shadow_root_valid_def element_ptr_kinds_eq_h3 shadow_root_eq_h3 shadow_root_ptr_kinds_eq_h3 tag_name_eq2_h3)[1] by (smt \type_wf h3\ finite_set_in local.get_shadow_root_ok returns_result_select_result select_result_I2 shadow_root_eq_h3) have "a_host_shadow_root_rel h = a_host_shadow_root_rel h2" apply(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h shadow_root_eq_h)[1] apply (smt assms(3) case_prod_conv h2 image_iff local.get_shadow_root_ok mem_Collect_eq new_element_ptr new_element_ptr_not_in_heap returns_result_select_result select_result_I2 shadow_root_eq_h) using shadow_root_none apply auto[1] by (metis (no_types, lifting) Collect_cong assms(3) case_prodE case_prodI h2 host_shadow_root_rel_def host_shadow_root_rel_shadow_root local.a_host_shadow_root_rel_def local.get_shadow_root_impl local.get_shadow_root_ok local.new_element_no_shadow_root new_element_ptr option.distinct(1) returns_result_select_result select_result_I2 shadow_root_eq_h) have "a_host_shadow_root_rel h2 = a_host_shadow_root_rel h3" apply(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h2 shadow_root_eq_h2)[1] apply (smt Collect_cong Shadow_DOM.a_host_shadow_root_rel_def assms(3) h2 host_shadow_root_rel_shadow_root is_OK_returns_result_E local.get_shadow_root_impl local.get_shadow_root_ok local.new_element_types_preserved select_result_I2 shadow_root_eq_h2 split_cong) apply (metis (no_types, lifting) Collect_cong \type_wf h3\ element_ptr_kinds_eq_h2 host_shadow_root_rel_def host_shadow_root_rel_shadow_root local.a_host_shadow_root_rel_def local.get_shadow_root_impl local.get_shadow_root_ok returns_result_select_result shadow_root_eq_h2 split_cong) done have "a_host_shadow_root_rel h3 = a_host_shadow_root_rel h'" apply(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h2 shadow_root_eq_h2)[1] apply (metis (no_types, lifting) Collect_cong \type_wf h3\ case_prodE case_prodI element_ptr_kinds_eq_h2 host_shadow_root_rel_def host_shadow_root_rel_shadow_root local.a_host_shadow_root_rel_def local.get_shadow_root_impl local.get_shadow_root_ok returns_result_select_result shadow_root_eq_h3) apply (smt Collect_cong \type_wf h'\ \type_wf h2\ case_prodD case_prodI2 host_shadow_root_rel_def host_shadow_root_rel_shadow_root is_OK_returns_result_E local.a_host_shadow_root_rel_def local.get_shadow_root_impl local.get_shadow_root_ok select_result_I2 shadow_root_eq_h2 shadow_root_eq_h3) done have "acyclic (parent_child_rel h \ a_host_shadow_root_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) have "parent_child_rel h \ a_host_shadow_root_rel h = parent_child_rel h2 \ a_host_shadow_root_rel h2" using \local.a_host_shadow_root_rel h = local.a_host_shadow_root_rel h2\ \parent_child_rel h = parent_child_rel h2\ by auto have "parent_child_rel h2 \ a_host_shadow_root_rel h2 = parent_child_rel h3 \ a_host_shadow_root_rel h3" using \local.a_host_shadow_root_rel h2 = local.a_host_shadow_root_rel h3\ \parent_child_rel h2 = parent_child_rel h3\ by auto have "parent_child_rel h' \ a_host_shadow_root_rel h' = parent_child_rel h3 \ a_host_shadow_root_rel h3" by (simp add: \local.a_host_shadow_root_rel h3 = local.a_host_shadow_root_rel h'\ \parent_child_rel h3 = parent_child_rel h'\) have "acyclic (parent_child_rel h3 \ a_host_shadow_root_rel h3)" using \acyclic (parent_child_rel h \ local.a_host_shadow_root_rel h)\ \parent_child_rel h \ local.a_host_shadow_root_rel h = parent_child_rel h2 \ local.a_host_shadow_root_rel h2\ \parent_child_rel h2 \ local.a_host_shadow_root_rel h2 = parent_child_rel h3 \ local.a_host_shadow_root_rel h3\ by auto then have "acyclic (parent_child_rel h' \ a_host_shadow_root_rel h')" by(simp add: \parent_child_rel h' \ a_host_shadow_root_rel h' = parent_child_rel h3 \ a_host_shadow_root_rel h3\) show " heap_is_wellformed h' " using \acyclic (parent_child_rel h' \ local.a_host_shadow_root_rel h')\ by(simp add: heap_is_wellformed_def CD.heap_is_wellformed_impl \local.CD.a_heap_is_wellformed h'\ \local.a_all_ptrs_in_heap h'\ \local.a_distinct_lists h'\ \local.a_shadow_root_valid h'\) qed end interpretation i_create_element_wf?: l_create_element_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr known_ptrs type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed parent_child_rel set_tag_name set_tag_name_locs set_disconnected_nodes set_disconnected_nodes_locs create_element get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs DocumentClass.known_ptr DocumentClass.type_wf by(auto simp add: l_create_element_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_create_element_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \create\_character\_data\ locale l_create_character_data_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs + l_create_character_data\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs set_val set_val_locs create_character_data known_ptr type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_new_character_data_get_disconnected_nodes get_disconnected_nodes get_disconnected_nodes_locs + l_set_val_get_disconnected_nodes type_wf set_val set_val_locs get_disconnected_nodes get_disconnected_nodes_locs + l_new_character_data_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs + l_set_val_get_child_nodes type_wf set_val set_val_locs known_ptr get_child_nodes get_child_nodes_locs + l_set_disconnected_nodes_get_child_nodes set_disconnected_nodes set_disconnected_nodes_locs get_child_nodes get_child_nodes_locs + l_set_disconnected_nodes type_wf set_disconnected_nodes set_disconnected_nodes_locs + l_set_disconnected_nodes_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs + l_set_val_get_shadow_root type_wf set_val set_val_locs get_shadow_root get_shadow_root_locs + l_set_disconnected_nodes_get_shadow_root set_disconnected_nodes set_disconnected_nodes_locs get_shadow_root get_shadow_root_locs + l_new_character_data_get_tag_name get_tag_name get_tag_name_locs + l_set_val_get_tag_name type_wf set_val set_val_locs get_tag_name get_tag_name_locs + l_get_tag_name type_wf get_tag_name get_tag_name_locs + l_set_disconnected_nodes_get_tag_name type_wf set_disconnected_nodes set_disconnected_nodes_locs get_tag_name get_tag_name_locs + l_new_character_data type_wf + l_known_ptrs known_ptr known_ptrs for known_ptr :: "(_::linorder) object_ptr \ bool" and known_ptrs :: "(_) heap \ bool" and type_wf :: "(_) heap \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and set_tag_name :: "(_) element_ptr \ char list \ ((_) heap, exception, unit) prog" and set_tag_name_locs :: "(_) element_ptr \ ((_) heap, exception, unit) prog set" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" and create_element :: "(_) document_ptr \ char list \ ((_) heap, exception, (_) element_ptr) prog" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_disconnected_document :: "(_) node_ptr \ ((_) heap, exception, (_) document_ptr) prog" and get_disconnected_document_locs :: "((_) heap \ (_) heap \ bool) set" and set_val :: "(_) character_data_ptr \ char list \ ((_) heap, exception, unit) prog" and set_val_locs :: "(_) character_data_ptr \ ((_) heap, exception, unit) prog set" and create_character_data :: "(_) document_ptr \ char list \ ((_) heap, exception, (_) character_data_ptr) prog" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_::linorder) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" begin lemma create_character_data_preserves_wellformedness: assumes "heap_is_wellformed h" and "h \ create_character_data document_ptr text \\<^sub>h h'" and "type_wf h" and "known_ptrs h" shows "heap_is_wellformed h'" and "type_wf h'" and "known_ptrs h'" proof - obtain new_character_data_ptr h2 h3 disc_nodes_h3 where new_character_data_ptr: "h \ new_character_data \\<^sub>r new_character_data_ptr" and h2: "h \ new_character_data \\<^sub>h h2" and h3: "h2 \ set_val new_character_data_ptr text \\<^sub>h h3" and disc_nodes_h3: "h3 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" and h': "h3 \ set_disconnected_nodes document_ptr (cast new_character_data_ptr # disc_nodes_h3) \\<^sub>h h'" using assms(2) by(auto simp add: CD.create_character_data_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF CD.get_disconnected_nodes_pure, rotated] ) then have "h \ create_character_data document_ptr text \\<^sub>r new_character_data_ptr" apply(auto simp add: CD.create_character_data_def intro!: bind_returns_result_I)[1] apply (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust) apply (metis is_OK_returns_heap_E is_OK_returns_result_I local.CD.get_disconnected_nodes_pure pure_returns_heap_eq) by (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust) have "new_character_data_ptr \ set |h \ character_data_ptr_kinds_M|\<^sub>r" using new_character_data_ptr CharacterDataMonad.ptr_kinds_ptr_kinds_M h2 using new_character_data_ptr_not_in_heap by blast then have "cast new_character_data_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r" by simp then have "cast new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r" by simp have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\| {|cast new_character_data_ptr|}" using new_character_data_new_ptr h2 new_character_data_ptr by blast then have node_ptr_kinds_eq_h: "node_ptr_kinds h2 = node_ptr_kinds h |\| {|cast new_character_data_ptr|}" apply(simp add: node_ptr_kinds_def) by force then have character_data_ptr_kinds_eq_h: "character_data_ptr_kinds h2 = character_data_ptr_kinds h |\| {|new_character_data_ptr|}" apply(simp add: character_data_ptr_kinds_def) by force have element_ptr_kinds_eq_h: "element_ptr_kinds h2 = element_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: node_ptr_kinds_def element_ptr_kinds_def) have document_ptr_kinds_eq_h: "document_ptr_kinds h2 = document_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: document_ptr_kinds_def) have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF CD.set_val_writes h3]) using CD.set_val_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2" using object_ptr_kinds_eq_h2 by(auto simp add: node_ptr_kinds_def) have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_disconnected_nodes_writes h']) using set_disconnected_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3" using object_ptr_kinds_eq_h3 by(auto simp add: node_ptr_kinds_def) have "known_ptr (cast new_character_data_ptr)" using \h \ create_character_data document_ptr text \\<^sub>r new_character_data_ptr\ local.create_character_data_known_ptr by blast then have "known_ptrs h2" using known_ptrs_new_ptr object_ptr_kinds_eq_h \known_ptrs h\ h2 by blast then have "known_ptrs h3" using known_ptrs_preserved object_ptr_kinds_eq_h2 by blast then show "known_ptrs h'" using known_ptrs_preserved object_ptr_kinds_eq_h3 by blast have "document_ptr |\| document_ptr_kinds h" using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2 CD.get_disconnected_nodes_ptr_in_heap \type_wf h\ document_ptr_kinds_def by (metis is_OK_returns_result_I) have children_eq_h: "\(ptr'::(_) object_ptr) children. ptr' \ cast new_character_data_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads h2 get_child_nodes_new_character_data[rotated, OF new_character_data_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have children_eq2_h: "\ptr'. ptr' \ cast new_character_data_ptr \ |h \ get_child_nodes ptr'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\| {|cast new_character_data_ptr|}" using new_character_data_new_ptr h2 new_character_data_ptr by blast then have node_ptr_kinds_eq_h: "node_ptr_kinds h2 = node_ptr_kinds h |\| {|cast new_character_data_ptr|}" apply(simp add: node_ptr_kinds_def) by force then have character_data_ptr_kinds_eq_h: "character_data_ptr_kinds h2 = character_data_ptr_kinds h |\| {|new_character_data_ptr|}" apply(simp add: character_data_ptr_kinds_def) by force have element_ptr_kinds_eq_h: "element_ptr_kinds h2 = element_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: node_ptr_kinds_def element_ptr_kinds_def) have document_ptr_kinds_eq_h: "document_ptr_kinds h2 = document_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: document_ptr_kinds_def) have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF CD.set_val_writes h3]) using CD.set_val_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2" using object_ptr_kinds_eq_h2 by(auto simp add: node_ptr_kinds_def) then have character_data_ptr_kinds_eq_h2: "character_data_ptr_kinds h3 = character_data_ptr_kinds h2" by(simp add: character_data_ptr_kinds_def) have element_ptr_kinds_eq_h2: "element_ptr_kinds h3 = element_ptr_kinds h2" using node_ptr_kinds_eq_h2 by(simp add: element_ptr_kinds_def) have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_disconnected_nodes_writes h']) using set_disconnected_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3" using object_ptr_kinds_eq_h3 by(auto simp add: node_ptr_kinds_def) then have character_data_ptr_kinds_eq_h3: "character_data_ptr_kinds h' = character_data_ptr_kinds h3" by(simp add: character_data_ptr_kinds_def) have element_ptr_kinds_eq_h3: "element_ptr_kinds h' = element_ptr_kinds h3" using node_ptr_kinds_eq_h3 by(simp add: element_ptr_kinds_def) have "document_ptr |\| document_ptr_kinds h" using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2 CD.get_disconnected_nodes_ptr_in_heap \type_wf h\ document_ptr_kinds_def by (metis is_OK_returns_result_I) have children_eq_h: "\(ptr'::(_) object_ptr) children. ptr' \ cast new_character_data_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads h2 get_child_nodes_new_character_data[rotated, OF new_character_data_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have children_eq2_h: "\ptr'. ptr' \ cast new_character_data_ptr \ |h \ get_child_nodes ptr'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have "h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []" using new_character_data_ptr h2 new_character_data_ptr_in_heap[OF h2 new_character_data_ptr] new_character_data_is_character_data_ptr[OF new_character_data_ptr] new_character_data_no_child_nodes by blast have disconnected_nodes_eq_h: "\doc_ptr disc_nodes. h \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using CD.get_disconnected_nodes_reads h2 get_disconnected_nodes_new_character_data[OF new_character_data_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have disconnected_nodes_eq2_h: "\doc_ptr. |h \ get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have tag_name_eq_h: "\ptr' disc_nodes. h \ get_tag_name ptr' \\<^sub>r disc_nodes = h2 \ get_tag_name ptr' \\<^sub>r disc_nodes" using get_tag_name_reads h2 get_tag_name_new_character_data[OF new_character_data_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have tag_name_eq2_h: "\ptr'. |h \ get_tag_name ptr'|\<^sub>r = |h2 \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. h2 \ get_child_nodes ptr' \\<^sub>r children = h3 \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads CD.set_val_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_val_get_child_nodes) then have children_eq2_h2: "\ptr'. |h2 \ get_child_nodes ptr'|\<^sub>r = |h3 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using CD.get_disconnected_nodes_reads CD.set_val_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_val_get_disconnected_nodes) then have disconnected_nodes_eq2_h2: "\doc_ptr. |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have tag_name_eq_h2: "\ptr' disc_nodes. h2 \ get_tag_name ptr' \\<^sub>r disc_nodes = h3 \ get_tag_name ptr' \\<^sub>r disc_nodes" using get_tag_name_reads CD.set_val_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_val_get_tag_name) then have tag_name_eq2_h2: "\ptr'. |h2 \ get_tag_name ptr'|\<^sub>r = |h3 \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have "type_wf h2" using \type_wf h\ new_character_data_types_preserved h2 by blast then have "type_wf h3" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF CD.set_val_writes h3] using set_val_types_preserved by(auto simp add: reflp_def transp_def) then show "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h'] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) have children_eq_h3: "\ptr' children. h3 \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_disconnected_nodes_get_child_nodes) then have children_eq2_h3: " \ptr'. |h3 \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h3: "\doc_ptr disc_nodes. document_ptr \ doc_ptr \ h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using CD.get_disconnected_nodes_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers) then have disconnected_nodes_eq2_h3: "\doc_ptr. document_ptr \ doc_ptr \ |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have tag_name_eq_h3: "\ptr' disc_nodes. h3 \ get_tag_name ptr' \\<^sub>r disc_nodes = h' \ get_tag_name ptr' \\<^sub>r disc_nodes" using get_tag_name_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_disconnected_nodes_get_tag_name) then have tag_name_eq2_h3: "\ptr'. |h3 \ get_tag_name ptr'|\<^sub>r = |h' \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have disc_nodes_document_ptr_h2: "h2 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" using disconnected_nodes_eq_h2 disc_nodes_h3 by auto then have disc_nodes_document_ptr_h: "h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" using disconnected_nodes_eq_h by auto then have "cast new_character_data_ptr \ set disc_nodes_h3" using \heap_is_wellformed h\ using \cast new_character_data_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ a_all_ptrs_in_heap_def heap_is_wellformed_def using NodeMonad.ptr_kinds_ptr_kinds_M local.heap_is_wellformed_disc_nodes_in_heap by blast have "acyclic (parent_child_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def) also have "parent_child_rel h = parent_child_rel h2" proof(auto simp add: CD.parent_child_rel_def)[1] fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h2" by (simp add: object_ptr_kinds_eq_h) next fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "x \ set |h2 \ get_child_nodes a|\<^sub>r" by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2_h) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h" using object_ptr_kinds_eq_h \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ by(auto) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "x \ set |h \ get_child_nodes a|\<^sub>r" by (metis (no_types, lifting) \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ children_eq2_h empty_iff empty_set image_eqI select_result_I2) qed also have "\ = parent_child_rel h3" by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2) also have "\ = parent_child_rel h'" by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3) finally have "CD.a_acyclic_heap h'" by (simp add: CD.acyclic_heap_def) have "CD.a_all_ptrs_in_heap h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_all_ptrs_in_heap h2" apply(auto simp add: CD.a_all_ptrs_in_heap_def)[1] using node_ptr_kinds_eq_h \cast new_character_data_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ apply (metis (no_types, lifting) NodeMonad.ptr_kinds_ptr_kinds_M \parent_child_rel h = parent_child_rel h2\ children_eq2_h finite_set_in finsert_iff funion_finsert_right CD.parent_child_rel_child CD.parent_child_rel_parent_in_heap node_ptr_kinds_commutes object_ptr_kinds_eq_h select_result_I2 subsetD sup_bot.right_neutral) by (metis (no_types, lifting) CD.get_child_nodes_ok CD.get_child_nodes_ptr_in_heap \h2 \ get_child_nodes (cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_character_data_ptr) \\<^sub>r []\ assms(3) assms(4) children_eq_h disconnected_nodes_eq2_h document_ptr_kinds_eq_h finite_set_in is_OK_returns_result_I local.known_ptrs_known_ptr node_ptr_kinds_commutes returns_result_select_result subset_code(1)) then have "CD.a_all_ptrs_in_heap h3" by (simp add: children_eq2_h2 disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 object_ptr_kinds_eq_h2) then have "CD.a_all_ptrs_in_heap h'" by (smt character_data_ptr_kinds_commutes character_data_ptr_kinds_eq_h2 children_eq2_h3 disc_nodes_h3 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 h' h2 local.CD.a_all_ptrs_in_heap_def local.set_disconnected_nodes_get_disconnected_nodes new_character_data_ptr new_character_data_ptr_in_heap node_ptr_kinds_eq_h3 notin_fset object_ptr_kinds_eq_h3 select_result_I2 set_ConsD subset_code(1)) have "\p. p |\| object_ptr_kinds h \ cast new_character_data_ptr \ set |h \ get_child_nodes p|\<^sub>r" using \heap_is_wellformed h\ \cast new_character_data_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ heap_is_wellformed_children_in_heap by (meson NodeMonad.ptr_kinds_ptr_kinds_M CD.a_all_ptrs_in_heap_def assms(3) assms(4) fset_mp fset_of_list_elem CD.get_child_nodes_ok known_ptrs_known_ptr returns_result_select_result) then have "\p. p |\| object_ptr_kinds h2 \ cast new_character_data_ptr \ set |h2 \ get_child_nodes p|\<^sub>r" using children_eq2_h apply(auto simp add: object_ptr_kinds_eq_h)[1] using \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ apply auto[1] by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\) then have "\p. p |\| object_ptr_kinds h3 \ cast new_character_data_ptr \ set |h3 \ get_child_nodes p|\<^sub>r" using object_ptr_kinds_eq_h2 children_eq2_h2 by auto then have new_character_data_ptr_not_in_any_children: "\p. p |\| object_ptr_kinds h' \ cast new_character_data_ptr \ set |h' \ get_child_nodes p|\<^sub>r" using object_ptr_kinds_eq_h3 children_eq2_h3 by auto have "CD.a_distinct_lists h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_distinct_lists h2" using \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ apply(auto simp add: CD.a_distinct_lists_def object_ptr_kinds_eq_h document_ptr_kinds_eq_h disconnected_nodes_eq2_h intro!: distinct_concat_map_I)[1] - apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert) + apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert_remove) apply(case_tac "x=cast new_character_data_ptr") apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply (metis IntI assms(1) assms(3) assms(4) empty_iff CD.get_child_nodes_ok local.heap_is_wellformed_one_parent local.known_ptrs_known_ptr returns_result_select_result) apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] thm children_eq2_h using \CD.a_distinct_lists h\ \type_wf h2\ disconnected_nodes_eq_h document_ptr_kinds_eq_h CD.distinct_lists_no_parent get_disconnected_nodes_ok returns_result_select_result by metis then have "CD.a_distinct_lists h3" by(auto simp add: CD.a_distinct_lists_def disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 children_eq2_h2 object_ptr_kinds_eq_h2)[1] then have "CD.a_distinct_lists h'" proof(auto simp add: CD.a_distinct_lists_def disconnected_nodes_eq2_h3 children_eq2_h3 object_ptr_kinds_eq_h3 document_ptr_kinds_eq_h3 intro!: distinct_concat_map_I)[1] fix x assume "distinct (concat (map (\document_ptr. |h3 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" and "x |\| document_ptr_kinds h3" then show "distinct |h' \ get_disconnected_nodes x|\<^sub>r" using document_ptr_kinds_eq_h3 disconnected_nodes_eq_h3 h' set_disconnected_nodes_get_disconnected_nodes by (metis (no_types, hide_lams) \cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_character_data_ptr \ set disc_nodes_h3\ \type_wf h2\ assms(1) disc_nodes_document_ptr_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disconnected_nodes_eq_h distinct.simps(2) document_ptr_kinds_eq_h2 local.get_disconnected_nodes_ok local.heap_is_wellformed_disconnected_nodes_distinct returns_result_select_result select_result_I2) next fix x y xa assume "distinct (concat (map (\document_ptr. |h3 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" and "x |\| document_ptr_kinds h3" and "y |\| document_ptr_kinds h3" and "x \ y" and "xa \ set |h' \ get_disconnected_nodes x|\<^sub>r" and "xa \ set |h' \ get_disconnected_nodes y|\<^sub>r" moreover have "set |h3 \ get_disconnected_nodes x|\<^sub>r \ set |h3 \ get_disconnected_nodes y|\<^sub>r = {}" using calculation by(auto dest: distinct_concat_map_E(1)) ultimately show "False" using NodeMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_character_data_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ by (smt local.CD.a_all_ptrs_in_heap_def \CD.a_all_ptrs_in_heap h\ disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 finite_set_in h' l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms select_result_I2 set_ConsD subsetD) next fix x xa xb assume 2: "(\x\fset (object_ptr_kinds h3). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h3). set |h3 \ get_disconnected_nodes x|\<^sub>r) = {}" and 3: "xa |\| object_ptr_kinds h3" and 4: "x \ set |h' \ get_child_nodes xa|\<^sub>r" and 5: "xb |\| document_ptr_kinds h3" and 6: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" show "False" using disc_nodes_document_ptr_h disconnected_nodes_eq2_h3 apply(cases "document_ptr = xb") apply (metis (no_types, lifting) "3" "4" "5" "6" CD.distinct_lists_no_parent \local.CD.a_distinct_lists h2\ \type_wf h'\ children_eq2_h2 children_eq2_h3 disc_nodes_document_ptr_h2 document_ptr_kinds_eq_h3 h' local.get_disconnected_nodes_ok local.set_disconnected_nodes_get_disconnected_nodes new_character_data_ptr_not_in_any_children object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 returns_result_eq returns_result_select_result set_ConsD) by (metis "3" "4" "5" "6" CD.distinct_lists_no_parent \local.CD.a_distinct_lists h3\ \type_wf h3\ children_eq2_h3 local.get_disconnected_nodes_ok returns_result_select_result) qed have "CD.a_owner_document_valid h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_owner_document_valid h'" using disc_nodes_h3 \document_ptr |\| document_ptr_kinds h\ apply(simp add: CD.a_owner_document_valid_def) apply(simp add: object_ptr_kinds_eq_h object_ptr_kinds_eq_h3 ) apply(simp add: object_ptr_kinds_eq_h2) apply(simp add: document_ptr_kinds_eq_h document_ptr_kinds_eq_h3 ) apply(simp add: document_ptr_kinds_eq_h2) apply(simp add: node_ptr_kinds_eq_h node_ptr_kinds_eq_h3 ) apply(simp add: node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h ) apply(auto simp add: children_eq2_h2[symmetric] children_eq2_h3[symmetric] disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3)[1] apply (metis (no_types, lifting) document_ptr_kinds_eq_h h' list.set_intros(1) local.set_disconnected_nodes_get_disconnected_nodes select_result_I2) apply(simp add: object_ptr_kinds_eq_h) by (metis (mono_tags, lifting) \cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2_h disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h finite_set_in h' l_ptr_kinds_M.ptr_kinds_ptr_kinds_M l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes list.set_intros(2) local.l_set_disconnected_nodes_get_disconnected_nodes_axioms object_ptr_kinds_M_def select_result_I2) have shadow_root_ptr_kinds_eq_h: "shadow_root_ptr_kinds h2 = shadow_root_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: shadow_root_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h2: "shadow_root_ptr_kinds h3 = shadow_root_ptr_kinds h2" using object_ptr_kinds_eq_h2 by(auto simp add: shadow_root_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h3: "shadow_root_ptr_kinds h' = shadow_root_ptr_kinds h3" using object_ptr_kinds_eq_h3 by(auto simp add: shadow_root_ptr_kinds_def) have shadow_root_eq_h: "\character_data_ptr shadow_root_opt. h \ get_shadow_root character_data_ptr \\<^sub>r shadow_root_opt = h2 \ get_shadow_root character_data_ptr \\<^sub>r shadow_root_opt" using get_shadow_root_reads h2 get_shadow_root_new_character_data[rotated, OF h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] using local.get_shadow_root_locs_impl new_character_data_ptr apply blast using local.get_shadow_root_locs_impl new_character_data_ptr by blast have shadow_root_eq_h2: "\ptr' children. h2 \ get_shadow_root ptr' \\<^sub>r children = h3 \ get_shadow_root ptr' \\<^sub>r children" using get_shadow_root_reads set_val_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_val_get_shadow_root) have shadow_root_eq_h3: "\ptr' children. h3 \ get_shadow_root ptr' \\<^sub>r children = h' \ get_shadow_root ptr' \\<^sub>r children" using get_shadow_root_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) using set_disconnected_nodes_get_shadow_root by(auto simp add: set_disconnected_nodes_get_shadow_root) have "a_all_ptrs_in_heap h" by (simp add: assms(1) local.a_all_ptrs_in_heap_def local.get_shadow_root_shadow_root_ptr_in_heap) then have "a_all_ptrs_in_heap h2" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h)[1] using returns_result_eq shadow_root_eq_h by fastforce then have "a_all_ptrs_in_heap h3" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h2)[1] using shadow_root_eq_h2 by blast then have "a_all_ptrs_in_heap h'" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h3)[1] by (simp add: shadow_root_eq_h3) have "a_distinct_lists h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_distinct_lists h2" apply(auto simp add: a_distinct_lists_def character_data_ptr_kinds_eq_h)[1] apply(auto simp add: distinct_insort intro!: distinct_concat_map_I split: option.splits)[1] by (metis \type_wf h2\ assms(1) assms(3) local.get_shadow_root_ok local.shadow_root_same_host returns_result_select_result shadow_root_eq_h) then have "a_distinct_lists h3" by(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h2 select_result_eq[OF shadow_root_eq_h2]) then have "a_distinct_lists h'" by(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h3 select_result_eq[OF shadow_root_eq_h3]) have "a_shadow_root_valid h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_shadow_root_valid h2" by(auto simp add: a_shadow_root_valid_def shadow_root_ptr_kinds_eq_h element_ptr_kinds_eq_h select_result_eq[OF shadow_root_eq_h] tag_name_eq2_h) then have "a_shadow_root_valid h3" by(auto simp add: a_shadow_root_valid_def shadow_root_ptr_kinds_eq_h2 element_ptr_kinds_eq_h2 select_result_eq[OF shadow_root_eq_h2] tag_name_eq2_h2) then have "a_shadow_root_valid h'" by(auto simp add: a_shadow_root_valid_def shadow_root_ptr_kinds_eq_h3 element_ptr_kinds_eq_h3 select_result_eq[OF shadow_root_eq_h3] tag_name_eq2_h3) have "a_host_shadow_root_rel h = a_host_shadow_root_rel h2" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h select_result_eq[OF shadow_root_eq_h]) have "a_host_shadow_root_rel h2 = a_host_shadow_root_rel h3" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h2 select_result_eq[OF shadow_root_eq_h2]) have "a_host_shadow_root_rel h3 = a_host_shadow_root_rel h'" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h3 select_result_eq[OF shadow_root_eq_h3]) have "acyclic (parent_child_rel h \ a_host_shadow_root_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) have "parent_child_rel h \ a_host_shadow_root_rel h = parent_child_rel h2 \ a_host_shadow_root_rel h2" using \local.a_host_shadow_root_rel h = local.a_host_shadow_root_rel h2\ \parent_child_rel h = parent_child_rel h2\ by auto have "parent_child_rel h2 \ a_host_shadow_root_rel h2 = parent_child_rel h3 \ a_host_shadow_root_rel h3" using \local.a_host_shadow_root_rel h2 = local.a_host_shadow_root_rel h3\ \parent_child_rel h2 = parent_child_rel h3\ by auto have "parent_child_rel h' \ a_host_shadow_root_rel h' = parent_child_rel h3 \ a_host_shadow_root_rel h3" by (simp add: \local.a_host_shadow_root_rel h3 = local.a_host_shadow_root_rel h'\ \parent_child_rel h3 = parent_child_rel h'\) have "acyclic (parent_child_rel h3 \ a_host_shadow_root_rel h3)" using \acyclic (parent_child_rel h \ local.a_host_shadow_root_rel h)\ \parent_child_rel h \ local.a_host_shadow_root_rel h = parent_child_rel h2 \ local.a_host_shadow_root_rel h2\ \parent_child_rel h2 \ local.a_host_shadow_root_rel h2 = parent_child_rel h3 \ local.a_host_shadow_root_rel h3\ by auto then have "acyclic (parent_child_rel h' \ a_host_shadow_root_rel h')" by(simp add: \parent_child_rel h' \ a_host_shadow_root_rel h' = parent_child_rel h3 \ a_host_shadow_root_rel h3\) have "CD.a_heap_is_wellformed h'" apply(simp add: CD.a_heap_is_wellformed_def) by (simp add: \local.CD.a_acyclic_heap h'\ \local.CD.a_all_ptrs_in_heap h'\ \local.CD.a_distinct_lists h'\ \local.CD.a_owner_document_valid h'\) show "heap_is_wellformed h' " using \acyclic (parent_child_rel h' \ local.a_host_shadow_root_rel h')\ by(simp add: heap_is_wellformed_def CD.heap_is_wellformed_impl \local.CD.a_heap_is_wellformed h'\ \local.a_all_ptrs_in_heap h'\ \local.a_distinct_lists h'\ \local.a_shadow_root_valid h'\) qed end subsubsection \create\_document\ locale l_create_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs + l_new_document_get_disconnected_nodes get_disconnected_nodes get_disconnected_nodes_locs + l_create_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M create_document + l_new_document_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs + l_get_tag_name type_wf get_tag_name get_tag_name_locs + l_new_document_get_tag_name get_tag_name get_tag_name_locs + l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs + l_new_document type_wf + l_known_ptrs known_ptr known_ptrs for known_ptr :: "(_::linorder) object_ptr \ bool" and type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_disconnected_document :: "(_) node_ptr \ ((_) heap, exception, (_) document_ptr) prog" and get_disconnected_document_locs :: "((_) heap \ (_) heap \ bool) set" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and set_val :: "(_) character_data_ptr \ char list \ ((_) heap, exception, unit) prog" and set_val_locs :: "(_) character_data_ptr \ ((_) heap, exception, unit) prog set" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" and create_document :: "((_) heap, exception, (_) document_ptr) prog" and known_ptrs :: "(_) heap \ bool" begin lemma create_document_preserves_wellformedness: assumes "heap_is_wellformed h" and "h \ create_document \\<^sub>h h'" and "type_wf h" and "known_ptrs h" shows "heap_is_wellformed h'" proof - obtain new_document_ptr where new_document_ptr: "h \ new_document \\<^sub>r new_document_ptr" and h': "h \ new_document \\<^sub>h h'" using assms(2) apply(simp add: create_document_def) using new_document_ok by blast have "new_document_ptr \ set |h \ document_ptr_kinds_M|\<^sub>r" using new_document_ptr DocumentMonad.ptr_kinds_ptr_kinds_M using new_document_ptr_not_in_heap h' by blast then have "cast new_document_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r" by simp have "new_document_ptr |\| document_ptr_kinds h" using new_document_ptr DocumentMonad.ptr_kinds_ptr_kinds_M using new_document_ptr_not_in_heap h' by blast then have "cast new_document_ptr |\| object_ptr_kinds h" by simp have object_ptr_kinds_eq: "object_ptr_kinds h' = object_ptr_kinds h |\| {|cast new_document_ptr|}" using new_document_new_ptr h' new_document_ptr by blast then have node_ptr_kinds_eq: "node_ptr_kinds h' = node_ptr_kinds h" apply(simp add: node_ptr_kinds_def) by force then have character_data_ptr_kinds_eq_h: "character_data_ptr_kinds h' = character_data_ptr_kinds h" by(simp add: character_data_ptr_kinds_def) have element_ptr_kinds_eq_h: "element_ptr_kinds h' = element_ptr_kinds h" using object_ptr_kinds_eq by(auto simp add: node_ptr_kinds_def element_ptr_kinds_def) have document_ptr_kinds_eq_h: "document_ptr_kinds h' = document_ptr_kinds h |\| {|new_document_ptr|}" using object_ptr_kinds_eq apply(auto simp add: document_ptr_kinds_def)[1] by (metis (no_types, lifting) document_ptr_kinds_commutes document_ptr_kinds_def finsertI1 fset.map_comp) have shadow_root_ptr_kinds_eq: "shadow_root_ptr_kinds h' = shadow_root_ptr_kinds h" using object_ptr_kinds_eq apply(simp add: shadow_root_ptr_kinds_def) by force have children_eq: "\(ptr'::(_) object_ptr) children. ptr' \ cast new_document_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads h' get_child_nodes_new_document[rotated, OF new_document_ptr h'] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have children_eq2: "\ptr'. ptr' \ cast new_document_ptr \ |h \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have "h' \ get_child_nodes (cast new_document_ptr) \\<^sub>r []" using new_document_ptr h' new_document_ptr_in_heap[OF h' new_document_ptr] new_document_is_document_ptr[OF new_document_ptr] new_document_no_child_nodes by blast have disconnected_nodes_eq_h: "\doc_ptr disc_nodes. doc_ptr \ new_document_ptr \ h \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using CD.get_disconnected_nodes_reads h' get_disconnected_nodes_new_document_different_pointers new_document_ptr apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by (metis(full_types) \\thesis. (\new_document_ptr. \h \ new_document \\<^sub>r new_document_ptr; h \ new_document \\<^sub>h h'\ \ thesis) \ thesis\ local.get_disconnected_nodes_new_document_different_pointers new_document_ptr)+ then have disconnected_nodes_eq2_h: "\doc_ptr. doc_ptr \ new_document_ptr \ |h \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have "h' \ get_disconnected_nodes new_document_ptr \\<^sub>r []" using h' local.new_document_no_disconnected_nodes new_document_ptr by blast have "type_wf h'" using \type_wf h\ new_document_types_preserved h' by blast have "acyclic (parent_child_rel h)" using \heap_is_wellformed h\ by (auto simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def) also have "parent_child_rel h = parent_child_rel h'" proof(auto simp add: CD.parent_child_rel_def)[1] fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h'" by (simp add: object_ptr_kinds_eq) next fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "x \ set |h' \ get_child_nodes a|\<^sub>r" by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast new_document_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2) next fix a x assume 0: "a |\| object_ptr_kinds h'" and 1: "x \ set |h' \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h" using object_ptr_kinds_eq \h' \ get_child_nodes (cast new_document_ptr) \\<^sub>r []\ by(auto) next fix a x assume 0: "a |\| object_ptr_kinds h'" and 1: "x \ set |h' \ get_child_nodes a|\<^sub>r" then show "x \ set |h \ get_child_nodes a|\<^sub>r" by (metis (no_types, lifting) \h' \ get_child_nodes (cast new_document_ptr) \\<^sub>r []\ children_eq2 empty_iff empty_set image_eqI select_result_I2) qed finally have "CD.a_acyclic_heap h'" by (simp add: CD.acyclic_heap_def) have "CD.a_all_ptrs_in_heap h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def ) then have "CD.a_all_ptrs_in_heap h'" apply(auto simp add: CD.a_all_ptrs_in_heap_def)[1] using ObjectMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ \parent_child_rel h = parent_child_rel h'\ assms(1) children_eq fset_of_list_elem local.heap_is_wellformed_children_in_heap CD.parent_child_rel_child CD.parent_child_rel_parent_in_heap node_ptr_kinds_eq apply (metis (no_types, lifting) \h' \ get_child_nodes (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr) \\<^sub>r []\ children_eq2 finite_set_in finsert_iff funion_finsert_right object_ptr_kinds_eq select_result_I2 subsetD sup_bot.right_neutral) by (metis (no_types, lifting) \h' \ get_disconnected_nodes new_document_ptr \\<^sub>r []\ \type_wf h'\ assms(1) disconnected_nodes_eq_h empty_iff empty_set local.get_disconnected_nodes_ok local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq returns_result_select_result select_result_I2) have "CD.a_distinct_lists h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_distinct_lists h'" using \h' \ get_disconnected_nodes new_document_ptr \\<^sub>r []\ \h' \ get_child_nodes (cast new_document_ptr) \\<^sub>r []\ apply(auto simp add: children_eq2[symmetric] CD.a_distinct_lists_def insort_split object_ptr_kinds_eq document_ptr_kinds_eq_h disconnected_nodes_eq2_h intro!: distinct_concat_map_I)[1] - apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert) + apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert_remove) apply(auto simp add: dest: distinct_concat_map_E)[1] apply(auto simp add: dest: distinct_concat_map_E)[1] using \new_document_ptr |\| document_ptr_kinds h\ apply(auto simp add: distinct_insort dest: distinct_concat_map_E)[1] apply (metis assms(1) assms(3) disconnected_nodes_eq2_h get_disconnected_nodes_ok local.heap_is_wellformed_disconnected_nodes_distinct returns_result_select_result) proof - fix x :: "(_) document_ptr" and y :: "(_) document_ptr" and xa :: "(_) node_ptr" assume a1: "x \ y" assume a2: "x |\| document_ptr_kinds h" assume a3: "x \ new_document_ptr" assume a4: "y |\| document_ptr_kinds h" assume a5: "y \ new_document_ptr" assume a6: "distinct (concat (map (\document_ptr. |h \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h)))))" assume a7: "xa \ set |h' \ get_disconnected_nodes x|\<^sub>r" assume a8: "xa \ set |h' \ get_disconnected_nodes y|\<^sub>r" have f9: "xa \ set |h \ get_disconnected_nodes x|\<^sub>r" using a7 a3 disconnected_nodes_eq2_h by presburger have f10: "xa \ set |h \ get_disconnected_nodes y|\<^sub>r" using a8 a5 disconnected_nodes_eq2_h by presburger have f11: "y \ set (sorted_list_of_set (fset (document_ptr_kinds h)))" using a4 by simp have "x \ set (sorted_list_of_set (fset (document_ptr_kinds h)))" using a2 by simp then show False using f11 f10 f9 a6 a1 by (meson disjoint_iff_not_equal distinct_concat_map_E(1)) next fix x xa xb assume 0: "h' \ get_disconnected_nodes new_document_ptr \\<^sub>r []" and 1: "h' \ get_child_nodes (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr) \\<^sub>r []" and 2: "distinct (concat (map (\ptr. |h \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h)))))" and 3: "distinct (concat (map (\document_ptr. |h \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h)))))" and 4: "(\x\fset (object_ptr_kinds h). set |h \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h). set |h \ get_disconnected_nodes x|\<^sub>r) = {}" and 5: "x \ set |h \ get_child_nodes xa|\<^sub>r" and 6: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" and 7: "xa |\| object_ptr_kinds h" and 8: "xa \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr" and 9: "xb |\| document_ptr_kinds h" and 10: "xb \ new_document_ptr" then show "False" by (metis \CD.a_distinct_lists h\ assms(3) disconnected_nodes_eq2_h CD.distinct_lists_no_parent get_disconnected_nodes_ok returns_result_select_result) qed have "CD.a_owner_document_valid h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_owner_document_valid h'" apply(auto simp add: CD.a_owner_document_valid_def)[1] by (metis \cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr |\| object_ptr_kinds h\ children_eq2 disconnected_nodes_eq2_h document_ptr_kinds_commutes finite_set_in funion_iff node_ptr_kinds_eq object_ptr_kinds_eq) have shadow_root_eq_h: "\character_data_ptr shadow_root_opt. h \ get_shadow_root character_data_ptr \\<^sub>r shadow_root_opt = h' \ get_shadow_root character_data_ptr \\<^sub>r shadow_root_opt" using get_shadow_root_reads assms(2) get_shadow_root_new_document[rotated, OF h'] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] using local.get_shadow_root_locs_impl new_document_ptr apply blast using local.get_shadow_root_locs_impl new_document_ptr by blast have "a_all_ptrs_in_heap h" by (simp add: assms(1) local.a_all_ptrs_in_heap_def local.get_shadow_root_shadow_root_ptr_in_heap) then have "a_all_ptrs_in_heap h'" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq document_ptr_kinds_eq_h)[1] using shadow_root_eq_h by fastforce have "a_distinct_lists h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_distinct_lists h'" apply(auto simp add: a_distinct_lists_def character_data_ptr_kinds_eq_h)[1] apply(auto simp add: distinct_insort intro!: distinct_concat_map_I split: option.splits)[1] by (metis \type_wf h'\ assms(1) assms(3) local.get_shadow_root_ok local.shadow_root_same_host returns_result_select_result shadow_root_eq_h) have tag_name_eq_h: "\ptr' disc_nodes. h \ get_tag_name ptr' \\<^sub>r disc_nodes = h' \ get_tag_name ptr' \\<^sub>r disc_nodes" using get_tag_name_reads h' get_tag_name_new_document[OF new_document_ptr h'] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ have "a_shadow_root_valid h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_shadow_root_valid h'" using new_document_is_document_ptr[OF new_document_ptr] by(auto simp add: a_shadow_root_valid_def element_ptr_kinds_eq_h document_ptr_kinds_eq_h shadow_root_ptr_kinds_eq select_result_eq[OF shadow_root_eq_h] select_result_eq[OF tag_name_eq_h]) have "a_host_shadow_root_rel h = a_host_shadow_root_rel h'" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h select_result_eq[OF shadow_root_eq_h]) have "acyclic (parent_child_rel h \ a_host_shadow_root_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) moreover have "parent_child_rel h \ a_host_shadow_root_rel h = parent_child_rel h' \ a_host_shadow_root_rel h'" by (simp add: \local.a_host_shadow_root_rel h = local.a_host_shadow_root_rel h'\ \parent_child_rel h = parent_child_rel h'\) ultimately have "acyclic (parent_child_rel h' \ a_host_shadow_root_rel h')" by simp have "CD.a_heap_is_wellformed h'" apply(simp add: CD.a_heap_is_wellformed_def) by (simp add: \local.CD.a_acyclic_heap h'\ \local.CD.a_all_ptrs_in_heap h'\ \local.CD.a_distinct_lists h'\ \local.CD.a_owner_document_valid h'\) show "heap_is_wellformed h'" using CD.heap_is_wellformed_impl \acyclic (parent_child_rel h' \ local.a_host_shadow_root_rel h')\ \local.CD.a_heap_is_wellformed h'\ \local.a_all_ptrs_in_heap h'\ \local.a_distinct_lists h'\ \local.a_shadow_root_valid h'\ local.heap_is_wellformed_def by auto qed end interpretation l_create_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf DocumentClass.type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs heap_is_wellformed parent_child_rel set_val set_val_locs set_disconnected_nodes set_disconnected_nodes_locs create_document known_ptrs by(auto simp add: l_create_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) subsubsection \attach\_shadow\_root\ locale l_attach_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs + l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr set_shadow_root set_shadow_root_locs set_mode set_mode_locs attach_shadow_root type_wf get_tag_name get_tag_name_locs get_shadow_root get_shadow_root_locs + l_new_shadow_root_get_disconnected_nodes get_disconnected_nodes get_disconnected_nodes_locs + l_set_mode_get_disconnected_nodes type_wf set_mode set_mode_locs get_disconnected_nodes get_disconnected_nodes_locs + l_new_shadow_root_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs + l_new_shadow_root_get_tag_name type_wf get_tag_name get_tag_name_locs + l_set_mode_get_child_nodes type_wf set_mode set_mode_locs known_ptr get_child_nodes get_child_nodes_locs + l_set_shadow_root_get_child_nodes type_wf set_shadow_root set_shadow_root_locs known_ptr get_child_nodes get_child_nodes_locs + l_set_shadow_root type_wf set_shadow_root set_shadow_root_locs + l_set_shadow_root_get_disconnected_nodes set_shadow_root set_shadow_root_locs get_disconnected_nodes get_disconnected_nodes_locs + l_set_mode_get_shadow_root type_wf set_mode set_mode_locs get_shadow_root get_shadow_root_locs + l_set_shadow_root_get_shadow_root type_wf set_shadow_root set_shadow_root_locs get_shadow_root get_shadow_root_locs + l_new_character_data_get_tag_name get_tag_name get_tag_name_locs + l_set_mode_get_tag_name type_wf set_mode set_mode_locs get_tag_name get_tag_name_locs + l_get_tag_name type_wf get_tag_name get_tag_name_locs + l_set_shadow_root_get_tag_name set_shadow_root set_shadow_root_locs get_tag_name get_tag_name_locs + l_new_shadow_root type_wf + l_known_ptrs known_ptr known_ptrs for known_ptr :: "(_::linorder) object_ptr \ bool" and known_ptrs :: "(_) heap \ bool" and type_wf :: "(_) heap \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and set_tag_name :: "(_) element_ptr \ char list \ ((_) heap, exception, unit) prog" and set_tag_name_locs :: "(_) element_ptr \ ((_) heap, exception, unit) prog set" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" and create_element :: "(_) document_ptr \ char list \ ((_) heap, exception, (_) element_ptr) prog" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_disconnected_document :: "(_) node_ptr \ ((_) heap, exception, (_) document_ptr) prog" and get_disconnected_document_locs :: "((_) heap \ (_) heap \ bool) set" and set_val :: "(_) character_data_ptr \ char list \ ((_) heap, exception, unit) prog" and set_val_locs :: "(_) character_data_ptr \ ((_) heap, exception, unit) prog set" and create_character_data :: "(_) document_ptr \ char list \ ((_) heap, exception, (_) character_data_ptr) prog" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_::linorder) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and set_shadow_root :: "(_) element_ptr \ (_) shadow_root_ptr option \ (_, unit) dom_prog" and set_shadow_root_locs :: "(_) element_ptr \ (_, unit) dom_prog set" and set_mode :: "(_) shadow_root_ptr \ shadow_root_mode \ (_, unit) dom_prog" and set_mode_locs :: "(_) shadow_root_ptr \ (_, unit) dom_prog set" and attach_shadow_root :: "(_) element_ptr \ shadow_root_mode \ (_, (_) shadow_root_ptr) dom_prog" begin lemma attach_shadow_root_child_preserves: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ attach_shadow_root element_ptr new_mode \\<^sub>h h'" shows "type_wf h'" and "known_ptrs h'" and "heap_is_wellformed h'" proof - obtain h2 h3 new_shadow_root_ptr element_tag_name where element_tag_name: "h \ get_tag_name element_ptr \\<^sub>r element_tag_name" and "element_tag_name \ safe_shadow_root_element_types" and prev_shadow_root: "h \ get_shadow_root element_ptr \\<^sub>r None" and h2: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h2" and new_shadow_root_ptr: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr" and h3: "h2 \ set_mode new_shadow_root_ptr new_mode \\<^sub>h h3" and h': "h3 \ set_shadow_root element_ptr (Some new_shadow_root_ptr) \\<^sub>h h'" using assms(4) by(auto simp add: attach_shadow_root_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_tag_name_pure, rotated] bind_returns_heap_E2[rotated, OF get_shadow_root_pure, rotated] split: if_splits) have "h \ attach_shadow_root element_ptr new_mode \\<^sub>r new_shadow_root_ptr" thm bind_pure_returns_result_I[OF get_tag_name_pure] apply(unfold attach_shadow_root_def)[1] using element_tag_name apply(rule bind_pure_returns_result_I[OF get_tag_name_pure]) apply(rule bind_pure_returns_result_I) using \element_tag_name \ safe_shadow_root_element_types\ apply(simp) using \element_tag_name \ safe_shadow_root_element_types\ apply(simp) using prev_shadow_root apply(rule bind_pure_returns_result_I[OF get_shadow_root_pure]) apply(rule bind_pure_returns_result_I) apply(simp) apply(simp) using h2 new_shadow_root_ptr h3 h' by(auto intro!: bind_returns_result_I intro: is_OK_returns_result_E[OF is_OK_returns_heap_I[OF h3]] is_OK_returns_result_E[OF is_OK_returns_heap_I[OF h']]) have "new_shadow_root_ptr \ set |h \ shadow_root_ptr_kinds_M|\<^sub>r" using new_shadow_root_ptr ShadowRootMonad.ptr_kinds_ptr_kinds_M h2 using new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ptr_not_in_heap by blast then have "cast new_shadow_root_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r" by simp have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\| {|cast new_shadow_root_ptr|}" using new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_new_ptr h2 new_shadow_root_ptr by blast then have document_ptr_kinds_eq_h: "document_ptr_kinds h2 = document_ptr_kinds h" apply(simp add: document_ptr_kinds_def) by force have shadow_root_ptr_kinds_eq_h: "shadow_root_ptr_kinds h2 = shadow_root_ptr_kinds h |\| {|new_shadow_root_ptr|}" using object_ptr_kinds_eq_h apply(simp add: shadow_root_ptr_kinds_def) by force have element_ptr_kinds_eq_h: "element_ptr_kinds h2 = element_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: node_ptr_kinds_def element_ptr_kinds_def) have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_mode_writes h3]) using set_mode_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2" by (auto simp add: document_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h2: "shadow_root_ptr_kinds h3 = shadow_root_ptr_kinds h2" using object_ptr_kinds_eq_h2 by (auto simp add: shadow_root_ptr_kinds_def) have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2" using object_ptr_kinds_eq_h2 by(auto simp add: node_ptr_kinds_def) have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_shadow_root_writes h']) using set_shadow_root_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3" by (auto simp add: document_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h3: "shadow_root_ptr_kinds h' = shadow_root_ptr_kinds h3" using object_ptr_kinds_eq_h3 by (auto simp add: shadow_root_ptr_kinds_def) have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3" using object_ptr_kinds_eq_h3 by(auto simp add: node_ptr_kinds_def) have "known_ptr (cast new_shadow_root_ptr)" using \h \ attach_shadow_root element_ptr new_mode \\<^sub>r new_shadow_root_ptr\ create_shadow_root_known_ptr by blast then have "known_ptrs h2" using known_ptrs_new_ptr object_ptr_kinds_eq_h \known_ptrs h\ h2 by blast then have "known_ptrs h3" using known_ptrs_preserved object_ptr_kinds_eq_h2 by blast then show "known_ptrs h'" using known_ptrs_preserved object_ptr_kinds_eq_h3 by blast have "element_ptr |\| element_ptr_kinds h" by (meson \h \ attach_shadow_root element_ptr new_mode \\<^sub>r new_shadow_root_ptr\ is_OK_returns_result_I local.attach_shadow_root_element_ptr_in_heap) have children_eq_h: "\(ptr'::(_) object_ptr) children. ptr' \ cast new_shadow_root_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads h2 get_child_nodes_new_shadow_root[rotated, OF new_shadow_root_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have children_eq2_h: "\ptr'. ptr' \ cast new_shadow_root_ptr \ |h \ get_child_nodes ptr'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\| {|cast new_shadow_root_ptr|}" using new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_new_ptr h2 new_shadow_root_ptr object_ptr_kinds_eq_h by blast then have node_ptr_kinds_eq_h: "node_ptr_kinds h2 = node_ptr_kinds h" apply(simp add: node_ptr_kinds_def) by force then have character_data_ptr_kinds_eq_h: "character_data_ptr_kinds h2 = character_data_ptr_kinds h" apply(simp add: character_data_ptr_kinds_def) done have element_ptr_kinds_eq_h: "element_ptr_kinds h2 = element_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: node_ptr_kinds_def element_ptr_kinds_def) have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_mode_writes h3]) using set_mode_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2" using object_ptr_kinds_eq_h2 by(auto simp add: node_ptr_kinds_def) then have character_data_ptr_kinds_eq_h2: "character_data_ptr_kinds h3 = character_data_ptr_kinds h2" by(simp add: character_data_ptr_kinds_def) have element_ptr_kinds_eq_h2: "element_ptr_kinds h3 = element_ptr_kinds h2" using node_ptr_kinds_eq_h2 by(simp add: element_ptr_kinds_def) have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_shadow_root_writes h']) using set_shadow_root_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3" using object_ptr_kinds_eq_h3 by(auto simp add: node_ptr_kinds_def) then have character_data_ptr_kinds_eq_h3: "character_data_ptr_kinds h' = character_data_ptr_kinds h3" by(simp add: character_data_ptr_kinds_def) have element_ptr_kinds_eq_h3: "element_ptr_kinds h' = element_ptr_kinds h3" using node_ptr_kinds_eq_h3 by(simp add: element_ptr_kinds_def) have children_eq_h: "\(ptr'::(_) object_ptr) children. ptr' \ cast new_shadow_root_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads h2 get_child_nodes_new_shadow_root[rotated, OF new_shadow_root_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have children_eq2_h: "\ptr'. ptr' \ cast new_shadow_root_ptr \ |h \ get_child_nodes ptr'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have "h2 \ get_child_nodes (cast new_shadow_root_ptr) \\<^sub>r []" using h2 local.new_shadow_root_no_child_nodes new_shadow_root_ptr by auto have disconnected_nodes_eq_h: "\doc_ptr disc_nodes. h \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads h2 get_disconnected_nodes_new_shadow_root[rotated, OF h2,rotated,OF new_shadow_root_ptr] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by (metis (no_types, lifting))+ then have disconnected_nodes_eq2_h: "\doc_ptr. |h \ get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have tag_name_eq_h: "\ptr' disc_nodes. h \ get_tag_name ptr' \\<^sub>r disc_nodes = h2 \ get_tag_name ptr' \\<^sub>r disc_nodes" using get_tag_name_reads h2 get_tag_name_new_shadow_root[OF new_shadow_root_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have tag_name_eq2_h: "\ptr'. |h \ get_tag_name ptr'|\<^sub>r = |h2 \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. h2 \ get_child_nodes ptr' \\<^sub>r children = h3 \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads set_mode_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_mode_get_child_nodes) then have children_eq2_h2: "\ptr'. |h2 \ get_child_nodes ptr'|\<^sub>r = |h3 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_mode_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_mode_get_disconnected_nodes) then have disconnected_nodes_eq2_h2: "\doc_ptr. |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have tag_name_eq_h2: "\ptr' disc_nodes. h2 \ get_tag_name ptr' \\<^sub>r disc_nodes = h3 \ get_tag_name ptr' \\<^sub>r disc_nodes" using get_tag_name_reads set_mode_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_mode_get_tag_name) then have tag_name_eq2_h2: "\ptr'. |h2 \ get_tag_name ptr'|\<^sub>r = |h3 \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have "type_wf h2" using \type_wf h\ new_shadow_root_types_preserved h2 by blast then have "type_wf h3" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_mode_writes h3] using set_mode_types_preserved by(auto simp add: reflp_def transp_def) then show "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_shadow_root_writes h'] using set_shadow_root_types_preserved by(auto simp add: reflp_def transp_def) have children_eq_h3: "\ptr' children. h3 \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads set_shadow_root_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_shadow_root_get_child_nodes) then have children_eq2_h3: " \ptr'. |h3 \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h3: "\doc_ptr disc_nodes. h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_shadow_root_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_shadow_root_get_disconnected_nodes) then have disconnected_nodes_eq2_h3: "\doc_ptr. |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have tag_name_eq_h3: "\ptr' disc_nodes. h3 \ get_tag_name ptr' \\<^sub>r disc_nodes = h' \ get_tag_name ptr' \\<^sub>r disc_nodes" using get_tag_name_reads set_shadow_root_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_shadow_root_get_tag_name) then have tag_name_eq2_h3: "\ptr'. |h3 \ get_tag_name ptr'|\<^sub>r = |h' \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have "acyclic (parent_child_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def) also have "parent_child_rel h = parent_child_rel h2" proof(auto simp add: CD.parent_child_rel_def)[1] fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h2" by (simp add: object_ptr_kinds_eq_h) next fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "x \ set |h2 \ get_child_nodes a|\<^sub>r" by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast new_shadow_root_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2_h) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h" using object_ptr_kinds_eq_h \h2 \ get_child_nodes (cast new_shadow_root_ptr) \\<^sub>r []\ by(auto) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "x \ set |h \ get_child_nodes a|\<^sub>r" by (metis (no_types, lifting) \h2 \ get_child_nodes (cast new_shadow_root_ptr) \\<^sub>r []\ children_eq2_h empty_iff empty_set image_eqI select_result_I2) qed also have "\ = parent_child_rel h3" by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2) also have "\ = parent_child_rel h'" by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3) finally have "CD.a_acyclic_heap h'" by (simp add: CD.acyclic_heap_def) have "CD.a_all_ptrs_in_heap h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_all_ptrs_in_heap h2" apply(auto simp add: CD.a_all_ptrs_in_heap_def)[1] using node_ptr_kinds_eq_h \h2 \ get_child_nodes (cast new_shadow_root_ptr) \\<^sub>r []\ apply (metis (no_types, lifting) CD.get_child_nodes_ok CD.l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms \known_ptrs h2\ \parent_child_rel h = parent_child_rel h2\ \type_wf h2\ assms(1) assms(2) l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_child local.known_ptrs_known_ptr local.parent_child_rel_child_in_heap node_ptr_kinds_commutes returns_result_select_result) by (metis assms(1) assms(2) disconnected_nodes_eq2_h document_ptr_kinds_eq_h local.get_disconnected_nodes_ok local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq_h returns_result_select_result) then have "CD.a_all_ptrs_in_heap h3" by (simp add: children_eq2_h2 disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 object_ptr_kinds_eq_h2) then have "CD.a_all_ptrs_in_heap h'" by (simp add: children_eq2_h3 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h3 object_ptr_kinds_eq_h3) have "CD.a_distinct_lists h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_distinct_lists h2" using \h2 \ get_child_nodes (cast new_shadow_root_ptr) \\<^sub>r []\ children_eq2_h apply(auto simp add: select_result_eq[OF disconnected_nodes_eq_h] CD.a_distinct_lists_def insort_split object_ptr_kinds_eq_h document_ptr_kinds_eq_h disconnected_nodes_eq2_h intro!: distinct_concat_map_I dest: distinct_concat_map_E)[1] - apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert) + apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert_remove) apply(auto simp add: dest: distinct_concat_map_E)[1] apply(case_tac "x = cast new_shadow_root_ptr") using \h2 \ get_child_nodes (cast new_shadow_root_ptr) \\<^sub>r []\ children_eq2_h apply blast apply(case_tac "y = cast new_shadow_root_ptr") using \h2 \ get_child_nodes (cast new_shadow_root_ptr) \\<^sub>r []\ children_eq2_h apply blast proof - fix x y :: "(_) object_ptr" fix xa :: "(_) node_ptr" assume a1: "distinct (concat (map (\ptr. |h \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h)))))" assume "x \ y" assume "xa \ set |h2 \ get_child_nodes x|\<^sub>r" assume "xa \ set |h2 \ get_child_nodes y|\<^sub>r" assume "x |\| object_ptr_kinds h" assume "x \ cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_shadow_root_ptr" assume "y |\| object_ptr_kinds h" assume "y \ cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_shadow_root_ptr" show False using distinct_concat_map_E(1)[OF a1, of x y] using \x |\| object_ptr_kinds h\ \y |\| object_ptr_kinds h\ using \xa \ set |h2 \ get_child_nodes x|\<^sub>r\ \xa \ set |h2 \ get_child_nodes y|\<^sub>r\ using \x \ y\ by(auto simp add: children_eq2_h[OF \x \ cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_shadow_root_ptr\] children_eq2_h[OF \y \ cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_shadow_root_ptr\]) next fix x :: "(_) node_ptr" fix xa :: "(_) object_ptr" fix xb :: "(_) document_ptr" assume "(\x\fset (object_ptr_kinds h). set |h \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" assume "x \ set |h2 \ get_child_nodes xa|\<^sub>r" assume "xb |\| document_ptr_kinds h" assume "x \ set |h2 \ get_disconnected_nodes xb|\<^sub>r" assume "xa |\| object_ptr_kinds h" assume "xa \ cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_shadow_root_ptr" have "set |h \ get_child_nodes xa|\<^sub>r \ set |h2 \ get_disconnected_nodes xb|\<^sub>r = {}" by (metis (no_types, lifting) CD.get_child_nodes_ok \xa |\| object_ptr_kinds h\ \xb |\| document_ptr_kinds h\ assms(1) assms(2) assms(3) disconnected_nodes_eq2_h is_OK_returns_result_E local.get_disconnected_nodes_ok local.heap_is_wellformed_children_disc_nodes_different local.known_ptrs_known_ptr select_result_I2) then show "False" using \x \ set |h2 \ get_child_nodes xa|\<^sub>r\ \x \ set |h2 \ get_disconnected_nodes xb|\<^sub>r\ \xa \ cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_shadow_root_ptr\ children_eq2_h by auto qed then have "CD.a_distinct_lists h3" by(auto simp add: CD.a_distinct_lists_def disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 children_eq2_h2 object_ptr_kinds_eq_h2)[1] then have "CD.a_distinct_lists h'" by(auto simp add: CD.a_distinct_lists_def disconnected_nodes_eq2_h3 children_eq2_h3 object_ptr_kinds_eq_h3 document_ptr_kinds_eq_h3 intro!: distinct_concat_map_I) have "CD.a_owner_document_valid h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_owner_document_valid h'" (* using disc_nodes_h3 \document_ptr |\| document_ptr_kinds h\ *) apply(simp add: CD.a_owner_document_valid_def) apply(simp add: object_ptr_kinds_eq_h object_ptr_kinds_eq_h3 ) apply(simp add: object_ptr_kinds_eq_h2) apply(simp add: document_ptr_kinds_eq_h document_ptr_kinds_eq_h3 ) apply(simp add: document_ptr_kinds_eq_h2) apply(simp add: node_ptr_kinds_eq_h node_ptr_kinds_eq_h3 ) apply(simp add: node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h ) apply(auto simp add: children_eq2_h2[symmetric] children_eq2_h3[symmetric] disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3)[1] by (metis CD.get_child_nodes_ok CD.get_child_nodes_ptr_in_heap \cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_shadow_root_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ assms(2) assms(3) children_eq2_h children_eq_h document_ptr_kinds_eq_h finite_set_in is_OK_returns_result_I l_ptr_kinds_M.ptr_kinds_ptr_kinds_M local.known_ptrs_known_ptr object_ptr_kinds_M_def returns_result_select_result) have shadow_root_eq_h: "\character_data_ptr shadow_root_opt. h \ get_shadow_root character_data_ptr \\<^sub>r shadow_root_opt = h2 \ get_shadow_root character_data_ptr \\<^sub>r shadow_root_opt" using get_shadow_root_reads h2 get_shadow_root_new_shadow_root[rotated, OF h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] using local.get_shadow_root_locs_impl new_shadow_root_ptr apply blast using local.get_shadow_root_locs_impl new_shadow_root_ptr by blast have shadow_root_eq_h2: "\ptr' children. h2 \ get_shadow_root ptr' \\<^sub>r children = h3 \ get_shadow_root ptr' \\<^sub>r children" using get_shadow_root_reads set_mode_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_mode_get_shadow_root) have shadow_root_eq_h3: "\ptr' children. element_ptr \ ptr' \ h3 \ get_shadow_root ptr' \\<^sub>r children = h' \ get_shadow_root ptr' \\<^sub>r children" using get_shadow_root_reads set_shadow_root_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_shadow_root_get_shadow_root_different_pointers) have shadow_root_h3: "h' \ get_shadow_root element_ptr \\<^sub>r Some new_shadow_root_ptr" using \type_wf h3\ h' local.set_shadow_root_get_shadow_root by blast have "a_all_ptrs_in_heap h" by (simp add: assms(1) local.a_all_ptrs_in_heap_def local.get_shadow_root_shadow_root_ptr_in_heap) then have "a_all_ptrs_in_heap h2" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h)[1] using returns_result_eq shadow_root_eq_h by fastforce then have "a_all_ptrs_in_heap h3" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h2)[1] using shadow_root_eq_h2 by blast then have "a_all_ptrs_in_heap h'" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h3)[1] apply(case_tac "shadow_root_ptr = new_shadow_root_ptr") using h2 new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ptr_in_heap new_shadow_root_ptr shadow_root_ptr_kinds_eq_h2 apply blast using \type_wf h3\ h' local.set_shadow_root_get_shadow_root returns_result_eq shadow_root_eq_h3 apply fastforce done have "a_distinct_lists h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_distinct_lists h2" apply(auto simp add: a_distinct_lists_def character_data_ptr_kinds_eq_h)[1] apply(auto simp add: distinct_insort intro!: distinct_concat_map_I split: option.splits)[1] by (metis \type_wf h2\ assms(1) assms(2) local.get_shadow_root_ok local.shadow_root_same_host returns_result_select_result shadow_root_eq_h) then have "a_distinct_lists h3" by(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h2 select_result_eq[OF shadow_root_eq_h2]) then have "a_distinct_lists h'" apply(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h3 select_result_eq[OF shadow_root_eq_h3])[1] apply(auto simp add: distinct_insort intro!: distinct_concat_map_I split: option.splits)[1] by (smt \type_wf h3\ assms(1) assms(2) h' h2 local.get_shadow_root_ok local.get_shadow_root_shadow_root_ptr_in_heap local.set_shadow_root_get_shadow_root local.shadow_root_same_host new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ptr_not_in_heap new_shadow_root_ptr returns_result_select_result select_result_I2 shadow_root_eq_h shadow_root_eq_h2 shadow_root_eq_h3) have "a_shadow_root_valid h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_shadow_root_valid h'" proof(unfold a_shadow_root_valid_def; safe) fix shadow_root_ptr assume "\shadow_root_ptr\fset (shadow_root_ptr_kinds h). \host\fset (element_ptr_kinds h). |h \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types \ |h \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr" assume "a_shadow_root_valid h" assume "shadow_root_ptr \ fset (shadow_root_ptr_kinds h')" show "\host\fset (element_ptr_kinds h'). |h' \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types \ |h' \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr" proof (cases "shadow_root_ptr = new_shadow_root_ptr") case True have "element_ptr \ fset (element_ptr_kinds h')" by (simp add: \element_ptr |\| element_ptr_kinds h\ element_ptr_kinds_eq_h element_ptr_kinds_eq_h2 element_ptr_kinds_eq_h3) moreover have "|h' \ get_tag_name element_ptr|\<^sub>r \ safe_shadow_root_element_types" by (smt \\thesis. (\h2 h3 new_shadow_root_ptr element_tag_name. \h \ get_tag_name element_ptr \\<^sub>r element_tag_name; element_tag_name \ safe_shadow_root_element_types; h \ get_shadow_root element_ptr \\<^sub>r None; h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h2; h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr; h2 \ set_mode new_shadow_root_ptr new_mode \\<^sub>h h3; h3 \ set_shadow_root element_ptr (Some new_shadow_root_ptr) \\<^sub>h h'\ \ thesis) \ thesis\ select_result_I2 tag_name_eq2_h tag_name_eq2_h2 tag_name_eq2_h3) moreover have "|h' \ get_shadow_root element_ptr|\<^sub>r = Some shadow_root_ptr" using shadow_root_h3 by (simp add: True) ultimately show ?thesis by meson next case False then obtain host where host: "host \ fset (element_ptr_kinds h)" and "|h \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types" and "|h \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr" using \shadow_root_ptr \ fset (shadow_root_ptr_kinds h')\ using \\shadow_root_ptr\fset (shadow_root_ptr_kinds h). \host\fset (element_ptr_kinds h). |h \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types \ |h \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr\ apply(simp add: shadow_root_ptr_kinds_eq_h3 shadow_root_ptr_kinds_eq_h2 shadow_root_ptr_kinds_eq_h) by (meson finite_set_in) moreover have "host \ element_ptr" using calculation(3) prev_shadow_root by auto ultimately show ?thesis using element_ptr_kinds_eq_h3 element_ptr_kinds_eq_h2 element_ptr_kinds_eq_h by (smt \type_wf h'\ assms(2) finite_set_in local.get_shadow_root_ok returns_result_eq returns_result_select_result shadow_root_eq_h shadow_root_eq_h2 shadow_root_eq_h3 tag_name_eq2_h tag_name_eq2_h2 tag_name_eq2_h3) qed qed have "a_host_shadow_root_rel h = a_host_shadow_root_rel h2" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h select_result_eq[OF shadow_root_eq_h]) have "a_host_shadow_root_rel h2 = a_host_shadow_root_rel h3" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h2 select_result_eq[OF shadow_root_eq_h2]) have "a_host_shadow_root_rel h' = {(cast element_ptr, cast new_shadow_root_ptr)} \ a_host_shadow_root_rel h3" apply(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h3 )[1] apply(case_tac "element_ptr \ aa") using select_result_eq[OF shadow_root_eq_h3] apply (simp add: image_iff) using select_result_eq[OF shadow_root_eq_h3] apply (metis (no_types, lifting) \local.a_host_shadow_root_rel h = local.a_host_shadow_root_rel h2\ \local.a_host_shadow_root_rel h2 = local.a_host_shadow_root_rel h3\ \type_wf h3\ host_shadow_root_rel_def local.get_shadow_root_impl local.get_shadow_root_ok option.distinct(1) prev_shadow_root returns_result_select_result) apply (metis (mono_tags, lifting) \\ptr'. (\x. element_ptr \ ptr') \ |h3 \ get_shadow_root ptr'|\<^sub>r = |h' \ get_shadow_root ptr'|\<^sub>r\ case_prod_conv image_iff is_OK_returns_result_I mem_Collect_eq option.inject returns_result_eq returns_result_select_result shadow_root_h3) using element_ptr_kinds_eq_h3 local.get_shadow_root_ptr_in_heap shadow_root_h3 apply fastforce using Shadow_DOM.a_host_shadow_root_rel_def \\ptr'. (\x. element_ptr \ ptr') \ |h3 \ get_shadow_root ptr'|\<^sub>r = |h' \ get_shadow_root ptr'|\<^sub>r\ \type_wf h3\ case_prodE case_prodI host_shadow_root_rel_shadow_root image_iff local.get_shadow_root_impl local.get_shadow_root_ok mem_Collect_eq option.discI prev_shadow_root returns_result_select_result select_result_I2 shadow_root_eq_h shadow_root_eq_h2 apply(auto)[1] by (smt case_prodI mem_Collect_eq option.distinct(1) pair_imageI returns_result_eq returns_result_select_result) have "acyclic (parent_child_rel h \ a_host_shadow_root_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) have "parent_child_rel h \ a_host_shadow_root_rel h = parent_child_rel h2 \ a_host_shadow_root_rel h2" using \local.a_host_shadow_root_rel h = local.a_host_shadow_root_rel h2\ \parent_child_rel h = parent_child_rel h2\ by auto have "parent_child_rel h2 \ a_host_shadow_root_rel h2 = parent_child_rel h3 \ a_host_shadow_root_rel h3" using \local.a_host_shadow_root_rel h2 = local.a_host_shadow_root_rel h3\ \parent_child_rel h2 = parent_child_rel h3\ by auto have "parent_child_rel h' \ a_host_shadow_root_rel h' = {(cast element_ptr, cast new_shadow_root_ptr)} \ parent_child_rel h3 \ a_host_shadow_root_rel h3" by (simp add: \local.a_host_shadow_root_rel h' = {(cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r element_ptr, cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_shadow_root_ptr)} \ local.a_host_shadow_root_rel h3\ \parent_child_rel h3 = parent_child_rel h'\) have "\a b. (a, b) \ parent_child_rel h3 \ a \ cast new_shadow_root_ptr" using CD.parent_child_rel_parent_in_heap \parent_child_rel h = parent_child_rel h2\ \parent_child_rel h2 = parent_child_rel h3\ document_ptr_kinds_commutes by (metis h2 new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ptr_not_in_heap new_shadow_root_ptr shadow_root_ptr_kinds_commutes) moreover have "\a b. (a, b) \ a_host_shadow_root_rel h3 \ a \ cast new_shadow_root_ptr" using shadow_root_eq_h2 by(auto simp add: a_host_shadow_root_rel_def) moreover have "cast new_shadow_root_ptr \ {x. (x, cast element_ptr) \ (parent_child_rel h3 \ a_host_shadow_root_rel h3)\<^sup>*}" by (metis (no_types, lifting) UnE calculation(1) calculation(2) cast_shadow_root_ptr_not_node_ptr(1) converse_rtranclE mem_Collect_eq) moreover have "acyclic (parent_child_rel h3 \ a_host_shadow_root_rel h3)" using \acyclic (parent_child_rel h \ local.a_host_shadow_root_rel h)\ \parent_child_rel h \ local.a_host_shadow_root_rel h = parent_child_rel h2 \ local.a_host_shadow_root_rel h2\ \parent_child_rel h2 \ local.a_host_shadow_root_rel h2 = parent_child_rel h3 \ local.a_host_shadow_root_rel h3\ by auto ultimately have "acyclic (parent_child_rel h' \ a_host_shadow_root_rel h')" by(simp add: \parent_child_rel h' \ a_host_shadow_root_rel h' = {(cast element_ptr, cast new_shadow_root_ptr)} \ parent_child_rel h3 \ a_host_shadow_root_rel h3\) have "CD.a_heap_is_wellformed h'" apply(simp add: CD.a_heap_is_wellformed_def) by (simp add: \local.CD.a_acyclic_heap h'\ \local.CD.a_all_ptrs_in_heap h'\ \local.CD.a_distinct_lists h'\ \local.CD.a_owner_document_valid h'\) show "heap_is_wellformed h' " using \acyclic (parent_child_rel h' \ local.a_host_shadow_root_rel h')\ by(simp add: heap_is_wellformed_def CD.heap_is_wellformed_impl \local.CD.a_heap_is_wellformed h'\ \local.a_all_ptrs_in_heap h'\ \local.a_distinct_lists h'\ \local.a_shadow_root_valid h'\) qed end interpretation l_attach_shadow_root_wf?: l_attach_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr known_ptrs type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed parent_child_rel set_tag_name set_tag_name_locs set_disconnected_nodes set_disconnected_nodes_locs create_element get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs set_val set_val_locs create_character_data DocumentClass.known_ptr DocumentClass.type_wf set_shadow_root set_shadow_root_locs set_mode set_mode_locs attach_shadow_root by(auto simp add: l_attach_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_attach_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] end diff --git a/thys/Shadow_SC_DOM/Shadow_DOM.thy b/thys/Shadow_SC_DOM/Shadow_DOM.thy --- a/thys/Shadow_SC_DOM/Shadow_DOM.thy +++ b/thys/Shadow_SC_DOM/Shadow_DOM.thy @@ -1,12913 +1,12913 @@ (*********************************************************************************** * Copyright (c) 2016-2020 The University of Sheffield, UK * 2019-2020 University of Exeter, UK * * 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. * * 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 HOLDER 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. * * SPDX-License-Identifier: BSD-2-Clause ***********************************************************************************) section\The Shadow DOM\ theory Shadow_DOM imports "monads/ShadowRootMonad" Core_SC_DOM.Core_DOM begin abbreviation "safe_shadow_root_element_types \ {''article'', ''aside'', ''blockquote'', ''body'', ''div'', ''footer'', ''h1'', ''h2'', ''h3'', ''h4'', ''h5'', ''h6'', ''header'', ''main'', ''nav'', ''p'', ''section'', ''span''}" subsection \Function Definitions\ subsubsection \get\_child\_nodes\ locale l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = CD: l_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs begin definition get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) shadow_root_ptr \ unit \ (_, (_) node_ptr list) dom_prog" where "get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr _ = get_M shadow_root_ptr RShadowRoot.child_nodes" definition a_get_child_nodes_tups :: "(((_) object_ptr \ bool) \ ((_) object_ptr \ unit \ (_, (_) node_ptr list) dom_prog)) list" where "a_get_child_nodes_tups \ [(is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r, get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast)]" definition a_get_child_nodes :: "(_) object_ptr \ (_, (_) node_ptr list) dom_prog" where "a_get_child_nodes ptr = invoke (CD.a_get_child_nodes_tups @ a_get_child_nodes_tups) ptr ()" definition a_get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" where "a_get_child_nodes_locs ptr \ (if is_shadow_root_ptr_kind ptr then {preserved (get_M (the (cast ptr)) RShadowRoot.child_nodes)} else {}) \ CD.a_get_child_nodes_locs ptr" definition first_child :: "(_) object_ptr \ (_, (_) node_ptr option) dom_prog" where "first_child ptr = do { children \ a_get_child_nodes ptr; return (case children of [] \ None | child#_ \ Some child)}" end global_interpretation l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs defines get_child_nodes = l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_child_nodes and get_child_nodes_locs = l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_child_nodes_locs . locale l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_type_wf type_wf + l_known_ptr known_ptr + l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_get_child_nodes_defs get_child_nodes get_child_nodes_locs + CD: l_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M for type_wf :: "(_) heap \ bool" and known_ptr :: "(_) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ bool" and get_child_nodes :: "(_) object_ptr \ (_, (_) node_ptr list) dom_prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ (_, (_) node_ptr list) dom_prog" and get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" + assumes known_ptr_impl: "known_ptr = ShadowRootClass.known_ptr" assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" assumes get_child_nodes_impl: "get_child_nodes = a_get_child_nodes" assumes get_child_nodes_locs_impl: "get_child_nodes_locs = a_get_child_nodes_locs" begin lemmas get_child_nodes_def = get_child_nodes_impl[unfolded a_get_child_nodes_def get_child_nodes_def] lemmas get_child_nodes_locs_def = get_child_nodes_locs_impl[unfolded a_get_child_nodes_locs_def get_child_nodes_locs_def, folded CD.get_child_nodes_locs_impl] lemma get_child_nodes_ok: assumes "known_ptr ptr" assumes "type_wf h" assumes "ptr |\| object_ptr_kinds h" shows "h \ ok (get_child_nodes ptr)" using assms[unfolded known_ptr_impl type_wf_impl] apply(auto simp add: get_child_nodes_def)[1] apply(split CD.get_child_nodes_splits, rule conjI)+ using ShadowRootClass.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t CD.get_child_nodes_ok CD.known_ptr_impl CD.type_wf_impl apply blast apply(auto simp add: CD.known_ptr_impl a_get_child_nodes_tups_def get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ok dest!: known_ptr_new_shadow_root_ptr intro!: bind_is_OK_I2)[1] by(auto dest: get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ok split: option.splits) lemma get_child_nodes_ptr_in_heap: assumes "h \ get_child_nodes ptr \\<^sub>r children" shows "ptr |\| object_ptr_kinds h" using assms by(auto simp add: get_child_nodes_def invoke_ptr_in_heap dest: is_OK_returns_result_I) lemma get_child_nodes_pure [simp]: "pure (get_child_nodes ptr) h" apply (auto simp add: get_child_nodes_def a_get_child_nodes_tups_def)[1] apply(split CD.get_child_nodes_splits, rule conjI)+ apply(simp) apply(split invoke_splits, rule conjI)+ apply(simp) by(auto simp add: get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_pure_I) lemma get_child_nodes_reads: "reads (get_child_nodes_locs ptr) (get_child_nodes ptr) h h'" apply (simp add: get_child_nodes_def a_get_child_nodes_tups_def get_child_nodes_locs_def CD.get_child_nodes_locs_def) apply(split CD.get_child_nodes_splits, rule conjI)+ apply(auto intro!: reads_subset[OF CD.get_child_nodes_reads[unfolded CD.get_child_nodes_locs_def]] split: if_splits)[1] apply(split invoke_splits, rule conjI)+ apply(auto)[1] apply(auto simp add: get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro: reads_subset[OF reads_singleton] reads_subset[OF check_in_heap_reads] intro!: reads_bind_pure reads_subset[OF return_reads] split: option.splits)[1] done end interpretation i_get_child_nodes?: l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs by(simp add: l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_child_nodes_is_l_get_child_nodes [instances]: "l_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs" apply(auto simp add: l_get_child_nodes_def instances)[1] using get_child_nodes_reads get_child_nodes_ok get_child_nodes_ptr_in_heap get_child_nodes_pure by blast+ paragraph \new\_document\ locale l_new_document_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD: l_new_document_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M for type_wf :: "(_) heap \ bool" and known_ptr :: "(_) object_ptr \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ bool" and get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_child_nodes_new_document: "ptr' \ cast new_document_ptr \ h \ new_document \\<^sub>r new_document_ptr \ h \ new_document \\<^sub>h h' \ r \ get_child_nodes_locs ptr' \ r h h'" apply(auto simp add: get_child_nodes_locs_def)[1] using CD.get_child_nodes_new_document apply (metis document_ptr_casts_commute3 empty_iff is_document_ptr_kind_none new_document_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t option.case_eq_if shadow_root_ptr_casts_commute3 singletonD) by (simp add: CD.get_child_nodes_new_document) lemma new_document_no_child_nodes: "h \ new_document \\<^sub>r new_document_ptr \ h \ new_document \\<^sub>h h' \ h' \ get_child_nodes (cast new_document_ptr) \\<^sub>r []" apply(auto simp add: get_child_nodes_def)[1] apply(split CD.get_child_nodes_splits, rule conjI)+ using CD.new_document_no_child_nodes apply auto[1] by(auto simp add: DocumentClass.a_known_ptr_def CD.known_ptr_impl known_ptr_def dest!: new_document_is_document_ptr) end interpretation i_new_document_get_child_nodes?: l_new_document_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr get_child_nodes get_child_nodes_locs DocumentClass.type_wf DocumentClass.known_ptr Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs by(unfold_locales) declare l_new_document_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma new_document_get_child_nodes_is_l_new_document_get_child_nodes [instances]: "l_new_document_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs" using new_document_is_l_new_document get_child_nodes_is_l_get_child_nodes apply(simp add: l_new_document_get_child_nodes_def l_new_document_get_child_nodes_axioms_def) using get_child_nodes_new_document new_document_no_child_nodes by fast paragraph \new\_shadow\_root\ locale l_new_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M for type_wf :: "(_) heap \ bool" and known_ptr :: "(_) object_ptr \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ bool" and get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_child_nodes_new_shadow_root: "ptr' \ cast new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ r \ get_child_nodes_locs ptr' \ r h h'" apply(auto simp add: get_child_nodes_locs_def)[1] apply (metis document_ptr_casts_commute3 insert_absorb insert_not_empty is_document_ptr_kind_none new_shadow_root_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t option.case_eq_if shadow_root_ptr_casts_commute3 singletonD) apply(auto simp add: CD.get_child_nodes_locs_def)[1] using new_shadow_root_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t apply blast apply (smt insertCI new_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t singleton_iff) apply (metis document_ptr_casts_commute3 empty_iff new_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t singletonD) done lemma new_shadow_root_no_child_nodes: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ h' \ get_child_nodes (cast new_shadow_root_ptr) \\<^sub>r []" apply(auto simp add: get_child_nodes_def )[1] apply(split CD.get_child_nodes_splits, rule conjI)+ apply(auto simp add: CD.get_child_nodes_def CD.a_get_child_nodes_tups_def)[1] apply(split invoke_splits, rule conjI)+ using NodeClass.a_known_ptr_def known_ptr_not_character_data_ptr known_ptr_not_document_ptr known_ptr_not_element_ptr local.CD.known_ptr_impl apply blast apply(auto simp add: is_document_ptr_def cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def split: option.splits document_ptr.splits)[1] apply(auto simp add: is_character_data_ptr_def cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def split: option.splits document_ptr.splits)[1] apply(auto simp add: is_element_ptr_def cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def split: option.splits document_ptr.splits)[1] apply(auto simp add: a_get_child_nodes_tups_def)[1] apply(split invoke_splits, rule conjI)+ apply(auto simp add: is_shadow_root_ptr_def split: shadow_root_ptr.splits dest!: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_is_shadow_root_ptr)[1] apply(auto intro!: bind_pure_returns_result_I)[1] apply(drule(1) new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ptr_in_heap) apply(auto simp add: shadow_root_ptr_kinds_def document_ptr_kinds_def)[1] apply (metis check_in_heap_ptr_in_heap is_OK_returns_result_E old.unit.exhaust) using new_shadow_root_children by (simp add: new_shadow_root_children get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def) end interpretation i_new_shadow_root_get_child_nodes?: l_new_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr get_child_nodes get_child_nodes_locs DocumentClass.type_wf DocumentClass.known_ptr Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs by(unfold_locales) declare l_new_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def[instances] locale l_new_shadow_root_get_child_nodes = l_get_child_nodes + assumes get_child_nodes_new_shadow_root: "ptr' \ cast new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ r \ get_child_nodes_locs ptr' \ r h h'" assumes new_shadow_root_no_child_nodes: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ h' \ get_child_nodes (cast new_shadow_root_ptr) \\<^sub>r []" lemma new_shadow_root_get_child_nodes_is_l_new_shadow_root_get_child_nodes [instances]: "l_new_shadow_root_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs" apply(simp add: l_new_shadow_root_get_child_nodes_def l_new_shadow_root_get_child_nodes_axioms_def instances) using get_child_nodes_new_shadow_root new_shadow_root_no_child_nodes by fast paragraph \new\_element\ locale l_new_element_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_new_element_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_child_nodes_new_element: "ptr' \ cast new_element_ptr \ h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ r \ get_child_nodes_locs ptr' \ r h h'" by (auto simp add: get_child_nodes_locs_def CD.get_child_nodes_locs_def new_element_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t new_element_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t new_element_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t new_element_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t split: prod.splits if_splits option.splits elim!: bind_returns_result_E bind_returns_heap_E intro: is_element_ptr_kind_obtains) lemma new_element_no_child_nodes: "h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ h' \ get_child_nodes (cast new_element_ptr) \\<^sub>r []" apply(auto simp add: get_child_nodes_def a_get_child_nodes_tups_def split: prod.splits elim!: bind_returns_result_E bind_returns_heap_E)[1] apply(split CD.get_child_nodes_splits, rule conjI)+ using local.new_element_no_child_nodes apply auto[1] apply(auto simp add: invoke_def)[1] using case_optionE apply fastforce apply(auto simp add: new_element_ptr_in_heap get_child_nodes\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def check_in_heap_def new_element_child_nodes intro!: bind_pure_returns_result_I intro: new_element_is_element_ptr elim!: new_element_ptr_in_heap)[1] proof - assume " h \ new_element \\<^sub>r new_element_ptr" assume "h \ new_element \\<^sub>h h'" assume "\ is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr)" assume "\ known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr)" moreover have "known_ptr (cast new_element_ptr)" using new_element_is_element_ptr \h \ new_element \\<^sub>r new_element_ptr\ by(auto simp add: known_ptr_impl ShadowRootClass.a_known_ptr_def DocumentClass.a_known_ptr_def CharacterDataClass.a_known_ptr_def ElementClass.a_known_ptr_def) ultimately show "False" by(simp add: known_ptr_impl CD.known_ptr_impl ShadowRootClass.a_known_ptr_def is_document_ptr_kind_none) qed end interpretation i_new_element_get_child_nodes?: l_new_element_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs by(unfold_locales) declare l_new_element_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma new_element_get_child_nodes_is_l_new_element_get_child_nodes [instances]: "l_new_element_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs" using new_element_is_l_new_element get_child_nodes_is_l_get_child_nodes apply(auto simp add: l_new_element_get_child_nodes_def l_new_element_get_child_nodes_axioms_def)[1] using get_child_nodes_new_element new_element_no_child_nodes by fast+ subsubsection \delete\_shadow\_root\ locale l_delete_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_child_nodes_delete_shadow_root: "ptr' \ cast shadow_root_ptr \ h \ delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M shadow_root_ptr \\<^sub>h h' \ r \ get_child_nodes_locs ptr' \ r h h'" by(auto simp add: get_child_nodes_locs_def CD.get_child_nodes_locs_def delete_shadow_root_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t delete_shadow_root_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t delete_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t delete_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t split: if_splits intro: is_shadow_root_ptr_kind_obtains intro: is_shadow_root_ptr_kind_obtains delete_shadow_root_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t delete_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t simp add: shadow_root_ptr_casts_commute3 delete_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t intro!: delete_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t dest: document_ptr_casts_commute3 split: option.splits) end locale l_delete_shadow_root_get_child_nodes = l_get_child_nodes_defs + assumes get_child_nodes_delete_shadow_root: "ptr' \ cast shadow_root_ptr \ h \ delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M shadow_root_ptr \\<^sub>h h' \ r \ get_child_nodes_locs ptr' \ r h h'" interpretation l_delete_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs by(auto simp add: l_delete_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) lemma l_delete_shadow_root_get_child_nodes_get_child_nodes_locs [instances]: "l_delete_shadow_root_get_child_nodes get_child_nodes_locs" apply(auto simp add: l_delete_shadow_root_get_child_nodes_def)[1] using get_child_nodes_delete_shadow_root apply fast done subsubsection \set\_child\_nodes\ locale l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = CD: l_set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs begin definition set_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) shadow_root_ptr \ (_) node_ptr list \ (_, unit) dom_prog" where "set_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr = put_M shadow_root_ptr RShadowRoot.child_nodes_update" definition a_set_child_nodes_tups :: "(((_) object_ptr \ bool) \ ((_) object_ptr \ (_) node_ptr list \ (_, unit) dom_prog)) list" where "a_set_child_nodes_tups \ [(is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r, set_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast)]" definition a_set_child_nodes :: "(_) object_ptr \ (_) node_ptr list \ (_, unit) dom_prog" where "a_set_child_nodes ptr children = invoke (CD.a_set_child_nodes_tups @ a_set_child_nodes_tups) ptr children" definition a_set_child_nodes_locs :: "(_) object_ptr \ (_, unit) dom_prog set" where "a_set_child_nodes_locs ptr \ (if is_shadow_root_ptr_kind ptr then all_args (put_M (the (cast ptr)) RShadowRoot.child_nodes_update) else {}) \ CD.a_set_child_nodes_locs ptr" end global_interpretation l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs defines set_child_nodes = l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_set_child_nodes and set_child_nodes_locs = l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_set_child_nodes_locs . locale l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_type_wf type_wf + l_known_ptr known_ptr + l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_set_child_nodes_defs set_child_nodes set_child_nodes_locs + CD: l_set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M for type_wf :: "(_) heap \ bool" and known_ptr :: "(_) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ bool" and set_child_nodes :: "(_) object_ptr \ (_) node_ptr list \ (_, unit) dom_prog" and set_child_nodes_locs :: "(_) object_ptr \ (_, unit) dom_prog set" and set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ (_) node_ptr list \ (_, unit) dom_prog" and set_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ (_, unit) dom_prog set" + assumes known_ptr_impl: "known_ptr = ShadowRootClass.known_ptr" assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" assumes set_child_nodes_impl: "set_child_nodes = a_set_child_nodes" assumes set_child_nodes_locs_impl: "set_child_nodes_locs = a_set_child_nodes_locs" begin lemmas set_child_nodes_def = set_child_nodes_impl[unfolded a_set_child_nodes_def set_child_nodes_def] lemmas set_child_nodes_locs_def =set_child_nodes_locs_impl[unfolded a_set_child_nodes_locs_def set_child_nodes_locs_def, folded CD.set_child_nodes_locs_impl] lemma set_child_nodes_writes: "writes (set_child_nodes_locs ptr) (set_child_nodes ptr children) h h'" apply (simp add: set_child_nodes_def a_set_child_nodes_tups_def set_child_nodes_locs_def) apply(split CD.set_child_nodes_splits, rule conjI)+ apply (simp add: CD.set_child_nodes_writes writes_union_right_I) apply(split invoke_splits, rule conjI)+ apply(auto simp add: a_set_child_nodes_def)[1] apply(auto simp add: set_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: writes_bind_pure intro: writes_union_right_I writes_union_left_I split: list.splits)[1] by (metis is_shadow_root_ptr_implies_kind option.case_eq_if) lemma set_child_nodes_pointers_preserved: assumes "w \ set_child_nodes_locs object_ptr" assumes "h \ w \\<^sub>h h'" shows "object_ptr_kinds h = object_ptr_kinds h'" using assms(1) object_ptr_kinds_preserved[OF writes_singleton2 assms(2)] by(auto simp add: all_args_def set_child_nodes_locs_def CD.set_child_nodes_locs_def split: if_splits) lemma set_child_nodes_types_preserved: assumes "w \ set_child_nodes_locs object_ptr" assumes "h \ w \\<^sub>h h'" shows "type_wf h = type_wf h'" using assms(1) type_wf_preserved[OF writes_singleton2 assms(2)] type_wf_impl by(auto simp add: all_args_def a_set_child_nodes_tups_def set_child_nodes_locs_def CD.set_child_nodes_locs_def split: if_splits option.splits) end interpretation i_set_child_nodes?: l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr set_child_nodes set_child_nodes_locs Core_DOM_Functions.set_child_nodes Core_DOM_Functions.set_child_nodes_locs apply(unfold_locales) by (auto simp add: set_child_nodes_def set_child_nodes_locs_def) declare l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_child_nodes_is_l_set_child_nodes [instances]: "l_set_child_nodes type_wf set_child_nodes set_child_nodes_locs" apply(auto simp add: l_set_child_nodes_def instances)[1] using set_child_nodes_writes apply fast using set_child_nodes_pointers_preserved apply(fast, fast) using set_child_nodes_types_preserved apply(fast, fast) done paragraph \get\_child\_nodes\ locale l_set_child_nodes_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_child_nodes set_child_nodes_locs set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + CD: l_set_child_nodes_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M for type_wf :: "(_) heap \ bool" and known_ptr :: "(_) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and set_child_nodes :: "(_) object_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_child_nodes_locs :: "(_) object_ptr \ ((_) heap, exception, unit) prog set" and set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap, exception, unit) prog set" begin lemma set_child_nodes_get_child_nodes: assumes "known_ptr ptr" assumes "type_wf h" assumes "h \ set_child_nodes ptr children \\<^sub>h h'" shows "h' \ get_child_nodes ptr \\<^sub>r children" proof - have "h \ check_in_heap ptr \\<^sub>r ()" using assms set_child_nodes_def invoke_ptr_in_heap by (metis (full_types) check_in_heap_ptr_in_heap is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust) then have ptr_in_h: "ptr |\| object_ptr_kinds h" by (simp add: check_in_heap_ptr_in_heap is_OK_returns_result_I) have "type_wf h'" apply(unfold type_wf_impl) apply(rule subst[where P=id, OF type_wf_preserved[OF set_child_nodes_writes assms(3), unfolded all_args_def], simplified]) by(auto simp add: all_args_def assms(2)[unfolded type_wf_impl] set_child_nodes_locs_def CD.set_child_nodes_locs_def split: if_splits) have "h' \ check_in_heap ptr \\<^sub>r ()" using check_in_heap_reads set_child_nodes_writes assms(3) \h \ check_in_heap ptr \\<^sub>r ()\ apply(rule reads_writes_separate_forwards) apply(auto simp add: all_args_def set_child_nodes_locs_def CD.set_child_nodes_locs_def)[1] done then have "ptr |\| object_ptr_kinds h'" using check_in_heap_ptr_in_heap by blast with assms ptr_in_h \type_wf h'\ show ?thesis apply(auto simp add: type_wf_impl known_ptr_impl get_child_nodes_def a_get_child_nodes_tups_def set_child_nodes_def a_set_child_nodes_tups_def del: bind_pure_returns_result_I2 intro!: bind_pure_returns_result_I2)[1] apply(split CD.get_child_nodes_splits, (rule conjI impI)+)+ apply(split CD.set_child_nodes_splits)+ apply(auto simp add: CD.set_child_nodes_get_child_nodes type_wf_impl CD.type_wf_impl dest: ShadowRootClass.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t)[1] apply(auto simp add: CD.set_child_nodes_get_child_nodes type_wf_impl CD.type_wf_impl dest: ShadowRootClass.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t)[1] apply(split CD.set_child_nodes_splits)+ by(auto simp add: known_ptr_impl CD.known_ptr_impl set_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def CD.type_wf_impl ShadowRootClass.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t dest: known_ptr_new_shadow_root_ptr)[2] qed lemma set_child_nodes_get_child_nodes_different_pointers: assumes "ptr \ ptr'" assumes "w \ set_child_nodes_locs ptr" assumes "h \ w \\<^sub>h h'" assumes "r \ get_child_nodes_locs ptr'" shows "r h h'" using assms apply(auto simp add: set_child_nodes_locs_def CD.set_child_nodes_locs_def get_child_nodes_locs_def CD.get_child_nodes_locs_def)[1] by(auto simp add: all_args_def elim!: is_document_ptr_kind_obtains is_shadow_root_ptr_kind_obtains is_element_ptr_kind_obtains split: if_splits option.splits) end interpretation i_set_child_nodes_get_child_nodes?: l_set_child_nodes_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs set_child_nodes set_child_nodes_locs Core_DOM_Functions.set_child_nodes Core_DOM_Functions.set_child_nodes_locs using instances by(auto simp add: l_set_child_nodes_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def ) declare l_set_child_nodes_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_child_nodes_get_child_nodes_is_l_set_child_nodes_get_child_nodes [instances]: "l_set_child_nodes_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs set_child_nodes set_child_nodes_locs" apply(auto simp add: instances l_set_child_nodes_get_child_nodes_def l_set_child_nodes_get_child_nodes_axioms_def)[1] using set_child_nodes_get_child_nodes apply fast using set_child_nodes_get_child_nodes_different_pointers apply fast done subsubsection \set\_tag\_type\ locale l_set_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD: l_set_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_tag_name set_tag_name_locs + l_type_wf type_wf for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and set_tag_name :: "(_) element_ptr \ tag_name \ (_, unit) dom_prog" and set_tag_name_locs :: "(_) element_ptr \ (_, unit) dom_prog set" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" begin lemmas set_tag_name_def = CD.set_tag_name_impl[unfolded CD.a_set_tag_name_def set_tag_name_def] lemmas set_tag_name_locs_def = CD.set_tag_name_locs_impl[unfolded CD.a_set_tag_name_locs_def set_tag_name_locs_def] lemma set_tag_name_ok: "type_wf h \ element_ptr |\| element_ptr_kinds h \ h \ ok (set_tag_name element_ptr tag)" apply(unfold type_wf_impl) unfolding set_tag_name_impl[unfolded a_set_tag_name_def] using get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ok put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ok using CD.set_tag_name_ok CD.type_wf_impl ShadowRootClass.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t by blast lemma set_tag_name_writes: "writes (set_tag_name_locs element_ptr) (set_tag_name element_ptr tag) h h'" using CD.set_tag_name_writes . lemma set_tag_name_pointers_preserved: assumes "w \ set_tag_name_locs element_ptr" assumes "h \ w \\<^sub>h h'" shows "object_ptr_kinds h = object_ptr_kinds h'" using assms by(simp add: CD.set_tag_name_pointers_preserved) lemma set_tag_name_typess_preserved: assumes "w \ set_tag_name_locs element_ptr" assumes "h \ w \\<^sub>h h'" shows "type_wf h = type_wf h'" apply(unfold type_wf_impl) apply(rule type_wf_preserved[OF writes_singleton2 assms(2)]) using assms(1) set_tag_name_locs_def by(auto simp add: all_args_def set_tag_name_locs_def split: if_splits) end interpretation i_set_tag_name?: l_set_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_tag_name set_tag_name_locs by(auto simp add: l_set_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_set_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_set_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma set_tag_name_is_l_set_tag_name [instances]: "l_set_tag_name type_wf set_tag_name set_tag_name_locs" apply(auto simp add: l_set_tag_name_def)[1] using set_tag_name_writes apply fast using set_tag_name_ok apply fast using set_tag_name_pointers_preserved apply (fast, fast) using set_tag_name_typess_preserved apply (fast, fast) done paragraph \get\_child\_nodes\ locale l_set_tag_name_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + CD: l_set_tag_name_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_tag_name set_tag_name_locs known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_tag_name_get_child_nodes: "\w \ set_tag_name_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_child_nodes_locs ptr'. r h h'))" apply(auto simp add: get_child_nodes_locs_def)[1] apply(auto simp add: set_tag_name_locs_def all_args_def)[1] using CD.set_tag_name_get_child_nodes apply(blast) using CD.set_tag_name_get_child_nodes apply(blast) done end interpretation i_set_tag_name_get_child_nodes?: l_set_tag_name_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_tag_name set_tag_name_locs known_ptr DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs by unfold_locales declare l_set_tag_name_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_tag_name_get_child_nodes_is_l_set_tag_name_get_child_nodes [instances]: "l_set_tag_name_get_child_nodes type_wf set_tag_name set_tag_name_locs known_ptr get_child_nodes get_child_nodes_locs" using set_tag_name_is_l_set_tag_name get_child_nodes_is_l_get_child_nodes apply(simp add: l_set_tag_name_get_child_nodes_def l_set_tag_name_get_child_nodes_axioms_def) using set_tag_name_get_child_nodes by fast subsubsection \get\_shadow\_root\ locale l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs begin definition a_get_shadow_root :: "(_) element_ptr \ (_, (_) shadow_root_ptr option) dom_prog" where "a_get_shadow_root element_ptr = get_M element_ptr shadow_root_opt" definition a_get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" where "a_get_shadow_root_locs element_ptr \ {preserved (get_M element_ptr shadow_root_opt)}" end global_interpretation l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs defines get_shadow_root = a_get_shadow_root and get_shadow_root_locs = a_get_shadow_root_locs . locale l_get_shadow_root_defs = fixes get_shadow_root :: "(_) element_ptr \ (_, (_) shadow_root_ptr option) dom_prog" fixes get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" locale l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_get_shadow_root_defs get_shadow_root get_shadow_root_locs + l_type_wf type_wf for type_wf :: "(_) heap \ bool" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" assumes get_shadow_root_impl: "get_shadow_root = a_get_shadow_root" assumes get_shadow_root_locs_impl: "get_shadow_root_locs = a_get_shadow_root_locs" begin lemmas get_shadow_root_def = get_shadow_root_impl[unfolded get_shadow_root_def a_get_shadow_root_def] lemmas get_shadow_root_locs_def = get_shadow_root_locs_impl[unfolded get_shadow_root_locs_def a_get_shadow_root_locs_def] lemma get_shadow_root_ok: "type_wf h \ element_ptr |\| element_ptr_kinds h \ h \ ok (get_shadow_root element_ptr)" unfolding get_shadow_root_def type_wf_impl using ShadowRootMonad.get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ok by blast lemma get_shadow_root_pure [simp]: "pure (get_shadow_root element_ptr) h" unfolding get_shadow_root_def by simp lemma get_shadow_root_ptr_in_heap: assumes "h \ get_shadow_root element_ptr \\<^sub>r children" shows "element_ptr |\| element_ptr_kinds h" using assms by(auto simp add: get_shadow_root_def get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap dest: is_OK_returns_result_I) lemma get_shadow_root_reads: "reads (get_shadow_root_locs element_ptr) (get_shadow_root element_ptr) h h'" by(simp add: get_shadow_root_def get_shadow_root_locs_def reads_bind_pure reads_insert_writes_set_right) end interpretation i_get_shadow_root?: l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs using instances by (auto simp add: l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_get_shadow_root = l_type_wf + l_get_shadow_root_defs + assumes get_shadow_root_reads: "reads (get_shadow_root_locs element_ptr) (get_shadow_root element_ptr) h h'" assumes get_shadow_root_ok: "type_wf h \ element_ptr |\| element_ptr_kinds h \ h \ ok (get_shadow_root element_ptr)" assumes get_shadow_root_ptr_in_heap: "h \ ok (get_shadow_root element_ptr) \ element_ptr |\| element_ptr_kinds h" assumes get_shadow_root_pure [simp]: "pure (get_shadow_root element_ptr) h" lemma get_shadow_root_is_l_get_shadow_root [instances]: "l_get_shadow_root type_wf get_shadow_root get_shadow_root_locs" using instances apply(auto simp add: l_get_shadow_root_def)[1] using get_shadow_root_reads apply blast using get_shadow_root_ok apply blast using get_shadow_root_ptr_in_heap apply blast done paragraph \set\_disconnected\_nodes\ locale l_set_disconnected_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_disconnected_nodes set_disconnected_nodes_locs + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma set_disconnected_nodes_get_shadow_root: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" by(auto simp add: set_disconnected_nodes_locs_def get_shadow_root_locs_def all_args_def) end locale l_set_disconnected_nodes_get_shadow_root = l_set_disconnected_nodes_defs + l_get_shadow_root_defs + assumes set_disconnected_nodes_get_shadow_root: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" interpretation i_set_disconnected_nodes_get_shadow_root?: l_set_disconnected_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_disconnected_nodes set_disconnected_nodes_locs get_shadow_root get_shadow_root_locs by(auto simp add: l_set_disconnected_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_disconnected_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_disconnected_nodes_get_shadow_root_is_l_set_disconnected_nodes_get_shadow_root [instances]: "l_set_disconnected_nodes_get_shadow_root set_disconnected_nodes_locs get_shadow_root_locs" apply(auto simp add: l_set_disconnected_nodes_get_shadow_root_def)[1] using set_disconnected_nodes_get_shadow_root apply fast done paragraph \set\_tag\_type\ locale l_set_tag_name_get_shadow_root\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_tag_name_get_shadow_root: "\w \ set_tag_name_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" by(auto simp add: set_tag_name_locs_def get_shadow_root_locs_def all_args_def intro: element_put_get_preserved[where setter=tag_name_update and getter=shadow_root_opt]) end locale l_set_tag_name_get_shadow_root = l_set_tag_name + l_get_shadow_root + assumes set_tag_name_get_shadow_root: "\w \ set_tag_name_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" interpretation i_set_tag_name_get_shadow_root?: l_set_tag_name_get_shadow_root\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_tag_name set_tag_name_locs get_shadow_root get_shadow_root_locs apply(auto simp add: l_set_tag_name_get_shadow_root\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)[1] using l_set_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms by unfold_locales declare l_set_tag_name_get_shadow_root\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_tag_name_get_shadow_root_is_l_set_tag_name_get_shadow_root [instances]: "l_set_tag_name_get_shadow_root type_wf set_tag_name set_tag_name_locs get_shadow_root get_shadow_root_locs" using set_tag_name_is_l_set_tag_name get_shadow_root_is_l_get_shadow_root apply(simp add: l_set_tag_name_get_shadow_root_def l_set_tag_name_get_shadow_root_axioms_def) using set_tag_name_get_shadow_root by fast paragraph \set\_child\_nodes\ locale l_set_child_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_child_nodes set_child_nodes_locs set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs for type_wf :: "(_) heap \ bool" and known_ptr :: "(_) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ bool" and set_child_nodes :: "(_) object_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_child_nodes_locs :: "(_) object_ptr \ ((_) heap, exception, unit) prog set" and set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap, exception, unit) prog set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma set_child_nodes_get_shadow_root: "\w \ set_child_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" apply(auto simp add: set_child_nodes_locs_def get_shadow_root_locs_def CD.set_child_nodes_locs_def all_args_def)[1] by(auto intro!: element_put_get_preserved[where getter=shadow_root_opt and setter=RElement.child_nodes_update]) end locale l_set_child_nodes_get_shadow_root = l_set_child_nodes_defs + l_get_shadow_root_defs + assumes set_child_nodes_get_shadow_root: "\w \ set_child_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" interpretation i_set_child_nodes_get_shadow_root?: l_set_child_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr set_child_nodes set_child_nodes_locs Core_DOM_Functions.set_child_nodes Core_DOM_Functions.set_child_nodes_locs get_shadow_root get_shadow_root_locs by(auto simp add: l_set_child_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_child_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_child_nodes_get_shadow_root_is_l_set_child_nodes_get_shadow_root [instances]: "l_set_child_nodes_get_shadow_root set_child_nodes_locs get_shadow_root_locs" apply(auto simp add: l_set_child_nodes_get_shadow_root_def)[1] using set_child_nodes_get_shadow_root apply fast done paragraph \delete\_shadow\_root\ locale l_delete_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_shadow_root_delete_shadow_root: "h \ delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M shadow_root_ptr \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" by(auto simp add: get_shadow_root_locs_def delete_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t) end locale l_delete_shadow_root_get_shadow_root = l_get_shadow_root_defs + assumes get_shadow_root_delete_shadow_root: "h \ delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M shadow_root_ptr \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" interpretation l_delete_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs by(auto simp add: l_delete_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) lemma l_delete_shadow_root_get_shadow_root_get_shadow_root_locs [instances]: "l_delete_shadow_root_get_shadow_root get_shadow_root_locs" apply(auto simp add: l_delete_shadow_root_get_shadow_root_def)[1] using get_shadow_root_delete_shadow_root apply fast done paragraph \new\_character\_data\ locale l_new_character_data_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs for type_wf :: "(_) heap \ bool" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_shadow_root_new_character_data: "h \ new_character_data \\<^sub>r new_character_data_ptr \ h \ new_character_data \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" by (auto simp add: get_shadow_root_locs_def new_character_data_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t new_character_data_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t split: prod.splits if_splits option.splits elim!: bind_returns_result_E bind_returns_heap_E intro: is_element_ptr_kind_obtains) end locale l_new_character_data_get_shadow_root = l_new_character_data + l_get_shadow_root + assumes get_shadow_root_new_character_data: "h \ new_character_data \\<^sub>r new_character_data_ptr \ h \ new_character_data \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" interpretation i_new_character_data_get_shadow_root?: l_new_character_data_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs by(unfold_locales) declare l_new_character_data_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma new_character_data_get_shadow_root_is_l_new_character_data_get_shadow_root [instances]: "l_new_character_data_get_shadow_root type_wf get_shadow_root get_shadow_root_locs" using new_character_data_is_l_new_character_data get_shadow_root_is_l_get_shadow_root apply(auto simp add: l_new_character_data_get_shadow_root_def l_new_character_data_get_shadow_root_axioms_def instances)[1] using get_shadow_root_new_character_data by fast paragraph \new\_document\ locale l_new_document_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs for type_wf :: "(_) heap \ bool" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_shadow_root_new_document: "h \ new_document \\<^sub>r new_document_ptr \ h \ new_document \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" by (auto simp add: get_shadow_root_locs_def new_document_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t new_document_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t split: prod.splits if_splits option.splits elim!: bind_returns_result_E bind_returns_heap_E intro: is_element_ptr_kind_obtains) end locale l_new_document_get_shadow_root = l_new_document + l_get_shadow_root + assumes get_shadow_root_new_document: "h \ new_document \\<^sub>r new_document_ptr \ h \ new_document \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" interpretation i_new_document_get_shadow_root?: l_new_document_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs by(unfold_locales) declare l_new_document_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma new_document_get_shadow_root_is_l_new_document_get_shadow_root [instances]: "l_new_document_get_shadow_root type_wf get_shadow_root get_shadow_root_locs" using new_document_is_l_new_document get_shadow_root_is_l_get_shadow_root apply(auto simp add: l_new_document_get_shadow_root_def l_new_document_get_shadow_root_axioms_def instances)[1] using get_shadow_root_new_document by fast paragraph \new\_element\ locale l_new_element_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs for type_wf :: "(_) heap \ bool" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_shadow_root_new_element: "ptr' \ new_element_ptr \ h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" by (auto simp add: get_shadow_root_locs_def new_element_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t new_element_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t new_element_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t split: prod.splits if_splits option.splits elim!: bind_returns_result_E bind_returns_heap_E intro: is_element_ptr_kind_obtains) lemma new_element_no_shadow_root: "h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ h' \ get_shadow_root new_element_ptr \\<^sub>r None" by(simp add: get_shadow_root_def new_element_shadow_root_opt) end locale l_new_element_get_shadow_root = l_new_element + l_get_shadow_root + assumes get_shadow_root_new_element: "ptr' \ new_element_ptr \ h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" assumes new_element_no_shadow_root: "h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ h' \ get_shadow_root new_element_ptr \\<^sub>r None" interpretation i_new_element_get_shadow_root?: l_new_element_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs by(unfold_locales) declare l_new_element_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma new_element_get_shadow_root_is_l_new_element_get_shadow_root [instances]: "l_new_element_get_shadow_root type_wf get_shadow_root get_shadow_root_locs" using new_element_is_l_new_element get_shadow_root_is_l_get_shadow_root apply(auto simp add: l_new_element_get_shadow_root_def l_new_element_get_shadow_root_axioms_def instances)[1] using get_shadow_root_new_element new_element_no_shadow_root by fast+ paragraph \new\_shadow\_root\ locale l_new_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs for type_wf :: "(_) heap \ bool" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_shadow_root_new_shadow_root: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" by (auto simp add: get_shadow_root_locs_def new_shadow_root_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t new_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t split: prod.splits if_splits option.splits elim!: bind_returns_result_E bind_returns_heap_E intro: is_element_ptr_kind_obtains) end locale l_new_shadow_root_get_shadow_root = l_get_shadow_root + assumes get_shadow_root_new_shadow_root: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" interpretation i_new_shadow_root_get_shadow_root?: l_new_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs by(unfold_locales) declare l_new_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma new_shadow_root_get_shadow_root_is_l_new_shadow_root_get_shadow_root [instances]: "l_new_shadow_root_get_shadow_root type_wf get_shadow_root get_shadow_root_locs" using get_shadow_root_is_l_get_shadow_root apply(auto simp add: l_new_shadow_root_get_shadow_root_def l_new_shadow_root_get_shadow_root_axioms_def instances)[1] using get_shadow_root_new_shadow_root by fast subsubsection \set\_shadow\_root\ locale l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs begin definition a_set_shadow_root :: "(_) element_ptr \ (_) shadow_root_ptr option \ (_, unit) dom_prog" where "a_set_shadow_root element_ptr = put_M element_ptr shadow_root_opt_update" definition a_set_shadow_root_locs :: "(_) element_ptr \ ((_, unit) dom_prog) set" where "a_set_shadow_root_locs element_ptr \ all_args (put_M element_ptr shadow_root_opt_update)" end global_interpretation l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs defines set_shadow_root = a_set_shadow_root and set_shadow_root_locs = a_set_shadow_root_locs . locale l_set_shadow_root_defs = fixes set_shadow_root :: "(_) element_ptr \ (_) shadow_root_ptr option \ (_, unit) dom_prog" fixes set_shadow_root_locs :: "(_) element_ptr \ (_, unit) dom_prog set" locale l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_type_wf type_wf + l_set_shadow_root_defs set_shadow_root set_shadow_root_locs + l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs for type_wf :: "(_) heap \ bool" and set_shadow_root :: "(_) element_ptr \ (_) shadow_root_ptr option \ (_, unit) dom_prog" and set_shadow_root_locs :: "(_) element_ptr \ (_, unit) dom_prog set" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" assumes set_shadow_root_impl: "set_shadow_root = a_set_shadow_root" assumes set_shadow_root_locs_impl: "set_shadow_root_locs = a_set_shadow_root_locs" begin lemmas set_shadow_root_def = set_shadow_root_impl[unfolded set_shadow_root_def a_set_shadow_root_def] lemmas set_shadow_root_locs_def = set_shadow_root_locs_impl[unfolded set_shadow_root_locs_def a_set_shadow_root_locs_def] lemma set_shadow_root_ok: "type_wf h \ element_ptr |\| element_ptr_kinds h \ h \ ok (set_shadow_root element_ptr tag)" apply(unfold type_wf_impl) unfolding set_shadow_root_def using get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ok put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ok by (simp add: ShadowRootMonad.put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ok) lemma set_shadow_root_ptr_in_heap: "h \ ok (set_shadow_root element_ptr shadow_root) \ element_ptr |\| element_ptr_kinds h" by(simp add: set_shadow_root_def ElementMonad.put_M_ptr_in_heap) lemma set_shadow_root_writes: "writes (set_shadow_root_locs element_ptr) (set_shadow_root element_ptr tag) h h'" by(auto simp add: set_shadow_root_def set_shadow_root_locs_def intro: writes_bind_pure) lemma set_shadow_root_pointers_preserved: assumes "w \ set_shadow_root_locs element_ptr" assumes "h \ w \\<^sub>h h'" shows "object_ptr_kinds h = object_ptr_kinds h'" using assms(1) object_ptr_kinds_preserved[OF writes_singleton2 assms(2)] by(auto simp add: all_args_def set_shadow_root_locs_def split: if_splits) lemma set_shadow_root_types_preserved: assumes "w \ set_shadow_root_locs element_ptr" assumes "h \ w \\<^sub>h h'" shows "type_wf h = type_wf h'" apply(unfold type_wf_impl) using assms(1) type_wf_preserved[OF writes_singleton2 assms(2)] by(auto simp add: all_args_def set_shadow_root_locs_def split: if_splits) end interpretation i_set_shadow_root?: l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_shadow_root set_shadow_root_locs by (auto simp add: l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_set_shadow_root = l_type_wf +l_set_shadow_root_defs + assumes set_shadow_root_writes: "writes (set_shadow_root_locs element_ptr) (set_shadow_root element_ptr disc_nodes) h h'" assumes set_shadow_root_ok: "type_wf h \ element_ptr |\| element_ptr_kinds h \ h \ ok (set_shadow_root element_ptr shadow_root)" assumes set_shadow_root_ptr_in_heap: "h \ ok (set_shadow_root element_ptr shadow_root) \ element_ptr |\| element_ptr_kinds h" assumes set_shadow_root_pointers_preserved: "w \ set_shadow_root_locs element_ptr \ h \ w \\<^sub>h h' \ object_ptr_kinds h = object_ptr_kinds h'" assumes set_shadow_root_types_preserved: "w \ set_shadow_root_locs element_ptr \ h \ w \\<^sub>h h' \ type_wf h = type_wf h'" lemma set_shadow_root_is_l_set_shadow_root [instances]: "l_set_shadow_root type_wf set_shadow_root set_shadow_root_locs" apply(auto simp add: l_set_shadow_root_def instances)[1] using set_shadow_root_writes apply blast using set_shadow_root_ok apply (blast) using set_shadow_root_ptr_in_heap apply blast using set_shadow_root_pointers_preserved apply(blast, blast) using set_shadow_root_types_preserved apply(blast, blast) done paragraph \get\_shadow\_root\ locale l_set_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_shadow_root_get_shadow_root: "type_wf h \ h \ set_shadow_root ptr shadow_root_ptr_opt \\<^sub>h h' \ h' \ get_shadow_root ptr \\<^sub>r shadow_root_ptr_opt" by(auto simp add: set_shadow_root_def get_shadow_root_def) lemma set_shadow_root_get_shadow_root_different_pointers: "ptr \ ptr' \ \w \ set_shadow_root_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" by(auto simp add: set_shadow_root_locs_def get_shadow_root_locs_def all_args_def) end interpretation i_set_shadow_root_get_shadow_root?: l_set_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_shadow_root set_shadow_root_locs get_shadow_root get_shadow_root_locs apply(auto simp add: l_set_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)[1] by(unfold_locales) declare l_set_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_set_shadow_root_get_shadow_root = l_type_wf + l_set_shadow_root_defs + l_get_shadow_root_defs + assumes set_shadow_root_get_shadow_root: "type_wf h \ h \ set_shadow_root ptr shadow_root_ptr_opt \\<^sub>h h' \ h' \ get_shadow_root ptr \\<^sub>r shadow_root_ptr_opt" assumes set_shadow_root_get_shadow_root_different_pointers: "ptr \ ptr' \ w \ set_shadow_root_locs ptr \ h \ w \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" lemma set_shadow_root_get_shadow_root_is_l_set_shadow_root_get_shadow_root [instances]: "l_set_shadow_root_get_shadow_root type_wf set_shadow_root set_shadow_root_locs get_shadow_root get_shadow_root_locs" apply(auto simp add: l_set_shadow_root_get_shadow_root_def instances)[1] using set_shadow_root_get_shadow_root apply fast using set_shadow_root_get_shadow_root_different_pointers apply fast done subsubsection \set\_mode\ locale l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs begin definition a_set_mode :: "(_) shadow_root_ptr \ shadow_root_mode \ (_, unit) dom_prog" where "a_set_mode shadow_root_ptr = put_M shadow_root_ptr mode_update" definition a_set_mode_locs :: "(_) shadow_root_ptr \ ((_, unit) dom_prog) set" where "a_set_mode_locs shadow_root_ptr \ all_args (put_M shadow_root_ptr mode_update)" end global_interpretation l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs defines set_mode = a_set_mode and set_mode_locs = a_set_mode_locs . locale l_set_mode_defs = fixes set_mode :: "(_) shadow_root_ptr \ shadow_root_mode \ (_, unit) dom_prog" fixes set_mode_locs :: "(_) shadow_root_ptr \ (_, unit) dom_prog set" locale l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_type_wf type_wf + l_set_mode_defs set_mode set_mode_locs + l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs for type_wf :: "(_) heap \ bool" and set_mode :: "(_) shadow_root_ptr \ shadow_root_mode \ (_, unit) dom_prog" and set_mode_locs :: "(_) shadow_root_ptr \ (_, unit) dom_prog set" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" assumes set_mode_impl: "set_mode = a_set_mode" assumes set_mode_locs_impl: "set_mode_locs = a_set_mode_locs" begin lemmas set_mode_def = set_mode_impl[unfolded set_mode_def a_set_mode_def] lemmas set_mode_locs_def = set_mode_locs_impl[unfolded set_mode_locs_def a_set_mode_locs_def] lemma set_mode_ok: "type_wf h \ shadow_root_ptr |\| shadow_root_ptr_kinds h \ h \ ok (set_mode shadow_root_ptr shadow_root_mode)" apply(unfold type_wf_impl) unfolding set_mode_def using get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ok put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ok by (simp add: ShadowRootMonad.put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ok) lemma set_mode_ptr_in_heap: "h \ ok (set_mode shadow_root_ptr shadow_root_mode) \ shadow_root_ptr |\| shadow_root_ptr_kinds h" by(simp add: set_mode_def put_M_ptr_in_heap) lemma set_mode_writes: "writes (set_mode_locs shadow_root_ptr) (set_mode shadow_root_ptr shadow_root_mode) h h'" by(auto simp add: set_mode_def set_mode_locs_def intro: writes_bind_pure) lemma set_mode_pointers_preserved: assumes "w \ set_mode_locs element_ptr" assumes "h \ w \\<^sub>h h'" shows "object_ptr_kinds h = object_ptr_kinds h'" using assms(1) object_ptr_kinds_preserved[OF writes_singleton2 assms(2)] by(auto simp add: all_args_def set_mode_locs_def split: if_splits) lemma set_mode_types_preserved: assumes "w \ set_mode_locs element_ptr" assumes "h \ w \\<^sub>h h'" shows "type_wf h = type_wf h'" apply(unfold type_wf_impl) using assms(1) type_wf_preserved[OF writes_singleton2 assms(2)] by(auto simp add: all_args_def set_mode_locs_def split: if_splits) end interpretation i_set_mode?: l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_mode set_mode_locs by (auto simp add: l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_set_mode = l_type_wf +l_set_mode_defs + assumes set_mode_writes: "writes (set_mode_locs shadow_root_ptr) (set_mode shadow_root_ptr shadow_root_mode) h h'" assumes set_mode_ok: "type_wf h \ shadow_root_ptr |\| shadow_root_ptr_kinds h \ h \ ok (set_mode shadow_root_ptr shadow_root_mode)" assumes set_mode_ptr_in_heap: "h \ ok (set_mode shadow_root_ptr shadow_root_mode) \ shadow_root_ptr |\| shadow_root_ptr_kinds h" assumes set_mode_pointers_preserved: "w \ set_mode_locs shadow_root_ptr \ h \ w \\<^sub>h h' \ object_ptr_kinds h = object_ptr_kinds h'" assumes set_mode_types_preserved: "w \ set_mode_locs shadow_root_ptr \ h \ w \\<^sub>h h' \ type_wf h = type_wf h'" lemma set_mode_is_l_set_mode [instances]: "l_set_mode type_wf set_mode set_mode_locs" apply(auto simp add: l_set_mode_def instances)[1] using set_mode_writes apply blast using set_mode_ok apply (blast) using set_mode_ptr_in_heap apply blast using set_mode_pointers_preserved apply(blast, blast) using set_mode_types_preserved apply(blast, blast) done paragraph \get\_child\_nodes\ locale l_set_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_shadow_root_get_child_nodes: "\w \ set_shadow_root_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_child_nodes_locs ptr'. r h h'))" by(auto simp add: get_child_nodes_locs_def set_shadow_root_locs_def CD.get_child_nodes_locs_def all_args_def intro: element_put_get_preserved[where setter=shadow_root_opt_update]) end interpretation i_set_shadow_root_get_child_nodes?: l_set_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs set_shadow_root set_shadow_root_locs by(unfold_locales) declare l_set_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_set_shadow_root_get_child_nodes = l_set_shadow_root + l_get_child_nodes + assumes set_shadow_root_get_child_nodes: "\w \ set_shadow_root_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_child_nodes_locs ptr'. r h h'))" lemma set_shadow_root_get_child_nodes_is_l_set_shadow_root_get_child_nodes [instances]: "l_set_shadow_root_get_child_nodes type_wf set_shadow_root set_shadow_root_locs known_ptr get_child_nodes get_child_nodes_locs" apply(auto simp add: l_set_shadow_root_get_child_nodes_def l_set_shadow_root_get_child_nodes_axioms_def instances)[1] using set_shadow_root_get_child_nodes apply blast done paragraph \get\_shadow\_root\ locale l_set_mode_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_mode_get_shadow_root: "\w \ set_mode_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" by(auto simp add: set_mode_locs_def get_shadow_root_locs_def all_args_def) end interpretation i_set_mode_get_shadow_root?: l_set_mode_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_mode set_mode_locs get_shadow_root get_shadow_root_locs by unfold_locales declare l_set_mode_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_set_mode_get_shadow_root = l_set_mode + l_get_shadow_root + assumes set_mode_get_shadow_root: "\w \ set_mode_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" lemma set_mode_get_shadow_root_is_l_set_mode_get_shadow_root [instances]: "l_set_mode_get_shadow_root type_wf set_mode set_mode_locs get_shadow_root get_shadow_root_locs" using set_mode_is_l_set_mode get_shadow_root_is_l_get_shadow_root apply(simp add: l_set_mode_get_shadow_root_def l_set_mode_get_shadow_root_axioms_def) using set_mode_get_shadow_root by fast paragraph \get\_child\_nodes\ locale l_set_mode_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_mode_get_child_nodes: "\w \ set_mode_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_child_nodes_locs ptr'. r h h'))" by(auto simp add: get_child_nodes_locs_def CD.get_child_nodes_locs_def set_mode_locs_def all_args_def)[1] end interpretation i_set_mode_get_child_nodes?: l_set_mode_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_mode set_mode_locs known_ptr DocumentClass.type_wf DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs by unfold_locales declare l_set_mode_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_set_mode_get_child_nodes = l_set_mode + l_get_child_nodes + assumes set_mode_get_child_nodes: "\w \ set_mode_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_child_nodes_locs ptr'. r h h'))" lemma set_mode_get_child_nodes_is_l_set_mode_get_child_nodes [instances]: "l_set_mode_get_child_nodes type_wf set_mode set_mode_locs known_ptr get_child_nodes get_child_nodes_locs" using set_mode_is_l_set_mode get_child_nodes_is_l_get_child_nodes apply(simp add: l_set_mode_get_child_nodes_def l_set_mode_get_child_nodes_axioms_def) using set_mode_get_child_nodes by fast subsubsection \get\_host\ locale l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_shadow_root_defs get_shadow_root get_shadow_root_locs for get_shadow_root :: "(_::linorder) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin definition a_get_host :: "(_) shadow_root_ptr \ (_, (_) element_ptr) dom_prog" where "a_get_host shadow_root_ptr = do { host_ptrs \ element_ptr_kinds_M \ filter_M (\element_ptr. do { shadow_root_opt \ get_shadow_root element_ptr; return (shadow_root_opt = Some shadow_root_ptr) }); (case host_ptrs of host_ptr#[] \ return host_ptr | _ \ error HierarchyRequestError) }" definition "a_get_host_locs \ (\element_ptr. (get_shadow_root_locs element_ptr)) \ (\ptr. {preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr RObject.nothing)})" end global_interpretation l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_shadow_root get_shadow_root_locs defines get_host = "a_get_host" and get_host_locs = "a_get_host_locs" . locale l_get_host_defs = fixes get_host :: "(_) shadow_root_ptr \ (_, (_) element_ptr) dom_prog" fixes get_host_locs :: "((_) heap \ (_) heap \ bool) set" locale l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_get_host_defs + l_get_shadow_root + assumes get_host_impl: "get_host = a_get_host" assumes get_host_locs_impl: "get_host_locs = a_get_host_locs" begin lemmas get_host_def = get_host_impl[unfolded a_get_host_def] lemmas get_host_locs_def = get_host_locs_impl[unfolded a_get_host_locs_def] lemma get_host_pure [simp]: "pure (get_host element_ptr) h" by(auto simp add: get_host_def intro!: bind_pure_I filter_M_pure_I split: list.splits) lemma get_host_reads: "reads get_host_locs (get_host element_ptr) h h'" using get_shadow_root_reads[unfolded reads_def] by(auto simp add: get_host_def get_host_locs_def intro!: reads_bind_pure reads_subset[OF check_in_heap_reads] reads_subset[OF error_reads] reads_subset[OF get_shadow_root_reads] reads_subset[OF return_reads] reads_subset[OF element_ptr_kinds_M_reads] filter_M_reads filter_M_pure_I bind_pure_I split: list.splits) end locale l_get_host = l_get_host_defs + assumes get_host_pure [simp]: "pure (get_host element_ptr) h" assumes get_host_reads: "reads get_host_locs (get_host node_ptr) h h'" interpretation i_get_host?: l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_shadow_root get_shadow_root_locs get_host get_host_locs type_wf using instances by (simp add: l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def get_host_def get_host_locs_def) declare l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_host_is_l_get_host [instances]: "l_get_host get_host get_host_locs" apply(auto simp add: l_get_host_def)[1] using get_host_reads apply fast done subsubsection \get\_mode\ locale l_get_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs begin definition a_get_mode :: "(_) shadow_root_ptr \ (_, shadow_root_mode) dom_prog" where "a_get_mode shadow_root_ptr = get_M shadow_root_ptr mode" definition a_get_mode_locs :: "(_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" where "a_get_mode_locs shadow_root_ptr \ {preserved (get_M shadow_root_ptr mode)}" end global_interpretation l_get_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs defines get_mode = a_get_mode and get_mode_locs = a_get_mode_locs . locale l_get_mode_defs = fixes get_mode :: "(_) shadow_root_ptr \ (_, shadow_root_mode) dom_prog" fixes get_mode_locs :: "(_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" locale l_get_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_get_mode_defs get_mode get_mode_locs + l_type_wf type_wf for get_mode :: "(_) shadow_root_ptr \ ((_) heap, exception, shadow_root_mode) prog" and get_mode_locs :: "(_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" and type_wf :: "(_) heap \ bool" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" assumes get_mode_impl: "get_mode = a_get_mode" assumes get_mode_locs_impl: "get_mode_locs = a_get_mode_locs" begin lemmas get_mode_def = get_mode_impl[unfolded get_mode_def a_get_mode_def] lemmas get_mode_locs_def = get_mode_locs_impl[unfolded get_mode_locs_def a_get_mode_locs_def] lemma get_mode_ok: "type_wf h \ shadow_root_ptr |\| shadow_root_ptr_kinds h \ h \ ok (get_mode shadow_root_ptr)" unfolding get_mode_def type_wf_impl using ShadowRootMonad.get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ok by blast lemma get_mode_pure [simp]: "pure (get_mode element_ptr) h" unfolding get_mode_def by simp lemma get_mode_ptr_in_heap: assumes "h \ get_mode shadow_root_ptr \\<^sub>r children" shows "shadow_root_ptr |\| shadow_root_ptr_kinds h" using assms by(auto simp add: get_mode_def get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ptr_in_heap dest: is_OK_returns_result_I) lemma get_mode_reads: "reads (get_mode_locs element_ptr) (get_mode element_ptr) h h'" by(simp add: get_mode_def get_mode_locs_def reads_bind_pure reads_insert_writes_set_right) end interpretation i_get_mode?: l_get_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_mode get_mode_locs type_wf using instances by (auto simp add: l_get_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_get_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_get_mode = l_type_wf + l_get_mode_defs + assumes get_mode_reads: "reads (get_mode_locs shadow_root_ptr) (get_mode shadow_root_ptr) h h'" assumes get_mode_ok: "type_wf h \ shadow_root_ptr |\| shadow_root_ptr_kinds h \ h \ ok (get_mode shadow_root_ptr)" assumes get_mode_ptr_in_heap: "h \ ok (get_mode shadow_root_ptr) \ shadow_root_ptr |\| shadow_root_ptr_kinds h" assumes get_mode_pure [simp]: "pure (get_mode shadow_root_ptr) h" lemma get_mode_is_l_get_mode [instances]: "l_get_mode type_wf get_mode get_mode_locs" apply(auto simp add: l_get_mode_def instances)[1] using get_mode_reads apply blast using get_mode_ok apply blast using get_mode_ptr_in_heap apply blast done subsubsection \get\_shadow\_root\_safe\ locale l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_shadow_root_defs get_shadow_root get_shadow_root_locs + l_get_mode_defs get_mode get_mode_locs for get_shadow_root :: "(_) element_ptr \ (_, (_) shadow_root_ptr option) dom_prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_mode :: "(_) shadow_root_ptr \ (_, shadow_root_mode) dom_prog" and get_mode_locs :: "(_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" begin definition a_get_shadow_root_safe :: "(_) element_ptr \ (_, (_) shadow_root_ptr option) dom_prog" where "a_get_shadow_root_safe element_ptr = do { shadow_root_ptr_opt \ get_shadow_root element_ptr; (case shadow_root_ptr_opt of Some shadow_root_ptr \ do { mode \ get_mode shadow_root_ptr; (if mode = Open then return (Some shadow_root_ptr) else return None ) } | None \ return None) }" definition a_get_shadow_root_safe_locs :: "(_) element_ptr \ (_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" where "a_get_shadow_root_safe_locs element_ptr shadow_root_ptr \ (get_shadow_root_locs element_ptr) \ (get_mode_locs shadow_root_ptr)" end global_interpretation l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_shadow_root get_shadow_root_locs get_mode get_mode_locs defines get_shadow_root_safe = a_get_shadow_root_safe and get_shadow_root_safe_locs = a_get_shadow_root_safe_locs . locale l_get_shadow_root_safe_defs = fixes get_shadow_root_safe :: "(_) element_ptr \ (_, (_) shadow_root_ptr option) dom_prog" fixes get_shadow_root_safe_locs :: "(_) element_ptr \ (_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" locale l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_shadow_root get_shadow_root_locs get_mode get_mode_locs + l_get_shadow_root_safe_defs get_shadow_root_safe get_shadow_root_safe_locs + l_get_shadow_root type_wf get_shadow_root get_shadow_root_locs + l_get_mode type_wf get_mode get_mode_locs + l_type_wf type_wf for type_wf :: "(_) heap \ bool" and get_shadow_root_safe :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_safe_locs :: "(_) element_ptr \ (_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ (_, (_) shadow_root_ptr option) dom_prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_mode :: "(_) shadow_root_ptr \ (_, shadow_root_mode) dom_prog" and get_mode_locs :: "(_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" assumes get_shadow_root_safe_impl: "get_shadow_root_safe = a_get_shadow_root_safe" assumes get_shadow_root_safe_locs_impl: "get_shadow_root_safe_locs = a_get_shadow_root_safe_locs" begin lemmas get_shadow_root_safe_def = get_shadow_root_safe_impl[unfolded get_shadow_root_safe_def a_get_shadow_root_safe_def] lemmas get_shadow_root_safe_locs_def = get_shadow_root_safe_locs_impl[unfolded get_shadow_root_safe_locs_def a_get_shadow_root_safe_locs_def] lemma get_shadow_root_safe_pure [simp]: "pure (get_shadow_root_safe element_ptr) h" apply(auto simp add: get_shadow_root_safe_def)[1] by (smt bind_returns_heap_E is_OK_returns_heap_E local.get_mode_pure local.get_shadow_root_pure option.case_eq_if pure_def pure_returns_heap_eq return_pure) end interpretation i_get_shadow_root_safe?: l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root_safe get_shadow_root_safe_locs get_shadow_root get_shadow_root_locs get_mode get_mode_locs using instances by (auto simp add: l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def get_shadow_root_safe_def get_shadow_root_safe_locs_def) declare l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_get_shadow_root_safe = l_get_shadow_root_safe_defs + assumes get_shadow_root_safe_pure [simp]: "pure (get_shadow_root_safe element_ptr) h" lemma get_shadow_root_safe_is_l_get_shadow_root_safe [instances]: "l_get_shadow_root_safe get_shadow_root_safe" using instances apply(auto simp add: l_get_shadow_root_safe_def)[1] done subsubsection \set\_disconnected\_nodes\ locale l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD: l_set_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_disconnected_nodes set_disconnected_nodes_locs + l_type_wf type_wf for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" begin lemma set_disconnected_nodes_ok: "type_wf h \ document_ptr |\| document_ptr_kinds h \ h \ ok (set_disconnected_nodes document_ptr node_ptrs)" using CD.set_disconnected_nodes_ok CD.type_wf_impl ShadowRootClass.type_wf_defs local.type_wf_impl by blast lemma set_disconnected_nodes_typess_preserved: assumes "w \ set_disconnected_nodes_locs object_ptr" assumes "h \ w \\<^sub>h h'" shows "type_wf h = type_wf h'" using assms(1) type_wf_preserved[OF writes_singleton2 assms(2)] apply(unfold type_wf_impl) by(auto simp add: all_args_def CD.set_disconnected_nodes_locs_def intro: put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_disconnected_nodes_type_wf_preserved split: if_splits) end interpretation i_set_disconnected_nodes?: l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_disconnected_nodes set_disconnected_nodes_locs by(auto simp add: l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma set_disconnected_nodes_is_l_set_disconnected_nodes [instances]: "l_set_disconnected_nodes type_wf set_disconnected_nodes set_disconnected_nodes_locs" apply(auto simp add: l_set_disconnected_nodes_def)[1] apply (simp add: i_set_disconnected_nodes.set_disconnected_nodes_writes) using set_disconnected_nodes_ok apply blast apply (simp add: i_set_disconnected_nodes.set_disconnected_nodes_ptr_in_heap) using i_set_disconnected_nodes.set_disconnected_nodes_pointers_preserved apply (blast, blast) using set_disconnected_nodes_typess_preserved apply(blast, blast) done paragraph \get\_child\_nodes\ locale l_set_disconnected_nodes_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_disconnected_nodes set_disconnected_nodes_locs + l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M for type_wf :: "(_) heap \ bool" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" and known_ptr :: "(_) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma set_disconnected_nodes_get_child_nodes: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_child_nodes_locs ptr'. r h h'))" by(auto simp add: set_disconnected_nodes_locs_def get_child_nodes_locs_def CD.get_child_nodes_locs_def all_args_def) end interpretation i_set_disconnected_nodes_get_child_nodes?: l_set_disconnected_nodes_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_disconnected_nodes set_disconnected_nodes_locs known_ptr DocumentClass.type_wf DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs by(auto simp add: l_set_disconnected_nodes_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_disconnected_nodes_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_disconnected_nodes_get_child_nodes_is_l_set_disconnected_nodes_get_child_nodes [instances]: "l_set_disconnected_nodes_get_child_nodes set_disconnected_nodes_locs get_child_nodes_locs" apply(auto simp add: l_set_disconnected_nodes_get_child_nodes_def)[1] using set_disconnected_nodes_get_child_nodes apply fast done paragraph \get\_host\ locale l_set_disconnected_nodes_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_disconnected_nodes_get_host: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_host_locs. r h h'))" by(auto simp add: CD.set_disconnected_nodes_locs_def get_shadow_root_locs_def get_host_locs_def all_args_def) end interpretation i_set_disconnected_nodes_get_host?: l_set_disconnected_nodes_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_disconnected_nodes set_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_host get_host_locs by(auto simp add: l_set_disconnected_nodes_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_disconnected_nodes_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_set_disconnected_nodes_get_host = l_set_disconnected_nodes_defs + l_get_host_defs + assumes set_disconnected_nodes_get_host: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_host_locs. r h h'))" lemma set_disconnected_nodes_get_host_is_l_set_disconnected_nodes_get_host [instances]: "l_set_disconnected_nodes_get_host set_disconnected_nodes_locs get_host_locs" apply(auto simp add: l_set_disconnected_nodes_get_host_def instances)[1] using set_disconnected_nodes_get_host by fast subsubsection \get\_tag\_name\ locale l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD: l_get_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_tag_name get_tag_name_locs + l_type_wf type_wf for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_tag_name :: "(_) element_ptr \ (_, tag_name) dom_prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" begin lemma get_tag_name_ok: "type_wf h \ element_ptr |\| element_ptr_kinds h \ h \ ok (get_tag_name element_ptr)" apply(unfold type_wf_impl get_tag_name_impl[unfolded a_get_tag_name_def]) using CD.get_tag_name_ok CD.type_wf_impl ShadowRootClass.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t by blast end interpretation i_get_tag_name?: l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_tag_name get_tag_name_locs by(auto simp add: l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_tag_name_is_l_get_tag_name [instances]: "l_get_tag_name type_wf get_tag_name get_tag_name_locs" apply(auto simp add: l_get_tag_name_def)[1] using get_tag_name_reads apply fast using get_tag_name_ok apply fast done paragraph \set\_disconnected\_nodes\ locale l_set_disconnected_nodes_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_disconnected_nodes_get_tag_name: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_tag_name_locs ptr'. r h h'))" by(auto simp add: CD.set_disconnected_nodes_locs_def CD.get_tag_name_locs_def all_args_def) end interpretation i_set_disconnected_nodes_get_tag_name?: l_set_disconnected_nodes_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_disconnected_nodes set_disconnected_nodes_locs get_tag_name get_tag_name_locs by(auto simp add: l_set_disconnected_nodes_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_disconnected_nodes_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma set_disconnected_nodes_get_tag_name_is_l_set_disconnected_nodes_get_tag_name [instances]: "l_set_disconnected_nodes_get_tag_name type_wf set_disconnected_nodes set_disconnected_nodes_locs get_tag_name get_tag_name_locs" apply(auto simp add: l_set_disconnected_nodes_get_tag_name_def l_set_disconnected_nodes_get_tag_name_axioms_def instances)[1] using set_disconnected_nodes_get_tag_name by fast paragraph \set\_child\_nodes\ locale l_set_child_nodes_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_child_nodes_get_tag_name: "\w \ set_child_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_tag_name_locs ptr'. r h h'))" by(auto simp add: CD.set_child_nodes_locs_def set_child_nodes_locs_def CD.get_tag_name_locs_def all_args_def intro: element_put_get_preserved[where getter=tag_name and setter=RElement.child_nodes_update]) end interpretation i_set_child_nodes_get_tag_name?: l_set_child_nodes_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr set_child_nodes set_child_nodes_locs Core_DOM_Functions.set_child_nodes Core_DOM_Functions.set_child_nodes_locs get_tag_name get_tag_name_locs by(auto simp add: l_set_child_nodes_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_child_nodes_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma set_child_nodes_get_tag_name_is_l_set_child_nodes_get_tag_name [instances]: "l_set_child_nodes_get_tag_name type_wf set_child_nodes set_child_nodes_locs get_tag_name get_tag_name_locs" apply(auto simp add: l_set_child_nodes_get_tag_name_def l_set_child_nodes_get_tag_name_axioms_def instances)[1] using set_child_nodes_get_tag_name by fast paragraph \delete\_shadow\_root\ locale l_delete_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_tag_name_delete_shadow_root: "h \ delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M shadow_root_ptr \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" by(auto simp add: CD.get_tag_name_locs_def delete_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t) end locale l_delete_shadow_root_get_tag_name = l_get_tag_name_defs + assumes get_tag_name_delete_shadow_root: "h \ delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M shadow_root_ptr \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" interpretation l_delete_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_tag_name get_tag_name_locs by(auto simp add: l_delete_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) lemma l_delete_shadow_root_get_tag_name_get_tag_name_locs [instances]: "l_delete_shadow_root_get_tag_name get_tag_name_locs" apply(auto simp add: l_delete_shadow_root_get_tag_name_def)[1] using get_tag_name_delete_shadow_root apply fast done paragraph \set\_shadow\_root\ locale l_set_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_shadow_root_get_tag_name: "\w \ set_shadow_root_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_tag_name_locs ptr'. r h h'))" by(auto simp add: set_shadow_root_locs_def CD.get_tag_name_locs_def all_args_def element_put_get_preserved[where setter=shadow_root_opt_update]) end interpretation i_set_shadow_root_get_tag_name?: l_set_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_shadow_root set_shadow_root_locs DocumentClass.type_wf get_tag_name get_tag_name_locs apply(auto simp add: l_set_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)[1] by(unfold_locales) declare l_set_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_set_shadow_root_get_tag_name = l_set_shadow_root_defs + l_get_tag_name_defs + assumes set_shadow_root_get_tag_name: "\w \ set_shadow_root_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_tag_name_locs ptr'. r h h'))" lemma set_shadow_root_get_tag_name_is_l_set_shadow_root_get_tag_name [instances]: "l_set_shadow_root_get_tag_name set_shadow_root_locs get_tag_name_locs" using set_shadow_root_is_l_set_shadow_root get_tag_name_is_l_get_tag_name apply(simp add: l_set_shadow_root_get_tag_name_def ) using set_shadow_root_get_tag_name by fast paragraph \new\_element\ locale l_new_element_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_tag_name get_tag_name_locs for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_tag_name :: "(_) element_ptr \ (_, tag_name) dom_prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_tag_name_new_element: "ptr' \ new_element_ptr \ h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" by (auto simp add: CD.get_tag_name_locs_def new_element_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t new_element_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t new_element_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t split: prod.splits if_splits option.splits elim!: bind_returns_result_E bind_returns_heap_E intro: is_element_ptr_kind_obtains) lemma new_element_empty_tag_name: "h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ h' \ get_tag_name new_element_ptr \\<^sub>r ''''" by(simp add: CD.get_tag_name_def new_element_tag_name) end locale l_new_element_get_tag_name = l_new_element + l_get_tag_name + assumes get_tag_name_new_element: "ptr' \ new_element_ptr \ h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" assumes new_element_empty_tag_name: "h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ h' \ get_tag_name new_element_ptr \\<^sub>r ''''" interpretation i_new_element_get_tag_name?: l_new_element_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_tag_name get_tag_name_locs by(auto simp add: l_new_element_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_new_element_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma new_element_get_tag_name_is_l_new_element_get_tag_name [instances]: "l_new_element_get_tag_name type_wf get_tag_name get_tag_name_locs" using new_element_is_l_new_element get_tag_name_is_l_get_tag_name apply(auto simp add: l_new_element_get_tag_name_def l_new_element_get_tag_name_axioms_def instances)[1] using get_tag_name_new_element new_element_empty_tag_name by fast+ paragraph \get\_shadow\_root\ locale l_set_mode_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_mode_get_tag_name: "\w \ set_mode_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_tag_name_locs ptr'. r h h'))" by(auto simp add: set_mode_locs_def CD.get_tag_name_locs_def all_args_def) end interpretation i_set_mode_get_tag_name?: l_set_mode_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_mode set_mode_locs DocumentClass.type_wf get_tag_name get_tag_name_locs by unfold_locales declare l_set_mode_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_set_mode_get_tag_name = l_set_mode + l_get_tag_name + assumes set_mode_get_tag_name: "\w \ set_mode_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_tag_name_locs ptr'. r h h'))" lemma set_mode_get_tag_name_is_l_set_mode_get_tag_name [instances]: "l_set_mode_get_tag_name type_wf set_mode set_mode_locs get_tag_name get_tag_name_locs" using set_mode_is_l_set_mode get_tag_name_is_l_get_tag_name apply(simp add: l_set_mode_get_tag_name_def l_set_mode_get_tag_name_axioms_def) using set_mode_get_tag_name by fast paragraph \new\_document\ locale l_new_document_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_tag_name get_tag_name_locs for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, tag_name) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_tag_name_new_document: "h \ new_document \\<^sub>r new_document_ptr \ h \ new_document \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" by(auto simp add: CD.get_tag_name_locs_def new_document_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t) end locale l_new_document_get_tag_name = l_get_tag_name_defs + assumes get_tag_name_new_document: "h \ new_document \\<^sub>r new_document_ptr \ h \ new_document \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" interpretation i_new_document_get_tag_name?: l_new_document_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_tag_name get_tag_name_locs by unfold_locales declare l_new_document_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def[instances] lemma new_document_get_tag_name_is_l_new_document_get_tag_name [instances]: "l_new_document_get_tag_name get_tag_name_locs" unfolding l_new_document_get_tag_name_def unfolding get_tag_name_locs_def using new_document_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t by blast paragraph \new\_shadow\_root\ locale l_new_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_tag_name_new_shadow_root: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" by (auto simp add: CD.get_tag_name_locs_def new_shadow_root_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t new_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t split: prod.splits if_splits option.splits elim!: bind_returns_result_E bind_returns_heap_E intro: is_element_ptr_kind_obtains) end locale l_new_shadow_root_get_tag_name = l_get_tag_name + assumes get_tag_name_new_shadow_root: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" interpretation i_new_shadow_root_get_tag_name?: l_new_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_tag_name get_tag_name_locs by(unfold_locales) declare l_new_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma new_shadow_root_get_tag_name_is_l_new_shadow_root_get_tag_name [instances]: "l_new_shadow_root_get_tag_name type_wf get_tag_name get_tag_name_locs" using get_tag_name_is_l_get_tag_name apply(auto simp add: l_new_shadow_root_get_tag_name_def l_new_shadow_root_get_tag_name_axioms_def instances)[1] using get_tag_name_new_shadow_root by fast paragraph \new\_character\_data\ locale l_new_character_data_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_tag_name get_tag_name_locs for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, tag_name) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_tag_name_new_character_data: "h \ new_character_data \\<^sub>r new_character_data_ptr \ h \ new_character_data \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" by(auto simp add: CD.get_tag_name_locs_def new_character_data_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t) end locale l_new_character_data_get_tag_name = l_get_tag_name_defs + assumes get_tag_name_new_character_data: "h \ new_character_data \\<^sub>r new_character_data_ptr \ h \ new_character_data \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" interpretation i_new_character_data_get_tag_name?: l_new_character_data_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_tag_name get_tag_name_locs by unfold_locales declare l_new_character_data_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def[instances] lemma new_character_data_get_tag_name_is_l_new_character_data_get_tag_name [instances]: "l_new_character_data_get_tag_name get_tag_name_locs" unfolding l_new_character_data_get_tag_name_def unfolding get_tag_name_locs_def using new_character_data_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t by blast paragraph \get\_tag\_type\ locale l_set_tag_name_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_tag_name_get_tag_name: assumes "h \ CD.a_set_tag_name element_ptr tag \\<^sub>h h'" shows "h' \ CD.a_get_tag_name element_ptr \\<^sub>r tag" using assms by(auto simp add: CD.a_get_tag_name_def CD.a_set_tag_name_def) lemma set_tag_name_get_tag_name_different_pointers: assumes "ptr \ ptr'" assumes "w \ CD.a_set_tag_name_locs ptr" assumes "h \ w \\<^sub>h h'" assumes "r \ CD.a_get_tag_name_locs ptr'" shows "r h h'" using assms by(auto simp add: all_args_def CD.a_set_tag_name_locs_def CD.a_get_tag_name_locs_def split: if_splits option.splits ) end interpretation i_set_tag_name_get_tag_name?: l_set_tag_name_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_tag_name get_tag_name_locs set_tag_name set_tag_name_locs by unfold_locales declare l_set_tag_name_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_tag_name_get_tag_name_is_l_set_tag_name_get_tag_name [instances]: "l_set_tag_name_get_tag_name type_wf get_tag_name get_tag_name_locs set_tag_name set_tag_name_locs" using set_tag_name_is_l_set_tag_name get_tag_name_is_l_get_tag_name apply(simp add: l_set_tag_name_get_tag_name_def l_set_tag_name_get_tag_name_axioms_def) using set_tag_name_get_tag_name set_tag_name_get_tag_name_different_pointers by fast+ subsubsection \attach\_shadow\_root\ locale l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_set_shadow_root_defs set_shadow_root set_shadow_root_locs + l_set_mode_defs set_mode set_mode_locs + l_get_tag_name_defs get_tag_name get_tag_name_locs + l_get_shadow_root_defs get_shadow_root get_shadow_root_locs for set_shadow_root :: "(_) element_ptr \ (_) shadow_root_ptr option \ ((_) heap, exception, unit) prog" and set_shadow_root_locs :: "(_) element_ptr \ ((_) heap, exception, unit) prog set" and set_mode :: "(_) shadow_root_ptr \ shadow_root_mode \ ((_) heap, exception, unit) prog" and set_mode_locs :: "(_) shadow_root_ptr \ ((_) heap, exception, unit) prog set" and get_tag_name :: "(_) element_ptr \ (_, char list) dom_prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin definition a_attach_shadow_root :: "(_) element_ptr \ shadow_root_mode \ (_, (_) shadow_root_ptr) dom_prog" where "a_attach_shadow_root element_ptr shadow_root_mode = do { tag \ get_tag_name element_ptr; (if tag \ safe_shadow_root_element_types then error HierarchyRequestError else return ()); prev_shadow_root \ get_shadow_root element_ptr; (if prev_shadow_root \ None then error HierarchyRequestError else return ()); new_shadow_root_ptr \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M; set_mode new_shadow_root_ptr shadow_root_mode; set_shadow_root element_ptr (Some new_shadow_root_ptr); return new_shadow_root_ptr }" end locale l_attach_shadow_root_defs = fixes attach_shadow_root :: "(_) element_ptr \ shadow_root_mode \ (_, (_) shadow_root_ptr) dom_prog" global_interpretation l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs set_shadow_root set_shadow_root_locs set_mode set_mode_locs get_tag_name get_tag_name_locs get_shadow_root get_shadow_root_locs defines attach_shadow_root = a_attach_shadow_root . locale l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs set_shadow_root set_shadow_root_locs set_mode set_mode_locs get_tag_name get_tag_name_locs get_shadow_root get_shadow_root_locs + l_attach_shadow_root_defs attach_shadow_root + l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_shadow_root set_shadow_root_locs + l_set_mode type_wf set_mode set_mode_locs + l_get_tag_name type_wf get_tag_name get_tag_name_locs + l_get_shadow_root type_wf get_shadow_root get_shadow_root_locs + l_known_ptr known_ptr for known_ptr :: "(_) object_ptr \ bool" and set_shadow_root :: "(_) element_ptr \ (_) shadow_root_ptr option \ ((_) heap, exception, unit) prog" and set_shadow_root_locs :: "(_) element_ptr \ ((_) heap, exception, unit) prog set" and set_mode :: "(_) shadow_root_ptr \ shadow_root_mode \ ((_) heap, exception, unit) prog" and set_mode_locs :: "(_) shadow_root_ptr \ ((_) heap, exception, unit) prog set" and attach_shadow_root :: "(_) element_ptr \ shadow_root_mode \ ((_) heap, exception, (_) shadow_root_ptr) prog" and type_wf :: "(_) heap \ bool" and get_tag_name :: "(_) element_ptr \ (_, char list) dom_prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" + assumes known_ptr_impl: "known_ptr = a_known_ptr" assumes attach_shadow_root_impl: "attach_shadow_root = a_attach_shadow_root" begin lemmas attach_shadow_root_def = a_attach_shadow_root_def[folded attach_shadow_root_impl] lemma attach_shadow_root_element_ptr_in_heap: assumes "h \ ok (attach_shadow_root element_ptr shadow_root_mode)" shows "element_ptr |\| element_ptr_kinds h" proof - obtain h' where "h \ attach_shadow_root element_ptr shadow_root_mode \\<^sub>h h'" using assms by auto then obtain h2 h3 new_shadow_root_ptr where h2: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h2" and new_shadow_root_ptr: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr" and h3: "h2 \ set_mode new_shadow_root_ptr shadow_root_mode \\<^sub>h h3" and "h3 \ set_shadow_root element_ptr (Some new_shadow_root_ptr) \\<^sub>h h'" by(auto simp add: attach_shadow_root_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_tag_name_pure, rotated] bind_returns_heap_E2[rotated, OF get_shadow_root_pure, rotated] split: if_splits) then have "element_ptr |\| element_ptr_kinds h3" using set_shadow_root_ptr_in_heap by blast moreover have "object_ptr_kinds h2 = object_ptr_kinds h |\| {|cast new_shadow_root_ptr|}" using h2 new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_new_ptr new_shadow_root_ptr by auto moreover have "object_ptr_kinds h2 = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_mode_writes h3]) using set_mode_pointers_preserved apply blast by (auto simp add: reflp_def transp_def) ultimately show ?thesis by (metis (no_types, lifting) cast_document_ptr_not_node_ptr(2) element_ptr_kinds_commutes finsertE funion_finsert_right node_ptr_kinds_commutes sup_bot.right_neutral) qed lemma create_shadow_root_known_ptr: assumes "h \ attach_shadow_root element_ptr shadow_root_mode \\<^sub>r new_shadow_root_ptr" shows "known_ptr (cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_shadow_root_ptr)" using assms by(auto simp add: attach_shadow_root_def known_ptr_impl ShadowRootClass.a_known_ptr_def new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def Let_def elim!: bind_returns_result_E) end locale l_attach_shadow_root = l_attach_shadow_root_defs interpretation i_attach_shadow_root?: l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr set_shadow_root set_shadow_root_locs set_mode set_mode_locs attach_shadow_root type_wf get_tag_name get_tag_name_locs get_shadow_root get_shadow_root_locs by(auto simp add: l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def attach_shadow_root_def instances) declare l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_parent\ global_interpretation l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs defines get_parent = "l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_parent get_child_nodes" and get_parent_locs = "l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_parent_locs get_child_nodes_locs" . interpretation i_get_parent?: l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs get_parent get_parent_locs by(simp add: l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def get_parent_def get_parent_locs_def instances) declare l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_parent_is_l_get_parent [instances]: "l_get_parent type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs" apply(simp add: l_get_parent_def l_get_parent_axioms_def instances) using get_parent_reads get_parent_ok get_parent_ptr_in_heap get_parent_pure get_parent_parent_in_heap get_parent_child_dual using get_parent_reads_pointers by blast paragraph \set\_disconnected\_nodes\ locale l_set_disconnected_nodes_get_parent\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_disconnected_nodes_get_child_nodes + l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_disconnected_nodes_get_parent [simp]: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_parent_locs. r h h'))" by(auto simp add: get_parent_locs_def CD.set_disconnected_nodes_locs_def all_args_def) end interpretation i_set_disconnected_nodes_get_parent?: l_set_disconnected_nodes_get_parent\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_disconnected_nodes set_disconnected_nodes_locs get_child_nodes get_child_nodes_locs type_wf DocumentClass.type_wf known_ptr known_ptrs get_parent get_parent_locs by (simp add: l_set_disconnected_nodes_get_parent\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_disconnected_nodes_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_disconnected_nodes_get_parent_is_l_set_disconnected_nodes_get_parent [instances]: "l_set_disconnected_nodes_get_parent set_disconnected_nodes_locs get_parent_locs" by(simp add: l_set_disconnected_nodes_get_parent_def) subsubsection \get\_root\_node\ global_interpretation l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_parent get_parent_locs defines get_root_node = "l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_root_node get_parent" and get_root_node_locs = "l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_root_node_locs get_parent_locs" and get_ancestors = "l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_ancestors get_parent" and get_ancestors_locs = "l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_ancestors_locs get_parent_locs" . declare a_get_ancestors.simps [code] interpretation i_get_root_node?: l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_ancestors get_ancestors_locs get_root_node get_root_node_locs by(simp add: l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def get_root_node_def get_root_node_locs_def get_ancestors_def get_ancestors_locs_def instances) declare l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_ancestors_is_l_get_ancestors [instances]: "l_get_ancestors get_ancestors" apply(auto simp add: l_get_ancestors_def)[1] using get_ancestors_ptr_in_heap apply fast using get_ancestors_ptr apply fast done lemma get_root_node_is_l_get_root_node [instances]: "l_get_root_node get_root_node get_parent" by (simp add: l_get_root_node_def Shadow_DOM.i_get_root_node.get_root_node_no_parent) subsubsection \get\_root\_node\_si\ locale l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_parent_defs get_parent get_parent_locs + l_get_host_defs get_host get_host_locs for get_parent :: "(_) node_ptr \ ((_) heap, exception, (_::linorder) object_ptr option) prog" and get_parent_locs :: "((_) heap \ (_) heap \ bool) set" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" begin partial_function (dom_prog) a_get_ancestors_si :: "(_::linorder) object_ptr \ (_, (_) object_ptr list) dom_prog" where "a_get_ancestors_si ptr = do { check_in_heap ptr; ancestors \ (case cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr of Some node_ptr \ do { parent_ptr_opt \ get_parent node_ptr; (case parent_ptr_opt of Some parent_ptr \ a_get_ancestors_si parent_ptr | None \ return []) } | None \ (case cast ptr of Some shadow_root_ptr \ do { host \ get_host shadow_root_ptr; a_get_ancestors_si (cast host) } | None \ return [])); return (ptr # ancestors) }" definition "a_get_ancestors_si_locs = get_parent_locs \ get_host_locs" definition a_get_root_node_si :: "(_) object_ptr \ (_, (_) object_ptr) dom_prog" where "a_get_root_node_si ptr = do { ancestors \ a_get_ancestors_si ptr; return (last ancestors) }" definition "a_get_root_node_si_locs = a_get_ancestors_si_locs" end locale l_get_ancestors_si_defs = fixes get_ancestors_si :: "(_::linorder) object_ptr \ (_, (_) object_ptr list) dom_prog" fixes get_ancestors_si_locs :: "((_) heap \ (_) heap \ bool) set" locale l_get_root_node_si_defs = fixes get_root_node_si :: "(_) object_ptr \ (_, (_) object_ptr) dom_prog" fixes get_root_node_si_locs :: "((_) heap \ (_) heap \ bool) set" locale l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_parent + l_get_host + l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_get_ancestors_si_defs + l_get_root_node_si_defs + assumes get_ancestors_si_impl: "get_ancestors_si = a_get_ancestors_si" assumes get_ancestors_si_locs_impl: "get_ancestors_si_locs = a_get_ancestors_si_locs" assumes get_root_node_si_impl: "get_root_node_si = a_get_root_node_si" assumes get_root_node_si_locs_impl: "get_root_node_si_locs = a_get_root_node_si_locs" begin lemmas get_ancestors_si_def = a_get_ancestors_si.simps[folded get_ancestors_si_impl] lemmas get_ancestors_si_locs_def = a_get_ancestors_si_locs_def[folded get_ancestors_si_locs_impl] lemmas get_root_node_si_def = a_get_root_node_si_def[folded get_root_node_si_impl get_ancestors_si_impl] lemmas get_root_node_si_locs_def = a_get_root_node_si_locs_def[folded get_root_node_si_locs_impl get_ancestors_si_locs_impl] lemma get_ancestors_si_pure [simp]: "pure (get_ancestors_si ptr) h" proof - have "\ptr h h' x. h \ get_ancestors_si ptr \\<^sub>r x \ h \ get_ancestors_si ptr \\<^sub>h h' \ h = h'" proof (induct rule: a_get_ancestors_si.fixp_induct[folded get_ancestors_si_impl]) case 1 then show ?case by(rule admissible_dom_prog) next case 2 then show ?case by simp next case (3 f) then show ?case using get_parent_pure get_host_pure apply(auto simp add: pure_returns_heap_eq pure_def split: option.splits elim!: bind_returns_heap_E bind_returns_result_E dest!: pure_returns_heap_eq[rotated, OF check_in_heap_pure])[1] apply (meson option.simps(3) returns_result_eq) apply(metis get_parent_pure pure_returns_heap_eq) apply(metis get_host_pure pure_returns_heap_eq) done qed then show ?thesis by (meson pure_eq_iff) qed lemma get_root_node_si_pure [simp]: "pure (get_root_node_si ptr) h" by(auto simp add: get_root_node_si_def bind_pure_I) lemma get_ancestors_si_ptr_in_heap: assumes "h \ ok (get_ancestors_si ptr)" shows "ptr |\| object_ptr_kinds h" using assms by(auto simp add: get_ancestors_si_def check_in_heap_ptr_in_heap elim!: bind_is_OK_E dest: is_OK_returns_result_I) lemma get_ancestors_si_ptr: assumes "h \ get_ancestors_si ptr \\<^sub>r ancestors" shows "ptr \ set ancestors" using assms by(simp add: get_ancestors_si_def) (auto elim!: bind_returns_result_E2 split: option.splits intro!: bind_pure_I) lemma get_ancestors_si_never_empty: assumes "h \ get_ancestors_si child \\<^sub>r ancestors" shows "ancestors \ []" using assms apply(simp add: get_ancestors_si_def) by(auto elim!: bind_returns_result_E2 split: option.splits) (* lemma get_ancestors_si_not_node: assumes "h \ get_ancestors_si ptr \\<^sub>r ancestors" assumes "\is_node_ptr_kind ptr" shows "ancestors = [ptr]" using assms by (simp add: get_ancestors_si_def) (auto elim!: bind_returns_result_E2 split: option.splits) *) lemma get_root_node_si_no_parent: "h \ get_parent node_ptr \\<^sub>r None \ h \ get_root_node_si (cast node_ptr) \\<^sub>r cast node_ptr" apply(auto simp add: check_in_heap_def get_root_node_si_def get_ancestors_si_def intro!: bind_pure_returns_result_I )[1] using get_parent_ptr_in_heap by blast lemma get_root_node_si_root_not_shadow_root: assumes "h \ get_root_node_si ptr \\<^sub>r root" shows "\ is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root" using assms proof(auto simp add: get_root_node_si_def elim!: bind_returns_result_E2) fix y assume "h \ get_ancestors_si ptr \\<^sub>r y" and "is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r (last y)" and "root = last y" then show False proof(induct y arbitrary: ptr) case Nil then show ?case using assms(1) get_ancestors_si_never_empty by blast next case (Cons a x) then show ?case apply(auto simp add: get_ancestors_si_def[of ptr] elim!: bind_returns_result_E2 split: option.splits if_splits)[1] using get_ancestors_si_never_empty apply blast using Cons.prems(2) apply auto[1] using \is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r (last y)\ \root = last y\ by auto qed qed end global_interpretation l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_parent get_parent_locs get_host get_host_locs defines get_root_node_si = a_get_root_node_si and get_root_node_si_locs = a_get_root_node_si_locs and get_ancestors_si = a_get_ancestors_si and get_ancestors_si_locs = a_get_ancestors_si_locs . declare a_get_ancestors_si.simps [code] interpretation i_get_root_node_si?: l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_host get_host_locs get_ancestors_si get_ancestors_si_locs get_root_node_si get_root_node_si_locs apply(auto simp add: l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances)[1] by(auto simp add: get_root_node_si_def get_root_node_si_locs_def get_ancestors_si_def get_ancestors_si_locs_def) declare l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma get_ancestors_si_is_l_get_ancestors [instances]: "l_get_ancestors get_ancestors_si" unfolding l_get_ancestors_def using get_ancestors_si_pure get_ancestors_si_ptr get_ancestors_si_ptr_in_heap by blast lemma get_root_node_si_is_l_get_root_node [instances]: "l_get_root_node get_root_node_si get_parent" apply(simp add: l_get_root_node_def) using get_root_node_si_no_parent by fast paragraph \set\_disconnected\_nodes\ locale l_set_disconnected_nodes_get_ancestors_si\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_disconnected_nodes_get_parent + l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_disconnected_nodes_get_host begin lemma set_disconnected_nodes_get_ancestors_si: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_ancestors_si_locs. r h h'))" by(auto simp add: get_parent_locs_def set_disconnected_nodes_locs_def set_disconnected_nodes_get_host get_ancestors_si_locs_def all_args_def) end locale l_set_disconnected_nodes_get_ancestors_si = l_set_disconnected_nodes_defs + l_get_ancestors_si_defs + assumes set_disconnected_nodes_get_ancestors_si: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_ancestors_si_locs. r h h'))" interpretation i_set_disconnected_nodes_get_ancestors_si?: l_set_disconnected_nodes_get_ancestors_si\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_disconnected_nodes set_disconnected_nodes_locs get_parent get_parent_locs type_wf known_ptr known_ptrs get_child_nodes get_child_nodes_locs get_host get_host_locs get_ancestors_si get_ancestors_si_locs get_root_node_si get_root_node_si_locs DocumentClass.type_wf by (auto simp add: l_set_disconnected_nodes_get_ancestors_si\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_disconnected_nodes_get_ancestors_si\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_disconnected_nodes_get_ancestors_si_is_l_set_disconnected_nodes_get_ancestors_si [instances]: "l_set_disconnected_nodes_get_ancestors_si set_disconnected_nodes_locs get_ancestors_si_locs" using instances apply(simp add: l_set_disconnected_nodes_get_ancestors_si_def) using set_disconnected_nodes_get_ancestors_si by fast subsubsection \get\_attribute\ lemma get_attribute_is_l_get_attribute [instances]: "l_get_attribute type_wf get_attribute get_attribute_locs" apply(auto simp add: l_get_attribute_def)[1] using i_get_attribute.get_attribute_reads apply fast using type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t i_get_attribute.get_attribute_ok apply blast using i_get_attribute.get_attribute_ptr_in_heap apply fast done subsubsection \to\_tree\_order\ global_interpretation l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs defines to_tree_order = "l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_to_tree_order get_child_nodes" . declare a_to_tree_order.simps [code] interpretation i_to_tree_order?: l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M ShadowRootClass.known_ptr ShadowRootClass.type_wf Shadow_DOM.get_child_nodes Shadow_DOM.get_child_nodes_locs to_tree_order by(auto simp add: l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def to_tree_order_def instances) declare l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \to\_tree\_order\_si\ locale l_to_tree_order_si\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_child_nodes_defs get_child_nodes get_child_nodes_locs + l_get_shadow_root_defs get_shadow_root get_shadow_root_locs for get_child_nodes :: "(_::linorder) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin partial_function (dom_prog) a_to_tree_order_si :: "(_) object_ptr \ (_, (_) object_ptr list) dom_prog" where "a_to_tree_order_si ptr = (do { children \ get_child_nodes ptr; shadow_root_part \ (case cast ptr of Some element_ptr \ do { shadow_root_opt \ get_shadow_root element_ptr; (case shadow_root_opt of Some shadow_root_ptr \ return [cast shadow_root_ptr] | None \ return []) } | None \ return []); treeorders \ map_M a_to_tree_order_si ((map (cast) children) @ shadow_root_part); return (ptr # concat treeorders) })" end locale l_to_tree_order_si_defs = fixes to_tree_order_si :: "(_) object_ptr \ (_, (_) object_ptr list) dom_prog" global_interpretation l_to_tree_order_si\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_shadow_root get_shadow_root_locs defines to_tree_order_si = "a_to_tree_order_si" . declare a_to_tree_order_si.simps [code] locale l_to_tree_order_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_to_tree_order_si_defs + l_to_tree_order_si\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_get_child_nodes + l_get_shadow_root + assumes to_tree_order_si_impl: "to_tree_order_si = a_to_tree_order_si" begin lemmas to_tree_order_si_def = a_to_tree_order_si.simps[folded to_tree_order_si_impl] lemma to_tree_order_si_pure [simp]: "pure (to_tree_order_si ptr) h" proof - have "\ptr h h' x. h \ to_tree_order_si ptr \\<^sub>r x \ h \ to_tree_order_si ptr \\<^sub>h h' \ h = h'" proof (induct rule: a_to_tree_order_si.fixp_induct[folded to_tree_order_si_impl]) case 1 then show ?case by (rule admissible_dom_prog) next case 2 then show ?case by simp next case (3 f) then have "\x h. pure (f x) h" by (metis is_OK_returns_heap_E is_OK_returns_result_E pure_def) then have "\xs h. pure (map_M f xs) h" by(rule map_M_pure_I) then show ?case by(auto elim!: bind_returns_heap_E2 split: option.splits) qed then show ?thesis unfolding pure_def by (metis is_OK_returns_heap_E is_OK_returns_result_E) qed end interpretation i_to_tree_order_si?: l_to_tree_order_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M to_tree_order_si get_child_nodes get_child_nodes_locs get_shadow_root get_shadow_root_locs type_wf known_ptr by(auto simp add: l_to_tree_order_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_to_tree_order_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def to_tree_order_si_def instances) declare l_to_tree_order_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \first\_in\_tree\_order\ global_interpretation l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs to_tree_order defines first_in_tree_order = "l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_first_in_tree_order to_tree_order" . interpretation i_first_in_tree_order?: l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M to_tree_order first_in_tree_order by(auto simp add: l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def first_in_tree_order_def) declare l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma to_tree_order_is_l_to_tree_order [instances]: "l_to_tree_order to_tree_order" by(auto simp add: l_to_tree_order_def) subsubsection \first\_in\_tree\_order\ global_interpretation l_dummy defines first_in_tree_order_si = "l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_first_in_tree_order to_tree_order_si" . subsubsection \get\_element\_by\ global_interpretation l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs to_tree_order first_in_tree_order get_attribute get_attribute_locs defines get_element_by_id = "l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_element_by_id first_in_tree_order get_attribute" and get_elements_by_class_name = "l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_elements_by_class_name to_tree_order get_attribute" and get_elements_by_tag_name = "l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_elements_by_tag_name to_tree_order" . interpretation i_get_element_by?: l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M to_tree_order first_in_tree_order get_attribute get_attribute_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name type_wf by(auto simp add: l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def get_element_by_id_def get_elements_by_class_name_def get_elements_by_tag_name_def instances) declare l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma get_element_by_is_l_get_element_by [instances]: "l_get_element_by get_element_by_id get_elements_by_tag_name to_tree_order" apply(auto simp add: l_get_element_by_def)[1] using get_element_by_id_result_in_tree_order apply fast done subsubsection \get\_element\_by\_si\ global_interpretation l_dummy defines get_element_by_id_si = "l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_element_by_id first_in_tree_order_si get_attribute" and get_elements_by_class_name_si = "l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_elements_by_class_name to_tree_order_si get_attribute" and get_elements_by_tag_name_si = "l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_elements_by_tag_name to_tree_order_si" . subsubsection \find\_slot\ locale l_find_slot\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_parent_defs get_parent get_parent_locs + l_get_shadow_root_defs get_shadow_root get_shadow_root_locs + l_get_mode_defs get_mode get_mode_locs + l_get_attribute_defs get_attribute get_attribute_locs + l_get_tag_name_defs get_tag_name get_tag_name_locs + l_first_in_tree_order_defs first_in_tree_order for get_parent :: "(_) node_ptr \ ((_) heap, exception, (_::linorder) object_ptr option) prog" and get_parent_locs :: "((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_mode :: "(_) shadow_root_ptr \ ((_) heap, exception, shadow_root_mode) prog" and get_mode_locs :: "(_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" and get_attribute :: "(_) element_ptr \ char list \ ((_) heap, exception, char list option) prog" and get_attribute_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and first_in_tree_order :: "(_) object_ptr \ ((_) object_ptr \ ((_) heap, exception, (_) element_ptr option) prog) \ ((_) heap, exception, (_) element_ptr option) prog" begin definition a_find_slot :: "bool \ (_) node_ptr \ (_, (_) element_ptr option) dom_prog" where "a_find_slot open_flag slotable = do { parent_opt \ get_parent slotable; (case parent_opt of Some parent \ if is_element_ptr_kind parent then do { shadow_root_ptr_opt \ get_shadow_root (the (cast parent)); (case shadow_root_ptr_opt of Some shadow_root_ptr \ do { shadow_root_mode \ get_mode shadow_root_ptr; if open_flag \ shadow_root_mode \ Open then return None else first_in_tree_order (cast shadow_root_ptr) (\ptr. if is_element_ptr_kind ptr then do { tag \ get_tag_name (the (cast ptr)); name_attr \ get_attribute (the (cast ptr)) ''name''; slotable_name_attr \ (if is_element_ptr_kind slotable then get_attribute (the (cast slotable)) ''slot'' else return None); (if (tag = ''slot'' \ (name_attr = slotable_name_attr \ (name_attr = None \ slotable_name_attr = Some '''') \ (name_attr = Some '''' \ slotable_name_attr = None))) then return (Some (the (cast ptr))) else return None)} else return None)} | None \ return None)} else return None | _ \ return None)}" definition a_assigned_slot :: "(_) node_ptr \ (_, (_) element_ptr option) dom_prog" where "a_assigned_slot = a_find_slot True" end global_interpretation l_find_slot\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_parent get_parent_locs get_shadow_root get_shadow_root_locs get_mode get_mode_locs get_attribute get_attribute_locs get_tag_name get_tag_name_locs first_in_tree_order defines find_slot = a_find_slot and assigned_slot = a_assigned_slot . locale l_find_slot_defs = fixes find_slot :: "bool \ (_) node_ptr \ (_, (_) element_ptr option) dom_prog" and assigned_slot :: "(_) node_ptr \ (_, (_) element_ptr option) dom_prog" locale l_find_slot\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_find_slot\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_find_slot_defs + l_get_parent + l_get_shadow_root + l_get_mode + l_get_attribute + l_get_tag_name + l_to_tree_order + l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + assumes find_slot_impl: "find_slot = a_find_slot" assumes assigned_slot_impl: "assigned_slot = a_assigned_slot" begin lemmas find_slot_def = find_slot_impl[unfolded a_find_slot_def] lemmas assigned_slot_def = assigned_slot_impl[unfolded a_assigned_slot_def] lemma find_slot_ptr_in_heap: assumes "h \ find_slot open_flag slotable \\<^sub>r slot_opt" shows "slotable |\| node_ptr_kinds h" using assms apply(auto simp add: find_slot_def elim!: bind_returns_result_E2)[1] using get_parent_ptr_in_heap by blast lemma find_slot_slot_in_heap: assumes "h \ find_slot open_flag slotable \\<^sub>r Some slot" shows "slot |\| element_ptr_kinds h" using assms apply(auto simp add: find_slot_def first_in_tree_order_def elim!: bind_returns_result_E2 map_filter_M_pure_E[where y=slot] split: option.splits if_splits list.splits intro!: map_filter_M_pure bind_pure_I)[1] using get_tag_name_ptr_in_heap by blast+ lemma find_slot_pure [simp]: "pure (find_slot open_flag slotable) h" by(auto simp add: find_slot_def first_in_tree_order_def intro!: bind_pure_I map_filter_M_pure split: option.splits list.splits) end interpretation i_find_slot?: l_find_slot\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_parent get_parent_locs get_shadow_root get_shadow_root_locs get_mode get_mode_locs get_attribute get_attribute_locs get_tag_name get_tag_name_locs first_in_tree_order find_slot assigned_slot type_wf known_ptr known_ptrs get_child_nodes get_child_nodes_locs to_tree_order by (auto simp add: find_slot_def assigned_slot_def l_find_slot\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_find_slot\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_find_slot\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_find_slot = l_find_slot_defs + assumes find_slot_ptr_in_heap: "h \ find_slot open_flag slotable \\<^sub>r slot_opt \ slotable |\| node_ptr_kinds h" assumes find_slot_slot_in_heap: "h \ find_slot open_flag slotable \\<^sub>r Some slot \ slot |\| element_ptr_kinds h" assumes find_slot_pure [simp]: "pure (find_slot open_flag slotable) h" lemma find_slot_is_l_find_slot [instances]: "l_find_slot find_slot" apply(auto simp add: l_find_slot_def)[1] using find_slot_ptr_in_heap apply fast using find_slot_slot_in_heap apply fast done subsubsection \get\_disconnected\_nodes\ locale l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD: l_get_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs + l_type_wf type_wf for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_disconnected_nodes :: "(_) document_ptr \ (_, (_) node_ptr list) dom_prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" begin lemma get_disconnected_nodes_ok: "type_wf h \ document_ptr |\| document_ptr_kinds h \ h \ ok (get_disconnected_nodes document_ptr)" apply(unfold type_wf_impl get_disconnected_nodes_impl[unfolded a_get_disconnected_nodes_def]) using CD.get_disconnected_nodes_ok CD.type_wf_impl ShadowRootClass.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t by blast end interpretation i_get_disconnected_nodes?: l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_disconnected_nodes get_disconnected_nodes_locs by(auto simp add: l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_disconnected_nodes_is_l_get_disconnected_nodes [instances]: "l_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs" apply(auto simp add: l_get_disconnected_nodes_def)[1] using i_get_disconnected_nodes.get_disconnected_nodes_reads apply fast using get_disconnected_nodes_ok apply fast using i_get_disconnected_nodes.get_disconnected_nodes_ptr_in_heap apply fast done paragraph \set\_child\_nodes\ locale l_set_child_nodes_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_child_nodes_get_disconnected_nodes: "\w \ set_child_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_disconnected_nodes_locs ptr'. r h h'))" by(auto simp add: set_child_nodes_locs_def CD.set_child_nodes_locs_def CD.get_disconnected_nodes_locs_def all_args_def elim: get_M_document_put_M_shadow_root split: option.splits) end interpretation i_set_child_nodes_get_disconnected_nodes?: l_set_child_nodes_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr set_child_nodes set_child_nodes_locs Core_DOM_Functions.set_child_nodes Core_DOM_Functions.set_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs apply(auto simp add: l_set_child_nodes_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)[1] by(unfold_locales) declare l_set_child_nodes_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_child_nodes_get_disconnected_nodes_is_l_set_child_nodes_get_disconnected_nodes [instances]: "l_set_child_nodes_get_disconnected_nodes type_wf set_child_nodes set_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs" using set_child_nodes_is_l_set_child_nodes get_disconnected_nodes_is_l_get_disconnected_nodes apply(simp add: l_set_child_nodes_get_disconnected_nodes_def l_set_child_nodes_get_disconnected_nodes_axioms_def ) using set_child_nodes_get_disconnected_nodes by fast paragraph \set\_disconnected\_nodes\ lemma set_disconnected_nodes_get_disconnected_nodes_l_set_disconnected_nodes_get_disconnected_nodes [instances]: "l_set_disconnected_nodes_get_disconnected_nodes ShadowRootClass.type_wf get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs" apply(auto simp add: l_set_disconnected_nodes_get_disconnected_nodes_def l_set_disconnected_nodes_get_disconnected_nodes_axioms_def instances)[1] using i_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes apply fast using i_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes_different_pointers apply fast done paragraph \delete\_shadow\_root\ locale l_delete_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_disconnected_nodes_delete_shadow_root: "cast shadow_root_ptr \ ptr' \ h \ delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M shadow_root_ptr \\<^sub>h h' \ r \ get_disconnected_nodes_locs ptr' \ r h h'" by(auto simp add: CD.get_disconnected_nodes_locs_def delete_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t) end locale l_delete_shadow_root_get_disconnected_nodes = l_get_disconnected_nodes_defs + assumes get_disconnected_nodes_delete_shadow_root: "cast shadow_root_ptr \ ptr' \ h \ delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M shadow_root_ptr \\<^sub>h h' \ r \ get_disconnected_nodes_locs ptr' \ r h h'" interpretation l_delete_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_disconnected_nodes get_disconnected_nodes_locs by(auto simp add: l_delete_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) lemma l_delete_shadow_root_get_disconnected_nodes_get_disconnected_nodes_locs [instances]: "l_delete_shadow_root_get_disconnected_nodes get_disconnected_nodes_locs" apply(auto simp add: l_delete_shadow_root_get_disconnected_nodes_def)[1] using get_disconnected_nodes_delete_shadow_root apply fast done paragraph \set\_shadow\_root\ locale l_set_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_shadow_root_get_disconnected_nodes: "\w \ set_shadow_root_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_disconnected_nodes_locs ptr'. r h h'))" by(auto simp add: set_shadow_root_locs_def CD.get_disconnected_nodes_locs_def all_args_def) end interpretation i_set_shadow_root_get_disconnected_nodes?: l_set_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_shadow_root set_shadow_root_locs DocumentClass.type_wf get_disconnected_nodes get_disconnected_nodes_locs apply(auto simp add: l_set_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)[1] by(unfold_locales) declare l_set_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_set_shadow_root_get_disconnected_nodes = l_set_shadow_root_defs + l_get_disconnected_nodes_defs + assumes set_shadow_root_get_disconnected_nodes: "\w \ set_shadow_root_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_disconnected_nodes_locs ptr'. r h h'))" lemma set_shadow_root_get_disconnected_nodes_is_l_set_shadow_root_get_disconnected_nodes [instances]: "l_set_shadow_root_get_disconnected_nodes set_shadow_root_locs get_disconnected_nodes_locs" using set_shadow_root_is_l_set_shadow_root get_disconnected_nodes_is_l_get_disconnected_nodes apply(simp add: l_set_shadow_root_get_disconnected_nodes_def ) using set_shadow_root_get_disconnected_nodes by fast paragraph \set\_mode\ locale l_set_mode_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_mode_get_disconnected_nodes: "\w \ set_mode_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_disconnected_nodes_locs ptr'. r h h'))" by(auto simp add: set_mode_locs_def CD.get_disconnected_nodes_locs_impl[unfolded CD.a_get_disconnected_nodes_locs_def] all_args_def) end interpretation i_set_mode_get_disconnected_nodes?: l_set_mode_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_mode set_mode_locs DocumentClass.type_wf get_disconnected_nodes get_disconnected_nodes_locs by unfold_locales declare l_set_mode_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_set_mode_get_disconnected_nodes = l_set_mode + l_get_disconnected_nodes + assumes set_mode_get_disconnected_nodes: "\w \ set_mode_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_disconnected_nodes_locs ptr'. r h h'))" lemma set_mode_get_disconnected_nodes_is_l_set_mode_get_disconnected_nodes [instances]: "l_set_mode_get_disconnected_nodes type_wf set_mode set_mode_locs get_disconnected_nodes get_disconnected_nodes_locs" using set_mode_is_l_set_mode get_disconnected_nodes_is_l_get_disconnected_nodes apply(simp add: l_set_mode_get_disconnected_nodes_def l_set_mode_get_disconnected_nodes_axioms_def) using set_mode_get_disconnected_nodes by fast paragraph \new\_shadow\_root\ locale l_new_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_disconnected_nodes :: "(_) document_ptr \ (_, (_) node_ptr list) dom_prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_disconnected_nodes_new_shadow_root_different_pointers: "cast new_shadow_root_ptr \ ptr' \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ r \ get_disconnected_nodes_locs ptr' \ r h h'" by(auto simp add: CD.get_disconnected_nodes_locs_def new_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t) lemma new_shadow_root_no_disconnected_nodes: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ h' \ get_disconnected_nodes (cast new_shadow_root_ptr) \\<^sub>r []" by(simp add: CD.get_disconnected_nodes_def new_shadow_root_disconnected_nodes) end interpretation i_new_shadow_root_get_disconnected_nodes?: l_new_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_disconnected_nodes get_disconnected_nodes_locs by unfold_locales declare l_new_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_new_shadow_root_get_disconnected_nodes = l_get_disconnected_nodes_defs + assumes get_disconnected_nodes_new_shadow_root_different_pointers: "cast new_shadow_root_ptr \ ptr' \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ r \ get_disconnected_nodes_locs ptr' \ r h h'" assumes new_shadow_root_no_disconnected_nodes: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ h' \ get_disconnected_nodes (cast new_shadow_root_ptr) \\<^sub>r []" lemma new_shadow_root_get_disconnected_nodes_is_l_new_shadow_root_get_disconnected_nodes [instances]: "l_new_shadow_root_get_disconnected_nodes get_disconnected_nodes get_disconnected_nodes_locs" apply (auto simp add: l_new_shadow_root_get_disconnected_nodes_def)[1] using get_disconnected_nodes_new_shadow_root_different_pointers apply fast using new_shadow_root_no_disconnected_nodes apply blast done subsubsection \remove\_shadow\_root\ locale l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_child_nodes_defs get_child_nodes get_child_nodes_locs + l_get_shadow_root_defs get_shadow_root get_shadow_root_locs + l_set_shadow_root_defs set_shadow_root set_shadow_root_locs + l_get_disconnected_nodes_defs get_disconnected_nodes get_disconnected_nodes_locs for get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and set_shadow_root :: "(_) element_ptr \ (_) shadow_root_ptr option \ ((_) heap, exception, unit) prog" and set_shadow_root_locs :: "(_) element_ptr \ ((_) heap, exception, unit) prog set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" begin definition a_remove_shadow_root :: "(_) element_ptr \ (_, unit) dom_prog" where "a_remove_shadow_root element_ptr = do { shadow_root_ptr_opt \ get_shadow_root element_ptr; (case shadow_root_ptr_opt of Some shadow_root_ptr \ do { children \ get_child_nodes (cast shadow_root_ptr); disconnected_nodes \ get_disconnected_nodes (cast shadow_root_ptr); (if children = [] \ disconnected_nodes = [] then do { set_shadow_root element_ptr None; delete_M shadow_root_ptr } else do { error HierarchyRequestError }) } | None \ error HierarchyRequestError) }" definition a_remove_shadow_root_locs :: "(_) element_ptr \ (_) shadow_root_ptr \ ((_, unit) dom_prog) set" where "a_remove_shadow_root_locs element_ptr shadow_root_ptr \ set_shadow_root_locs element_ptr \ {delete_M shadow_root_ptr}" end global_interpretation l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_shadow_root get_shadow_root_locs set_shadow_root set_shadow_root_locs get_disconnected_nodes get_disconnected_nodes_locs defines remove_shadow_root = "a_remove_shadow_root" and remove_shadow_root_locs = a_remove_shadow_root_locs . locale l_remove_shadow_root_defs = fixes remove_shadow_root :: "(_) element_ptr \ (_, unit) dom_prog" fixes remove_shadow_root_locs :: "(_) element_ptr \ (_) shadow_root_ptr \ ((_, unit) dom_prog) set" locale l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_remove_shadow_root_defs + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_child_nodes + l_get_disconnected_nodes + assumes remove_shadow_root_impl: "remove_shadow_root = a_remove_shadow_root" assumes remove_shadow_root_locs_impl: "remove_shadow_root_locs = a_remove_shadow_root_locs" begin lemmas remove_shadow_root_def = remove_shadow_root_impl[unfolded remove_shadow_root_def a_remove_shadow_root_def] lemmas remove_shadow_root_locs_def = remove_shadow_root_locs_impl[unfolded remove_shadow_root_locs_def a_remove_shadow_root_locs_def] lemma remove_shadow_root_writes: "writes (remove_shadow_root_locs element_ptr (the |h \ get_shadow_root element_ptr|\<^sub>r)) (remove_shadow_root element_ptr) h h'" apply(auto simp add: remove_shadow_root_locs_def remove_shadow_root_def all_args_def writes_union_right_I writes_union_left_I set_shadow_root_writes intro!: writes_bind writes_bind_pure[OF get_shadow_root_pure] writes_bind_pure[OF get_child_nodes_pure] intro: writes_subset[OF set_shadow_root_writes] writes_subset[OF writes_singleton2] split: option.splits)[1] using writes_union_left_I[OF set_shadow_root_writes] apply (metis inf_sup_aci(5) insert_is_Un) using writes_union_right_I[OF writes_singleton[of delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M]] by (smt insert_is_Un writes_singleton2 writes_union_left_I) end interpretation i_remove_shadow_root?: l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_shadow_root get_shadow_root_locs set_shadow_root set_shadow_root_locs get_disconnected_nodes get_disconnected_nodes_locs remove_shadow_root remove_shadow_root_locs type_wf known_ptr by(auto simp add: l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def remove_shadow_root_def remove_shadow_root_locs_def instances) declare l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] paragraph \get\_child\_nodes\ locale l_remove_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma remove_shadow_root_get_child_nodes_different_pointers: assumes "ptr \ cast shadow_root_ptr" assumes "w \ remove_shadow_root_locs element_ptr shadow_root_ptr" assumes "h \ w \\<^sub>h h'" assumes "r \ get_child_nodes_locs ptr" shows "r h h'" using assms by(auto simp add: all_args_def get_child_nodes_locs_def CD.get_child_nodes_locs_def remove_shadow_root_locs_def set_shadow_root_locs_def delete_shadow_root_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t delete_shadow_root_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t intro: is_shadow_root_ptr_kind_obtains simp add: delete_shadow_root_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t delete_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t delete_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t[rotated] element_put_get_preserved[where setter=shadow_root_opt_update] elim: is_document_ptr_kind_obtains is_shadow_root_ptr_kind_obtains split: if_splits option.splits)[1] end locale l_remove_shadow_root_get_child_nodes = l_get_child_nodes_defs + l_remove_shadow_root_defs + assumes remove_shadow_root_get_child_nodes_different_pointers: "ptr \ cast shadow_root_ptr \ w \ remove_shadow_root_locs element_ptr shadow_root_ptr \ h \ w \\<^sub>h h' \ r \ get_child_nodes_locs ptr \ r h h'" interpretation i_remove_shadow_root_get_child_nodes?: l_remove_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs get_shadow_root get_shadow_root_locs set_shadow_root set_shadow_root_locs get_disconnected_nodes get_disconnected_nodes_locs remove_shadow_root remove_shadow_root_locs by(auto simp add: l_remove_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_remove_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma remove_shadow_root_get_child_nodes_is_l_remove_shadow_root_get_child_nodes [instances]: "l_remove_shadow_root_get_child_nodes get_child_nodes_locs remove_shadow_root_locs" apply(auto simp add: l_remove_shadow_root_get_child_nodes_def instances )[1] using remove_shadow_root_get_child_nodes_different_pointers apply fast done paragraph \get\_tag\_name\ locale l_remove_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma remove_shadow_root_get_tag_name: assumes "w \ remove_shadow_root_locs element_ptr shadow_root_ptr" assumes "h \ w \\<^sub>h h'" assumes "r \ get_tag_name_locs ptr" shows "r h h'" using assms by(auto simp add: all_args_def remove_shadow_root_locs_def set_shadow_root_locs_def CD.get_tag_name_locs_def delete_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_put_get_preserved[where setter=shadow_root_opt_update] split: if_splits option.splits) end locale l_remove_shadow_root_get_tag_name = l_get_tag_name_defs + l_remove_shadow_root_defs + assumes remove_shadow_root_get_tag_name: "w \ remove_shadow_root_locs element_ptr shadow_root_ptr \ h \ w \\<^sub>h h' \ r \ get_tag_name_locs ptr \ r h h'" interpretation i_remove_shadow_root_get_tag_name?: l_remove_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_tag_name get_tag_name_locs get_child_nodes get_child_nodes_locs get_shadow_root get_shadow_root_locs set_shadow_root set_shadow_root_locs get_disconnected_nodes get_disconnected_nodes_locs remove_shadow_root remove_shadow_root_locs known_ptr by(auto simp add: l_remove_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_remove_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma remove_shadow_root_get_tag_name_is_l_remove_shadow_root_get_tag_name [instances]: "l_remove_shadow_root_get_tag_name get_tag_name_locs remove_shadow_root_locs" apply(auto simp add: l_remove_shadow_root_get_tag_name_def instances )[1] using remove_shadow_root_get_tag_name apply fast done subsubsection \get\_owner\_document\ locale l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_host_defs get_host get_host_locs + CD: l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_root_node get_root_node_locs get_disconnected_nodes get_disconnected_nodes_locs for get_root_node :: "(_::linorder) object_ptr \ ((_) heap, exception, (_) object_ptr) prog" and get_root_node_locs :: "((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" begin definition a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) shadow_root_ptr \ unit \ (_, (_) document_ptr) dom_prog" where "a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr = CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r (cast shadow_root_ptr)" definition a_get_owner_document_tups :: "(((_) object_ptr \ bool) \ ((_) object_ptr \ unit \ (_, (_) document_ptr) dom_prog)) list" where "a_get_owner_document_tups = [(is_shadow_root_ptr, a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast)]" definition a_get_owner_document :: "(_::linorder) object_ptr \ (_, (_) document_ptr) dom_prog" where "a_get_owner_document ptr = invoke (CD.a_get_owner_document_tups @ a_get_owner_document_tups) ptr ()" end global_interpretation l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_root_node get_root_node_locs get_disconnected_nodes get_disconnected_nodes_locs get_host get_host_locs defines get_owner_document_tups = "l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_owner_document_tups" and get_owner_document = "l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_owner_document get_root_node get_disconnected_nodes" and get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r = "l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r" and get_owner_document_tups\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = "l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_owner_document_tups get_root_node get_disconnected_nodes" and get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r = "l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r get_root_node get_disconnected_nodes" . locale l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_root_node get_root_node_locs get_disconnected_nodes get_disconnected_nodes_locs get_host get_host_locs + l_get_owner_document_defs get_owner_document + l_get_host get_host get_host_locs + CD: l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_parent get_parent_locs known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs get_root_node get_root_node_locs get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M for known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_::linorder) object_ptr \ bool" and get_parent :: "(_) node_ptr \ ((_) heap, exception, (_) object_ptr option) prog" and get_parent_locs :: "((_) heap \ (_) heap \ bool) set" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and get_root_node :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr) prog" and get_root_node_locs :: "((_) heap \ (_) heap \ bool) set" and get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap, exception, (_) document_ptr) prog" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_owner_document :: "(_) object_ptr \ ((_) heap, exception, (_) document_ptr) prog" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" + assumes get_owner_document_impl: "get_owner_document = a_get_owner_document" begin lemmas get_owner_document_def = a_get_owner_document_def[folded get_owner_document_impl] lemma get_owner_document_pure [simp]: "pure (get_owner_document ptr) h" proof - have "\shadow_root_ptr. pure (a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr ()) h" apply(auto simp add: a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_pure_I filter_M_pure_I split: option.splits)[1] by(auto simp add: CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_pure_I filter_M_pure_I split: option.splits) then show ?thesis apply(auto simp add: get_owner_document_def)[1] apply(split CD.get_owner_document_splits, rule conjI)+ apply(simp) apply(auto simp add: a_get_owner_document_tups_def)[1] apply(split invoke_splits, rule conjI)+ apply simp by(auto intro!: bind_pure_I) qed lemma get_owner_document_ptr_in_heap: assumes "h \ ok (get_owner_document ptr)" shows "ptr |\| object_ptr_kinds h" using assms by(auto simp add: get_owner_document_def invoke_ptr_in_heap dest: is_OK_returns_heap_I) end interpretation i_get_owner_document?: l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M DocumentClass.known_ptr get_parent get_parent_locs DocumentClass.type_wf get_disconnected_nodes get_disconnected_nodes_locs get_root_node get_root_node_locs CD.a_get_owner_document get_host get_host_locs get_owner_document get_child_nodes get_child_nodes_locs using get_child_nodes_is_l_get_child_nodes[unfolded ShadowRootClass.known_ptr_defs] by(auto simp add: instances l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def get_owner_document_def Core_DOM_Functions.get_owner_document_def) declare l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_owner_document_is_l_get_owner_document [instances]: "l_get_owner_document get_owner_document" apply(auto simp add: l_get_owner_document_def)[1] using get_owner_document_ptr_in_heap apply fast done subsubsection \remove\_child\ global_interpretation l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs set_child_nodes set_child_nodes_locs get_parent get_parent_locs get_owner_document get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs defines remove = "l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_remove get_child_nodes set_child_nodes get_parent get_owner_document get_disconnected_nodes set_disconnected_nodes" and remove_child = "l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_remove_child get_child_nodes set_child_nodes get_owner_document get_disconnected_nodes set_disconnected_nodes" and remove_child_locs = "l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_remove_child_locs set_child_nodes_locs set_disconnected_nodes_locs" . interpretation i_remove_child?: l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M Shadow_DOM.get_child_nodes Shadow_DOM.get_child_nodes_locs Shadow_DOM.set_child_nodes Shadow_DOM.set_child_nodes_locs Shadow_DOM.get_parent Shadow_DOM.get_parent_locs Shadow_DOM.get_owner_document get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs remove_child remove_child_locs remove ShadowRootClass.type_wf ShadowRootClass.known_ptr ShadowRootClass.known_ptrs by(auto simp add: l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def remove_child_def remove_child_locs_def remove_def instances) declare l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_disconnected\_document\ locale l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_disconnected_nodes_defs get_disconnected_nodes get_disconnected_nodes_locs for get_disconnected_nodes :: "(_::linorder) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" begin definition a_get_disconnected_document :: "(_) node_ptr \ (_, (_) document_ptr) dom_prog" where "a_get_disconnected_document node_ptr = do { check_in_heap (cast node_ptr); ptrs \ document_ptr_kinds_M; candidates \ filter_M (\document_ptr. do { disconnected_nodes \ get_disconnected_nodes document_ptr; return (node_ptr \ set disconnected_nodes) }) ptrs; (case candidates of Cons document_ptr [] \ return document_ptr | _ \ error HierarchyRequestError ) }" definition "a_get_disconnected_document_locs = (\document_ptr. get_disconnected_nodes_locs document_ptr) \ (\ptr. {preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr RObject.nothing)})" end locale l_get_disconnected_document_defs = fixes get_disconnected_document :: "(_) node_ptr \ (_, (_::linorder) document_ptr) dom_prog" fixes get_disconnected_document_locs :: "((_) heap \ (_) heap \ bool) set" locale l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_get_disconnected_document_defs + l_get_disconnected_nodes + assumes get_disconnected_document_impl: "get_disconnected_document = a_get_disconnected_document" assumes get_disconnected_document_locs_impl: "get_disconnected_document_locs = a_get_disconnected_document_locs" begin lemmas get_disconnected_document_def = get_disconnected_document_impl[unfolded a_get_disconnected_document_def] lemmas get_disconnected_document_locs_def = get_disconnected_document_locs_impl[unfolded a_get_disconnected_document_locs_def] lemma get_disconnected_document_pure [simp]: "pure (get_disconnected_document ptr) h" using get_disconnected_nodes_pure by(auto simp add: get_disconnected_document_def intro!: bind_pure_I filter_M_pure_I split: list.splits) lemma get_disconnected_document_ptr_in_heap [simp]: "h \ ok (get_disconnected_document node_ptr) \ node_ptr |\| node_ptr_kinds h" using get_disconnected_document_def is_OK_returns_result_I check_in_heap_ptr_in_heap by (metis (no_types, lifting) bind_returns_heap_E get_disconnected_document_pure node_ptr_kinds_commutes pure_pure) lemma get_disconnected_document_disconnected_document_in_heap: assumes "h \ get_disconnected_document child_node \\<^sub>r disconnected_document" shows "disconnected_document |\| document_ptr_kinds h" using assms get_disconnected_nodes_pure by(auto simp add: get_disconnected_document_def elim!: bind_returns_result_E2 dest!: filter_M_not_more_elements[where x=disconnected_document] intro!: filter_M_pure_I bind_pure_I split: if_splits list.splits) lemma get_disconnected_document_reads: "reads get_disconnected_document_locs (get_disconnected_document node_ptr) h h'" using get_disconnected_nodes_reads[unfolded reads_def] by(auto simp add: get_disconnected_document_def get_disconnected_document_locs_def intro!: reads_bind_pure reads_subset[OF check_in_heap_reads] reads_subset[OF error_reads] reads_subset[OF get_disconnected_nodes_reads] reads_subset[OF return_reads] reads_subset[OF document_ptr_kinds_M_reads] filter_M_reads filter_M_pure_I bind_pure_I split: list.splits) end locale l_get_disconnected_document = l_get_disconnected_document_defs + assumes get_disconnected_document_reads: "reads get_disconnected_document_locs (get_disconnected_document node_ptr) h h'" assumes get_disconnected_document_ptr_in_heap: "h \ ok (get_disconnected_document node_ptr) \ node_ptr |\| node_ptr_kinds h" assumes get_disconnected_document_pure [simp]: "pure (get_disconnected_document node_ptr) h" assumes get_disconnected_document_disconnected_document_in_heap: "h \ get_disconnected_document child_node \\<^sub>r disconnected_document \ disconnected_document |\| document_ptr_kinds h" global_interpretation l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_disconnected_nodes get_disconnected_nodes_locs defines get_disconnected_document = "l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_disconnected_document get_disconnected_nodes" and get_disconnected_document_locs = "l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_disconnected_document_locs get_disconnected_nodes_locs" . interpretation i_get_disconnected_document?: l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs get_disconnected_document get_disconnected_document_locs type_wf by(auto simp add: l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def get_disconnected_document_def get_disconnected_document_locs_def instances) declare l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_disconnected_document_is_l_get_disconnected_document [instances]: "l_get_disconnected_document get_disconnected_document get_disconnected_document_locs" apply(auto simp add: l_get_disconnected_document_def instances)[1] using get_disconnected_document_ptr_in_heap get_disconnected_document_pure get_disconnected_document_disconnected_document_in_heap get_disconnected_document_reads by blast+ paragraph \get\_disconnected\_nodes\ locale l_set_tag_name_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_tag_name_get_disconnected_nodes: "\w \ set_tag_name_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_disconnected_nodes_locs ptr'. r h h'))" by(auto simp add: CD.set_tag_name_locs_impl[unfolded CD.a_set_tag_name_locs_def] CD.get_disconnected_nodes_locs_impl[unfolded CD.a_get_disconnected_nodes_locs_def] all_args_def) end interpretation i_set_tag_name_get_disconnected_nodes?: l_set_tag_name_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_tag_name set_tag_name_locs get_disconnected_nodes get_disconnected_nodes_locs by unfold_locales declare l_set_tag_name_get_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_tag_name_get_disconnected_nodes_is_l_set_tag_name_get_disconnected_nodes [instances]: "l_set_tag_name_get_disconnected_nodes type_wf set_tag_name set_tag_name_locs get_disconnected_nodes get_disconnected_nodes_locs" using set_tag_name_is_l_set_tag_name get_disconnected_nodes_is_l_get_disconnected_nodes apply(simp add: l_set_tag_name_get_disconnected_nodes_def l_set_tag_name_get_disconnected_nodes_axioms_def) using set_tag_name_get_disconnected_nodes by fast subsubsection \get\_ancestors\_di\ locale l_get_ancestors_di\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_parent_defs get_parent get_parent_locs + l_get_host_defs get_host get_host_locs + l_get_disconnected_document_defs get_disconnected_document get_disconnected_document_locs for get_parent :: "(_) node_ptr \ ((_) heap, exception, (_::linorder) object_ptr option) prog" and get_parent_locs :: "((_) heap \ (_) heap \ bool) set" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_disconnected_document :: "(_) node_ptr \ ((_) heap, exception, (_) document_ptr) prog" and get_disconnected_document_locs :: "((_) heap \ (_) heap \ bool) set" begin partial_function (dom_prog) a_get_ancestors_di :: "(_::linorder) object_ptr \ (_, (_) object_ptr list) dom_prog" where "a_get_ancestors_di ptr = do { check_in_heap ptr; ancestors \ (case cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr of Some node_ptr \ do { parent_ptr_opt \ get_parent node_ptr; (case parent_ptr_opt of Some parent_ptr \ a_get_ancestors_di parent_ptr | None \ do { document_ptr \ get_disconnected_document node_ptr; a_get_ancestors_di (cast document_ptr) }) } | None \ (case cast ptr of Some shadow_root_ptr \ do { host \ get_host shadow_root_ptr; a_get_ancestors_di (cast host) } | None \ return [])); return (ptr # ancestors) }" definition "a_get_ancestors_di_locs = get_parent_locs \ get_host_locs \ get_disconnected_document_locs" end locale l_get_ancestors_di_defs = fixes get_ancestors_di :: "(_::linorder) object_ptr \ (_, (_) object_ptr list) dom_prog" fixes get_ancestors_di_locs :: "((_) heap \ (_) heap \ bool) set" locale l_get_ancestors_di\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_parent + l_get_host + l_get_disconnected_document + l_get_ancestors_di\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_get_ancestors_di_defs + assumes get_ancestors_di_impl: "get_ancestors_di = a_get_ancestors_di" assumes get_ancestors_di_locs_impl: "get_ancestors_di_locs = a_get_ancestors_di_locs" begin lemmas get_ancestors_di_def = a_get_ancestors_di.simps[folded get_ancestors_di_impl] lemmas get_ancestors_di_locs_def = a_get_ancestors_di_locs_def[folded get_ancestors_di_locs_impl] lemma get_ancestors_di_pure [simp]: "pure (get_ancestors_di ptr) h" proof - have "\ptr h h' x. h \ get_ancestors_di ptr \\<^sub>r x \ h \ get_ancestors_di ptr \\<^sub>h h' \ h = h'" proof (induct rule: a_get_ancestors_di.fixp_induct[folded get_ancestors_di_impl]) case 1 then show ?case by(rule admissible_dom_prog) next case 2 then show ?case by simp next case (3 f) then show ?case using get_parent_pure get_host_pure get_disconnected_document_pure apply(auto simp add: pure_returns_heap_eq pure_def split: option.splits elim!: bind_returns_heap_E bind_returns_result_E dest!: pure_returns_heap_eq[rotated, OF check_in_heap_pure])[1] apply (metis is_OK_returns_result_I returns_heap_eq returns_result_eq) apply (meson option.simps(3) returns_result_eq) apply (meson option.simps(3) returns_result_eq) apply(metis get_parent_pure pure_returns_heap_eq) apply(metis get_host_pure pure_returns_heap_eq) done qed then show ?thesis by (meson pure_eq_iff) qed lemma get_ancestors_di_ptr: assumes "h \ get_ancestors_di ptr \\<^sub>r ancestors" shows "ptr \ set ancestors" using assms by(simp add: get_ancestors_di_def) (auto elim!: bind_returns_result_E2 split: option.splits intro!: bind_pure_I) lemma get_ancestors_di_ptr_in_heap: assumes "h \ ok (get_ancestors_di ptr)" shows "ptr |\| object_ptr_kinds h" using assms by(auto simp add: get_ancestors_di_def check_in_heap_ptr_in_heap elim!: bind_is_OK_E dest: is_OK_returns_result_I) lemma get_ancestors_di_never_empty: assumes "h \ get_ancestors_di child \\<^sub>r ancestors" shows "ancestors \ []" using assms apply(simp add: get_ancestors_di_def) by(auto elim!: bind_returns_result_E2 split: option.splits intro!: bind_pure_I) end global_interpretation l_get_ancestors_di\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_parent get_parent_locs get_host get_host_locs get_disconnected_document get_disconnected_document_locs defines get_ancestors_di = a_get_ancestors_di and get_ancestors_di_locs = a_get_ancestors_di_locs . declare a_get_ancestors_di.simps [code] interpretation i_get_ancestors_di?: l_get_ancestors_di\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_host get_host_locs get_disconnected_document get_disconnected_document_locs get_ancestors_di get_ancestors_di_locs by(auto simp add: l_get_ancestors_di\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_ancestors_di\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def get_ancestors_di_def get_ancestors_di_locs_def instances) declare l_get_ancestors_di\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_ancestors_di_is_l_get_ancestors [instances]: "l_get_ancestors get_ancestors_di" apply(auto simp add: l_get_ancestors_def)[1] using get_ancestors_di_ptr_in_heap apply fast using get_ancestors_di_ptr apply fast done subsubsection \adopt\_node\ locale l_adopt_node\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w_\<^sub>D\<^sub>O\<^sub>M_defs = CD: l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_owner_document get_parent get_parent_locs remove_child remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs + l_get_ancestors_di_defs get_ancestors_di get_ancestors_di_locs for get_owner_document :: "(_::linorder) object_ptr \ ((_) heap, exception, (_) document_ptr) prog" and get_parent :: "(_) node_ptr \ ((_) heap, exception, (_) object_ptr option) prog" and get_parent_locs :: "((_) heap \ (_) heap \ bool) set" and remove_child :: "(_) object_ptr \ (_) node_ptr \ ((_) heap, exception, unit) prog" and remove_child_locs :: "(_) object_ptr \ (_) document_ptr \ ((_) heap, exception, unit) prog set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" and get_ancestors_di :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr list) prog" and get_ancestors_di_locs :: "((_) heap \ (_) heap \ bool) set" begin definition a_adopt_node :: "(_) document_ptr \ (_) node_ptr \ (_, unit) dom_prog" where "a_adopt_node document node = do { ancestors \ get_ancestors_di (cast document); (if cast node \ set ancestors then error HierarchyRequestError else CD.a_adopt_node document node)}" definition a_adopt_node_locs :: "(_) object_ptr option \ (_) document_ptr \ (_) document_ptr \ (_, unit) dom_prog set" where "a_adopt_node_locs = CD.a_adopt_node_locs" end locale l_adopt_node\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_adopt_node\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w_\<^sub>D\<^sub>O\<^sub>M_defs get_owner_document get_parent get_parent_locs remove_child remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs get_ancestors_di get_ancestors_di_locs + l_adopt_node_defs adopt_node adopt_node_locs + l_get_ancestors_di\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_host get_host_locs get_disconnected_document get_disconnected_document_locs get_ancestors_di get_ancestors_di_locs + CD: l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_owner_document get_parent get_parent_locs remove_child remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M adopt_node_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs set_child_nodes set_child_nodes_locs remove for get_owner_document :: "(_::linorder) object_ptr \ ((_) heap, exception, (_) document_ptr) prog" and get_parent :: "(_) node_ptr \ ((_) heap, exception, (_) object_ptr option) prog" and get_parent_locs :: "((_) heap \ (_) heap \ bool) set" and remove_child :: "(_) object_ptr \ (_) node_ptr \ ((_) heap, exception, unit) prog" and remove_child_locs :: "(_) object_ptr \ (_) document_ptr \ ((_) heap, exception, unit) prog set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" and get_ancestors_di :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr list) prog" and get_ancestors_di_locs :: "((_) heap \ (_) heap \ bool) set" and adopt_node :: "(_) document_ptr \ (_) node_ptr \ ((_) heap, exception, unit) prog" and adopt_node_locs :: "(_) object_ptr option \ (_) document_ptr \ (_) document_ptr \ ((_) heap, exception, unit) prog set" and adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) document_ptr \ (_) node_ptr \ ((_) heap, exception, unit) prog" and adopt_node_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr option \ (_) document_ptr \ (_) document_ptr \ ((_) heap, exception, unit) prog set" and known_ptr :: "(_) object_ptr \ bool" and type_wf :: "(_) heap \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and known_ptrs :: "(_) heap \ bool" and set_child_nodes :: "(_) object_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_child_nodes_locs :: "(_) object_ptr \ ((_) heap, exception, unit) prog set" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_disconnected_document :: "(_) node_ptr \ ((_) heap, exception, (_) document_ptr) prog" and get_disconnected_document_locs :: "((_) heap \ (_) heap \ bool) set" and remove :: "(_) node_ptr \ ((_) heap, exception, unit) prog" + assumes adopt_node_impl: "adopt_node = a_adopt_node" assumes adopt_node_locs_impl: "adopt_node_locs = a_adopt_node_locs" begin lemmas adopt_node_def = a_adopt_node_def[folded adopt_node_impl CD.adopt_node_impl] lemmas adopt_node_locs_def = a_adopt_node_locs_def[folded adopt_node_locs_impl CD.adopt_node_locs_impl] lemma adopt_node_writes: "writes (adopt_node_locs |h \ get_parent node|\<^sub>r |h \ get_owner_document (cast node)|\<^sub>r document_ptr) (adopt_node document_ptr node) h h'" by(auto simp add: CD.adopt_node_writes adopt_node_def CD.adopt_node_impl[symmetric] adopt_node_locs_def CD.adopt_node_locs_impl[symmetric] intro!: writes_bind_pure[OF get_ancestors_di_pure]) lemma adopt_node_pointers_preserved: "w \ adopt_node_locs parent owner_document document_ptr \ h \ w \\<^sub>h h' \ object_ptr_kinds h = object_ptr_kinds h'" using CD.adopt_node_locs_impl CD.adopt_node_pointers_preserved local.adopt_node_locs_def by blast lemma adopt_node_types_preserved: "w \ adopt_node_locs parent owner_document document_ptr \ h \ w \\<^sub>h h' \ type_wf h = type_wf h'" using CD.adopt_node_locs_impl CD.adopt_node_types_preserved local.adopt_node_locs_def by blast lemma adopt_node_child_in_heap: "h \ ok (adopt_node document_ptr child) \ child |\| node_ptr_kinds h" by (smt CD.adopt_node_child_in_heap CD.adopt_node_impl bind_is_OK_E error_returns_heap is_OK_returns_heap_E l_adopt_node\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M.adopt_node_def l_adopt_node\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms local.get_ancestors_di_pure pure_returns_heap_eq) lemma adopt_node_children_subset: "h \ adopt_node owner_document node \\<^sub>h h' \ h \ get_child_nodes ptr \\<^sub>r children \ h' \ get_child_nodes ptr \\<^sub>r children' \ known_ptrs h \ type_wf h \ set children' \ set children" by (smt CD.adopt_node_children_subset CD.adopt_node_impl bind_returns_heap_E error_returns_heap local.adopt_node_def local.get_ancestors_di_pure pure_returns_heap_eq) end global_interpretation l_adopt_node\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w_\<^sub>D\<^sub>O\<^sub>M_defs get_owner_document get_parent get_parent_locs remove_child remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs get_ancestors_di get_ancestors_di_locs defines adopt_node = "a_adopt_node" and adopt_node_locs = "a_adopt_node_locs" and adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = "CD.a_adopt_node" and adopt_node_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = "CD.a_adopt_node_locs" . interpretation i_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M: l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_owner_document get_parent get_parent_locs remove_child remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M adopt_node_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs set_child_nodes set_child_nodes_locs remove by(auto simp add: l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def adopt_node_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare i_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] interpretation i_adopt_node?: l_adopt_node\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_owner_document get_parent get_parent_locs remove_child remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs get_ancestors_di get_ancestors_di_locs adopt_node adopt_node_locs CD.a_adopt_node CD.a_adopt_node_locs known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs set_child_nodes set_child_nodes_locs get_host get_host_locs get_disconnected_document get_disconnected_document_locs remove by(auto simp add: l_adopt_node\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_adopt_node\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def adopt_node_def adopt_node_locs_def instances) declare l_adopt_node\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma adopt_node_is_l_adopt_node [instances]: "l_adopt_node type_wf known_ptr known_ptrs get_parent adopt_node adopt_node_locs get_child_nodes get_owner_document" apply(auto simp add: l_adopt_node_def l_adopt_node_axioms_def instances)[1] using adopt_node_writes apply fast using adopt_node_pointers_preserved apply (fast, fast) using adopt_node_types_preserved apply (fast, fast) using adopt_node_child_in_heap apply fast using adopt_node_children_subset apply fast done paragraph \get\_shadow\_root\ locale l_adopt_node_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_child_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_disconnected_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_adopt_node\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma adopt_node_get_shadow_root: "\w \ adopt_node_locs parent owner_document document_ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" by(auto simp add: adopt_node_locs_def CD.adopt_node_locs_def CD.remove_child_locs_def all_args_def set_disconnected_nodes_get_shadow_root set_child_nodes_get_shadow_root) end locale l_adopt_node_get_shadow_root = l_adopt_node_defs + l_get_shadow_root_defs + assumes adopt_node_get_shadow_root: "\w \ adopt_node_locs parent owner_document document_ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" interpretation i_adopt_node_get_shadow_root?: l_adopt_node_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr set_child_nodes set_child_nodes_locs Core_DOM_Functions.set_child_nodes Core_DOM_Functions.set_child_nodes_locs get_shadow_root get_shadow_root_locs set_disconnected_nodes set_disconnected_nodes_locs get_owner_document get_parent get_parent_locs remove_child remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs get_ancestors_di get_ancestors_di_locs adopt_node adopt_node_locs adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M adopt_node_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs known_ptrs get_host get_host_locs get_disconnected_document get_disconnected_document_locs remove by(auto simp add: l_adopt_node_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_adopt_node_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] interpretation i_adopt_node_get_shadow_root?: l_adopt_node_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr set_child_nodes set_child_nodes_locs Core_DOM_Functions.set_child_nodes Core_DOM_Functions.set_child_nodes_locs get_shadow_root get_shadow_root_locs set_disconnected_nodes set_disconnected_nodes_locs get_owner_document get_parent get_parent_locs remove_child remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs get_ancestors_di get_ancestors_di_locs adopt_node adopt_node_locs adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M adopt_node_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs known_ptrs get_host get_host_locs get_disconnected_document get_disconnected_document_locs remove by(auto simp add: l_adopt_node_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_adopt_node_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma adopt_node_get_shadow_root_is_l_adopt_node_get_shadow_root [instances]: "l_adopt_node_get_shadow_root adopt_node_locs get_shadow_root_locs" apply(auto simp add: l_adopt_node_get_shadow_root_def)[1] using adopt_node_get_shadow_root apply fast done subsubsection \insert\_before\ global_interpretation l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_parent get_parent_locs get_child_nodes get_child_nodes_locs set_child_nodes set_child_nodes_locs get_ancestors_di get_ancestors_di_locs adopt_node adopt_node_locs set_disconnected_nodes set_disconnected_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_owner_document defines next_sibling = a_next_sibling and insert_node = a_insert_node and ensure_pre_insertion_validity = a_ensure_pre_insertion_validity and insert_before = a_insert_before and insert_before_locs = a_insert_before_locs . global_interpretation l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs insert_before defines append_child = "l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_append_child insert_before" . interpretation i_insert_before?: l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_parent get_parent_locs get_child_nodes get_child_nodes_locs set_child_nodes set_child_nodes_locs get_ancestors_di get_ancestors_di_locs adopt_node adopt_node_locs set_disconnected_nodes set_disconnected_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_owner_document insert_before insert_before_locs append_child type_wf known_ptr known_ptrs by(auto simp add: l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def insert_before_def insert_before_locs_def instances) declare l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] interpretation i_append_child?: l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M append_child insert_before insert_before_locs by(simp add: l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances append_child_def) declare l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] subsubsection \get\_assigned\_nodes\ fun map_filter_M2 :: "('x \ ('heap, 'e, 'y option) prog) \ 'x list \ ('heap, 'e, 'y list) prog" where "map_filter_M2 f [] = return []" | "map_filter_M2 f (x # xs) = do { res \ f x; remainder \ map_filter_M2 f xs; return ((case res of Some r \ [r] | None \ []) @ remainder) }" lemma map_filter_M2_pure [simp]: assumes "\x. x \ set xs \ pure (f x) h" shows "pure (map_filter_M2 f xs) h" using assms apply(induct xs arbitrary: h) by(auto elim!: bind_returns_result_E2 intro!: bind_pure_I) lemma map_filter_pure_no_monad: assumes "\x. x \ set xs \ pure (f x) h" assumes "h \ map_filter_M2 f xs \\<^sub>r ys" shows "ys = map the (filter (\x. x \ None) (map (\x. |h \ f x|\<^sub>r) xs))" and "\x. x \ set xs \ h \ ok (f x)" using assms apply(induct xs arbitrary: h ys) by(auto elim!: bind_returns_result_E2) lemma map_filter_pure_foo: assumes "\x. x \ set xs \ pure (f x) h" assumes "h \ map_filter_M2 f xs \\<^sub>r ys" assumes "y \ set ys" obtains x where "h \ f x \\<^sub>r Some y" and "x \ set xs" using assms apply(induct xs arbitrary: ys) by(auto elim!: bind_returns_result_E2) lemma map_filter_M2_in_result: assumes "h \ map_filter_M2 P xs \\<^sub>r ys" assumes "a \ set xs" assumes "\x. x \ set xs \ pure (P x) h" assumes "h \ P a \\<^sub>r Some b" shows "b \ set ys" using assms apply(induct xs arbitrary: h ys) by(auto elim!: bind_returns_result_E2 ) locale l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_tag_name_defs get_tag_name get_tag_name_locs + l_get_root_node_defs get_root_node get_root_node_locs + l_get_host_defs get_host get_host_locs + l_get_child_nodes_defs get_child_nodes get_child_nodes_locs + l_find_slot_defs find_slot assigned_slot + l_remove_defs remove + l_insert_before_defs insert_before insert_before_locs + l_append_child_defs append_child + l_remove_shadow_root_defs remove_shadow_root remove_shadow_root_locs for get_child_nodes :: "(_::linorder) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_root_node :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr) prog" and get_root_node_locs :: "((_) heap \ (_) heap \ bool) set" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and find_slot :: "bool \ (_) node_ptr \ ((_) heap, exception, (_) element_ptr option) prog" and assigned_slot :: "(_) node_ptr \ ((_) heap, exception, (_) element_ptr option) prog" and remove :: "(_) node_ptr \ ((_) heap, exception, unit) prog" and insert_before :: "(_) object_ptr \ (_) node_ptr \ (_) node_ptr option \ ((_) heap, exception, unit) prog" and insert_before_locs :: "(_) object_ptr \ (_) object_ptr option \ (_) document_ptr \ (_) document_ptr \ (_, unit) dom_prog set" and append_child :: "(_) object_ptr \ (_) node_ptr \ ((_) heap, exception, unit) prog" and remove_shadow_root :: "(_) element_ptr \ ((_) heap, exception, unit) prog" and remove_shadow_root_locs :: "(_) element_ptr \ (_) shadow_root_ptr \ ((_) heap, exception, unit) prog set" begin definition a_assigned_nodes :: "(_) element_ptr \ (_, (_) node_ptr list) dom_prog" where "a_assigned_nodes slot = do { tag \ get_tag_name slot; (if tag \ ''slot'' then error HierarchyRequestError else return ()); root \ get_root_node (cast slot); if is_shadow_root_ptr_kind root then do { host \ get_host (the (cast root)); children \ get_child_nodes (cast host); filter_M (\slotable. do { found_slot \ find_slot False slotable; return (found_slot = Some slot)}) children} else return []}" partial_function (dom_prog) a_assigned_nodes_flatten :: "(_) element_ptr \ (_, (_) node_ptr list) dom_prog" where "a_assigned_nodes_flatten slot = do { tag \ get_tag_name slot; (if tag \ ''slot'' then error HierarchyRequestError else return ()); root \ get_root_node (cast slot); (if is_shadow_root_ptr_kind root then do { slotables \ a_assigned_nodes slot; slotables_or_child_nodes \ (if slotables = [] then do { get_child_nodes (cast slot) } else do { return slotables }); list_of_lists \ map_M (\node_ptr. do { (case cast node_ptr of Some element_ptr \ do { tag \ get_tag_name element_ptr; (if tag = ''slot'' then do { root \ get_root_node (cast element_ptr); (if is_shadow_root_ptr_kind root then do { a_assigned_nodes_flatten element_ptr } else do { return [node_ptr] }) } else do { return [node_ptr] }) } | None \ return [node_ptr]) }) slotables_or_child_nodes; return (concat list_of_lists) } else return []) }" definition a_flatten_dom :: "(_, unit) dom_prog" where "a_flatten_dom = do { tups \ element_ptr_kinds_M \ map_filter_M2 (\element_ptr. do { tag \ get_tag_name element_ptr; assigned_nodes \ a_assigned_nodes element_ptr; (if tag = ''slot'' \ assigned_nodes \ [] then return (Some (element_ptr, assigned_nodes)) else return None)}); forall_M (\(slot, assigned_nodes). do { get_child_nodes (cast slot) \ forall_M remove; forall_M (append_child (cast slot)) assigned_nodes }) tups; shadow_root_ptr_kinds_M \ forall_M (\shadow_root_ptr. do { host \ get_host shadow_root_ptr; get_child_nodes (cast host) \ forall_M remove; get_child_nodes (cast shadow_root_ptr) \ forall_M (append_child (cast host)); remove_shadow_root host }); return () }" end global_interpretation l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_tag_name get_tag_name_locs get_root_node get_root_node_locs get_host get_host_locs find_slot assigned_slot remove insert_before insert_before_locs append_child remove_shadow_root remove_shadow_root_locs defines assigned_nodes = "l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_assigned_nodes get_child_nodes get_tag_name get_root_node get_host find_slot" and assigned_nodes_flatten = "l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_assigned_nodes_flatten get_child_nodes get_tag_name get_root_node get_host find_slot" and flatten_dom = "l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_flatten_dom get_child_nodes get_tag_name get_root_node get_host find_slot remove append_child remove_shadow_root" . declare a_assigned_nodes_flatten.simps [code] locale l_assigned_nodes_defs = fixes assigned_nodes :: "(_) element_ptr \ (_, (_) node_ptr list) dom_prog" fixes assigned_nodes_flatten :: "(_) element_ptr \ (_, (_) node_ptr list) dom_prog" fixes flatten_dom :: "(_, unit) dom_prog" locale l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_assigned_nodes_defs assigned_nodes assigned_nodes_flatten flatten_dom + l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_tag_name get_tag_name_locs get_root_node get_root_node_locs get_host get_host_locs find_slot assigned_slot remove insert_before insert_before_locs append_child remove_shadow_root remove_shadow_root_locs (* + l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M *) + l_get_shadow_root type_wf get_shadow_root get_shadow_root_locs + l_set_shadow_root type_wf set_shadow_root set_shadow_root_locs + l_remove + l_insert_before insert_before insert_before_locs + l_find_slot find_slot assigned_slot + l_get_tag_name type_wf get_tag_name get_tag_name_locs + l_get_root_node get_root_node get_root_node_locs get_parent get_parent_locs + l_get_host get_host get_host_locs + l_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs + l_to_tree_order to_tree_order for known_ptr :: "(_::linorder) object_ptr \ bool" and assigned_nodes :: "(_) element_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and assigned_nodes_flatten :: "(_) element_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and flatten_dom :: "((_) heap, exception, unit) prog" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_root_node :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr) prog" and get_root_node_locs :: "((_) heap \ (_) heap \ bool) set" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and find_slot :: "bool \ (_) node_ptr \ ((_) heap, exception, (_) element_ptr option) prog" and assigned_slot :: "(_) node_ptr \ ((_) heap, exception, (_) element_ptr option) prog" and remove :: "(_) node_ptr \ ((_) heap, exception, unit) prog" and insert_before :: "(_) object_ptr \ (_) node_ptr \ (_) node_ptr option \ ((_) heap, exception, unit) prog" and insert_before_locs :: "(_) object_ptr \ (_) object_ptr option \ (_) document_ptr \ (_) document_ptr \ (_, unit) dom_prog set" and append_child :: "(_) object_ptr \ (_) node_ptr \ ((_) heap, exception, unit) prog" and remove_shadow_root :: "(_) element_ptr \ ((_) heap, exception, unit) prog" and remove_shadow_root_locs :: "(_) element_ptr \ (_) shadow_root_ptr \ ((_) heap, exception, unit) prog set" and type_wf :: "(_) heap \ bool" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and set_shadow_root :: "(_) element_ptr \ (_) shadow_root_ptr option \ ((_) heap, exception, unit) prog" and set_shadow_root_locs :: "(_) element_ptr \ ((_) heap, exception, unit) prog set" and get_parent :: "(_) node_ptr \ ((_) heap, exception, (_) object_ptr option) prog" and get_parent_locs :: "((_) heap \ (_) heap \ bool) set" and to_tree_order :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr list) prog" + assumes assigned_nodes_impl: "assigned_nodes = a_assigned_nodes" assumes flatten_dom_impl: "flatten_dom = a_flatten_dom" begin lemmas assigned_nodes_def = assigned_nodes_impl[unfolded a_assigned_nodes_def] lemmas flatten_dom_def = flatten_dom_impl[unfolded a_flatten_dom_def, folded assigned_nodes_impl] lemma assigned_nodes_pure [simp]: "pure (assigned_nodes slot) h" by(auto simp add: assigned_nodes_def intro!: bind_pure_I filter_M_pure_I) lemma assigned_nodes_ptr_in_heap: assumes "h \ ok (assigned_nodes slot)" shows "slot |\| element_ptr_kinds h" using assms apply(auto simp add: assigned_nodes_def)[1] by (meson bind_is_OK_E is_OK_returns_result_I local.get_tag_name_ptr_in_heap) lemma assigned_nodes_slot_is_slot: assumes "h \ ok (assigned_nodes slot)" shows "h \ get_tag_name slot \\<^sub>r ''slot''" using assms by(auto simp add: assigned_nodes_def elim!: bind_is_OK_E split: if_splits) lemma assigned_nodes_different_ptr: assumes "h \ assigned_nodes slot \\<^sub>r nodes" assumes "h \ assigned_nodes slot' \\<^sub>r nodes'" assumes "slot \ slot'" shows "set nodes \ set nodes' = {}" proof (rule ccontr) assume "set nodes \ set nodes' \ {} " then obtain common_ptr where "common_ptr \ set nodes" and "common_ptr \ set nodes'" by auto have "h \ find_slot False common_ptr \\<^sub>r Some slot" using \common_ptr \ set nodes\ using assms(1) by(auto simp add: assigned_nodes_def elim!: bind_returns_result_E2 split: if_splits dest!: filter_M_holds_for_result[where x=common_ptr] intro!: bind_pure_I) moreover have "h \ find_slot False common_ptr \\<^sub>r Some slot'" using \common_ptr \ set nodes'\ using assms(2) by(auto simp add: assigned_nodes_def elim!: bind_returns_result_E2 split: if_splits dest!: filter_M_holds_for_result[where x=common_ptr] intro!: bind_pure_I) ultimately show False using assms(3) by (meson option.inject returns_result_eq) qed end interpretation i_assigned_nodes?: l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr assigned_nodes assigned_nodes_flatten flatten_dom get_child_nodes get_child_nodes_locs get_tag_name get_tag_name_locs get_root_node get_root_node_locs get_host get_host_locs find_slot assigned_slot remove insert_before insert_before_locs append_child remove_shadow_root remove_shadow_root_locs type_wf get_shadow_root get_shadow_root_locs set_shadow_root set_shadow_root_locs get_parent get_parent_locs to_tree_order by(auto simp add: instances l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def assigned_nodes_def flatten_dom_def) declare l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_assigned_nodes = l_assigned_nodes_defs + assumes assigned_nodes_pure [simp]: "pure (assigned_nodes slot) h" assumes assigned_nodes_ptr_in_heap: "h \ ok (assigned_nodes slot) \ slot |\| element_ptr_kinds h" assumes assigned_nodes_slot_is_slot: "h \ ok (assigned_nodes slot) \ h \ get_tag_name slot \\<^sub>r ''slot''" assumes assigned_nodes_different_ptr: "h \ assigned_nodes slot \\<^sub>r nodes \ h \ assigned_nodes slot' \\<^sub>r nodes' \ slot \ slot' \ set nodes \ set nodes' = {}" lemma assigned_nodes_is_l_assigned_nodes [instances]: "l_assigned_nodes assigned_nodes" apply(auto simp add: l_assigned_nodes_def)[1] using assigned_nodes_ptr_in_heap apply fast using assigned_nodes_slot_is_slot apply fast using assigned_nodes_different_ptr apply fast done subsubsection \set\_val\ locale l_set_val\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD: l_set_val\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_val set_val_locs + l_type_wf type_wf for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and set_val :: "(_) character_data_ptr \ char list \ (_, unit) dom_prog" and set_val_locs :: "(_) character_data_ptr \ (_, unit) dom_prog set" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" begin lemma set_val_ok: "type_wf h \ character_data_ptr |\| character_data_ptr_kinds h \ h \ ok (set_val character_data_ptr tag)" using CD.set_val_ok CD.type_wf_impl ShadowRootClass.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t local.type_wf_impl by blast lemma set_val_writes: "writes (set_val_locs character_data_ptr) (set_val character_data_ptr tag) h h'" using CD.set_val_writes . lemma set_val_pointers_preserved: assumes "w \ set_val_locs character_data_ptr" assumes "h \ w \\<^sub>h h'" shows "object_ptr_kinds h = object_ptr_kinds h'" using assms CD.set_val_pointers_preserved by simp lemma set_val_typess_preserved: assumes "w \ set_val_locs character_data_ptr" assumes "h \ w \\<^sub>h h'" shows "type_wf h = type_wf h'" apply(unfold type_wf_impl) using assms(1) type_wf_preserved[OF writes_singleton2 assms(2)] by(auto simp add: all_args_def CD.set_val_locs_impl[unfolded CD.a_set_val_locs_def] split: if_splits) end interpretation i_set_val?: l_set_val\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_val set_val_locs apply(unfold_locales) by (auto simp add: set_val_def set_val_locs_def) declare l_set_val\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_val_is_l_set_val [instances]: "l_set_val type_wf set_val set_val_locs" apply(simp add: l_set_val_def) using set_val_ok set_val_writes set_val_pointers_preserved set_val_typess_preserved by blast paragraph \get\_shadow\_root\ locale l_set_val_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_val\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_val_get_shadow_root: "\w \ set_val_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" by(auto simp add: CD.set_val_locs_impl[unfolded CD.a_set_val_locs_def] get_shadow_root_locs_def all_args_def) end locale l_set_val_get_shadow_root = l_set_val + l_get_shadow_root + assumes set_val_get_shadow_root: "\w \ set_val_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" interpretation i_set_val_get_shadow_root?: l_set_val_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_val set_val_locs get_shadow_root get_shadow_root_locs apply(auto simp add: l_set_val_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)[1] using l_set_val\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms by unfold_locales declare l_set_val_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_val_get_shadow_root_is_l_set_val_get_shadow_root [instances]: "l_set_val_get_shadow_root type_wf set_val set_val_locs get_shadow_root get_shadow_root_locs" using set_val_is_l_set_val get_shadow_root_is_l_get_shadow_root apply(simp add: l_set_val_get_shadow_root_def l_set_val_get_shadow_root_axioms_def) using set_val_get_shadow_root by fast paragraph \get\_tag\_type\ locale l_set_val_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_val\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_val_get_tag_name: "\w \ set_val_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_tag_name_locs ptr'. r h h'))" by(auto simp add: CD.set_val_locs_impl[unfolded CD.a_set_val_locs_def] CD.get_tag_name_locs_impl[unfolded CD.a_get_tag_name_locs_def] all_args_def) end locale l_set_val_get_tag_name = l_set_val + l_get_tag_name + assumes set_val_get_tag_name: "\w \ set_val_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_tag_name_locs ptr'. r h h'))" interpretation i_set_val_get_tag_name?: l_set_val_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_val set_val_locs get_tag_name get_tag_name_locs by unfold_locales declare l_set_val_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_val_get_tag_name_is_l_set_val_get_tag_name [instances]: "l_set_val_get_tag_name type_wf set_val set_val_locs get_tag_name get_tag_name_locs" using set_val_is_l_set_val get_tag_name_is_l_get_tag_name apply(simp add: l_set_val_get_tag_name_def l_set_val_get_tag_name_axioms_def) using set_val_get_tag_name by fast subsubsection \create\_character\_data\ locale l_create_character_data\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD: l_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M _ _ _ _ _ _ type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M _ known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_known_ptr known_ptr for known_ptr :: "(_) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ bool" + assumes known_ptr_impl: "known_ptr = a_known_ptr" begin lemma create_character_data_document_in_heap: assumes "h \ ok (create_character_data document_ptr text)" shows "document_ptr |\| document_ptr_kinds h" using assms CD.create_character_data_document_in_heap by simp lemma create_character_data_known_ptr: assumes "h \ create_character_data document_ptr text \\<^sub>r new_character_data_ptr" shows "known_ptr (cast new_character_data_ptr)" using assms CD.create_character_data_known_ptr by(simp add: known_ptr_impl CD.known_ptr_impl ShadowRootClass.a_known_ptr_def) end locale l_create_character_data = l_create_character_data_defs interpretation i_create_character_data?: l_create_character_data\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs set_val set_val_locs create_character_data known_ptr DocumentClass.type_wf DocumentClass.known_ptr by(auto simp add: l_create_character_data\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_create_character_data\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \create\_element\ locale l_create_element\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD: l_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs set_tag_name set_tag_name_locs type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M create_element known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_known_ptr known_ptr for get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" and set_tag_name :: "(_) element_ptr \ char list \ ((_) heap, exception, unit) prog" and set_tag_name_locs :: "(_) element_ptr \ ((_) heap, exception, unit) prog set" and type_wf :: "(_) heap \ bool" and create_element :: "(_) document_ptr \ char list \ ((_) heap, exception, (_) element_ptr) prog" and known_ptr :: "(_) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ bool" + assumes known_ptr_impl: "known_ptr = a_known_ptr" begin lemmas create_element_def = CD.create_element_def lemma create_element_document_in_heap: assumes "h \ ok (create_element document_ptr tag)" shows "document_ptr |\| document_ptr_kinds h" using CD.create_element_document_in_heap assms . lemma create_element_known_ptr: assumes "h \ create_element document_ptr tag \\<^sub>r new_element_ptr" shows "known_ptr (cast new_element_ptr)" proof - have "is_element_ptr new_element_ptr" using assms apply(auto simp add: create_element_def elim!: bind_returns_result_E)[1] using new_element_is_element_ptr by blast then show ?thesis by(auto simp add: known_ptr_impl known_ptr_defs DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs) qed end interpretation i_create_element?: l_create_element\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs set_tag_name set_tag_name_locs type_wf create_element known_ptr DocumentClass.type_wf DocumentClass.known_ptr by(auto simp add: l_create_element\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_create_element\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_create_element\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] subsection \A wellformed heap (Core DOM)\ subsubsection \wellformed\_heap\ locale l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = CD: l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs + l_get_shadow_root_defs get_shadow_root get_shadow_root_locs + l_get_tag_name_defs get_tag_name get_tag_name_locs for get_child_nodes :: "(_::linorder) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin definition a_host_shadow_root_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" where "a_host_shadow_root_rel h = (\(x, y). (cast x, cast y)) ` {(host, shadow_root). host |\| element_ptr_kinds h \ |h \ get_shadow_root host|\<^sub>r = Some shadow_root}" lemma a_host_shadow_root_rel_code [code]: "a_host_shadow_root_rel h = set (concat (map (\host. (case |h \ get_shadow_root host|\<^sub>r of Some shadow_root \ [(cast host, cast shadow_root)] | None \ [])) (sorted_list_of_fset (element_ptr_kinds h))) )" by(auto simp add: a_host_shadow_root_rel_def) definition a_ptr_disconnected_node_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" where "a_ptr_disconnected_node_rel h = (\(x, y). (cast x, cast y)) ` {(document_ptr, disconnected_node). document_ptr |\| document_ptr_kinds h \ disconnected_node \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r}" lemma a_ptr_disconnected_node_rel_code [code]: "a_ptr_disconnected_node_rel h = set (concat (map (\ptr. map (\node. (cast ptr, cast node)) |h \ get_disconnected_nodes ptr|\<^sub>r) (sorted_list_of_fset (document_ptr_kinds h))) )" by(auto simp add: a_ptr_disconnected_node_rel_def) definition a_all_ptrs_in_heap :: "(_) heap \ bool" where "a_all_ptrs_in_heap h = ((\host shadow_root_ptr. (h \ get_shadow_root host \\<^sub>r Some shadow_root_ptr) \ shadow_root_ptr |\| shadow_root_ptr_kinds h))" definition a_distinct_lists :: "(_) heap \ bool" where "a_distinct_lists h = distinct (concat ( map (\element_ptr. (case |h \ get_shadow_root element_ptr|\<^sub>r of Some shadow_root_ptr \ [shadow_root_ptr] | None \ [])) |h \ element_ptr_kinds_M|\<^sub>r ))" definition a_shadow_root_valid :: "(_) heap \ bool" where "a_shadow_root_valid h = (\shadow_root_ptr \ fset (shadow_root_ptr_kinds h). (\host \ fset(element_ptr_kinds h). |h \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types \ |h \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr))" definition a_heap_is_wellformed :: "(_) heap \ bool" where "a_heap_is_wellformed h \ CD.a_heap_is_wellformed h \ acyclic (CD.a_parent_child_rel h \ a_host_shadow_root_rel h \ a_ptr_disconnected_node_rel h) \ a_all_ptrs_in_heap h \ a_distinct_lists h \ a_shadow_root_valid h" end global_interpretation l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs defines heap_is_wellformed = a_heap_is_wellformed and parent_child_rel = CD.a_parent_child_rel and host_shadow_root_rel = a_host_shadow_root_rel and ptr_disconnected_node_rel = a_ptr_disconnected_node_rel and all_ptrs_in_heap = a_all_ptrs_in_heap and distinct_lists = a_distinct_lists and shadow_root_valid = a_shadow_root_valid and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD.a_heap_is_wellformed and parent_child_rel\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD.a_parent_child_rel and acyclic_heap\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD.a_acyclic_heap and all_ptrs_in_heap\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD.a_all_ptrs_in_heap and distinct_lists\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD.a_distinct_lists and owner_document_valid\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD.a_owner_document_valid . interpretation i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M: l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel by (auto simp add: l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def parent_child_rel_def instances) declare i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_is_l_heap_is_wellformed [instances]: "l_heap_is_wellformed type_wf known_ptr heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel get_child_nodes get_disconnected_nodes" apply(auto simp add: l_heap_is_wellformed_def)[1] using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_children_in_heap apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_disc_nodes_in_heap apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_one_parent apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_one_disc_parent apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_children_disc_nodes_different apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_disconnected_nodes_distinct apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_children_distinct apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_children_disc_nodes apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_child apply (blast, blast) using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_finite apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_acyclic apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_node_ptr apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_parent_in_heap apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_child_in_heap apply blast done locale l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs + CD: l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel + l_heap_is_wellformed_defs heap_is_wellformed parent_child_rel + l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_shadow_root get_shadow_root_locs get_host get_host_locs type_wf + l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs get_disconnected_document get_disconnected_document_locs type_wf + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs for get_child_nodes :: "(_::linorder) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and known_ptr :: "(_) object_ptr \ bool" and type_wf :: "(_) heap \ bool" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_disconnected_document :: "(_) node_ptr \ ((_) heap, exception, (_) document_ptr) prog" and get_disconnected_document_locs :: "((_) heap \ (_) heap \ bool) set" + assumes heap_is_wellformed_impl: "heap_is_wellformed = a_heap_is_wellformed" begin lemmas heap_is_wellformed_def = heap_is_wellformed_impl[unfolded a_heap_is_wellformed_def, folded CD.heap_is_wellformed_impl CD.parent_child_rel_impl] lemma a_distinct_lists_code [code]: "a_all_ptrs_in_heap h = ((\host \ fset (element_ptr_kinds h). h \ ok (get_shadow_root host) \ (case |h \ get_shadow_root host|\<^sub>r of Some shadow_root_ptr \ shadow_root_ptr |\| shadow_root_ptr_kinds h | None \ True)))" apply(auto simp add: a_all_ptrs_in_heap_def split: option.splits)[1] by (meson is_OK_returns_result_I local.get_shadow_root_ptr_in_heap notin_fset select_result_I2) lemma get_shadow_root_shadow_root_ptr_in_heap: assumes "heap_is_wellformed h" assumes "h \ get_shadow_root host \\<^sub>r Some shadow_root_ptr" shows "shadow_root_ptr |\| shadow_root_ptr_kinds h" using assms by(auto simp add: heap_is_wellformed_def a_all_ptrs_in_heap_def) lemma get_host_ptr_in_heap: assumes "heap_is_wellformed h" assumes "h \ get_host shadow_root_ptr \\<^sub>r host" shows "shadow_root_ptr |\| shadow_root_ptr_kinds h" using assms get_shadow_root_shadow_root_ptr_in_heap by(auto simp add: get_host_def elim!: bind_returns_result_E2 dest!: filter_M_holds_for_result intro!: bind_pure_I split: list.splits) lemma shadow_root_same_host: assumes "heap_is_wellformed h" and "type_wf h" assumes "h \ get_shadow_root host \\<^sub>r Some shadow_root_ptr" assumes "h \ get_shadow_root host' \\<^sub>r Some shadow_root_ptr" shows "host = host'" proof (rule ccontr) assume " host \ host'" have "host |\| element_ptr_kinds h" using assms(3) by (meson is_OK_returns_result_I local.get_shadow_root_ptr_in_heap) moreover have "host' |\| element_ptr_kinds h" using assms(4) by (meson is_OK_returns_result_I local.get_shadow_root_ptr_in_heap) ultimately show False using assms apply(auto simp add: heap_is_wellformed_def a_distinct_lists_def)[1] apply(drule distinct_concat_map_E(1)[where x=host and y=host']) apply(simp) apply(simp) using \host \ host'\ apply(simp) apply(auto)[1] done qed lemma shadow_root_host_dual: assumes "h \ get_host shadow_root_ptr \\<^sub>r host" shows "h \ get_shadow_root host \\<^sub>r Some shadow_root_ptr" using assms by(auto simp add: get_host_def dest: filter_M_holds_for_result elim!: bind_returns_result_E2 intro!: bind_pure_I split: list.splits) lemma disc_doc_disc_node_dual: assumes "h \ get_disconnected_document disc_node \\<^sub>r disc_doc" obtains disc_nodes where "h \ get_disconnected_nodes disc_doc \\<^sub>r disc_nodes" and "disc_node \ set disc_nodes" using assms get_disconnected_nodes_pure by(auto simp add: get_disconnected_document_def bind_pure_I dest!: filter_M_holds_for_result elim!: bind_returns_result_E2 intro!: filter_M_pure_I split: if_splits list.splits) lemma get_host_valid_tag_name: assumes "heap_is_wellformed h" and "type_wf h" assumes "h \ get_host shadow_root_ptr \\<^sub>r host" assumes "h \ get_tag_name host \\<^sub>r tag" shows "tag \ safe_shadow_root_element_types" proof - obtain host' where "host' |\| element_ptr_kinds h" and "|h \ get_tag_name host'|\<^sub>r \ safe_shadow_root_element_types" and "h \ get_shadow_root host' \\<^sub>r Some shadow_root_ptr" using assms apply(auto simp add: heap_is_wellformed_def a_shadow_root_valid_def)[1] by (smt assms(1) finite_set_in get_host_ptr_in_heap local.get_shadow_root_ok returns_result_select_result) then have "host = host'" by (meson assms(1) assms(2) assms(3) shadow_root_host_dual shadow_root_same_host) then show ?thesis by (smt \\thesis. (\host'. \host' |\| element_ptr_kinds h; |h \ get_tag_name host'|\<^sub>r \ safe_shadow_root_element_types; h \ get_shadow_root host' \\<^sub>r Some shadow_root_ptr\ \ thesis) \ thesis\ \h \ get_shadow_root host' \\<^sub>r Some shadow_root_ptr\ assms(1) assms(2) assms(4) select_result_I2 shadow_root_same_host) qed lemma a_host_shadow_root_rel_finite: "finite (a_host_shadow_root_rel h)" proof - have "a_host_shadow_root_rel h = (\host \ fset (element_ptr_kinds h). (case |h \ get_shadow_root host|\<^sub>r of Some shadow_root \ {(cast host, cast shadow_root)} | None \ {}))" by(auto simp add: a_host_shadow_root_rel_def split: option.splits) moreover have "finite (\host \ fset (element_ptr_kinds h). (case |h \ get_shadow_root host|\<^sub>r of Some shadow_root \ {(cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r host, cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root)} | None \ {}))" by(auto split: option.splits) ultimately show ?thesis by auto qed lemma a_ptr_disconnected_node_rel_finite: "finite (a_ptr_disconnected_node_rel h)" proof - have "a_ptr_disconnected_node_rel h = (\owner_document \ set |h \ document_ptr_kinds_M|\<^sub>r. (\disconnected_node \ set |h \ get_disconnected_nodes owner_document|\<^sub>r. {(cast owner_document, cast disconnected_node)}))" by(auto simp add: a_ptr_disconnected_node_rel_def) moreover have "finite (\owner_document \ set |h \ document_ptr_kinds_M|\<^sub>r. (\disconnected_node \ set |h \ get_disconnected_nodes owner_document|\<^sub>r. {(cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disconnected_node)}))" by simp ultimately show ?thesis by simp qed lemma heap_is_wellformed_children_in_heap: "heap_is_wellformed h \ h \ get_child_nodes ptr \\<^sub>r children \ child \ set children \ child |\| node_ptr_kinds h" using CD.heap_is_wellformed_children_in_heap local.heap_is_wellformed_def by blast lemma heap_is_wellformed_disc_nodes_in_heap: "heap_is_wellformed h \ h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes \ node \ set disc_nodes \ node |\| node_ptr_kinds h" using CD.heap_is_wellformed_disc_nodes_in_heap local.heap_is_wellformed_def by blast lemma heap_is_wellformed_one_parent: "heap_is_wellformed h \ h \ get_child_nodes ptr \\<^sub>r children \ h \ get_child_nodes ptr' \\<^sub>r children' \ set children \ set children' \ {} \ ptr = ptr'" using CD.heap_is_wellformed_one_parent local.heap_is_wellformed_def by blast lemma heap_is_wellformed_one_disc_parent: "heap_is_wellformed h \ h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes \ h \ get_disconnected_nodes document_ptr' \\<^sub>r disc_nodes' \ set disc_nodes \ set disc_nodes' \ {} \ document_ptr = document_ptr'" using CD.heap_is_wellformed_one_disc_parent local.heap_is_wellformed_def by blast lemma heap_is_wellformed_children_disc_nodes_different: "heap_is_wellformed h \ h \ get_child_nodes ptr \\<^sub>r children \ h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes \ set children \ set disc_nodes = {}" using CD.heap_is_wellformed_children_disc_nodes_different local.heap_is_wellformed_def by blast lemma heap_is_wellformed_disconnected_nodes_distinct: "heap_is_wellformed h \ h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes \ distinct disc_nodes" using CD.heap_is_wellformed_disconnected_nodes_distinct local.heap_is_wellformed_def by blast lemma heap_is_wellformed_children_distinct: "heap_is_wellformed h \ h \ get_child_nodes ptr \\<^sub>r children \ distinct children" using CD.heap_is_wellformed_children_distinct local.heap_is_wellformed_def by blast lemma heap_is_wellformed_children_disc_nodes: "heap_is_wellformed h \ node_ptr |\| node_ptr_kinds h \ \(\parent \ fset (object_ptr_kinds h). node_ptr \ set |h \ get_child_nodes parent|\<^sub>r) \ (\document_ptr \ fset (document_ptr_kinds h). node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r)" using CD.heap_is_wellformed_children_disc_nodes local.heap_is_wellformed_def by blast lemma parent_child_rel_finite: "heap_is_wellformed h \ finite (parent_child_rel h)" using CD.parent_child_rel_finite by blast lemma parent_child_rel_acyclic: "heap_is_wellformed h \ acyclic (parent_child_rel h)" using CD.parent_child_rel_acyclic heap_is_wellformed_def by blast lemma parent_child_rel_child_in_heap: "heap_is_wellformed h \ type_wf h \ known_ptr parent \ (parent, child_ptr) \ parent_child_rel h \ child_ptr |\| object_ptr_kinds h" using CD.parent_child_rel_child_in_heap local.heap_is_wellformed_def by blast end interpretation i_heap_is_wellformed?: l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs by(auto simp add: l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def parent_child_rel_def heap_is_wellformed_def instances) declare l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma heap_is_wellformed_is_l_heap_is_wellformed [instances]: "l_heap_is_wellformed ShadowRootClass.type_wf ShadowRootClass.known_ptr Shadow_DOM.heap_is_wellformed Shadow_DOM.parent_child_rel Shadow_DOM.get_child_nodes get_disconnected_nodes" apply(auto simp add: l_heap_is_wellformed_def instances)[1] using heap_is_wellformed_children_in_heap apply metis using heap_is_wellformed_disc_nodes_in_heap apply metis using heap_is_wellformed_one_parent apply blast using heap_is_wellformed_one_disc_parent apply blast using heap_is_wellformed_children_disc_nodes_different apply blast using heap_is_wellformed_disconnected_nodes_distinct apply metis using heap_is_wellformed_children_distinct apply metis using heap_is_wellformed_children_disc_nodes apply metis using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_child apply(blast, blast) using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_finite apply blast using parent_child_rel_acyclic apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_node_ptr apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_parent_in_heap apply blast using parent_child_rel_child_in_heap apply metis done subsubsection \get\_parent\ interpretation i_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M: l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs get_parent get_parent_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel get_disconnected_nodes by(simp add: l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare i_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] interpretation i_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M: l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs get_parent get_parent_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel get_disconnected_nodes get_disconnected_nodes_locs by(auto simp add: l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare i_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_is_l_get_parent_wf [instances]: "l_get_parent_wf type_wf known_ptr known_ptrs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel get_child_nodes get_parent" apply(auto simp add: l_get_parent_wf_def l_get_parent_wf_axioms_def instances)[1] using i_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.child_parent_dual apply fast using i_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_wellformed_induct apply metis using i_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_wellformed_induct_rev apply metis using i_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_parent apply fast done subsubsection \get\_disconnected\_nodes\ subsubsection \set\_disconnected\_nodes\ paragraph \get\_disconnected\_nodes\ interpretation i_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M: l_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel get_child_nodes by (simp add: l_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare i_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.l_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_is_l_set_disconnected_nodes_get_disconnected_nodes_wf [instances]: "l_set_disconnected_nodes_get_disconnected_nodes_wf type_wf known_ptr heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel get_child_nodes get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs" apply(auto simp add: l_set_disconnected_nodes_get_disconnected_nodes_wf_def l_set_disconnected_nodes_get_disconnected_nodes_wf_axioms_def instances)[1] using i_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.remove_from_disconnected_nodes_removes apply fast done paragraph \get\_root\_node\ interpretation i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M: l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf known_ptrs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_parent get_parent_locs get_ancestors get_ancestors_locs get_root_node get_root_node_locs by(simp add: l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma get_ancestors_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_is_l_get_ancestors_wf [instances]: "l_get_ancestors_wf heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel known_ptr known_ptrs type_wf get_ancestors get_ancestors_locs get_child_nodes get_parent" apply(auto simp add: l_get_ancestors_wf_def l_get_ancestors_wf_axioms_def instances)[1] using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_never_empty apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_ok apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_reads apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_ptrs_in_heap apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_remains_not_in_ancestors apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_also_parent apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_obtains_children apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_parent_child_rel apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_parent_child_rel apply blast done lemma get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_is_l_get_root_node_wf [instances]: "l_get_root_node_wf heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_root_node type_wf known_ptr known_ptrs get_ancestors get_parent" apply(auto simp add: l_get_root_node_wf_def l_get_root_node_wf_axioms_def instances)[1] using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_root_node_ok apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_root_node_ptr_in_heap apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_root_node_root_in_heap apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_same_root_node apply(blast, blast) using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_root_node_same_no_parent apply blast (* using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_root_node_not_node_same apply blast *) using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_root_node_parent_same apply (blast, blast) done subsubsection \to\_tree\_order\ interpretation i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M: l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs to_tree_order known_ptrs get_parent get_parent_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel get_disconnected_nodes get_disconnected_nodes_locs apply(simp add: l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) done declare i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_is_l_to_tree_order_wf [instances]: "l_to_tree_order_wf heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent get_child_nodes" apply(auto simp add: l_to_tree_order_wf_def l_to_tree_order_wf_axioms_def instances)[1] using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_ok apply blast using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_ptrs_in_heap apply blast using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_parent_child_rel apply(blast, blast) using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_child2 apply blast using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_node_ptrs apply blast using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_child apply blast using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_ptr_in_result apply blast using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_parent apply blast using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_subset apply blast done paragraph \get\_root\_node\ interpretation i_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M: l_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf known_ptrs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_parent get_parent_locs get_ancestors get_ancestors_locs get_root_node get_root_node_locs to_tree_order by(auto simp add: l_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare i_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.l_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_is_l_to_tree_order_wf_get_root_node_wf [instances]: "l_to_tree_order_wf_get_root_node_wf type_wf known_ptr known_ptrs to_tree_order get_root_node heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M" apply(auto simp add: l_to_tree_order_wf_get_root_node_wf_def l_to_tree_order_wf_get_root_node_wf_axioms_def instances)[1] using i_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_get_root_node apply blast using i_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_same_root apply blast done subsubsection \remove\_child\ interpretation i_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M: l_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs set_child_nodes set_child_nodes_locs get_parent get_parent_locs get_owner_document get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs remove_child remove_child_locs remove type_wf known_ptr known_ptrs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel by unfold_locales declare i_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.l_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_is_l_remove_child_wf2 [instances]: "l_remove_child_wf2 type_wf known_ptr known_ptrs remove_child heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes remove" apply(auto simp add: l_remove_child_wf2_def l_remove_child_wf2_axioms_def instances)[1] using i_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.remove_child_heap_is_wellformed_preserved apply(fast, fast, fast) using i_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.remove_heap_is_wellformed_preserved apply(fast, fast, fast) using i_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.remove_child_removes_child apply fast using i_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.remove_child_removes_first_child apply fast using i_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.remove_removes_child apply fast using i_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.remove_for_all_empty_children apply fast done subsection \A wellformed heap\ subsubsection \get\_parent\ interpretation i_get_parent_wf?: l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs get_parent get_parent_locs heap_is_wellformed parent_child_rel get_disconnected_nodes using instances by(simp add: l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma get_parent_wf_is_l_get_parent_wf [instances]: "l_get_parent_wf ShadowRootClass.type_wf ShadowRootClass.known_ptr ShadowRootClass.known_ptrs heap_is_wellformed parent_child_rel Shadow_DOM.get_child_nodes Shadow_DOM.get_parent" apply(auto simp add: l_get_parent_wf_def l_get_parent_wf_axioms_def instances)[1] using child_parent_dual apply blast using heap_wellformed_induct apply metis using heap_wellformed_induct_rev apply metis using parent_child_rel_parent apply metis done subsubsection \remove\_shadow\_root\ locale l_remove_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_tag_name + l_get_disconnected_nodes + l_set_shadow_root_get_tag_name + l_get_child_nodes + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_delete_shadow_root_get_disconnected_nodes + l_delete_shadow_root_get_child_nodes + l_set_shadow_root_get_disconnected_nodes + l_set_shadow_root_get_child_nodes + l_delete_shadow_root_get_tag_name + l_set_shadow_root_get_shadow_root + l_delete_shadow_root_get_shadow_root + l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma remove_shadow_root_preserves: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ remove_shadow_root ptr \\<^sub>h h'" shows "known_ptrs h'" and "type_wf h'" "heap_is_wellformed h'" proof - obtain shadow_root_ptr h2 where "h \ get_shadow_root ptr \\<^sub>r Some shadow_root_ptr" and "h \ get_child_nodes (cast shadow_root_ptr) \\<^sub>r []" and "h \ get_disconnected_nodes (cast shadow_root_ptr) \\<^sub>r []" and h2: "h \ set_shadow_root ptr None \\<^sub>h h2" and h': "h2 \ delete_M shadow_root_ptr \\<^sub>h h'" using assms(4) by(auto simp add: remove_shadow_root_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_shadow_root_pure, rotated] bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated] bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] split: option.splits if_splits) have "type_wf h2" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_shadow_root_writes h2] using \type_wf h\ set_shadow_root_types_preserved by(auto simp add: reflp_def transp_def) then show "type_wf h'" using h' delete_shadow_root_type_wf_preserved local.type_wf_impl by blast have object_ptr_kinds_eq_h: "object_ptr_kinds h = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_shadow_root_writes h2]) using set_shadow_root_pointers_preserved apply blast by (auto simp add: reflp_def transp_def) have node_ptr_kinds_eq_h: "node_ptr_kinds h = node_ptr_kinds h2" using object_ptr_kinds_eq_h by (simp add: node_ptr_kinds_def) have element_ptr_kinds_eq_h: "element_ptr_kinds h = element_ptr_kinds h2" using node_ptr_kinds_eq_h by (simp add: element_ptr_kinds_def) have document_ptr_kinds_eq_h: "document_ptr_kinds h = document_ptr_kinds h2" using object_ptr_kinds_eq_h by (simp add: document_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h: "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h2" using object_ptr_kinds_eq_h by (simp add: document_ptr_kinds_eq_h shadow_root_ptr_kinds_def) have "known_ptrs h2" using \known_ptrs h\ object_ptr_kinds_eq_h known_ptrs_subset by blast have object_ptr_kinds_eq_h2: "object_ptr_kinds h' |\| object_ptr_kinds h2" using h' delete_shadow_root_pointers by auto have object_ptr_kinds_eq2_h2: "object_ptr_kinds h2 = object_ptr_kinds h' |\| {|cast shadow_root_ptr|}" using h' delete_shadow_root_pointers by auto have node_ptr_kinds_eq_h2: "node_ptr_kinds h2 = node_ptr_kinds h'" using object_ptr_kinds_eq_h2 by(auto simp add: node_ptr_kinds_def delete_shadow_root_pointers[OF h']) have element_ptr_kinds_eq_h2: "element_ptr_kinds h2 = element_ptr_kinds h'" using node_ptr_kinds_eq_h2 by (simp add: element_ptr_kinds_def) have document_ptr_kinds_eq_h2: "document_ptr_kinds h2 = document_ptr_kinds h' |\| {|cast shadow_root_ptr|}" using object_ptr_kinds_eq_h2 apply(auto simp add: document_ptr_kinds_def delete_shadow_root_pointers[OF h'])[1] using document_ptr_kinds_def by fastforce then have document_ptr_kinds_eq2_h2: "document_ptr_kinds h' |\| document_ptr_kinds h2" using h' delete_shadow_root_pointers by auto have shadow_root_ptr_kinds_eq_h2: "shadow_root_ptr_kinds h' |\| shadow_root_ptr_kinds h2" using object_ptr_kinds_eq_h2 apply(auto simp add: shadow_root_ptr_kinds_def document_ptr_kinds_def)[1] by auto have shadow_root_ptr_kinds_eq2_h2: "shadow_root_ptr_kinds h2 = shadow_root_ptr_kinds h' |\| {|shadow_root_ptr|}" using object_ptr_kinds_eq2_h2 apply (auto simp add: shadow_root_ptr_kinds_def)[1] using document_ptr_kinds_eq_h2 apply auto[1] apply (metis \h \ get_shadow_root ptr \\<^sub>r Some shadow_root_ptr\ assms(1) document_ptr_kinds_eq_h fset.map_comp local.get_shadow_root_shadow_root_ptr_in_heap shadow_root_ptr_kinds_def) using document_ptr_kinds_eq_h2 by auto show "known_ptrs h'" using object_ptr_kinds_eq_h2 \known_ptrs h2\ known_ptrs_subset by blast have disconnected_nodes_eq_h: "\doc_ptr disc_nodes. h \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_shadow_root_writes h2 set_shadow_root_get_disconnected_nodes by(rule reads_writes_preserved) then have disconnected_nodes_eq2_h: "\doc_ptr. |h \ get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. doc_ptr \ cast shadow_root_ptr \ h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads get_disconnected_nodes_delete_shadow_root[rotated, OF h'] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by(metis (no_types, lifting))+ then have disconnected_nodes_eq2_h2: "\doc_ptr. doc_ptr \ cast shadow_root_ptr \ |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have tag_name_eq_h: "\doc_ptr disc_nodes. h \ get_tag_name doc_ptr \\<^sub>r disc_nodes = h2 \ get_tag_name doc_ptr \\<^sub>r disc_nodes" using get_tag_name_reads set_shadow_root_writes h2 set_shadow_root_get_tag_name by(rule reads_writes_preserved) then have tag_name_eq2_h: "\doc_ptr. |h \ get_tag_name doc_ptr|\<^sub>r = |h2 \ get_tag_name doc_ptr|\<^sub>r" using select_result_eq by force have tag_name_eq_h2: "\doc_ptr disc_nodes. h2 \ get_tag_name doc_ptr \\<^sub>r disc_nodes = h' \ get_tag_name doc_ptr \\<^sub>r disc_nodes" using get_tag_name_reads get_tag_name_delete_shadow_root[OF h'] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have tag_name_eq2_h2: "\doc_ptr. |h2 \ get_tag_name doc_ptr|\<^sub>r = |h' \ get_tag_name doc_ptr|\<^sub>r" using select_result_eq by force have children_eq_h: "\ptr' children. h \ get_child_nodes ptr' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads set_shadow_root_writes h2 set_shadow_root_get_child_nodes by(rule reads_writes_preserved) then have children_eq2_h: "\ptr'. |h \ get_child_nodes ptr'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. ptr' \ cast shadow_root_ptr \ h2 \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads h' get_child_nodes_delete_shadow_root apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have children_eq2_h2: "\ptr'. ptr' \ cast shadow_root_ptr \ |h2 \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have "cast shadow_root_ptr |\| object_ptr_kinds h'" using h' delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ptr_not_in_heap by auto have get_shadow_root_eq_h: "\shadow_root_opt ptr'. ptr \ ptr' \ h \ get_shadow_root ptr' \\<^sub>r shadow_root_opt = h2 \ get_shadow_root ptr' \\<^sub>r shadow_root_opt" using get_shadow_root_reads set_shadow_root_writes h2 apply(rule reads_writes_preserved) using set_shadow_root_get_shadow_root_different_pointers by fast have get_shadow_root_eq_h2: "\shadow_root_opt ptr'. h2 \ get_shadow_root ptr' \\<^sub>r shadow_root_opt = h' \ get_shadow_root ptr' \\<^sub>r shadow_root_opt" using get_shadow_root_reads get_shadow_root_delete_shadow_root[OF h'] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have get_shadow_root_eq2_h2: "\ptr'. |h2 \ get_shadow_root ptr'|\<^sub>r = |h' \ get_shadow_root ptr'|\<^sub>r" using select_result_eq by force have "acyclic (parent_child_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def) moreover have "parent_child_rel h = parent_child_rel h2" by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq_h children_eq2_h) moreover have "parent_child_rel h' \ parent_child_rel h2" using object_ptr_kinds_eq_h2 apply(auto simp add: CD.parent_child_rel_def)[1] by (metis \cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr |\| object_ptr_kinds h'\ children_eq2_h2) ultimately have "CD.a_acyclic_heap h'" using acyclic_subset by (auto simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def) moreover have "CD.a_all_ptrs_in_heap h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_all_ptrs_in_heap h2" apply(auto simp add: CD.a_all_ptrs_in_heap_def object_ptr_kinds_eq_h node_ptr_kinds_def children_eq_h disconnected_nodes_eq_h)[1] apply (metis (no_types, lifting) children_eq2_h finite_set_in subsetD) by (metis (no_types, lifting) disconnected_nodes_eq2_h document_ptr_kinds_eq_h finite_set_in in_mono) then have "CD.a_all_ptrs_in_heap h'" apply(auto simp add: CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 children_eq_h2 disconnected_nodes_eq_h2)[1] apply(case_tac "ptr = cast shadow_root_ptr") using object_ptr_kinds_eq_h2 children_eq_h2 apply (meson \cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr |\| object_ptr_kinds h'\ is_OK_returns_result_I local.get_child_nodes_ptr_in_heap) apply(auto dest!: children_eq_h2)[1] using assms(1) children_eq_h local.heap_is_wellformed_children_in_heap node_ptr_kinds_eq_h node_ptr_kinds_eq_h2 apply blast apply (meson \known_ptrs h'\ \type_wf h'\ local.get_child_nodes_ok local.known_ptrs_known_ptr returns_result_select_result) by (metis (no_types, lifting) \cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr |\| object_ptr_kinds h'\ \type_wf h2\ assms(1) disconnected_nodes_eq2_h2 disconnected_nodes_eq_h document_ptr_kinds_commutes document_ptr_kinds_eq2_h2 fin_mono local.get_disconnected_nodes_ok local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq_h node_ptr_kinds_eq_h2 returns_result_select_result) moreover have "CD.a_distinct_lists h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_distinct_lists h2" by(auto simp add: CD.a_distinct_lists_def object_ptr_kinds_eq_h document_ptr_kinds_eq_h children_eq2_h disconnected_nodes_eq2_h) then have "CD.a_distinct_lists h'" apply(auto simp add: CD.a_distinct_lists_def document_ptr_kinds_eq_h2 disconnected_nodes_eq2_h2)[1] apply(auto simp add: intro!: distinct_concat_map_I)[1] apply(case_tac "x = cast shadow_root_ptr") using \cast shadow_root_ptr |\| object_ptr_kinds h'\ apply simp using children_eq_h2 concat_map_all_distinct[of "(\ptr. |h2 \ get_child_nodes ptr|\<^sub>r)"] apply (metis (no_types, lifting) children_eq2_h2 finite_fset fmember.rep_eq fset_mp object_ptr_kinds_eq_h2 set_sorted_list_of_set) apply(case_tac "x = cast shadow_root_ptr") using \cast shadow_root_ptr |\| object_ptr_kinds h'\ apply simp apply(case_tac "y = cast shadow_root_ptr") using \cast shadow_root_ptr |\| object_ptr_kinds h'\ apply simp using children_eq_h2 distinct_concat_map_E(1)[of "(\ptr. |h2 \ get_child_nodes ptr|\<^sub>r)"] apply (smt IntI children_eq2_h2 empty_iff finite_fset fmember.rep_eq fset_mp object_ptr_kinds_eq_h2 set_sorted_list_of_set) apply(auto simp add: intro!: distinct_concat_map_I)[1] apply(case_tac "x = cast shadow_root_ptr") using \cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr |\| object_ptr_kinds h'\ document_ptr_kinds_commutes apply blast apply (metis (mono_tags, lifting) \local.CD.a_distinct_lists h2\ \type_wf h'\ disconnected_nodes_eq_h2 is_OK_returns_result_E local.CD.distinct_lists_disconnected_nodes local.get_disconnected_nodes_ok select_result_I2) apply(case_tac "x = cast shadow_root_ptr") using \cast shadow_root_ptr |\| object_ptr_kinds h'\ apply simp apply(case_tac "y = cast shadow_root_ptr") using \cast shadow_root_ptr |\| object_ptr_kinds h'\ apply simp proof - fix x and y and xa assume a1: "x |\| document_ptr_kinds h'" assume a2: "y |\| document_ptr_kinds h'" assume a3: "x \ y" assume a4: "x \ cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr" assume a5: "y \ cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr" assume a6: "xa \ set |h' \ get_disconnected_nodes x|\<^sub>r" assume a7: "xa \ set |h' \ get_disconnected_nodes y|\<^sub>r" assume "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>r) (insort (cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr) (sorted_list_of_set (fset (document_ptr_kinds h') - {cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr})))))" then show False using a7 a6 a5 a4 a3 a2 a1 by (metis (no_types) IntI distinct_concat_map_E(1)[of "(\ptr. |h2 \ get_disconnected_nodes ptr|\<^sub>r)"] disconnected_nodes_eq2_h2 empty_iff finite_fset finsert.rep_eq fmember.rep_eq insert_iff set_sorted_list_of_set - sorted_list_of_set.insert_remove) + sorted_list_of_set_insert_remove) next fix x xa xb assume 0: "distinct (concat (map (\ptr. |h2 \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h2)))))" and 1: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>r) (insort (cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr) (sorted_list_of_set (fset (document_ptr_kinds h') - {cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr})))))" and 2: "(\x\fset (object_ptr_kinds h2). set |h2 \ get_child_nodes x|\<^sub>r) \ (\x\set (insort (cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr) (sorted_list_of_set (fset (document_ptr_kinds h') - {cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr}))). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" and 3: "xa |\| object_ptr_kinds h'" and 4: "x \ set |h' \ get_child_nodes xa|\<^sub>r" and 5: "xb |\| document_ptr_kinds h'" and 6: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" then show "False" apply(cases "xa = cast shadow_root_ptr") using \cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr |\| object_ptr_kinds h'\ apply blast apply(cases "xb = cast shadow_root_ptr") using \cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr |\| object_ptr_kinds h'\ document_ptr_kinds_commutes apply blast by (metis (no_types, hide_lams) \local.CD.a_distinct_lists h2\ \type_wf h'\ children_eq2_h2 disconnected_nodes_eq_h2 fset_rev_mp is_OK_returns_result_E local.CD.distinct_lists_no_parent local.get_disconnected_nodes_ok object_ptr_kinds_eq_h2 select_result_I2) qed moreover have "CD.a_owner_document_valid h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_owner_document_valid h2" by(auto simp add: CD.a_owner_document_valid_def object_ptr_kinds_eq_h document_ptr_kinds_eq_h node_ptr_kinds_eq_h children_eq2_h disconnected_nodes_eq2_h) then have "CD.a_owner_document_valid h'" apply(auto simp add: CD.a_owner_document_valid_def document_ptr_kinds_eq_h2 node_ptr_kinds_eq_h2 disconnected_nodes_eq2_h2)[1] by (smt \h \ get_child_nodes (cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr) \\<^sub>r []\ \h \ get_disconnected_nodes (cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr) \\<^sub>r []\ \local.CD.a_distinct_lists h\ children_eq2_h children_eq2_h2 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 finite_set_in finsert_iff funion_finsert_right local.CD.distinct_lists_no_parent object_ptr_kinds_eq2_h2 object_ptr_kinds_eq_h select_result_I2 sup_bot.comm_neutral) ultimately have "heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'" by(simp add: CD.heap_is_wellformed_def) moreover have "acyclic (parent_child_rel h \ a_host_shadow_root_rel h \ a_ptr_disconnected_node_rel h)" using \heap_is_wellformed h\ by(simp add: heap_is_wellformed_def) then have "acyclic (parent_child_rel h2 \ a_host_shadow_root_rel h2 \ a_ptr_disconnected_node_rel h2)" proof - have "a_host_shadow_root_rel h2 \ a_host_shadow_root_rel h" apply(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h)[1] apply(case_tac "aa = ptr") apply(simp) apply (metis (no_types, lifting) \type_wf h2\ assms(2) h2 local.get_shadow_root_ok local.type_wf_impl option.distinct(1) returns_result_eq returns_result_select_result set_shadow_root_get_shadow_root) using get_shadow_root_eq_h by (metis (mono_tags, lifting) \type_wf h2\ image_eqI is_OK_returns_result_E local.get_shadow_root_ok mem_Collect_eq prod.simps(2) select_result_I2) moreover have "a_ptr_disconnected_node_rel h = a_ptr_disconnected_node_rel h2" by (simp add: a_ptr_disconnected_node_rel_def disconnected_nodes_eq2_h document_ptr_kinds_eq_h) ultimately show ?thesis using \parent_child_rel h = parent_child_rel h2\ by (smt \acyclic (parent_child_rel h \ local.a_host_shadow_root_rel h \ local.a_ptr_disconnected_node_rel h)\ acyclic_subset subset_refl sup_mono) qed then have "acyclic (parent_child_rel h' \ a_host_shadow_root_rel h' \ a_ptr_disconnected_node_rel h')" proof - have "a_host_shadow_root_rel h' \ a_host_shadow_root_rel h2" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h2 get_shadow_root_eq2_h2) moreover have "a_ptr_disconnected_node_rel h2 = a_ptr_disconnected_node_rel h'" apply(simp add: a_ptr_disconnected_node_rel_def disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2) by (metis (no_types, lifting) \cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr |\| object_ptr_kinds h'\ \h \ get_child_nodes (cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr) \\<^sub>r []\ \h \ get_disconnected_nodes (cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr) \\<^sub>r []\ \local.CD.a_distinct_lists h\ disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 document_ptr_kinds_commutes is_OK_returns_result_I local.CD.distinct_lists_no_parent local.get_disconnected_nodes_ptr_in_heap select_result_I2) ultimately show ?thesis using \parent_child_rel h' \ parent_child_rel h2\ \acyclic (parent_child_rel h2 \ a_host_shadow_root_rel h2 \ a_ptr_disconnected_node_rel h2)\ using acyclic_subset order_refl sup_mono by (metis (no_types, hide_lams)) qed moreover have "a_all_ptrs_in_heap h" using \heap_is_wellformed h\ by(simp add: heap_is_wellformed_def) then have "a_all_ptrs_in_heap h2" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h)[1] apply(case_tac "host = ptr") apply(simp) apply (metis assms(2) h2 local.type_wf_impl option.distinct(1) returns_result_eq set_shadow_root_get_shadow_root) using get_shadow_root_eq_h by fastforce then have "a_all_ptrs_in_heap h'" apply(auto simp add: a_all_ptrs_in_heap_def get_shadow_root_eq_h2)[1] apply(auto simp add: shadow_root_ptr_kinds_eq2_h2)[1] by (metis (no_types, lifting) \h \ get_shadow_root ptr \\<^sub>r Some shadow_root_ptr\ assms(1) assms(2) get_shadow_root_eq_h get_shadow_root_eq_h2 h2 local.shadow_root_same_host local.type_wf_impl option.distinct(1) select_result_I2 set_shadow_root_get_shadow_root) moreover have "a_distinct_lists h" using \heap_is_wellformed h\ by(simp add: heap_is_wellformed_def) then have "a_distinct_lists h2" apply(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h)[1] apply(auto intro!: distinct_concat_map_I split: option.splits)[1] apply(case_tac "x = ptr") apply(simp) apply (metis (no_types, hide_lams) assms(2) h2 is_OK_returns_result_I l_set_shadow_root_get_shadow_root.set_shadow_root_get_shadow_root l_set_shadow_root_get_shadow_root_axioms local.type_wf_impl option.discI returns_result_eq returns_result_select_result) apply(case_tac "y = ptr") apply(simp) apply (metis (no_types, hide_lams) assms(2) h2 is_OK_returns_result_I l_set_shadow_root_get_shadow_root.set_shadow_root_get_shadow_root l_set_shadow_root_get_shadow_root_axioms local.type_wf_impl option.discI returns_result_eq returns_result_select_result) by (metis \type_wf h2\ assms(1) assms(2) get_shadow_root_eq_h local.get_shadow_root_ok local.shadow_root_same_host returns_result_select_result) then have "a_distinct_lists h'" by(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h2 get_shadow_root_eq2_h2) moreover have "a_shadow_root_valid h" using \heap_is_wellformed h\ by(simp add: heap_is_wellformed_def) then have "a_shadow_root_valid h'" apply(auto simp add: a_shadow_root_valid_def shadow_root_ptr_kinds_eq_h element_ptr_kinds_eq_h tag_name_eq2_h)[1] apply(simp add: shadow_root_ptr_kinds_eq2_h2 element_ptr_kinds_eq_h2 tag_name_eq2_h2) using get_shadow_root_eq_h get_shadow_root_eq_h2 by (smt \cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr |\| object_ptr_kinds h'\ \h \ get_shadow_root ptr \\<^sub>r Some shadow_root_ptr\ assms(2) document_ptr_kinds_commutes element_ptr_kinds_eq_h element_ptr_kinds_eq_h2 finite_set_in local.get_shadow_root_ok option.inject returns_result_select_result select_result_I2 shadow_root_ptr_kinds_commutes) ultimately show "heap_is_wellformed h'" by(simp add: heap_is_wellformed_def) qed end interpretation i_remove_shadow_root_wf?: l_remove_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_tag_name get_tag_name_locs get_disconnected_nodes get_disconnected_nodes_locs set_shadow_root set_shadow_root_locs known_ptr get_child_nodes get_child_nodes_locs get_shadow_root get_shadow_root_locs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs remove_shadow_root remove_shadow_root_locs known_ptrs get_parent get_parent_locs by(auto simp add: l_remove_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_remove_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_root\_node\ interpretation i_get_root_node_wf?: l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf known_ptrs heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_parent get_parent_locs get_ancestors get_ancestors_locs get_root_node get_root_node_locs by(simp add: l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma get_ancestors_wf_is_l_get_ancestors_wf [instances]: "l_get_ancestors_wf heap_is_wellformed parent_child_rel known_ptr known_ptrs type_wf get_ancestors get_ancestors_locs get_child_nodes get_parent" apply(auto simp add: l_get_ancestors_wf_def l_get_ancestors_wf_axioms_def instances)[1] using get_ancestors_never_empty apply blast using get_ancestors_ok apply blast using get_ancestors_reads apply blast using get_ancestors_ptrs_in_heap apply blast using get_ancestors_remains_not_in_ancestors apply blast using get_ancestors_also_parent apply blast using get_ancestors_obtains_children apply blast using get_ancestors_parent_child_rel apply blast using get_ancestors_parent_child_rel apply blast done lemma get_root_node_wf_is_l_get_root_node_wf [instances]: "l_get_root_node_wf heap_is_wellformed get_root_node type_wf known_ptr known_ptrs get_ancestors get_parent" using known_ptrs_is_l_known_ptrs apply(auto simp add: l_get_root_node_wf_def l_get_root_node_wf_axioms_def)[1] using get_root_node_ok apply blast using get_root_node_ptr_in_heap apply blast using get_root_node_root_in_heap apply blast using get_ancestors_same_root_node apply(blast, blast) using get_root_node_same_no_parent apply blast (* using get_root_node_not_node_same apply blast *) using get_root_node_parent_same apply (blast, blast) done subsubsection \get\_parent\_get\_host\_get\_disconnected\_document\ locale l_get_parent_get_host_get_disconnected_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs + l_get_disconnected_document get_disconnected_document get_disconnected_document_locs + l_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs + l_get_parent_wf type_wf known_ptr known_ptrs heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs get_parent get_parent_locs + l_get_shadow_root type_wf get_shadow_root get_shadow_root_locs + l_get_host get_host get_host_locs + l_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs for get_child_nodes :: "(_::linorder) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and known_ptr :: "(_) object_ptr \ bool" and type_wf :: "(_) heap \ bool" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_disconnected_document :: "(_) node_ptr \ ((_) heap, exception, (_) document_ptr) prog" and get_disconnected_document_locs :: "((_) heap \ (_) heap \ bool) set" and known_ptrs :: "(_) heap \ bool" and get_parent :: "(_) node_ptr \ ((_) heap, exception, (_) object_ptr option) prog" and get_parent_locs :: "((_) heap \ (_) heap \ bool) set" begin lemma a_host_shadow_root_rel_shadow_root: "h \ get_shadow_root host \\<^sub>r shadow_root_option \ shadow_root_option = Some shadow_root \ ((cast host, cast shadow_root) \ a_host_shadow_root_rel h)" apply(auto simp add: a_host_shadow_root_rel_def)[1] by(metis (mono_tags, lifting) case_prodI is_OK_returns_result_I l_get_shadow_root.get_shadow_root_ptr_in_heap local.l_get_shadow_root_axioms mem_Collect_eq pair_imageI select_result_I2) lemma a_host_shadow_root_rel_host: "heap_is_wellformed h \ h \ get_host shadow_root \\<^sub>r host \ ((cast host, cast shadow_root) \ a_host_shadow_root_rel h)" apply(auto simp add: a_host_shadow_root_rel_def)[1] using shadow_root_host_dual by (metis (no_types, lifting) Collect_cong a_host_shadow_root_rel_shadow_root local.a_host_shadow_root_rel_def split_cong) lemma a_ptr_disconnected_node_rel_disconnected_node: "h \ get_disconnected_nodes document \\<^sub>r disc_nodes \ node_ptr \ set disc_nodes \ (cast document, cast node_ptr) \ a_ptr_disconnected_node_rel h" apply(auto simp add: a_ptr_disconnected_node_rel_def)[1] by (smt CD.get_disconnected_nodes_ptr_in_heap case_prodI is_OK_returns_result_I mem_Collect_eq pair_imageI select_result_I2) lemma a_ptr_disconnected_node_rel_document: "heap_is_wellformed h \ h \ get_disconnected_document node_ptr \\<^sub>r document \ (cast document, cast node_ptr) \ a_ptr_disconnected_node_rel h" apply(auto simp add: a_ptr_disconnected_node_rel_def)[1] using disc_doc_disc_node_dual by (metis (no_types, lifting) local.a_ptr_disconnected_node_rel_def a_ptr_disconnected_node_rel_disconnected_node) lemma heap_wellformed_induct_si [consumes 1, case_names step]: assumes "heap_is_wellformed h" assumes "\parent. (\children child. h \ get_child_nodes parent \\<^sub>r children \ child \ set children \ P (cast child)) \ (\shadow_root host. parent = cast host \ h \ get_shadow_root host \\<^sub>r Some shadow_root \ P (cast shadow_root)) \ (\owner_document disc_nodes node_ptr. parent = cast owner_document \ h \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes \ node_ptr \ set disc_nodes \ P (cast node_ptr)) \ P parent" shows "P ptr" proof - fix ptr have "finite (parent_child_rel h \ a_host_shadow_root_rel h \ a_ptr_disconnected_node_rel h)" using a_host_shadow_root_rel_finite a_ptr_disconnected_node_rel_finite using local.CD.parent_child_rel_finite local.CD.parent_child_rel_impl by auto then have "wf ((parent_child_rel h \ a_host_shadow_root_rel h \ a_ptr_disconnected_node_rel h)\)" using assms(1) apply(simp add: heap_is_wellformed_def) by (simp add: finite_acyclic_wf_converse local.CD.parent_child_rel_impl) then show "?thesis" proof (induct rule: wf_induct_rule) case (less parent) then show ?case apply(auto)[1] using assms a_ptr_disconnected_node_rel_disconnected_node a_host_shadow_root_rel_shadow_root local.CD.parent_child_rel_child by blast qed qed lemma heap_wellformed_induct_rev_si [consumes 1, case_names step]: assumes "heap_is_wellformed h" assumes "\child. (\parent child_node. child = cast child_node \ h \ get_parent child_node \\<^sub>r Some parent \ P parent) \ (\host shadow_root. child = cast shadow_root \ h \ get_host shadow_root \\<^sub>r host \ P (cast host)) \ (\disc_doc disc_node. child = cast disc_node \ h \ get_disconnected_document disc_node \\<^sub>r disc_doc\ P (cast disc_doc)) \ P child" shows "P ptr" proof - fix ptr have "finite (parent_child_rel h \ a_host_shadow_root_rel h \ a_ptr_disconnected_node_rel h)" using a_host_shadow_root_rel_finite a_ptr_disconnected_node_rel_finite using local.CD.parent_child_rel_finite local.CD.parent_child_rel_impl by auto then have "wf (parent_child_rel h \ a_host_shadow_root_rel h \ a_ptr_disconnected_node_rel h)" using assms(1) apply(simp add: heap_is_wellformed_def) by (simp add: finite_acyclic_wf) then show "?thesis" proof (induct rule: wf_induct_rule) case (less parent) then show ?case apply(auto)[1] using parent_child_rel_parent a_host_shadow_root_rel_host a_ptr_disconnected_node_rel_document using assms(1) assms(2) by auto qed qed end interpretation i_get_parent_get_host_get_disconnected_document_wf?: l_get_parent_get_host_get_disconnected_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs known_ptrs get_parent get_parent_locs by(auto simp add: l_get_parent_get_host_get_disconnected_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_parent_get_host_get_disconnected_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_get_parent_get_host_wf = l_heap_is_wellformed_defs + l_get_parent_defs + l_get_shadow_root_defs + l_get_host_defs + l_get_child_nodes_defs + l_get_disconnected_document_defs + l_get_disconnected_nodes_defs + assumes heap_wellformed_induct_si [consumes 1, case_names step]: "heap_is_wellformed h \ (\parent. (\children child. h \ get_child_nodes parent \\<^sub>r children \ child \ set children \ P (cast child)) \ (\shadow_root host. parent = cast host \ h \ get_shadow_root host \\<^sub>r Some shadow_root \ P (cast shadow_root)) \ (\owner_document disc_nodes node_ptr. parent = cast owner_document \ h \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes \ node_ptr \ set disc_nodes \ P (cast node_ptr)) \ P parent) \ P ptr" assumes heap_wellformed_induct_rev_si [consumes 1, case_names step]: "heap_is_wellformed h \ (\child. (\parent child_node. child = cast child_node \ h \ get_parent child_node \\<^sub>r Some parent \ P parent) \ (\host shadow_root. child = cast shadow_root \ h \ get_host shadow_root \\<^sub>r host \ P (cast host)) \ (\disc_doc disc_node. child = cast disc_node \ h \ get_disconnected_document disc_node \\<^sub>r disc_doc \ P (cast disc_doc)) \ P child) \ P ptr" lemma l_get_parent_get_host_wf_is_get_parent_get_host_wf [instances]: "l_get_parent_get_host_wf heap_is_wellformed get_parent get_shadow_root get_host get_child_nodes get_disconnected_document get_disconnected_nodes" using heap_wellformed_induct_si heap_wellformed_induct_rev_si using l_get_parent_get_host_wf_def by blast subsubsection \get\_host\ locale l_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs + l_type_wf type_wf + l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_shadow_root get_shadow_root_locs get_host get_host_locs type_wf + l_get_shadow_root type_wf get_shadow_root get_shadow_root_locs for known_ptr :: "(_::linorder) object_ptr \ bool" and known_ptrs :: "(_) heap \ bool" and type_wf :: "(_) heap \ bool" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_child_nodes :: "(_::linorder) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" begin lemma get_host_ok [simp]: assumes "heap_is_wellformed h" assumes "type_wf h" assumes "known_ptrs h" assumes "shadow_root_ptr |\| shadow_root_ptr_kinds h" shows "h \ ok (get_host shadow_root_ptr)" proof - obtain host where host: "host |\| element_ptr_kinds h" and "|h \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types" and shadow_root: "h \ get_shadow_root host \\<^sub>r Some shadow_root_ptr" using assms(1) assms(4) get_shadow_root_ok assms(2) apply(auto simp add: heap_is_wellformed_def a_shadow_root_valid_def)[1] by (smt finite_set_in returns_result_select_result) obtain host_candidates where host_candidates: "h \ filter_M (\element_ptr. Heap_Error_Monad.bind (get_shadow_root element_ptr) (\shadow_root_opt. return (shadow_root_opt = Some shadow_root_ptr))) (sorted_list_of_set (fset (element_ptr_kinds h))) \\<^sub>r host_candidates" apply(drule is_OK_returns_result_E[rotated]) using get_shadow_root_ok assms(2) by(auto intro!: filter_M_is_OK_I bind_pure_I bind_is_OK_I2) then have "host_candidates = [host]" apply(rule filter_M_ex1) using host apply(auto)[1] apply (smt assms(1) assms(2) bind_pure_returns_result_I2 bind_returns_result_E finite_set_in host local.get_shadow_root_ok local.get_shadow_root_pure local.shadow_root_same_host return_returns_result returns_result_eq shadow_root sorted_list_of_fset.rep_eq sorted_list_of_fset_simps(1)) apply (simp add: bind_pure_I) apply(auto intro!: bind_pure_returns_result_I)[1] apply (smt assms(2) bind_pure_returns_result_I2 host local.get_shadow_root_ok local.get_shadow_root_pure return_returns_result returns_result_eq shadow_root) done then show ?thesis using host_candidates host assms(1) get_shadow_root_ok apply(auto simp add: get_host_def known_ptrs_known_ptr intro!: bind_is_OK_pure_I filter_M_pure_I filter_M_is_OK_I bind_pure_I split: list.splits)[1] using assms(2) apply blast apply (meson list.distinct(1) returns_result_eq) by (meson list.distinct(1) list.inject returns_result_eq) qed end interpretation i_get_host_wf?: l_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_document get_disconnected_document_locs known_ptr known_ptrs type_wf get_host get_host_locs get_shadow_root get_shadow_root_locs get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M by(auto simp add: l_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_get_host_wf = l_heap_is_wellformed_defs + l_known_ptrs + l_type_wf + l_get_host_defs + assumes get_host_ok: "heap_is_wellformed h \ known_ptrs h \ type_wf h \ shadow_root_ptr |\| shadow_root_ptr_kinds h \ h \ ok (get_host shadow_root_ptr)" lemma get_host_wf_is_l_get_host_wf [instances]: "l_get_host_wf heap_is_wellformed known_ptr known_ptrs type_wf get_host" by(auto simp add: l_get_host_wf_def l_get_host_wf_axioms_def instances) subsubsection \get\_root\_node\_si\ locale l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent_wf + l_get_parent_get_host_wf + l_get_host_wf begin lemma get_root_node_ptr_in_heap: assumes "h \ ok (get_root_node_si ptr)" shows "ptr |\| object_ptr_kinds h" using assms unfolding get_root_node_si_def using get_ancestors_si_ptr_in_heap by auto lemma get_ancestors_si_ok: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" and "ptr |\| object_ptr_kinds h" shows "h \ ok (get_ancestors_si ptr)" proof (insert assms(1) assms(4), induct rule: heap_wellformed_induct_rev_si) case (step child) then show ?case using assms(2) assms(3) apply(auto simp add: get_ancestors_si_def[of child] assms(1) get_parent_parent_in_heap intro!: bind_is_OK_pure_I split: option.splits)[1] using local.get_parent_ok apply blast using get_host_ok assms(1) apply blast by (meson assms(1) is_OK_returns_result_I local.get_shadow_root_ptr_in_heap local.shadow_root_host_dual) qed lemma get_ancestors_si_remains_not_in_ancestors: assumes "heap_is_wellformed h" and "heap_is_wellformed h'" and "h \ get_ancestors_si ptr \\<^sub>r ancestors" and "h' \ get_ancestors_si ptr \\<^sub>r ancestors'" and "\p children children'. h \ get_child_nodes p \\<^sub>r children \ h' \ get_child_nodes p \\<^sub>r children' \ set children' \ set children" and "\p shadow_root_option shadow_root_option'. h \ get_shadow_root p \\<^sub>r shadow_root_option \ h' \ get_shadow_root p \\<^sub>r shadow_root_option' \ (if shadow_root_option = None then shadow_root_option' = None else shadow_root_option' = None \ shadow_root_option' = shadow_root_option)" and "node \ set ancestors" and object_ptr_kinds_eq3: "object_ptr_kinds h = object_ptr_kinds h'" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" and type_wf': "type_wf h'" shows "node \ set ancestors'" proof - have object_ptr_kinds_M_eq: "\ptrs. h \ object_ptr_kinds_M \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" using object_ptr_kinds_eq3 by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_eq: "|h \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" by(simp) show ?thesis proof (insert assms(1) assms(3) assms(4) assms(7), induct ptr arbitrary: ancestors ancestors' rule: heap_wellformed_induct_rev_si) case (step child) obtain ancestors_remains where ancestors_remains: "ancestors = child # ancestors_remains" using \h \ get_ancestors_si child \\<^sub>r ancestors\ get_ancestors_si_never_empty by(auto simp add: get_ancestors_si_def[of child] elim!: bind_returns_result_E2 split: option.splits) obtain ancestors_remains' where ancestors_remains': "ancestors' = child # ancestors_remains'" using \h' \ get_ancestors_si child \\<^sub>r ancestors'\ get_ancestors_si_never_empty by(auto simp add: get_ancestors_si_def[of child] elim!: bind_returns_result_E2 split: option.splits) have "child |\| object_ptr_kinds h" using local.get_ancestors_si_ptr_in_heap object_ptr_kinds_eq3 step.prems(2) by fastforce have "node \ child" using ancestors_remains step.prems(3) by auto have 1: "\p parent. h' \ get_parent p \\<^sub>r Some parent \ h \ get_parent p \\<^sub>r Some parent" proof - fix p parent assume "h' \ get_parent p \\<^sub>r Some parent" then obtain children' where children': "h' \ get_child_nodes parent \\<^sub>r children'" and p_in_children': "p \ set children'" using get_parent_child_dual by blast obtain children where children: "h \ get_child_nodes parent \\<^sub>r children" using get_child_nodes_ok assms(1) get_child_nodes_ptr_in_heap object_ptr_kinds_eq children' known_ptrs using type_wf type_wf' by (metis \h' \ get_parent p \\<^sub>r Some parent\ get_parent_parent_in_heap is_OK_returns_result_E local.known_ptrs_known_ptr object_ptr_kinds_eq3) have "p \ set children" using assms(5) children children' p_in_children' by blast then show "h \ get_parent p \\<^sub>r Some parent" using child_parent_dual assms(1) children known_ptrs type_wf by blast qed have 2: "\p host. h' \ get_host p \\<^sub>r host \ h \ get_host p \\<^sub>r host" proof - fix p host assume "h' \ get_host p \\<^sub>r host" then have "h' \ get_shadow_root host \\<^sub>r Some p" using local.shadow_root_host_dual by blast then have "h \ get_shadow_root host \\<^sub>r Some p" by (metis assms(6) element_ptr_kinds_commutes is_OK_returns_result_I local.get_shadow_root_ok local.get_shadow_root_ptr_in_heap node_ptr_kinds_commutes object_ptr_kinds_eq3 option.distinct(1) returns_result_select_result type_wf) then show "h \ get_host p \\<^sub>r host" by (metis assms(1) is_OK_returns_result_E known_ptrs local.get_host_ok local.get_shadow_root_shadow_root_ptr_in_heap local.shadow_root_host_dual local.shadow_root_same_host type_wf) qed show ?case proof (cases "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child") case None then show ?thesis using step(4) step(5) \node \ child\ apply(auto simp add: get_ancestors_si_def[of child] elim!: bind_returns_result_E2 split: option.splits)[1] by (metis "2" assms(1) shadow_root_same_host list.set_intros(2) shadow_root_host_dual step.hyps(2) step.prems(3) type_wf) next case (Some node_child) then show ?thesis using step(4) step(5) \node \ child\ apply(auto simp add: get_ancestors_si_def[of child] elim!: bind_returns_result_E2 split: option.splits)[1] apply (meson "1" option.distinct(1) returns_result_eq) by (metis "1" list.set_intros(2) option.inject returns_result_eq step.hyps(1) step.prems(3)) qed qed qed lemma get_ancestors_si_ptrs_in_heap: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_ancestors_si ptr \\<^sub>r ancestors" assumes "ptr' \ set ancestors" shows "ptr' |\| object_ptr_kinds h" proof (insert assms(4) assms(5), induct ancestors arbitrary: ptr) case Nil then show ?case by(auto) next case (Cons a ancestors) then obtain x where x: "h \ get_ancestors_si x \\<^sub>r a # ancestors" by(auto simp add: get_ancestors_si_def[of a] elim!: bind_returns_result_E2 split: option.splits) then have "x = a" by(auto simp add: get_ancestors_si_def[of x] elim!: bind_returns_result_E2 split: option.splits) then show ?case proof (cases "ptr' = a") case True then show ?thesis using Cons.hyps Cons.prems(2) get_ancestors_si_ptr_in_heap x using \x = a\ by blast next case False obtain ptr'' where ptr'': "h \ get_ancestors_si ptr'' \\<^sub>r ancestors" using \ h \ get_ancestors_si x \\<^sub>r a # ancestors\ Cons.prems(2) False apply(auto simp add: get_ancestors_si_def elim!: bind_returns_result_E2)[1] apply(auto elim!: bind_returns_result_E2 split: option.splits intro!: bind_pure_I)[1] apply(auto elim!: bind_returns_result_E2 split: option.splits intro!: bind_pure_I)[1] apply (metis local.get_ancestors_si_def) by (simp add: local.get_ancestors_si_def) then show ?thesis using Cons.hyps Cons.prems(2) False by auto qed qed lemma get_ancestors_si_reads: assumes "heap_is_wellformed h" shows "reads get_ancestors_si_locs (get_ancestors_si node_ptr) h h'" proof (insert assms(1), induct rule: heap_wellformed_induct_rev_si) case (step child) then show ?case using [[simproc del: Product_Type.unit_eq]] get_parent_reads[unfolded reads_def] get_host_reads[unfolded reads_def] apply(simp (no_asm) add: get_ancestors_si_def) by(auto simp add: get_ancestors_si_locs_def get_parent_reads_pointers intro!: reads_bind_pure reads_subset[OF check_in_heap_reads] reads_subset[OF return_reads] reads_subset[OF get_parent_reads] reads_subset[OF get_child_nodes_reads] reads_subset[OF get_host_reads] split: option.splits) qed lemma get_ancestors_si_subset: assumes "heap_is_wellformed h" and "h \ get_ancestors_si ptr \\<^sub>r ancestors" and "ancestor \ set ancestors" and "h \ get_ancestors_si ancestor \\<^sub>r ancestor_ancestors" and type_wf: "type_wf h" and known_ptrs: "known_ptrs h" shows "set ancestor_ancestors \ set ancestors" proof (insert assms(1) assms(2) assms(3), induct ptr arbitrary: ancestors rule: heap_wellformed_induct_rev_si) case (step child) have "child |\| object_ptr_kinds h" using get_ancestors_si_ptr_in_heap step(4) by auto (* then have "h \ check_in_heap child \\<^sub>r ()" using returns_result_select_result by force *) obtain tl_ancestors where tl_ancestors: "ancestors = child # tl_ancestors" using step(4) by(auto simp add: get_ancestors_si_def[of child] intro!: bind_pure_I elim!: bind_returns_result_E2 split: option.splits) show ?case proof (induct "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child") case None show ?case proof (induct "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child") case None then show ?case using step(4) \None = cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child\ apply(auto simp add: get_ancestors_si_def[of child] elim!: bind_returns_result_E2)[1] by (metis (no_types, lifting) assms(4) empty_iff empty_set select_result_I2 set_ConsD step.prems(1) step.prems(2)) next case (Some shadow_root_child) then have "cast shadow_root_child |\| document_ptr_kinds h" using \child |\| object_ptr_kinds h\ apply(auto simp add: document_ptr_kinds_def split: option.splits)[1] by (metis (mono_tags, lifting) document_ptr_casts_commute3 document_ptr_kinds_commutes document_ptr_kinds_def fset.map_comp shadow_root_ptr_casts_commute) then have "shadow_root_child |\| shadow_root_ptr_kinds h" using shadow_root_ptr_kinds_commutes by blast obtain host where host: "h \ get_host shadow_root_child \\<^sub>r host" using get_host_ok assms by (meson \shadow_root_child |\| shadow_root_ptr_kinds h\ is_OK_returns_result_E) then have "h \ get_ancestors_si (cast host) \\<^sub>r tl_ancestors" using Some step(4) tl_ancestors None by(auto simp add: get_ancestors_si_def[of child] intro!: bind_pure_returns_result_I elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq) then show ?case using step(2) Some host step(5) tl_ancestors using assms(4) dual_order.trans eq_iff returns_result_eq set_ConsD set_subset_Cons shadow_root_ptr_casts_commute document_ptr_casts_commute step.prems(1) by (smt case_optionE local.shadow_root_host_dual option.case_distrib option.distinct(1)) qed next case (Some child_node) note s1 = Some obtain parent_opt where parent_opt: "h \ get_parent child_node \\<^sub>r parent_opt" using \child |\| object_ptr_kinds h\ assms(1) Some[symmetric] get_parent_ok[OF type_wf known_ptrs] by (metis (no_types, lifting) is_OK_returns_result_E known_ptrs get_parent_ok l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms node_ptr_casts_commute node_ptr_kinds_commutes) then show ?case proof (induct parent_opt) case None then have "ancestors = [child]" using step(4) s1 apply(simp add: get_ancestors_si_def) by(auto elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq) show ?case using step(4) step(5) apply(auto simp add: \ancestors = [child]\)[1] using assms(4) returns_result_eq by fastforce next case (Some parent) then have "h \ get_ancestors_si parent \\<^sub>r tl_ancestors" using s1 tl_ancestors step(4) by(auto simp add: get_ancestors_si_def[of child] elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq) show ?case by (metis (no_types, lifting) Some.prems \h \ get_ancestors_si parent \\<^sub>r tl_ancestors\ assms(4) eq_iff node_ptr_casts_commute order_trans s1 select_result_I2 set_ConsD set_subset_Cons step.hyps(1) step.prems(1) step.prems(2) tl_ancestors) qed qed qed lemma get_ancestors_si_also_parent: assumes "heap_is_wellformed h" and "h \ get_ancestors_si some_ptr \\<^sub>r ancestors" and "cast child \ set ancestors" and "h \ get_parent child \\<^sub>r Some parent" and type_wf: "type_wf h" and known_ptrs: "known_ptrs h" shows "parent \ set ancestors" proof - obtain child_ancestors where child_ancestors: "h \ get_ancestors_si (cast child) \\<^sub>r child_ancestors" by (meson assms(1) assms(4) get_ancestors_si_ok is_OK_returns_result_I known_ptrs local.get_parent_ptr_in_heap node_ptr_kinds_commutes returns_result_select_result type_wf) then have "parent \ set child_ancestors" apply(simp add: get_ancestors_si_def) by(auto elim!: bind_returns_result_E2 split: option.splits dest!: returns_result_eq[OF assms(4)] get_ancestors_si_ptr) then show ?thesis using assms child_ancestors get_ancestors_si_subset by blast qed lemma get_ancestors_si_also_host: assumes "heap_is_wellformed h" and "h \ get_ancestors_si some_ptr \\<^sub>r ancestors" and "cast shadow_root \ set ancestors" and "h \ get_host shadow_root \\<^sub>r host" and type_wf: "type_wf h" and known_ptrs: "known_ptrs h" shows "cast host \ set ancestors" proof - obtain child_ancestors where child_ancestors: "h \ get_ancestors_si (cast shadow_root) \\<^sub>r child_ancestors" by (meson assms(1) assms(2) assms(3) get_ancestors_si_ok get_ancestors_si_ptrs_in_heap is_OK_returns_result_E known_ptrs type_wf) then have "cast host \ set child_ancestors" apply(simp add: get_ancestors_si_def) by(auto elim!: bind_returns_result_E2 split: option.splits dest!: returns_result_eq[OF assms(4)] get_ancestors_si_ptr) then show ?thesis using assms child_ancestors get_ancestors_si_subset by blast qed lemma get_ancestors_si_obtains_children_or_shadow_root: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" and "h \ get_ancestors_si ptr \\<^sub>r ancestors" and "ancestor \ ptr" and "ancestor \ set ancestors" shows "((\children ancestor_child. h \ get_child_nodes ancestor \\<^sub>r children \ ancestor_child \ set children \ cast ancestor_child \ set ancestors \ thesis) \ thesis) \ ((\ancestor_element shadow_root. ancestor = cast ancestor_element \ h \ get_shadow_root ancestor_element \\<^sub>r Some shadow_root \ cast shadow_root \ set ancestors \ thesis) \ thesis)" proof (insert assms(4) assms(5) assms(6), induct ptr arbitrary: ancestors rule: heap_wellformed_induct_rev_si[OF assms(1)]) case (1 child) then show ?case proof (cases "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child") case None then obtain shadow_root where shadow_root: "child = cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root" using 1(4) 1(5) 1(6) by(auto simp add: get_ancestors_si_def[of child] elim!: bind_returns_result_E2 split: option.splits) then obtain host where host: "h \ get_host shadow_root \\<^sub>r host" by (metis "1.prems"(1) assms(1) assms(2) assms(3) document_ptr_kinds_commutes get_ancestors_si_ptrs_in_heap is_OK_returns_result_E local.get_ancestors_si_ptr local.get_host_ok shadow_root_ptr_kinds_commutes) then obtain host_ancestors where host_ancestors: "h \ get_ancestors_si (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r host) \\<^sub>r host_ancestors" by (metis "1.prems"(1) assms(1) assms(2) assms(3) get_ancestors_si_also_host get_ancestors_si_ok get_ancestors_si_ptrs_in_heap is_OK_returns_result_E local.get_ancestors_si_ptr shadow_root) then have "ancestors = cast shadow_root # host_ancestors" using 1(4) 1(5) 1(3) None shadow_root host by(auto simp add: get_ancestors_si_def[of child, simplified shadow_root] elim!: bind_returns_result_E2 dest!: returns_result_eq[OF host] split: option.splits) then show ?thesis proof (cases "ancestor = cast host") case True then show ?thesis using "1.prems"(1) host local.get_ancestors_si_ptr local.shadow_root_host_dual shadow_root by blast next case False have "ancestor \ set ancestors" using host host_ancestors 1(3) get_ancestors_si_also_host assms(1) assms(2) assms(3) using "1.prems"(3) by blast then have "((\children ancestor_child. h \ get_child_nodes ancestor \\<^sub>r children \ ancestor_child \ set children \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ancestor_child \ set host_ancestors \ thesis) \ thesis) \ ((\ancestor_element shadow_root. ancestor = cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ancestor_element \ h \ get_shadow_root ancestor_element \\<^sub>r Some shadow_root \ cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root \ set host_ancestors \ thesis) \ thesis)" using "1.hyps"(2) "1.prems"(2) False \ancestors = cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root # host_ancestors\ host host_ancestors shadow_root by auto then show ?thesis using \ancestors = cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root # host_ancestors\ by auto qed next case (Some child_node) then obtain parent where parent: "h \ get_parent child_node \\<^sub>r Some parent" using 1(4) 1(5) 1(6) by(auto simp add: get_ancestors_si_def[of child] elim!: bind_returns_result_E2 split: option.splits) then obtain parent_ancestors where parent_ancestors: "h \ get_ancestors_si parent \\<^sub>r parent_ancestors" by (meson assms(1) assms(2) assms(3) get_ancestors_si_ok is_OK_returns_result_E local.get_parent_parent_in_heap) then have "ancestors = cast child_node # parent_ancestors" using 1(4) 1(5) 1(3) Some by(auto simp add: get_ancestors_si_def[of child, simplified Some] elim!: bind_returns_result_E2 dest!: returns_result_eq[OF parent] split: option.splits) then show ?thesis proof (cases "ancestor = parent") case True then show ?thesis by (metis (no_types, lifting) "1.prems"(1) Some local.get_ancestors_si_ptr local.get_parent_child_dual node_ptr_casts_commute parent) next case False have "ancestor \ set ancestors" by (simp add: "1.prems"(3)) then have "((\children ancestor_child. h \ get_child_nodes ancestor \\<^sub>r children \ ancestor_child \ set children \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ancestor_child \ set parent_ancestors \ thesis) \ thesis) \ ((\ancestor_element shadow_root. ancestor = cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ancestor_element \ h \ get_shadow_root ancestor_element \\<^sub>r Some shadow_root \ cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root \ set parent_ancestors \ thesis) \ thesis)" using "1.hyps"(1) "1.prems"(2) False Some \ancestors = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child_node # parent_ancestors\ parent parent_ancestors by auto then show ?thesis using \ancestors = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child_node # parent_ancestors\ by auto qed qed qed lemma a_host_shadow_root_rel_shadow_root: "h \ get_shadow_root host \\<^sub>r Some shadow_root \ (cast host, cast shadow_root) \ a_host_shadow_root_rel h" by(auto simp add: is_OK_returns_result_I get_shadow_root_ptr_in_heap a_host_shadow_root_rel_def) lemma get_ancestors_si_parent_child_a_host_shadow_root_rel: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "h \ get_ancestors_si child \\<^sub>r ancestors" shows "(ptr, child) \ (parent_child_rel h \ a_host_shadow_root_rel h)\<^sup>* \ ptr \ set ancestors" proof assume "(ptr, child) \ (parent_child_rel h \ local.a_host_shadow_root_rel h)\<^sup>* " then show "ptr \ set ancestors" proof (induct ptr rule: heap_wellformed_induct_si[OF assms(1)]) case (1 ptr) then show ?case proof (cases "ptr = child") case True then show ?thesis using assms(4) local.get_ancestors_si_ptr by blast next case False obtain ptr_child where ptr_child: "(ptr, ptr_child) \ (parent_child_rel h \ local.a_host_shadow_root_rel h) \ (ptr_child, child) \ (parent_child_rel h \ local.a_host_shadow_root_rel h)\<^sup>*" using converse_rtranclE[OF 1(4)] \ptr \ child\ by metis then show ?thesis proof(cases "(ptr, ptr_child) \ parent_child_rel h") case True then obtain ptr_child_node where ptr_child_ptr_child_node: "ptr_child = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node" using ptr_child node_ptr_casts_commute3 CD.parent_child_rel_node_ptr by (metis) then obtain children where children: "h \ get_child_nodes ptr \\<^sub>r children" and ptr_child_node: "ptr_child_node \ set children" proof - assume a1: "\children. \h \ get_child_nodes ptr \\<^sub>r children; ptr_child_node \ set children\ \ thesis" have "ptr |\| object_ptr_kinds h" using CD.parent_child_rel_parent_in_heap True by blast moreover have "ptr_child_node \ set |h \ get_child_nodes ptr|\<^sub>r" by (metis True assms(2) assms(3) calculation local.CD.parent_child_rel_child local.get_child_nodes_ok local.known_ptrs_known_ptr ptr_child_ptr_child_node returns_result_select_result) ultimately show ?thesis using a1 get_child_nodes_ok \type_wf h\ \known_ptrs h\ by (meson local.known_ptrs_known_ptr returns_result_select_result) qed moreover have "(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node, child) \ (parent_child_rel h \ local.a_host_shadow_root_rel h)\<^sup>*" using ptr_child True ptr_child_ptr_child_node by auto ultimately have "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node \ set ancestors" using 1 by auto moreover have "h \ get_parent ptr_child_node \\<^sub>r Some ptr" using assms(1) children ptr_child_node child_parent_dual using \known_ptrs h\ \type_wf h\ by blast ultimately show ?thesis using get_ancestors_si_also_parent assms \type_wf h\ by blast next case False then obtain host where host: "ptr = cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r host" using ptr_child by(auto simp add: a_host_shadow_root_rel_def) then obtain shadow_root where shadow_root: "h \ get_shadow_root host \\<^sub>r Some shadow_root" and ptr_child_shadow_root: "ptr_child = cast shadow_root" using ptr_child False apply(auto simp add: a_host_shadow_root_rel_def)[1] by (metis (no_types, lifting) assms(3) local.get_shadow_root_ok select_result_I) moreover have "(cast shadow_root, child) \ (parent_child_rel h \ local.a_host_shadow_root_rel h)\<^sup>*" using ptr_child ptr_child_shadow_root by blast ultimately have "cast shadow_root \ set ancestors" using "1.hyps"(2) host by blast moreover have "h \ get_host shadow_root \\<^sub>r host" by (metis assms(1) assms(2) assms(3) is_OK_returns_result_E local.get_host_ok local.get_shadow_root_shadow_root_ptr_in_heap local.shadow_root_host_dual local.shadow_root_same_host shadow_root) ultimately show ?thesis using get_ancestors_si_also_host assms(1) assms(2) assms(3) assms(4) host by blast qed qed qed next assume "ptr \ set ancestors" then show "(ptr, child) \ (parent_child_rel h \ local.a_host_shadow_root_rel h)\<^sup>*" proof (induct ptr rule: heap_wellformed_induct_si[OF assms(1)]) case (1 ptr) then show ?case proof (cases "ptr = child") case True then show ?thesis by simp next case False have "\thesis. ((\children ancestor_child. h \ get_child_nodes ptr \\<^sub>r children \ ancestor_child \ set children \ cast ancestor_child \ set ancestors \ thesis) \ thesis) \ ((\ancestor_element shadow_root. ptr = cast ancestor_element \ h \ get_shadow_root ancestor_element \\<^sub>r Some shadow_root \ cast shadow_root \ set ancestors \ thesis) \ thesis)" using "1.prems" False assms(1) assms(2) assms(3) assms(4) get_ancestors_si_obtains_children_or_shadow_root by blast then show ?thesis proof (cases "\thesis. ((\children ancestor_child. h \ get_child_nodes ptr \\<^sub>r children \ ancestor_child \ set children \ cast ancestor_child \ set ancestors \ thesis) \ thesis)") case True then obtain children ancestor_child where "h \ get_child_nodes ptr \\<^sub>r children" and "ancestor_child \ set children" and "cast ancestor_child \ set ancestors" by blast then show ?thesis by (meson "1.hyps"(1) in_rtrancl_UnI local.CD.parent_child_rel_child r_into_rtrancl rtrancl_trans) next case False obtain ancestor_element shadow_root where "ptr = cast ancestor_element" and "h \ get_shadow_root ancestor_element \\<^sub>r Some shadow_root" and "cast shadow_root \ set ancestors" using False \\thesis. ((\children ancestor_child. h \ get_child_nodes ptr \\<^sub>r children \ ancestor_child \ set children \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ancestor_child \ set ancestors \ thesis) \ thesis) \ ((\ancestor_element shadow_root. ptr = cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ancestor_element \ h \ get_shadow_root ancestor_element \\<^sub>r Some shadow_root \ cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root \ set ancestors \ thesis) \ thesis)\ by blast then show ?thesis using 1(2) a_host_shadow_root_rel_shadow_root apply(simp) by (meson Un_iff converse_rtrancl_into_rtrancl) qed qed qed qed lemma get_root_node_si_root_in_heap: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_root_node_si ptr \\<^sub>r root" shows "root |\| object_ptr_kinds h" using assms apply(auto simp add: get_root_node_si_def elim!: bind_returns_result_E2)[1] by (simp add: get_ancestors_si_never_empty get_ancestors_si_ptrs_in_heap) lemma get_root_node_si_same_no_parent: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_root_node_si ptr \\<^sub>r cast child" shows "h \ get_parent child \\<^sub>r None" proof (insert assms(1) assms(4), induct ptr rule: heap_wellformed_induct_rev_si) case (step c) then show ?case proof (cases "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r c") case None then show ?thesis using step(4) by(auto simp add: get_root_node_si_def get_ancestors_si_def[of c] elim!: bind_returns_result_E2 split: if_splits option.splits intro!: step(2) bind_pure_returns_result_I) next case (Some child_node) note s = this then obtain parent_opt where parent_opt: "h \ get_parent child_node \\<^sub>r parent_opt" using step(4) apply(auto simp add: get_root_node_si_def get_ancestors_si_def intro!: bind_pure_I elim!: bind_returns_result_E2)[1] by(auto split: option.splits) then show ?thesis proof(induct parent_opt) case None then show ?case using Some get_root_node_si_no_parent returns_result_eq step.prems by fastforce next case (Some parent) then show ?case using step(4) s apply(auto simp add: get_root_node_si_def get_ancestors_si_def[of c] elim!: bind_returns_result_E2 split: option.splits list.splits if_splits)[1] using assms(1) get_ancestors_si_never_empty apply blast by(auto simp add: get_root_node_si_def dest: returns_result_eq intro!: step(1) bind_pure_returns_result_I) qed qed qed lemma get_root_node_si_parent_child_a_host_shadow_root_rel: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "h \ get_root_node_si ptr \\<^sub>r root" shows "(root, ptr) \ (parent_child_rel h \ a_host_shadow_root_rel h)\<^sup>*" using assms using get_ancestors_si_parent_child_a_host_shadow_root_rel get_ancestors_si_never_empty by(auto simp add: get_root_node_si_def elim!: bind_returns_result_E2 intro!: bind_pure_returns_result_I) end interpretation i_get_root_node_si_wf?: l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_host get_host_locs get_ancestors_si get_ancestors_si_locs get_root_node_si get_root_node_si_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_document get_disconnected_document_locs by(auto simp add: l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_disconnected\_document\ locale l_get_disconnected_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent_wf + l_get_parent begin lemma get_disconnected_document_ok: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_parent node_ptr \\<^sub>r None" shows "h \ ok (get_disconnected_document node_ptr)" proof - have "node_ptr |\| node_ptr_kinds h" by (meson assms(4) is_OK_returns_result_I local.get_parent_ptr_in_heap) have "\(\parent \ fset (object_ptr_kinds h). node_ptr \ set |h \ get_child_nodes parent|\<^sub>r)" apply(auto)[1] using assms(4) child_parent_dual[OF assms(1)] assms(1) assms(2) assms(3) known_ptrs_known_ptr option.simps(3) returns_result_eq returns_result_select_result by (metis (no_types, lifting) CD.get_child_nodes_ok) then have "(\document_ptr \ fset (document_ptr_kinds h). node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r)" using heap_is_wellformed_children_disc_nodes using \node_ptr |\| node_ptr_kinds h\ assms(1) by blast then obtain some_owner_document where "some_owner_document \ set (sorted_list_of_set (fset (document_ptr_kinds h)))" and "node_ptr \ set |h \ get_disconnected_nodes some_owner_document|\<^sub>r" by auto have h5: "\!x. x \ set (sorted_list_of_set (fset (document_ptr_kinds h))) \ h \ Heap_Error_Monad.bind (get_disconnected_nodes x) (\children. return (node_ptr \ set children)) \\<^sub>r True" apply(auto intro!: bind_pure_returns_result_I)[1] apply (smt CD.get_disconnected_nodes_ok CD.get_disconnected_nodes_pure \\document_ptr\fset (document_ptr_kinds h). node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r\ assms(2) bind_pure_returns_result_I2 notin_fset return_returns_result select_result_I2) apply(auto elim!: bind_returns_result_E2 intro!: bind_pure_returns_result_I)[1] using heap_is_wellformed_one_disc_parent assms(1) by blast let ?filter_M = "filter_M (\document_ptr. Heap_Error_Monad.bind (get_disconnected_nodes document_ptr) (\disconnected_nodes. return (node_ptr \ set disconnected_nodes))) (sorted_list_of_set (fset (document_ptr_kinds h)))" have "h \ ok (?filter_M)" using CD.get_disconnected_nodes_ok by (smt CD.get_disconnected_nodes_pure DocumentMonad.ptr_kinds_M_ptr_kinds DocumentMonad.ptr_kinds_ptr_kinds_M assms(2) bind_is_OK_pure_I bind_pure_I document_ptr_kinds_M_def filter_M_is_OK_I l_ptr_kinds_M.ptr_kinds_M_ok return_ok return_pure returns_result_select_result) then obtain candidates where candidates: "h \ filter_M (\document_ptr. Heap_Error_Monad.bind (get_disconnected_nodes document_ptr) (\disconnected_nodes. return (node_ptr \ set disconnected_nodes))) (sorted_list_of_set (fset (document_ptr_kinds h))) \\<^sub>r candidates" by auto have "candidates = [some_owner_document]" apply(rule filter_M_ex1[OF candidates \some_owner_document \ set (sorted_list_of_set (fset (document_ptr_kinds h)))\ h5]) using \node_ptr \ set |h \ get_disconnected_nodes some_owner_document|\<^sub>r\ \some_owner_document \ set (sorted_list_of_set (fset (document_ptr_kinds h)))\ by(auto simp add: CD.get_disconnected_nodes_ok assms(2) intro!: bind_pure_I intro!: bind_pure_returns_result_I) then show ?thesis using candidates \node_ptr |\| node_ptr_kinds h\ apply(auto simp add: get_disconnected_document_def intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I split: list.splits)[1] apply (meson not_Cons_self2 returns_result_eq) by (meson list.distinct(1) list.inject returns_result_eq) qed end interpretation i_get_disconnected_document_wf?: l_get_disconnected_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs known_ptrs get_parent get_parent_locs by(auto simp add: l_get_disconnected_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_disconnected_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_ancestors\_di\ locale l_get_ancestors_di_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_ancestors_di\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent_wf + l_get_parent_get_host_wf + l_get_host_wf + l_get_disconnected_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent_get_host_get_disconnected_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_ancestors_di_ok: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" and "ptr |\| object_ptr_kinds h" shows "h \ ok (get_ancestors_di ptr)" proof (insert assms(1) assms(4), induct rule: heap_wellformed_induct_rev_si) case (step child) then show ?case using assms(2) assms(3) apply(auto simp add: get_ancestors_di_def[of child] assms(1) get_parent_parent_in_heap intro!: bind_is_OK_pure_I bind_pure_I split: option.splits)[1] using local.get_parent_ok apply blast using assms(1) get_disconnected_document_ok apply blast apply(simp add: get_ancestors_di_def ) apply(auto intro!: bind_is_OK_pure_I split: option.splits)[1] apply (metis (no_types, lifting) bind_is_OK_E document_ptr_kinds_commutes is_OK_returns_heap_I local.get_ancestors_di_def local.get_disconnected_document_disconnected_document_in_heap step.hyps(3)) apply (metis (no_types, lifting) bind_is_OK_E document_ptr_kinds_commutes is_OK_returns_heap_I local.get_ancestors_di_def local.get_disconnected_document_disconnected_document_in_heap step.hyps(3)) using assms(1) local.get_disconnected_document_disconnected_document_in_heap local.get_host_ok shadow_root_ptr_kinds_commutes apply blast apply (smt assms(1) bind_returns_heap_E document_ptr_casts_commute2 is_OK_returns_heap_E is_OK_returns_heap_I l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M.shadow_root_same_host local.get_disconnected_document_disconnected_document_in_heap local.get_host_pure local.l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms local.shadow_root_host_dual option.simps(4) option.simps(5) pure_returns_heap_eq shadow_root_ptr_casts_commute2) using get_host_ok assms(1) apply blast by (meson assms(1) is_OK_returns_result_I local.get_shadow_root_ptr_in_heap local.shadow_root_host_dual) qed lemma get_ancestors_di_ptrs_in_heap: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_ancestors_di ptr \\<^sub>r ancestors" assumes "ptr' \ set ancestors" shows "ptr' |\| object_ptr_kinds h" proof (insert assms(4) assms(5), induct ancestors arbitrary: ptr) case Nil then show ?case by(auto) next case (Cons a ancestors) then obtain x where x: "h \ get_ancestors_di x \\<^sub>r a # ancestors" by(auto simp add: get_ancestors_di_def[of a] elim!: bind_returns_result_E2 split: option.splits) then have "x = a" by(auto simp add: get_ancestors_di_def[of x] intro!: bind_pure_I elim!: bind_returns_result_E2 split: option.splits) then show ?case proof (cases "ptr' = a") case True then show ?thesis using Cons.hyps Cons.prems(2) get_ancestors_di_ptr_in_heap x using \x = a\ by blast next case False obtain ptr'' where ptr'': "h \ get_ancestors_di ptr'' \\<^sub>r ancestors" using \ h \ get_ancestors_di x \\<^sub>r a # ancestors\ Cons.prems(2) False apply(auto simp add: get_ancestors_di_def elim!: bind_returns_result_E2)[1] apply(auto elim!: bind_returns_result_E2 split: option.splits intro!: bind_pure_I)[1] apply(auto elim!: bind_returns_result_E2 split: option.splits intro!: bind_pure_I)[1] apply (metis (no_types, lifting) local.get_ancestors_di_def) apply (metis (no_types, lifting) local.get_ancestors_di_def) by (simp add: local.get_ancestors_di_def) then show ?thesis using Cons.hyps Cons.prems(2) False by auto qed qed lemma get_ancestors_di_reads: assumes "heap_is_wellformed h" shows "reads get_ancestors_di_locs (get_ancestors_di node_ptr) h h'" proof (insert assms(1), induct rule: heap_wellformed_induct_rev_si) case (step child) then show ?case using (* [[simproc del: Product_Type.unit_eq]] *) get_parent_reads[unfolded reads_def] get_host_reads[unfolded reads_def] get_disconnected_document_reads[unfolded reads_def] apply(auto simp add: get_ancestors_di_def[of child])[1] by(auto simp add: get_ancestors_di_locs_def get_parent_reads_pointers intro!: bind_pure_I reads_bind_pure reads_subset[OF check_in_heap_reads] reads_subset[OF return_reads] reads_subset[OF get_parent_reads] reads_subset[OF get_child_nodes_reads] reads_subset[OF get_host_reads] reads_subset[OF get_disconnected_document_reads] split: option.splits list.splits ) qed lemma get_ancestors_di_subset: assumes "heap_is_wellformed h" and "h \ get_ancestors_di ptr \\<^sub>r ancestors" and "ancestor \ set ancestors" and "h \ get_ancestors_di ancestor \\<^sub>r ancestor_ancestors" and type_wf: "type_wf h" and known_ptrs: "known_ptrs h" shows "set ancestor_ancestors \ set ancestors" proof (insert assms(1) assms(2) assms(3), induct ptr arbitrary: ancestors rule: heap_wellformed_induct_rev_si) case (step child) have "child |\| object_ptr_kinds h" using get_ancestors_di_ptr_in_heap step(4) by auto (* then have "h \ check_in_heap child \\<^sub>r ()" using returns_result_select_result by force *) obtain tl_ancestors where tl_ancestors: "ancestors = child # tl_ancestors" using step(4) by(auto simp add: get_ancestors_di_def[of child] intro!: bind_pure_I elim!: bind_returns_result_E2 split: option.splits) show ?case proof (induct "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child") case None show ?case proof (induct "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child") case None then show ?case using step(4) \None = cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child\ apply(auto simp add: get_ancestors_di_def[of child] elim!: bind_returns_result_E2)[1] by (metis (no_types, lifting) assms(4) empty_iff empty_set select_result_I2 set_ConsD step.prems(1) step.prems(2)) next case (Some shadow_root_child) then have "cast shadow_root_child |\| document_ptr_kinds h" using \child |\| object_ptr_kinds h\ apply(auto simp add: document_ptr_kinds_def split: option.splits)[1] by (metis (mono_tags, lifting) document_ptr_casts_commute3 document_ptr_kinds_commutes document_ptr_kinds_def fset.map_comp shadow_root_ptr_casts_commute) then have "shadow_root_child |\| shadow_root_ptr_kinds h" using shadow_root_ptr_kinds_commutes by blast obtain host where host: "h \ get_host shadow_root_child \\<^sub>r host" using get_host_ok assms by (meson \shadow_root_child |\| shadow_root_ptr_kinds h\ is_OK_returns_result_E) then have "h \ get_ancestors_di (cast host) \\<^sub>r tl_ancestors" using Some step(4) tl_ancestors None by(auto simp add: get_ancestors_di_def[of child] intro!: bind_pure_returns_result_I elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq) then show ?case using step(2) Some host step(5) tl_ancestors using assms(4) dual_order.trans eq_iff returns_result_eq set_ConsD set_subset_Cons shadow_root_ptr_casts_commute document_ptr_casts_commute step.prems(1) by (smt case_optionE local.shadow_root_host_dual option.case_distrib option.distinct(1)) qed next case (Some child_node) note s1 = Some obtain parent_opt where parent_opt: "h \ get_parent child_node \\<^sub>r parent_opt" using \child |\| object_ptr_kinds h\ assms(1) Some[symmetric] get_parent_ok[OF type_wf known_ptrs] by (metis (no_types, lifting) is_OK_returns_result_E known_ptrs get_parent_ok l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms node_ptr_casts_commute node_ptr_kinds_commutes) then show ?case proof (induct parent_opt) case None then obtain disc_doc where disc_doc: "h \ get_disconnected_document child_node \\<^sub>r disc_doc" and "h \ get_ancestors_di (cast disc_doc) \\<^sub>r tl_ancestors" using step(4) s1 tl_ancestors apply(simp add: get_ancestors_di_def[of child]) by(auto elim!: bind_returns_result_E2 intro!: bind_pure_I split: option.splits dest: returns_result_eq) then show ?thesis using step(3) step(4) step(5) by (metis (no_types, lifting) assms(4) dual_order.trans eq_iff node_ptr_casts_commute s1 select_result_I2 set_ConsD set_subset_Cons tl_ancestors) next case (Some parent) then have "h \ get_ancestors_di parent \\<^sub>r tl_ancestors" using s1 tl_ancestors step(4) by(auto simp add: get_ancestors_di_def[of child] elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq) show ?case by (metis (no_types, lifting) Some.prems \h \ get_ancestors_di parent \\<^sub>r tl_ancestors\ assms(4) eq_iff node_ptr_casts_commute order_trans s1 select_result_I2 set_ConsD set_subset_Cons step.hyps(1) step.prems(1) step.prems(2) tl_ancestors) qed qed qed lemma get_ancestors_di_also_parent: assumes "heap_is_wellformed h" and "h \ get_ancestors_di some_ptr \\<^sub>r ancestors" and "cast child \ set ancestors" and "h \ get_parent child \\<^sub>r Some parent" and type_wf: "type_wf h" and known_ptrs: "known_ptrs h" shows "parent \ set ancestors" proof - obtain child_ancestors where child_ancestors: "h \ get_ancestors_di (cast child) \\<^sub>r child_ancestors" by (meson assms(1) assms(4) get_ancestors_di_ok is_OK_returns_result_I known_ptrs local.get_parent_ptr_in_heap node_ptr_kinds_commutes returns_result_select_result type_wf) then have "parent \ set child_ancestors" apply(simp add: get_ancestors_di_def) by(auto elim!: bind_returns_result_E2 split: option.splits dest!: returns_result_eq[OF assms(4)] get_ancestors_di_ptr) then show ?thesis using assms child_ancestors get_ancestors_di_subset by blast qed lemma get_ancestors_di_also_host: assumes "heap_is_wellformed h" and "h \ get_ancestors_di some_ptr \\<^sub>r ancestors" and "cast shadow_root \ set ancestors" and "h \ get_host shadow_root \\<^sub>r host" and type_wf: "type_wf h" and known_ptrs: "known_ptrs h" shows "cast host \ set ancestors" proof - obtain child_ancestors where child_ancestors: "h \ get_ancestors_di (cast shadow_root) \\<^sub>r child_ancestors" by (meson assms(1) assms(2) assms(3) get_ancestors_di_ok get_ancestors_di_ptrs_in_heap is_OK_returns_result_E known_ptrs type_wf) then have "cast host \ set child_ancestors" apply(simp add: get_ancestors_di_def) by(auto elim!: bind_returns_result_E2 split: option.splits dest!: returns_result_eq[OF assms(4)] get_ancestors_di_ptr) then show ?thesis using assms child_ancestors get_ancestors_di_subset by blast qed lemma get_ancestors_di_also_disconnected_document: assumes "heap_is_wellformed h" and "h \ get_ancestors_di some_ptr \\<^sub>r ancestors" and "cast disc_node \ set ancestors" and "h \ get_disconnected_document disc_node \\<^sub>r disconnected_document" and type_wf: "type_wf h" and known_ptrs: "known_ptrs h" and "h \ get_parent disc_node \\<^sub>r None" shows "cast disconnected_document \ set ancestors" proof - obtain child_ancestors where child_ancestors: "h \ get_ancestors_di (cast disc_node) \\<^sub>r child_ancestors" by (meson assms(1) assms(2) assms(3) get_ancestors_di_ok get_ancestors_di_ptrs_in_heap is_OK_returns_result_E known_ptrs type_wf) then have "cast disconnected_document \ set child_ancestors" apply(simp add: get_ancestors_di_def) by(auto elim!: bind_returns_result_E2 intro!: bind_pure_I split: option.splits dest!: returns_result_eq[OF assms(4)] returns_result_eq[OF assms(7)] get_ancestors_di_ptr) then show ?thesis using assms child_ancestors get_ancestors_di_subset by blast qed lemma disc_node_disc_doc_dual: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_parent node_ptr \\<^sub>r None" assumes "h \ get_disconnected_nodes disc_doc \\<^sub>r disc_nodes" assumes "node_ptr \ set disc_nodes" shows "h \ get_disconnected_document node_ptr \\<^sub>r disc_doc" proof - have "node_ptr |\| node_ptr_kinds h" by (meson assms(4) is_OK_returns_result_I local.get_parent_ptr_in_heap) then have "\(\parent \ fset (object_ptr_kinds h). node_ptr \ set |h \ get_child_nodes parent|\<^sub>r)" apply(auto)[1] using child_parent_dual[OF assms(1)] assms known_ptrs_known_ptr option.simps(3) returns_result_eq returns_result_select_result by (metis (no_types, lifting) get_child_nodes_ok) then have "(\document_ptr \ fset (document_ptr_kinds h). node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r)" using heap_is_wellformed_children_disc_nodes using \node_ptr |\| node_ptr_kinds h\ assms(1) by blast (* then obtain some_owner_document where "some_owner_document \ set (sorted_list_of_set (fset (document_ptr_kinds h)))" and "node_ptr \ set |h \ get_disconnected_nodes some_owner_document|\<^sub>r" by auto *) then have "disc_doc \ set (sorted_list_of_set (fset (document_ptr_kinds h)))" and "node_ptr \ set |h \ get_disconnected_nodes disc_doc|\<^sub>r" using CD.get_disconnected_nodes_ptr_in_heap assms(5) assms(6) by auto have h5: "\!x. x \ set (sorted_list_of_set (fset (document_ptr_kinds h))) \ h \ Heap_Error_Monad.bind (get_disconnected_nodes x) (\children. return (node_ptr \ set children)) \\<^sub>r True" apply(auto intro!: bind_pure_returns_result_I)[1] apply (smt CD.get_disconnected_nodes_ok CD.get_disconnected_nodes_pure \\document_ptr\fset (document_ptr_kinds h). node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r\ assms(2) bind_pure_returns_result_I2 notin_fset return_returns_result select_result_I2) apply(auto elim!: bind_returns_result_E2 intro!: bind_pure_returns_result_I)[1] using heap_is_wellformed_one_disc_parent assms(1) by blast let ?filter_M = "filter_M (\document_ptr. Heap_Error_Monad.bind (get_disconnected_nodes document_ptr) (\disconnected_nodes. return (node_ptr \ set disconnected_nodes))) (sorted_list_of_set (fset (document_ptr_kinds h)))" have "h \ ok (?filter_M)" using CD.get_disconnected_nodes_ok by (smt CD.get_disconnected_nodes_pure DocumentMonad.ptr_kinds_M_ptr_kinds DocumentMonad.ptr_kinds_ptr_kinds_M assms(2) bind_is_OK_pure_I bind_pure_I document_ptr_kinds_M_def filter_M_is_OK_I l_ptr_kinds_M.ptr_kinds_M_ok return_ok return_pure returns_result_select_result) then obtain candidates where candidates: "h \ ?filter_M \\<^sub>r candidates" by auto have "candidates = [disc_doc]" apply(rule filter_M_ex1[OF candidates \disc_doc \ set (sorted_list_of_set (fset (document_ptr_kinds h)))\ h5]) using \node_ptr \ set |h \ get_disconnected_nodes disc_doc|\<^sub>r\ \disc_doc \ set (sorted_list_of_set (fset (document_ptr_kinds h)))\ by(auto simp add: CD.get_disconnected_nodes_ok assms(2) intro!: bind_pure_I intro!: bind_pure_returns_result_I) then show ?thesis using \node_ptr |\| node_ptr_kinds h\ candidates by(auto simp add: bind_pure_I get_disconnected_document_def elim!: bind_returns_result_E2 intro!: bind_pure_returns_result_I filter_M_pure_I) qed lemma get_ancestors_di_obtains_children_or_shadow_root_or_disconnected_node: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" and "h \ get_ancestors_di ptr \\<^sub>r ancestors" and "ancestor \ ptr" and "ancestor \ set ancestors" shows "((\children ancestor_child. h \ get_child_nodes ancestor \\<^sub>r children \ ancestor_child \ set children \ cast ancestor_child \ set ancestors \ thesis) \ thesis) \ ((\ancestor_element shadow_root. ancestor = cast ancestor_element \ h \ get_shadow_root ancestor_element \\<^sub>r Some shadow_root \ cast shadow_root \ set ancestors \ thesis) \ thesis) \ ((\disc_doc disc_nodes disc_node. ancestor = cast disc_doc \ h \ get_disconnected_nodes disc_doc \\<^sub>r disc_nodes \ disc_node \ set disc_nodes \ cast disc_node \ set ancestors \ thesis) \ thesis)" proof (insert assms(4) assms(5) assms(6), induct ptr arbitrary: ancestors rule: heap_wellformed_induct_rev_si[OF assms(1)]) case (1 child) then show ?case proof (cases "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child") case None then obtain shadow_root where shadow_root: "child = cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root" using 1(4) 1(5) 1(6) by(auto simp add: get_ancestors_di_def[of child] elim!: bind_returns_result_E2 split: option.splits) then obtain host where host: "h \ get_host shadow_root \\<^sub>r host" by (metis "1.prems"(1) assms(1) assms(2) assms(3) document_ptr_kinds_commutes get_ancestors_di_ptrs_in_heap is_OK_returns_result_E local.get_ancestors_di_ptr local.get_host_ok shadow_root_ptr_kinds_commutes) then obtain host_ancestors where host_ancestors: "h \ get_ancestors_di (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r host) \\<^sub>r host_ancestors" by (metis "1.prems"(1) assms(1) assms(2) assms(3) get_ancestors_di_also_host get_ancestors_di_ok get_ancestors_di_ptrs_in_heap is_OK_returns_result_E local.get_ancestors_di_ptr shadow_root) then have "ancestors = cast shadow_root # host_ancestors" using 1(4) 1(5) 1(3) None shadow_root host by(auto simp add: get_ancestors_di_def[of child, simplified shadow_root] elim!: bind_returns_result_E2 dest!: returns_result_eq[OF host] split: option.splits) then show ?thesis proof (cases "ancestor = cast host") case True then show ?thesis using "1.prems"(1) host local.get_ancestors_di_ptr local.shadow_root_host_dual shadow_root by blast next case False have "ancestor \ set ancestors" using host host_ancestors 1(3) get_ancestors_di_also_host assms(1) assms(2) assms(3) using "1.prems"(3) by blast then have "((\children ancestor_child. h \ get_child_nodes ancestor \\<^sub>r children \ ancestor_child \ set children \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ancestor_child \ set host_ancestors \ thesis) \ thesis) \ ((\ancestor_element shadow_root. ancestor = cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ancestor_element \ h \ get_shadow_root ancestor_element \\<^sub>r Some shadow_root \ cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root \ set host_ancestors \ thesis) \ thesis) \ ((\disc_doc disc_nodes disc_node. ancestor = cast disc_doc \ h \ get_disconnected_nodes disc_doc \\<^sub>r disc_nodes \ disc_node \ set disc_nodes \ cast disc_node \ set host_ancestors \ thesis) \ thesis)" using "1.hyps"(2) "1.prems"(2) False \ancestors = cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root # host_ancestors\ host host_ancestors shadow_root by auto then show ?thesis using \ancestors = cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root # host_ancestors\ by auto qed next case (Some child_node) then obtain parent_opt where parent_opt: "h \ get_parent child_node \\<^sub>r parent_opt" by (metis (no_types, lifting) "1.prems"(1) assms(1) assms(2) assms(3) get_ancestors_di_ptrs_in_heap is_OK_returns_result_E local.get_ancestors_di_ptr local.get_parent_ok node_ptr_casts_commute node_ptr_kinds_commutes) then show ?thesis proof (induct parent_opt) case None then obtain disc_doc where disc_doc: "h \ get_disconnected_document child_node \\<^sub>r disc_doc" by (meson assms(1) assms(2) assms(3) is_OK_returns_result_E local.get_disconnected_document_ok) then obtain parent_ancestors where parent_ancestors: "h \ get_ancestors_di (cast disc_doc) \\<^sub>r parent_ancestors" by (meson assms(1) assms(2) assms(3) document_ptr_kinds_commutes is_OK_returns_result_E l_get_ancestors_di_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_di_ok l_get_ancestors_di_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms local.get_disconnected_document_disconnected_document_in_heap) then have "ancestors = cast child_node # parent_ancestors" using 1(4) 1(5) 1(6) Some \cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child = Some child_node\ apply(auto simp add: get_ancestors_di_def[of child, simplified \cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child = Some child_node\] intro!: bind_pure_I elim!: bind_returns_result_E2 dest!: returns_result_eq[OF disc_doc] split: option.splits)[1] apply (simp add: returns_result_eq) by (meson None.prems option.distinct(1) returns_result_eq) then show ?thesis proof (cases "ancestor = cast disc_doc") case True then show ?thesis by (metis \ancestors = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child_node # parent_ancestors\ disc_doc list.set_intros(1) local.disc_doc_disc_node_dual) next case False have "ancestor \ set ancestors" by (simp add: "1.prems"(3)) then have "((\children ancestor_child. h \ get_child_nodes ancestor \\<^sub>r children \ ancestor_child \ set children \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ancestor_child \ set parent_ancestors \ thesis) \ thesis) \ ((\ancestor_element shadow_root. ancestor = cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ancestor_element \ h \ get_shadow_root ancestor_element \\<^sub>r Some shadow_root \ cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root \ set parent_ancestors \ thesis) \ thesis) \ ((\disc_doc disc_nodes disc_node. ancestor = cast disc_doc \ h \ get_disconnected_nodes disc_doc \\<^sub>r disc_nodes \ disc_node \ set disc_nodes \ cast disc_node \ set parent_ancestors \ thesis) \ thesis)" using "1.hyps"(3) "1.prems"(2) False Some \cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child = Some child_node\ \ancestors = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child_node # parent_ancestors\ disc_doc parent_ancestors by auto then show ?thesis using \ancestors = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child_node # parent_ancestors\ by auto qed next case (Some option) then obtain parent where parent: "h \ get_parent child_node \\<^sub>r Some parent" using 1(4) 1(5) 1(6) by(auto simp add: get_ancestors_di_def[of child] intro!: bind_pure_I elim!: bind_returns_result_E2 split: option.splits) then obtain parent_ancestors where parent_ancestors: "h \ get_ancestors_di parent \\<^sub>r parent_ancestors" by (meson assms(1) assms(2) assms(3) get_ancestors_di_ok is_OK_returns_result_E local.get_parent_parent_in_heap) then have "ancestors = cast child_node # parent_ancestors" using 1(4) 1(5) 1(6) Some \cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child = Some child_node\ by(auto simp add: get_ancestors_di_def[of child, simplified \cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child = Some child_node\] dest!: elim!: bind_returns_result_E2 dest!: returns_result_eq[OF parent] split: option.splits) then show ?thesis proof (cases "ancestor = parent") case True then show ?thesis by (metis \ancestors = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child_node # parent_ancestors\ list.set_intros(1) local.get_parent_child_dual parent) next case False have "ancestor \ set ancestors" by (simp add: "1.prems"(3)) then have "((\children ancestor_child. h \ get_child_nodes ancestor \\<^sub>r children \ ancestor_child \ set children \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ancestor_child \ set parent_ancestors \ thesis) \ thesis) \ ((\ancestor_element shadow_root. ancestor = cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ancestor_element \ h \ get_shadow_root ancestor_element \\<^sub>r Some shadow_root \ cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root \ set parent_ancestors \ thesis) \ thesis) \ ((\disc_doc disc_nodes disc_node. ancestor = cast disc_doc \ h \ get_disconnected_nodes disc_doc \\<^sub>r disc_nodes \ disc_node \ set disc_nodes \ cast disc_node \ set parent_ancestors \ thesis) \ thesis)" using "1.hyps"(1) "1.prems"(2) False Some \cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child = Some child_node\ \ancestors = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child_node # parent_ancestors\ parent parent_ancestors by auto then show ?thesis using \ancestors = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child_node # parent_ancestors\ by auto qed qed qed qed lemma get_ancestors_di_parent_child_a_host_shadow_root_rel: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "h \ get_ancestors_di child \\<^sub>r ancestors" shows "(ptr, child) \ (parent_child_rel h \ a_host_shadow_root_rel h \ a_ptr_disconnected_node_rel h)\<^sup>* \ ptr \ set ancestors" proof assume "(ptr, child) \ (parent_child_rel h \ local.a_host_shadow_root_rel h \ a_ptr_disconnected_node_rel h)\<^sup>* " then show "ptr \ set ancestors" proof (induct ptr rule: heap_wellformed_induct_si[OF assms(1)]) case (1 ptr) then show ?case proof (cases "ptr = child") case True then show ?thesis using assms(4) get_ancestors_di_ptr by blast next case False obtain ptr_child where ptr_child: "(ptr, ptr_child) \ (parent_child_rel h \ local.a_host_shadow_root_rel h \ a_ptr_disconnected_node_rel h) \ (ptr_child, child) \ (parent_child_rel h \ local.a_host_shadow_root_rel h \ a_ptr_disconnected_node_rel h)\<^sup>*" using converse_rtranclE[OF 1(4)] \ptr \ child\ by metis then show ?thesis proof(cases "(ptr, ptr_child) \ parent_child_rel h") case True then obtain ptr_child_node where ptr_child_ptr_child_node: "ptr_child = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node" using ptr_child node_ptr_casts_commute3 CD.parent_child_rel_node_ptr by (metis) then obtain children where children: "h \ get_child_nodes ptr \\<^sub>r children" and ptr_child_node: "ptr_child_node \ set children" proof - assume a1: "\children. \h \ get_child_nodes ptr \\<^sub>r children; ptr_child_node \ set children\ \ thesis" have "ptr |\| object_ptr_kinds h" using CD.parent_child_rel_parent_in_heap True by blast moreover have "ptr_child_node \ set |h \ get_child_nodes ptr|\<^sub>r" by (metis True assms(2) assms(3) calculation local.CD.parent_child_rel_child local.get_child_nodes_ok local.known_ptrs_known_ptr ptr_child_ptr_child_node returns_result_select_result) ultimately show ?thesis using a1 get_child_nodes_ok \type_wf h\ \known_ptrs h\ by (meson local.known_ptrs_known_ptr returns_result_select_result) qed moreover have "(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node, child) \ (parent_child_rel h \ local.a_host_shadow_root_rel h \ a_ptr_disconnected_node_rel h)\<^sup>*" using ptr_child True ptr_child_ptr_child_node by auto ultimately have "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node \ set ancestors" using 1 by auto moreover have "h \ get_parent ptr_child_node \\<^sub>r Some ptr" using assms(1) children ptr_child_node child_parent_dual using \known_ptrs h\ \type_wf h\ by blast ultimately show ?thesis using get_ancestors_di_also_parent assms \type_wf h\ by blast next case False then show ?thesis proof (cases "(ptr, ptr_child) \ a_host_shadow_root_rel h") case True then obtain host where host: "ptr = cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r host" using ptr_child by(auto simp add: a_host_shadow_root_rel_def) then obtain shadow_root where shadow_root: "h \ get_shadow_root host \\<^sub>r Some shadow_root" and ptr_child_shadow_root: "ptr_child = cast shadow_root" using False True apply(auto simp add: a_host_shadow_root_rel_def)[1] by (metis (no_types, lifting) assms(3) local.get_shadow_root_ok select_result_I) moreover have "(cast shadow_root, child) \ (parent_child_rel h \ local.a_host_shadow_root_rel h \ a_ptr_disconnected_node_rel h)\<^sup>*" using ptr_child ptr_child_shadow_root by blast ultimately have "cast shadow_root \ set ancestors" using "1.hyps"(2) host by blast moreover have "h \ get_host shadow_root \\<^sub>r host" by (metis assms(1) assms(2) assms(3) is_OK_returns_result_E local.get_host_ok local.get_shadow_root_shadow_root_ptr_in_heap local.shadow_root_host_dual local.shadow_root_same_host shadow_root) ultimately show ?thesis using get_ancestors_di_also_host assms(1) assms(2) assms(3) assms(4) host by blast next case False then obtain disc_doc where disc_doc: "ptr = cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_doc" using ptr_child \(ptr, ptr_child) \ parent_child_rel h\ by(auto simp add: a_ptr_disconnected_node_rel_def) then obtain disc_node disc_nodes where disc_nodes: "h \ get_disconnected_nodes disc_doc \\<^sub>r disc_nodes" and disc_node: "disc_node \ set disc_nodes" and ptr_child_disc_node: "ptr_child = cast disc_node" using False \(ptr, ptr_child) \ parent_child_rel h\ ptr_child apply(auto simp add: a_ptr_disconnected_node_rel_def)[1] by (metis (no_types, lifting) CD.get_disconnected_nodes_ok assms(3) returns_result_select_result) moreover have "(cast disc_node, child) \ (parent_child_rel h \ local.a_host_shadow_root_rel h \ a_ptr_disconnected_node_rel h)\<^sup>*" using ptr_child ptr_child_disc_node by blast ultimately have "cast disc_node \ set ancestors" using "1.hyps"(3) disc_doc by blast moreover have "h \ get_parent disc_node \\<^sub>r None" using \(ptr, ptr_child) \ parent_child_rel h\ apply(auto simp add: parent_child_rel_def ptr_child_disc_node)[1] by (smt assms(1) assms(2) assms(3) assms(4) calculation disc_node disc_nodes get_ancestors_di_ptrs_in_heap is_OK_returns_result_E local.CD.a_heap_is_wellformed_def local.CD.distinct_lists_no_parent local.CD.heap_is_wellformed_impl local.get_parent_child_dual local.get_parent_ok local.get_parent_parent_in_heap local.heap_is_wellformed_def node_ptr_kinds_commutes select_result_I2 split_option_ex) then have "h \ get_disconnected_document disc_node \\<^sub>r disc_doc" using disc_node_disc_doc_dual disc_nodes disc_node assms(1) assms(2) assms(3) by blast ultimately show ?thesis using \h \ get_parent disc_node \\<^sub>r None\ assms(1) assms(2) assms(3) assms(4) disc_doc get_ancestors_di_also_disconnected_document by blast qed qed qed qed next assume "ptr \ set ancestors" then show "(ptr, child) \ (parent_child_rel h \ local.a_host_shadow_root_rel h \ a_ptr_disconnected_node_rel h)\<^sup>*" proof (induct ptr rule: heap_wellformed_induct_si[OF assms(1)]) case (1 ptr) then show ?case proof (cases "ptr = child") case True then show ?thesis by simp next case False have 2: "\thesis. ((\children ancestor_child. h \ get_child_nodes ptr \\<^sub>r children \ ancestor_child \ set children \ cast ancestor_child \ set ancestors \ thesis) \ thesis) \ ((\ancestor_element shadow_root. ptr = cast ancestor_element \ h \ get_shadow_root ancestor_element \\<^sub>r Some shadow_root \ cast shadow_root \ set ancestors \ thesis) \ thesis) \ ((\disc_doc disc_nodes disc_node. ptr = cast disc_doc \ h \ get_disconnected_nodes disc_doc \\<^sub>r disc_nodes \ disc_node \ set disc_nodes \ cast disc_node \ set ancestors \ thesis) \ thesis)" using "1.prems" False assms(1) assms(2) assms(3) assms(4) get_ancestors_di_obtains_children_or_shadow_root_or_disconnected_node by blast then show ?thesis proof (cases "\thesis. ((\children ancestor_child. h \ get_child_nodes ptr \\<^sub>r children \ ancestor_child \ set children \ cast ancestor_child \ set ancestors \ thesis) \ thesis)") case True then obtain children ancestor_child where "h \ get_child_nodes ptr \\<^sub>r children" and "ancestor_child \ set children" and "cast ancestor_child \ set ancestors" by blast then show ?thesis by (meson "1.hyps"(1) in_rtrancl_UnI local.CD.parent_child_rel_child r_into_rtrancl rtrancl_trans) next case False note f1 = this then show ?thesis proof (cases "\thesis. ((\disc_doc disc_nodes disc_node. ptr = cast disc_doc \ h \ get_disconnected_nodes disc_doc \\<^sub>r disc_nodes \ disc_node \ set disc_nodes \ cast disc_node \ set ancestors \ thesis) \ thesis)") case True then obtain disc_doc disc_nodes disc_node where "ptr = cast disc_doc" and "h \ get_disconnected_nodes disc_doc \\<^sub>r disc_nodes" and "disc_node \ set disc_nodes" and "cast disc_node \ set ancestors" by blast then show ?thesis by (meson "1.hyps"(3) in_rtrancl_UnI local.a_ptr_disconnected_node_rel_disconnected_node r_into_rtrancl rtrancl_trans) next case False then obtain ancestor_element shadow_root where "ptr = cast ancestor_element" and "h \ get_shadow_root ancestor_element \\<^sub>r Some shadow_root" and "cast shadow_root \ set ancestors" using f1 2 by smt then show ?thesis by (meson "1.hyps"(2) in_rtrancl_UnI local.a_host_shadow_root_rel_shadow_root r_into_rtrancl rtrancl_trans) qed qed qed qed qed end interpretation i_get_ancestors_di_wf?: l_get_ancestors_di_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_host get_host_locs get_disconnected_document get_disconnected_document_locs get_ancestors_di get_ancestors_di_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M by(auto simp add: l_get_ancestors_di_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_ancestors_di_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_owner\_document\ locale l_get_owner_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_disconnected_nodes + l_get_child_nodes + l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent_wf + l_known_ptrs + l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + assumes known_ptr_impl: "known_ptr = ShadowRootClass.known_ptr" begin lemma get_owner_document_disconnected_nodes: assumes "heap_is_wellformed h" assumes "h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes" assumes "node_ptr \ set disc_nodes" assumes known_ptrs: "known_ptrs h" assumes type_wf: "type_wf h" shows "h \ get_owner_document (cast node_ptr) \\<^sub>r document_ptr" proof - have 2: "node_ptr |\| node_ptr_kinds h" using assms apply(auto simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.a_all_ptrs_in_heap_def)[1] using assms(1) local.heap_is_wellformed_disc_nodes_in_heap by blast have 3: "document_ptr |\| document_ptr_kinds h" using assms(2) get_disconnected_nodes_ptr_in_heap by blast then have 4: "\(\parent_ptr. parent_ptr |\| object_ptr_kinds h \ node_ptr \ set |h \ get_child_nodes parent_ptr|\<^sub>r)" using CD.distinct_lists_no_parent assms unfolding heap_is_wellformed_def CD.heap_is_wellformed_def by simp moreover have "(\document_ptr. document_ptr |\| document_ptr_kinds h \ node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r) \ (\parent_ptr. parent_ptr |\| object_ptr_kinds h \ node_ptr \ set |h \ get_child_nodes parent_ptr|\<^sub>r)" using assms(1) 2 "3" assms(2) assms(3) by auto ultimately have 0: "\!document_ptr\set |h \ document_ptr_kinds_M|\<^sub>r. node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r" using concat_map_distinct assms(1) known_ptrs_implies by (smt CD.heap_is_wellformed_one_disc_parent DocumentMonad.ptr_kinds_ptr_kinds_M disjoint_iff_not_equal local.get_disconnected_nodes_ok local.heap_is_wellformed_def returns_result_select_result type_wf) have "h \ get_parent node_ptr \\<^sub>r None" using 4 2 apply(auto simp add: get_parent_def intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I )[1] apply(auto intro!: filter_M_empty_I bind_pure_I bind_pure_returns_result_I)[1] using get_child_nodes_ok assms(4) type_wf by (metis get_child_nodes_ok known_ptrs_known_ptr returns_result_select_result) then have 4: "h \ get_root_node (cast node_ptr) \\<^sub>r cast node_ptr" using get_root_node_no_parent by simp obtain document_ptrs where document_ptrs: "h \ document_ptr_kinds_M \\<^sub>r document_ptrs" by simp then have "h \ ok (filter_M (\document_ptr. do { disconnected_nodes \ get_disconnected_nodes document_ptr; return (((cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)) \ cast ` set disconnected_nodes) }) document_ptrs)" using assms(1) get_disconnected_nodes_ok type_wf by(auto intro!: bind_is_OK_I2 filter_M_is_OK_I bind_pure_I) then obtain candidates where candidates: "h \ filter_M (\document_ptr. do { disconnected_nodes \ get_disconnected_nodes document_ptr; return (((cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)) \ cast ` set disconnected_nodes) }) document_ptrs \\<^sub>r candidates" by auto have filter: "filter (\document_ptr. |h \ do { disconnected_nodes \ get_disconnected_nodes document_ptr; return (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr \ cast ` set disconnected_nodes) }|\<^sub>r) document_ptrs = [document_ptr]" apply(rule filter_ex1) using 0 document_ptrs apply(simp)[1] apply (smt "0" "3" assms bind_is_OK_pure_I bind_pure_returns_result_I bind_pure_returns_result_I2 bind_returns_result_E2 bind_returns_result_E3 document_ptr_kinds_M_def get_disconnected_nodes_ok get_disconnected_nodes_pure image_eqI is_OK_returns_result_E l_ptr_kinds_M.ptr_kinds_ptr_kinds_M return_ok return_returns_result returns_result_eq select_result_E select_result_I select_result_I2 select_result_I2) using assms(2) assms(3) apply (smt bind_is_OK_I2 bind_returns_result_E3 get_disconnected_nodes_pure image_eqI is_OK_returns_result_I return_ok return_returns_result select_result_E) using document_ptrs 3 apply(simp) using document_ptrs by simp have "h \ filter_M (\document_ptr. do { disconnected_nodes \ get_disconnected_nodes document_ptr; return (((cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)) \ cast ` set disconnected_nodes) }) document_ptrs \\<^sub>r [document_ptr]" apply(rule filter_M_filter2) using get_disconnected_nodes_ok document_ptrs 3 assms(1) type_wf filter by(auto intro: bind_pure_I bind_is_OK_I2) with 4 document_ptrs have "h \ CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r document_ptr" by(auto simp add: CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I split: option.splits) moreover have "known_ptr (cast node_ptr)" using known_ptrs_known_ptr[OF known_ptrs, where ptr="cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr"] 2 known_ptrs_implies by(simp) ultimately show ?thesis using 2 apply(auto simp add: CD.a_get_owner_document_tups_def get_owner_document_def a_get_owner_document_tups_def known_ptr_impl)[1] apply(split invoke_splits, (rule conjI | rule impI)+)+ apply(drule(1) known_ptr_not_shadow_root_ptr) apply(drule(1) known_ptr_not_document_ptr) apply(drule(1) known_ptr_not_character_data_ptr) apply(drule(1) known_ptr_not_element_ptr) apply(simp add: NodeClass.known_ptr_defs) by(auto split: option.splits intro!: bind_pure_returns_result_I) qed lemma in_disconnected_nodes_no_parent: assumes "heap_is_wellformed h" assumes "h \ get_parent node_ptr \\<^sub>r None" assumes "h \ get_owner_document (cast node_ptr) \\<^sub>r owner_document" assumes "h \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes" assumes "known_ptrs h" assumes "type_wf h" shows "node_ptr \ set disc_nodes" proof - have "\parent. parent |\| object_ptr_kinds h \ node_ptr \ set |h \ get_child_nodes parent|\<^sub>r" using assms(2) by (meson get_child_nodes_ok assms(1) assms(5) assms(6) local.child_parent_dual local.known_ptrs_known_ptr option.distinct(1) returns_result_eq returns_result_select_result) then show ?thesis by (smt assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) finite_set_in is_OK_returns_result_I local.get_disconnected_nodes_ok local.get_owner_document_disconnected_nodes local.get_parent_ptr_in_heap local.heap_is_wellformed_children_disc_nodes returns_result_select_result select_result_I2) qed lemma get_owner_document_owner_document_in_heap_node: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r owner_document" shows "owner_document |\| document_ptr_kinds h" proof - obtain root where root: "h \ get_root_node (cast node_ptr) \\<^sub>r root" using assms(4) by(auto simp add: CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 split: option.splits) then show ?thesis proof (cases "is_document_ptr root") case True then show ?thesis using assms(4) root apply(auto simp add: CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 intro!: filter_M_pure_I bind_pure_I split: option.splits)[1] apply(drule(1) returns_result_eq) apply(auto)[1] using assms document_ptr_kinds_commutes get_root_node_root_in_heap by blast next case False have "known_ptr root" using assms local.get_root_node_root_in_heap local.known_ptrs_known_ptr root by blast have "root |\| object_ptr_kinds h" using root using assms local.get_root_node_root_in_heap by blast show ?thesis proof (cases "is_shadow_root_ptr root") case True then show ?thesis using assms(4) root apply(auto simp add: CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 intro!: filter_M_pure_I bind_pure_I split: option.splits)[1] apply(drule(1) returns_result_eq) apply(auto)[1] using assms document_ptr_kinds_commutes get_root_node_root_in_heap by blast next case False then have "is_node_ptr_kind root" using \\ is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root\ \known_ptr root\ \root |\| object_ptr_kinds h\ apply(simp add: known_ptr_impl known_ptr_defs DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs) using is_node_ptr_kind_none by force then have "(\document_ptr \ fset (document_ptr_kinds h). root \ cast ` set |h \ get_disconnected_nodes document_ptr|\<^sub>r)" using local.child_parent_dual local.get_child_nodes_ok local.get_root_node_same_no_parent local.heap_is_wellformed_children_disc_nodes local.known_ptrs_known_ptr node_ptr_casts_commute3 node_ptr_inclusion node_ptr_kinds_commutes notin_fset option.distinct(1) returns_result_eq returns_result_select_result root by (metis (no_types, lifting) assms \root |\| object_ptr_kinds h\) then obtain some_owner_document where "some_owner_document |\| document_ptr_kinds h" and "root \ cast ` set |h \ get_disconnected_nodes some_owner_document|\<^sub>r" by auto then obtain candidates where candidates: "h \ filter_M (\document_ptr. Heap_Error_Monad.bind (get_disconnected_nodes document_ptr) (\disconnected_nodes. return (root \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ` set disconnected_nodes))) (sorted_list_of_set (fset (document_ptr_kinds h))) \\<^sub>r candidates" by (metis (no_types, lifting) assms bind_is_OK_I2 bind_pure_I filter_M_is_OK_I finite_fset is_OK_returns_result_E local.get_disconnected_nodes_ok local.get_disconnected_nodes_pure notin_fset return_ok return_pure sorted_list_of_set(1)) then have "some_owner_document \ set candidates" apply(rule filter_M_in_result_if_ok) using \some_owner_document |\| document_ptr_kinds h\ \root \ cast ` set |h \ get_disconnected_nodes some_owner_document|\<^sub>r\ apply(auto intro!: bind_pure_I bind_pure_returns_result_I)[1] using \some_owner_document |\| document_ptr_kinds h\ \root \ cast ` set |h \ get_disconnected_nodes some_owner_document|\<^sub>r\ apply(auto intro!: bind_pure_I bind_pure_returns_result_I)[1] using \some_owner_document |\| document_ptr_kinds h\ \root \ cast ` set |h \ get_disconnected_nodes some_owner_document|\<^sub>r\ apply(auto simp add: assms local.get_disconnected_nodes_ok intro!: bind_pure_I bind_pure_returns_result_I)[1] done then have "candidates \ []" by auto then have "owner_document \ set candidates" using assms(4) root apply(auto simp add: CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 intro!: filter_M_pure_I bind_pure_I split: option.splits)[1] apply (metis candidates list.set_sel(1) returns_result_eq) by (metis \is_node_ptr_kind root\ node_ptr_no_document_ptr_cast returns_result_eq) then show ?thesis using candidates by (meson bind_pure_I bind_returns_result_E2 filter_M_holds_for_result is_OK_returns_result_I local.get_disconnected_nodes_ptr_in_heap local.get_disconnected_nodes_pure return_pure) qed qed qed lemma get_owner_document_owner_document_in_heap: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_owner_document ptr \\<^sub>r owner_document" shows "owner_document |\| document_ptr_kinds h" using assms apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def CD.a_get_owner_document_tups_def)[1] apply(split invoke_split_asm)+ proof - assume "h \ invoke [] ptr () \\<^sub>r owner_document" then show "owner_document |\| document_ptr_kinds h" by (meson invoke_empty is_OK_returns_result_I) next assume "h \ Heap_Error_Monad.bind (check_in_heap ptr) (\_. (CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r) ptr ()) \\<^sub>r owner_document" then show "owner_document |\| document_ptr_kinds h" by(auto simp add: CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 split: if_splits) next assume 0: "heap_is_wellformed h" and 1: "type_wf h" and 2: "known_ptrs h" and 3: "\ is_element_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 4: "is_character_data_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 5: "h \ Heap_Error_Monad.bind (check_in_heap ptr) (\_. (CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r) ptr ()) \\<^sub>r owner_document" then show ?thesis by (metis bind_returns_result_E2 check_in_heap_pure comp_apply get_owner_document_owner_document_in_heap_node) next assume 0: "heap_is_wellformed h" and 1: "type_wf h" and 2: "known_ptrs h" and 3: "is_element_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 4: "h \ Heap_Error_Monad.bind (check_in_heap ptr) (\_. (CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r) ptr ()) \\<^sub>r owner_document" then show ?thesis by (metis bind_returns_result_E2 check_in_heap_pure comp_apply get_owner_document_owner_document_in_heap_node) next assume 0: "heap_is_wellformed h" and 1: "type_wf h" and 2: "known_ptrs h" and 3: "\ is_element_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 4: "\ is_character_data_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 5: "\ is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 6: "is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 7: "h \ Heap_Error_Monad.bind (check_in_heap ptr) (\_. (local.a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r) ptr ()) \\<^sub>r owner_document" then show "owner_document |\| document_ptr_kinds h" by(auto simp add: CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 split: if_splits) qed lemma get_owner_document_ok: assumes "heap_is_wellformed h" "known_ptrs h" "type_wf h" assumes "ptr |\| object_ptr_kinds h" shows "h \ ok (get_owner_document ptr)" proof - have "known_ptr ptr" using assms(2) assms(4) local.known_ptrs_known_ptr by blast then show ?thesis apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def CD.a_get_owner_document_tups_def)[1] apply(split invoke_splits, (rule conjI | rule impI)+)+ apply(auto simp add: known_ptr_impl)[1] using NodeClass.a_known_ptr_def known_ptr_not_character_data_ptr known_ptr_not_document_ptr known_ptr_not_shadow_root_ptr known_ptr_not_element_ptr apply blast using assms(4) apply(auto simp add: get_root_node_def CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I filter_M_is_OK_I split: option.splits)[1] using assms(4) apply(auto simp add: get_root_node_def CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I filter_M_is_OK_I split: option.splits)[1] using assms(4) apply(auto simp add: assms(1) assms(2) assms(3) local.get_ancestors_ok get_disconnected_nodes_ok get_root_node_def CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I filter_M_is_OK_I split: option.splits)[1] using assms(4) apply(auto simp add: assms(1) assms(2) assms(3) local.get_ancestors_ok get_disconnected_nodes_ok get_root_node_def CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I filter_M_is_OK_I split: option.splits)[1] done qed lemma get_owner_document_child_same: assumes "heap_is_wellformed h" "known_ptrs h" "type_wf h" assumes "h \ get_child_nodes ptr \\<^sub>r children" assumes "child \ set children" shows "h \ get_owner_document ptr \\<^sub>r owner_document \ h \ get_owner_document (cast child) \\<^sub>r owner_document" proof - have "ptr |\| object_ptr_kinds h" by (meson assms(4) is_OK_returns_result_I local.get_child_nodes_ptr_in_heap) then have "known_ptr ptr" using assms(2) local.known_ptrs_known_ptr by blast have "cast child |\| object_ptr_kinds h" using assms(1) assms(4) assms(5) local.heap_is_wellformed_children_in_heap node_ptr_kinds_commutes by blast then have "known_ptr (cast child)" using assms(2) local.known_ptrs_known_ptr by blast then have "is_character_data_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r (cast child) \ is_element_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r (cast child)" by(auto simp add: known_ptr_impl NodeClass.a_known_ptr_def ElementClass.a_known_ptr_def CharacterDataClass.a_known_ptr_def DocumentClass.a_known_ptr_def a_known_ptr_def split: option.splits) obtain root where root: "h \ get_root_node ptr \\<^sub>r root" by (meson \ptr |\| object_ptr_kinds h\ assms(1) assms(2) assms(3) is_OK_returns_result_E local.get_root_node_ok) then have "h \ get_root_node (cast child) \\<^sub>r root" using assms(1) assms(2) assms(3) assms(4) assms(5) local.child_parent_dual local.get_root_node_parent_same by blast have "h \ get_owner_document ptr \\<^sub>r owner_document \ h \ CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child () \\<^sub>r owner_document" proof (cases "is_document_ptr ptr") case True then obtain document_ptr where document_ptr: "cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr = ptr" using case_optionE document_ptr_casts_commute by blast then have "root = cast document_ptr" using root by(auto simp add: get_root_node_def get_ancestors_def elim!: bind_returns_result_E2 split: option.splits) then have "h \ CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr () \\<^sub>r owner_document \ h \ CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child () \\<^sub>r owner_document" using document_ptr \h \ get_root_node (cast child) \\<^sub>r root\[simplified \root = cast document_ptr\ document_ptr] apply(auto simp add: CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 dest!: bind_returns_result_E3[rotated, OF \h \ get_root_node (cast child) \\<^sub>r root\[simplified \root = cast document_ptr\ document_ptr], rotated] intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I split: if_splits option.splits)[1] using \ptr |\| object_ptr_kinds h\ document_ptr_kinds_commutes by blast then show ?thesis using \known_ptr ptr\ apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def CD.a_get_owner_document_tups_def known_ptr_impl)[1] apply(split invoke_splits, ((rule conjI | rule impI)+)?)+ apply(drule(1) known_ptr_not_shadow_root_ptr[folded known_ptr_impl]) apply(drule(1) known_ptr_not_document_ptr) apply(drule(1) known_ptr_not_character_data_ptr) apply(drule(1) known_ptr_not_element_ptr) apply(simp add: NodeClass.known_ptr_defs) using \ptr |\| object_ptr_kinds h\ True by(auto simp add: document_ptr[symmetric] intro!: bind_pure_returns_result_I split: option.splits) next case False then show ?thesis proof (cases "is_shadow_root_ptr ptr") case True then obtain shadow_root_ptr where shadow_root_ptr: "cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr = ptr" using case_optionE shadow_root_ptr_casts_commute by (metis (no_types, lifting) document_ptr_casts_commute3 is_document_ptr_kind_none option.case_eq_if) then have "root = cast shadow_root_ptr" using root by(auto simp add: get_root_node_def get_ancestors_def elim!: bind_returns_result_E2 split: option.splits) then have "h \ a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr () \\<^sub>r owner_document \ h \ CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child () \\<^sub>r owner_document" using shadow_root_ptr \h \ get_root_node (cast child) \\<^sub>r root\[simplified \root = cast shadow_root_ptr\ shadow_root_ptr] apply(auto simp add: CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 dest!: bind_returns_result_E3[rotated, OF \h \ get_root_node (cast child) \\<^sub>r root\[simplified \root = cast shadow_root_ptr\ shadow_root_ptr], rotated] intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I split: if_splits option.splits)[1] using \ptr |\| object_ptr_kinds h\ shadow_root_ptr_kinds_commutes document_ptr_kinds_commutes by blast then show ?thesis using \known_ptr ptr\ apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def CD.a_get_owner_document_tups_def known_ptr_impl)[1] apply(split invoke_splits, ((rule conjI | rule impI)+)?)+ apply(drule(1) known_ptr_not_shadow_root_ptr[folded known_ptr_impl]) apply(drule(1) known_ptr_not_document_ptr) apply(drule(1) known_ptr_not_character_data_ptr) apply(drule(1) known_ptr_not_element_ptr) apply(simp add: NodeClass.known_ptr_defs) using \ptr |\| object_ptr_kinds h\ True using False by(auto simp add: a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def shadow_root_ptr[symmetric] intro!: bind_pure_returns_result_I split: option.splits) next case False then obtain node_ptr where node_ptr: "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr = ptr" using \known_ptr ptr\ \\ is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr\ by(auto simp add: known_ptr_impl known_ptr_defs DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits) then have "h \ CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r owner_document \ h \ CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child () \\<^sub>r owner_document" using root \h \ get_root_node (cast child) \\<^sub>r root\ unfolding CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def by (meson bind_pure_returns_result_I bind_returns_result_E3 local.get_root_node_pure) then show ?thesis using \known_ptr ptr\ apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def CD.a_get_owner_document_tups_def known_ptr_impl)[1] apply(split invoke_splits, ((rule conjI | rule impI)+)?)+ apply(drule(1) known_ptr_not_shadow_root_ptr[folded known_ptr_impl]) apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl]) apply(drule(1) known_ptr_not_character_data_ptr) apply(drule(1) known_ptr_not_element_ptr) apply(simp add: NodeClass.known_ptr_defs) using \cast child |\| object_ptr_kinds h\ \ptr |\| object_ptr_kinds h\ False \\ is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr\ apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1] using \cast child |\| object_ptr_kinds h\ \ptr |\| object_ptr_kinds h\ False \\ is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr\ apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1] using \cast child |\| object_ptr_kinds h\ \ptr |\| object_ptr_kinds h\ False \\ is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr\ apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1] using \cast child |\| object_ptr_kinds h\ \ptr |\| object_ptr_kinds h\ False \\ is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr\ apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1] using \cast child |\| object_ptr_kinds h\ \ptr |\| object_ptr_kinds h\ False \\ is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr\ apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1] apply(split invoke_splits, ((rule conjI | rule impI)+)?)+ by(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I dest!: is_OK_returns_result_I) qed qed then show ?thesis using \is_character_data_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \ is_element_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child)\ using \cast child |\| object_ptr_kinds h\ by(auto simp add: get_owner_document_def[of "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child"] a_get_owner_document_tups_def CD.a_get_owner_document_tups_def split: invoke_splits) qed lemma get_owner_document_rel: assumes "heap_is_wellformed h" "known_ptrs h" "type_wf h" assumes "h \ get_owner_document ptr \\<^sub>r owner_document" assumes "ptr \ cast owner_document" shows "(cast owner_document, ptr) \ (parent_child_rel h \ a_ptr_disconnected_node_rel h)\<^sup>*" proof - have "ptr |\| object_ptr_kinds h" using assms by (meson is_OK_returns_result_I local.get_owner_document_ptr_in_heap) then have "known_ptr ptr" using known_ptrs_known_ptr[OF assms(2)] by simp have "is_node_ptr_kind ptr" proof (rule ccontr) assume "\ is_node_ptr_kind ptr" then show False using assms(4) \known_ptr ptr\ apply(auto simp add: known_ptr_impl get_owner_document_def a_get_owner_document_tups_def CD.a_get_owner_document_tups_def)[1] apply(split invoke_splits)+ apply(drule(1) known_ptr_not_shadow_root_ptr) apply(drule(1) known_ptr_not_document_ptr) apply(drule(1) known_ptr_not_character_data_ptr) apply(drule(1) known_ptr_not_element_ptr) apply(simp add: NodeClass.known_ptr_defs) using \ptr |\| object_ptr_kinds h\ assms(5) by(auto simp add: CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 split: if_splits option.splits) qed then obtain node_ptr where node_ptr: "ptr = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr" by (metis node_ptr_casts_commute3) then have "h \ CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r owner_document" using assms(4) \known_ptr ptr\ apply(auto simp add: known_ptr_impl get_owner_document_def a_get_owner_document_tups_def CD.a_get_owner_document_tups_def)[1] apply(split invoke_splits)+ apply(drule(1) known_ptr_not_shadow_root_ptr) apply(drule(1) known_ptr_not_document_ptr) apply(drule(1) known_ptr_not_character_data_ptr) apply(drule(1) known_ptr_not_element_ptr) apply(simp add: NodeClass.known_ptr_defs) using \ptr |\| object_ptr_kinds h\ by (auto simp add: is_document_ptr_kind_none) then obtain root where root: "h \ get_root_node (cast node_ptr) \\<^sub>r root" by(auto simp add: CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2) then have "root |\| object_ptr_kinds h" using assms(1) assms(2) assms(3) local.get_root_node_root_in_heap by blast then have "known_ptr root" using \known_ptrs h\ local.known_ptrs_known_ptr by blast have "(root, cast node_ptr) \ (parent_child_rel h \ a_ptr_disconnected_node_rel h)\<^sup>*" using root by (simp add: assms(1) assms(2) assms(3) in_rtrancl_UnI local.get_root_node_parent_child_rel) show ?thesis proof (cases "is_document_ptr_kind root") case True then have "root = cast owner_document" using \h \ CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r owner_document\ root by(auto simp add: CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def is_document_ptr_kind_def dest!: bind_returns_result_E3[rotated, OF root, rotated] split: option.splits) then have "(root, cast node_ptr) \ (parent_child_rel h \ a_ptr_disconnected_node_rel h)\<^sup>*" using assms(1) assms(2) assms(3) in_rtrancl_UnI local.get_root_node_parent_child_rel root by blast then show ?thesis using \root = cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\ node_ptr by blast next case False then obtain root_node where root_node: "root = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node" using assms(2) \root |\| object_ptr_kinds h\ by(auto simp add: known_ptr_impl ShadowRootClass.known_ptr_defs DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs dest!: known_ptrs_known_ptr split: option.splits) have "h \ CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node () \\<^sub>r owner_document" using \h \ CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r owner_document\ root False apply(auto simp add: root_node CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 dest!: bind_returns_result_E3[rotated, OF root, rotated] split: option.splits intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I)[1] by (simp add: assms(1) assms(2) assms(3) local.get_root_node_no_parent local.get_root_node_same_no_parent) then have "h \ get_owner_document root \\<^sub>r owner_document" using \known_ptr root\ apply(auto simp add: get_owner_document_def CD.a_get_owner_document_tups_def a_get_owner_document_tups_def known_ptr_impl)[1] apply(split invoke_splits, ((rule conjI | rule impI)+)?)+ apply(drule(1) known_ptr_not_shadow_root_ptr[folded known_ptr_impl]) apply(drule(1) known_ptr_not_document_ptr) apply(drule(1) known_ptr_not_character_data_ptr) apply(drule(1) known_ptr_not_element_ptr) apply(simp add: NodeClass.known_ptr_defs) using \h \ CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r owner_document\ root False \root |\| object_ptr_kinds h\ apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1] using \h \ CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r owner_document\ root False \root |\| object_ptr_kinds h\ apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1] using \h \ CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r owner_document\ root False \root |\| object_ptr_kinds h\ apply(auto simp add: root_node intro!: bind_pure_returns_result_I split: option.splits)[1] using \h \ CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r owner_document\ root False \root |\| object_ptr_kinds h\ apply(auto simp add: root_node intro!: bind_pure_returns_result_I split: option.splits)[1] done have "\ (\parent\fset (object_ptr_kinds h). root_node \ set |h \ get_child_nodes parent|\<^sub>r)" using root_node by (metis (no_types, lifting) assms(1) assms(2) assms(3) local.child_parent_dual local.get_child_nodes_ok local.get_root_node_same_no_parent local.known_ptrs_known_ptr notin_fset option.distinct(1) returns_result_eq returns_result_select_result root) have "root_node |\| node_ptr_kinds h" using assms(1) assms(2) assms(3) local.get_root_node_root_in_heap node_ptr_kinds_commutes root root_node by blast then have "\document_ptr\fset (document_ptr_kinds h). root_node \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r" using \\ (\parent\fset (object_ptr_kinds h). root_node \ set |h \ get_child_nodes parent|\<^sub>r)\ assms(1) local.heap_is_wellformed_children_disc_nodes by blast then obtain disc_nodes document_ptr where "h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes" and "root_node \ set disc_nodes" by (meson assms(3) local.get_disconnected_nodes_ok notin_fset returns_result_select_result) then have "document_ptr |\| document_ptr_kinds h" by (meson is_OK_returns_result_I local.get_disconnected_nodes_ptr_in_heap) then have "document_ptr = owner_document" by (metis \h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes\ \h \ get_owner_document root \\<^sub>r owner_document\ \root_node \ set disc_nodes\ assms(1) assms(2) assms(3) local.get_owner_document_disconnected_nodes returns_result_eq root_node) then have "(cast owner_document, cast root_node) \ a_ptr_disconnected_node_rel h" apply(auto simp add: a_ptr_disconnected_node_rel_def)[1] using \h \ local.CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r owner_document\ assms(1) assms(2) assms(3) get_owner_document_owner_document_in_heap_node by (metis (no_types, lifting) \h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes\ \root_node \ set disc_nodes\ case_prodI mem_Collect_eq pair_imageI select_result_I2) moreover have "(cast root_node, cast node_ptr) \ (parent_child_rel h \ a_host_shadow_root_rel h \ a_ptr_disconnected_node_rel h)\<^sup>*" by (metis assms(1) assms(2) assms(3) in_rtrancl_UnI local.get_root_node_parent_child_rel root root_node) ultimately show ?thesis by (metis (no_types, lifting) assms(1) assms(2) assms(3) in_rtrancl_UnI local.get_root_node_parent_child_rel node_ptr r_into_rtrancl root root_node rtrancl_trans) qed qed end interpretation i_get_owner_document_wf?: l_get_owner_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_disconnected_nodes get_disconnected_nodes_locs known_ptr get_child_nodes get_child_nodes_locs DocumentClass.known_ptr get_parent get_parent_locs DocumentClass.type_wf get_root_node get_root_node_locs CD.a_get_owner_document get_host get_host_locs get_owner_document get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_document get_disconnected_document_locs known_ptrs get_ancestors get_ancestors_locs by(auto simp add: l_get_owner_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_owner_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_get_owner_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_owner_document_wf_is_l_get_owner_document_wf [instances]: "l_get_owner_document_wf heap_is_wellformed type_wf known_ptr known_ptrs get_disconnected_nodes get_owner_document get_parent get_child_nodes" apply(auto simp add: l_get_owner_document_wf_def l_get_owner_document_wf_axioms_def instances)[1] using get_owner_document_disconnected_nodes apply fast using in_disconnected_nodes_no_parent apply fast using get_owner_document_owner_document_in_heap apply fast using get_owner_document_ok apply fast using get_owner_document_child_same apply (fast, fast) done paragraph \get\_owner\_document\ locale l_get_owner_document_wf_get_root_node_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_root_node_wf + l_get_owner_document_wf + assumes known_ptr_impl: "known_ptr = a_known_ptr" begin lemma get_root_node_document: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_root_node ptr \\<^sub>r root" assumes "is_document_ptr_kind root" shows "h \ get_owner_document ptr \\<^sub>r the (cast root)" proof - have "ptr |\| object_ptr_kinds h" using assms(4) by (meson is_OK_returns_result_I local.get_root_node_ptr_in_heap) then have "known_ptr ptr" using assms(3) local.known_ptrs_known_ptr by blast { assume "is_document_ptr_kind ptr" then have "ptr = root" using assms(4) by(auto simp add: get_root_node_def get_ancestors_def elim!: bind_returns_result_E2 split: option.splits) then have ?thesis using \is_document_ptr_kind ptr\ \known_ptr ptr\ \ptr |\| object_ptr_kinds h\ apply(auto simp add: known_ptr_impl get_owner_document_def a_get_owner_document_tups_def CD.a_get_owner_document_tups_def)[1] apply(split invoke_splits, (rule conjI | rule impI)+)+ apply(drule(1) known_ptr_not_shadow_root_ptr[folded known_ptr_impl]) apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl]) apply(drule(1) known_ptr_not_character_data_ptr) apply(drule(1) known_ptr_not_element_ptr) apply(simp add: NodeClass.known_ptr_defs) by(auto simp add: CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_pure_returns_result_I split: option.splits) } moreover { assume "is_node_ptr_kind ptr" then have ?thesis using \known_ptr ptr\ \ptr |\| object_ptr_kinds h\ apply(auto simp add: known_ptr_impl get_owner_document_def a_get_owner_document_tups_def CD.a_get_owner_document_tups_def)[1] apply(split invoke_splits, (rule conjI | rule impI)+)+ apply(drule(1) known_ptr_not_shadow_root_ptr[folded known_ptr_impl]) apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl]) apply(drule(1) known_ptr_not_character_data_ptr) apply(drule(1) known_ptr_not_element_ptr) apply(simp add: NodeClass.known_ptr_defs) apply(auto split: option.splits)[1] using \h \ get_root_node ptr \\<^sub>r root\ assms(5) by(auto simp add: CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def is_document_ptr_kind_def intro!: bind_pure_returns_result_I split: option.splits) } ultimately show ?thesis using \known_ptr ptr\ by(auto simp add: known_ptr_impl known_ptr_defs DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits) qed lemma get_root_node_same_owner_document: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_root_node ptr \\<^sub>r root" shows "h \ get_owner_document ptr \\<^sub>r owner_document \ h \ get_owner_document root \\<^sub>r owner_document" proof - have "ptr |\| object_ptr_kinds h" by (meson assms(4) is_OK_returns_result_I local.get_root_node_ptr_in_heap) have "root |\| object_ptr_kinds h" using assms(1) assms(2) assms(3) assms(4) local.get_root_node_root_in_heap by blast have "known_ptr ptr" using \ptr |\| object_ptr_kinds h\ assms(3) local.known_ptrs_known_ptr by blast have "known_ptr root" using \root |\| object_ptr_kinds h\ assms(3) local.known_ptrs_known_ptr by blast show ?thesis proof (cases "is_document_ptr_kind ptr") case True then have "ptr = root" using assms(4) apply(auto simp add: get_root_node_def elim!: bind_returns_result_E2)[1] by (metis document_ptr_casts_commute3 last_ConsL local.get_ancestors_not_node node_ptr_no_document_ptr_cast) then show ?thesis by auto next case False then have "is_node_ptr_kind ptr" using \known_ptr ptr\ by(auto simp add: known_ptr_impl known_ptr_defs DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits) then obtain node_ptr where node_ptr: "ptr = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr" by (metis node_ptr_casts_commute3) show ?thesis proof assume "h \ get_owner_document ptr \\<^sub>r owner_document" then have "h \ CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r owner_document" using node_ptr apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def CD.a_get_owner_document_tups_def)[1] apply(split invoke_splits)+ apply (meson invoke_empty is_OK_returns_result_I) by(auto elim!: bind_returns_result_E2 split: option.splits) show "h \ get_owner_document root \\<^sub>r owner_document" proof (cases "is_document_ptr_kind root") case True then show ?thesis proof (cases "is_shadow_root_ptr root") case True then have "is_shadow_root_ptr root" using True \known_ptr root\ by(auto simp add: known_ptr_impl known_ptr_defs DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits) have "root = cast owner_document" using \is_document_ptr_kind root\ by (smt \h \ get_owner_document ptr \\<^sub>r owner_document\ assms(1) assms(2) assms(3) assms(4) document_ptr_casts_commute3 get_root_node_document returns_result_eq) then show ?thesis apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def CD.a_get_owner_document_tups_def)[1] apply(split invoke_splits, (rule conjI | rule impI)+)+ using \is_shadow_root_ptr root\ apply blast using \root |\| object_ptr_kinds h\ apply(simp add: a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def is_node_ptr_kind_none) apply (metis \h \ get_owner_document ptr \\<^sub>r owner_document\ assms(1) assms(2) assms(3) case_optionE document_ptr_kinds_def is_shadow_root_ptr_kind_none l_get_owner_document_wf.get_owner_document_owner_document_in_heap local.l_get_owner_document_wf_axioms not_None_eq return_bind shadow_root_ptr_casts_commute3 shadow_root_ptr_kinds_commutes shadow_root_ptr_kinds_def) using \root |\| object_ptr_kinds h\ document_ptr_kinds_commutes apply(auto simp add: a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def is_node_ptr_kind_none intro!: bind_pure_returns_result_I)[1] using \root |\| object_ptr_kinds h\ document_ptr_kinds_commutes apply(auto simp add: a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def is_node_ptr_kind_none intro!: bind_pure_returns_result_I)[1] using \root |\| object_ptr_kinds h\ document_ptr_kinds_commutes apply(auto simp add: a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def is_node_ptr_kind_none intro!: bind_pure_returns_result_I)[1] done next case False then have "is_document_ptr root" using True \known_ptr root\ by(auto simp add: known_ptr_impl known_ptr_defs DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits) have "root = cast owner_document" using True by (smt \h \ get_owner_document ptr \\<^sub>r owner_document\ assms(1) assms(2) assms(3) assms(4) document_ptr_casts_commute3 get_root_node_document returns_result_eq) then show ?thesis apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def CD.a_get_owner_document_tups_def)[1] apply(split invoke_splits, (rule conjI | rule impI)+)+ using \is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root\ apply blast using \root |\| object_ptr_kinds h\ apply(auto simp add: a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def is_node_ptr_kind_none)[1] apply (metis \h \ get_owner_document ptr \\<^sub>r owner_document\ assms(1) assms(2) assms(3) case_optionE document_ptr_kinds_def is_shadow_root_ptr_kind_none l_get_owner_document_wf.get_owner_document_owner_document_in_heap local.l_get_owner_document_wf_axioms not_None_eq return_bind shadow_root_ptr_casts_commute3 shadow_root_ptr_kinds_commutes shadow_root_ptr_kinds_def) using \root |\| object_ptr_kinds h\ document_ptr_kinds_commutes apply(auto simp add: a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def is_node_ptr_kind_none intro!: bind_pure_returns_result_I)[1] using \root |\| object_ptr_kinds h\ document_ptr_kinds_commutes apply(auto simp add: a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def is_node_ptr_kind_none intro!: bind_pure_returns_result_I)[1] using \root |\| object_ptr_kinds h\ document_ptr_kinds_commutes apply(auto simp add: a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def is_node_ptr_kind_none intro!: bind_pure_returns_result_I)[1] done qed next case False then have "is_node_ptr_kind root" using \known_ptr root\ by(auto simp add: known_ptr_impl known_ptr_defs DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits) then obtain root_node_ptr where root_node_ptr: "root = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr" by (metis node_ptr_casts_commute3) then have "h \ CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr () \\<^sub>r owner_document" using \h \ CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r owner_document\ assms(4) apply(auto simp add: CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I split: option.splits)[1] apply (metis assms(1) assms(2) assms(3) local.get_root_node_no_parent local.get_root_node_same_no_parent node_ptr returns_result_eq) using \is_node_ptr_kind root\ node_ptr returns_result_eq by fastforce then show ?thesis apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def CD.a_get_owner_document_tups_def)[1] apply(split invoke_splits, (rule conjI | rule impI)+)+ using \is_node_ptr_kind root\ \known_ptr root\ apply(auto simp add: known_ptr_impl known_ptr_defs DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits)[1] using \is_node_ptr_kind root\ \known_ptr root\ apply(auto simp add: known_ptr_impl known_ptr_defs DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits)[1] using \is_node_ptr_kind root\ \known_ptr root\ apply(auto simp add: known_ptr_impl known_ptr_defs DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits)[1] using \root |\| object_ptr_kinds h\ by(auto simp add: root_node_ptr) qed next assume "h \ get_owner_document root \\<^sub>r owner_document" show "h \ get_owner_document ptr \\<^sub>r owner_document" proof (cases "is_document_ptr_kind root") case True have "root = cast owner_document" using \h \ get_owner_document root \\<^sub>r owner_document\ apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def CD.a_get_owner_document_tups_def)[1] apply(split invoke_splits)+ apply (meson invoke_empty is_OK_returns_result_I) apply(auto simp add: True CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 split: if_splits option.splits)[1] apply(auto simp add: True CD.a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 split: if_splits option.splits)[1] apply (metis True cast_document_ptr_not_node_ptr(2) is_document_ptr_kind_obtains is_node_ptr_kind_none node_ptr_casts_commute3 option.case_eq_if) by (metis True cast_document_ptr_not_node_ptr(1) document_ptr_casts_commute3 is_node_ptr_kind_none node_ptr_casts_commute3 option.case_eq_if) then show ?thesis using assms(1) assms(2) assms(3) assms(4) get_root_node_document by fastforce next case False then have "is_node_ptr_kind root" using \known_ptr root\ by(auto simp add: known_ptr_impl known_ptr_defs DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits) then obtain root_node_ptr where root_node_ptr: "root = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr" by (metis node_ptr_casts_commute3) then have "h \ CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr () \\<^sub>r owner_document" using \h \ get_owner_document root \\<^sub>r owner_document\ apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def CD.a_get_owner_document_tups_def)[1] apply(split invoke_splits)+ apply (meson invoke_empty is_OK_returns_result_I) by(auto simp add: is_document_ptr_kind_none elim!: bind_returns_result_E2) then have "h \ CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r owner_document" apply(auto simp add: CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I split: option.splits)[1] using assms(1) assms(2) assms(3) assms(4) local.get_root_node_no_parent local.get_root_node_same_no_parent node_ptr returns_result_eq root_node_ptr by fastforce+ then show ?thesis apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def CD.a_get_owner_document_tups_def)[1] apply(split invoke_splits, (rule conjI | rule impI)+)+ using node_ptr \known_ptr ptr\ \ptr |\| object_ptr_kinds h\ by(auto simp add: known_ptr_impl known_ptr_defs DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs intro!: bind_pure_returns_result_I split: option.splits) qed qed qed qed end interpretation i_get_owner_document_wf_get_root_node_wf?: l_get_owner_document_wf_get_root_node_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M DocumentClass.known_ptr get_parent get_parent_locs DocumentClass.type_wf get_disconnected_nodes get_disconnected_nodes_locs get_root_node get_root_node_locs CD.a_get_owner_document get_host get_host_locs get_owner_document get_child_nodes get_child_nodes_locs type_wf known_ptr known_ptrs get_ancestors get_ancestors_locs heap_is_wellformed parent_child_rel by(auto simp add: l_get_owner_document_wf_get_root_node_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_owner_document_wf_get_root_node_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_get_owner_document_wf_get_root_node_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_owner_document_wf_get_root_node_wf_is_l_get_owner_document_wf_get_root_node_wf [instances]: "l_get_owner_document_wf_get_root_node_wf heap_is_wellformed type_wf known_ptr known_ptrs get_root_node get_owner_document" apply(auto simp add: l_get_owner_document_wf_get_root_node_wf_def l_get_owner_document_wf_get_root_node_wf_axioms_def instances)[1] using get_root_node_document apply blast using get_root_node_same_owner_document apply (blast, blast) done subsubsection \remove\_child\ locale l_remove_child_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_disconnected_nodes_get_disconnected_nodes + l_get_child_nodes + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_owner_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_child_nodes_get_shadow_root + l_set_disconnected_nodes_get_shadow_root + l_set_child_nodes_get_tag_name + l_set_disconnected_nodes_get_tag_name + CD: l_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma remove_child_preserves_type_wf: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ remove_child ptr child \\<^sub>h h'" shows "type_wf h'" using CD.remove_child_heap_is_wellformed_preserved(1) assms unfolding heap_is_wellformed_def by auto lemma remove_child_preserves_known_ptrs: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ remove_child ptr child \\<^sub>h h'" shows "known_ptrs h'" using CD.remove_child_heap_is_wellformed_preserved(2) assms unfolding heap_is_wellformed_def by auto lemma remove_child_heap_is_wellformed_preserved: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ remove_child ptr child \\<^sub>h h'" shows "heap_is_wellformed h'" proof - obtain owner_document children_h h2 disconnected_nodes_h where owner_document: "h \ get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r owner_document" and children_h: "h \ get_child_nodes ptr \\<^sub>r children_h" and child_in_children_h: "child \ set children_h" and disconnected_nodes_h: "h \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h" and h2: "h \ set_disconnected_nodes owner_document (child # disconnected_nodes_h) \\<^sub>h h2" and h': "h2 \ set_child_nodes ptr (remove1 child children_h) \\<^sub>h h'" using assms(4) apply(auto simp add: remove_child_def elim!: bind_returns_heap_E dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure] pure_returns_heap_eq[rotated, OF get_child_nodes_pure] split: if_splits)[1] using pure_returns_heap_eq by fastforce have "heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'" using CD.remove_child_heap_is_wellformed_preserved(3) assms unfolding heap_is_wellformed_def by auto have "h \ get_owner_document ptr \\<^sub>r owner_document" using owner_document children_h child_in_children_h using local.get_owner_document_child_same assms by blast have shadow_root_eq: "\ptr' shadow_root_ptr_opt. h \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt = h' \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt" using get_shadow_root_reads remove_child_writes assms(4) apply(rule reads_writes_preserved) by(auto simp add: remove_child_locs_def set_child_nodes_get_shadow_root set_disconnected_nodes_get_shadow_root) then have shadow_root_eq2: "\ptr'. |h \ get_shadow_root ptr'|\<^sub>r = |h' \ get_shadow_root ptr'|\<^sub>r" by (meson select_result_eq) have tag_name_eq: "\ptr' tag. h \ get_tag_name ptr' \\<^sub>r tag = h' \ get_tag_name ptr' \\<^sub>r tag" using get_tag_name_reads remove_child_writes assms(4) apply(rule reads_writes_preserved) by(auto simp add: remove_child_locs_def set_child_nodes_get_tag_name set_disconnected_nodes_get_tag_name) then have tag_name_eq2: "\ptr'. |h \ get_tag_name ptr'|\<^sub>r = |h' \ get_tag_name ptr'|\<^sub>r" by (meson select_result_eq) have object_ptr_kinds_eq: "object_ptr_kinds h = object_ptr_kinds h'" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF remove_child_writes assms(4)]) unfolding remove_child_locs_def using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) have document_ptr_kinds_eq: "document_ptr_kinds h = document_ptr_kinds h'" using object_ptr_kinds_eq by(auto simp add: document_ptr_kinds_def document_ptr_kinds_def) have shadow_root_ptr_kinds_eq: "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'" using object_ptr_kinds_eq by(auto simp add: shadow_root_ptr_kinds_def document_ptr_kinds_def) have element_ptr_kinds_eq: "element_ptr_kinds h = element_ptr_kinds h'" using object_ptr_kinds_eq by(auto simp add: element_ptr_kinds_def node_ptr_kinds_def) have "parent_child_rel h' \ parent_child_rel h" using \heap_is_wellformed h\ heap_is_wellformed_def using CD.remove_child_parent_child_rel_subset using \known_ptrs h\ \type_wf h\ assms(4) by simp have "known_ptr ptr" using assms(3) using children_h get_child_nodes_ptr_in_heap local.known_ptrs_known_ptr by blast have "type_wf h2" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h2] using set_disconnected_nodes_types_preserved \type_wf h\ by(auto simp add: reflp_def transp_def) have children_eq: "\ptr' children. ptr \ ptr' \ h \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" apply(rule reads_writes_preserved[OF get_child_nodes_reads remove_child_writes assms(4)]) unfolding remove_child_locs_def using set_disconnected_nodes_get_child_nodes set_child_nodes_get_child_nodes_different_pointers by fast then have children_eq2: "\ptr' children. ptr \ ptr' \ |h \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have "h2 \ get_child_nodes ptr \\<^sub>r children_h" apply(rule reads_writes_separate_forwards[OF get_child_nodes_reads set_disconnected_nodes_writes h2 children_h] ) by (simp add: set_disconnected_nodes_get_child_nodes) have children_h': "h' \ get_child_nodes ptr \\<^sub>r remove1 child children_h" using assms(4) owner_document h2 disconnected_nodes_h children_h apply(auto simp add: remove_child_def split: if_splits)[1] apply(drule bind_returns_heap_E3) apply(auto split: if_splits)[1] apply(simp) apply(auto split: if_splits)[1] apply(drule bind_returns_heap_E3) apply(auto)[1] apply(simp) apply(drule bind_returns_heap_E3) apply(auto)[1] apply(simp) apply(drule bind_returns_heap_E4) apply(auto)[1] apply simp using \type_wf h2\ set_child_nodes_get_child_nodes \known_ptr ptr\ h' by blast have disconnected_nodes_eq: "\ptr' disc_nodes. ptr' \ owner_document \ h \ get_disconnected_nodes ptr' \\<^sub>r disc_nodes = h2 \ get_disconnected_nodes ptr' \\<^sub>r disc_nodes" using local.get_disconnected_nodes_reads set_disconnected_nodes_writes h2 apply(rule reads_writes_preserved) by (metis local.set_disconnected_nodes_get_disconnected_nodes_different_pointers) then have disconnected_nodes_eq2: "\ptr'. ptr' \ owner_document \ |h \ get_disconnected_nodes ptr'|\<^sub>r = |h2 \ get_disconnected_nodes ptr'|\<^sub>r" by (meson select_result_eq) have "h2 \ get_disconnected_nodes owner_document \\<^sub>r child # disconnected_nodes_h" using h2 local.set_disconnected_nodes_get_disconnected_nodes by blast have disconnected_nodes_eq_h2: "\ptr' disc_nodes. h2 \ get_disconnected_nodes ptr' \\<^sub>r disc_nodes = h' \ get_disconnected_nodes ptr' \\<^sub>r disc_nodes" using local.get_disconnected_nodes_reads set_child_nodes_writes h' apply(rule reads_writes_preserved) using local.set_child_nodes_get_disconnected_nodes by blast then have disconnected_nodes_eq2_h2: "\ptr'. |h2 \ get_disconnected_nodes ptr'|\<^sub>r = |h' \ get_disconnected_nodes ptr'|\<^sub>r" by (meson select_result_eq) have "a_host_shadow_root_rel h' = a_host_shadow_root_rel h" by(auto simp add: a_host_shadow_root_rel_def shadow_root_eq2 element_ptr_kinds_eq) moreover have "(ptr, cast child) \ parent_child_rel h" using child_in_children_h children_h local.CD.parent_child_rel_child by blast moreover have "a_ptr_disconnected_node_rel h' = insert (cast owner_document, cast child) (a_ptr_disconnected_node_rel h)" using \h2 \ get_disconnected_nodes owner_document \\<^sub>r child # disconnected_nodes_h\ disconnected_nodes_eq2 disconnected_nodes_h apply(auto simp add: a_ptr_disconnected_node_rel_def disconnected_nodes_eq2_h2[symmetric] document_ptr_kinds_eq[symmetric])[1] apply(case_tac "aa = owner_document") apply(auto)[1] apply(auto)[1] apply (metis (no_types, lifting) assms(4) case_prodI disconnected_nodes_eq_h2 h2 is_OK_returns_heap_I local.remove_child_in_disconnected_nodes local.set_disconnected_nodes_ptr_in_heap mem_Collect_eq owner_document pair_imageI select_result_I2) by (metis (no_types, lifting) case_prodI list.set_intros(2) mem_Collect_eq pair_imageI select_result_I2) then have "a_ptr_disconnected_node_rel h' = a_ptr_disconnected_node_rel h \ {(cast owner_document, cast child)}" by auto moreover have "acyclic (parent_child_rel h \ a_host_shadow_root_rel h \ a_ptr_disconnected_node_rel h)" using assms(1) local.heap_is_wellformed_def by blast moreover have "parent_child_rel h' = parent_child_rel h - {(ptr, cast child)}" apply(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq children_eq2)[1] apply (metis (no_types, lifting) children_eq2 children_h children_h' notin_set_remove1 select_result_I2) using \h2 \ get_disconnected_nodes owner_document \\<^sub>r child # disconnected_nodes_h\ \heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'\ disconnected_nodes_eq_h2 local.CD.distinct_lists_no_parent local.CD.heap_is_wellformed_def apply auto[1] by (metis (no_types, lifting) children_eq2 children_h children_h' in_set_remove1 select_result_I2) moreover have "(cast owner_document, ptr) \ (parent_child_rel h \ a_ptr_disconnected_node_rel h)\<^sup>*" using \h \ get_owner_document ptr \\<^sub>r owner_document\ get_owner_document_rel using assms(1) assms(2) assms(3) by blast then have "(cast owner_document, ptr) \ (parent_child_rel h \ a_host_shadow_root_rel h \ a_ptr_disconnected_node_rel h)\<^sup>*" by (metis (no_types, lifting) in_rtrancl_UnI inf_sup_aci(5) inf_sup_aci(7)) ultimately have "acyclic (parent_child_rel h' \ a_host_shadow_root_rel h' \ a_ptr_disconnected_node_rel h')" by (smt Un_assoc Un_insert_left Un_insert_right acyclic_insert insert_Diff_single insert_absorb2 mk_disjoint_insert prod.inject rtrancl_Un_separator_converseE rtrancl_trans singletonD sup_bot.comm_neutral) show ?thesis using \heap_is_wellformed h\ using \heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'\ using \acyclic (parent_child_rel h' \ a_host_shadow_root_rel h' \ a_ptr_disconnected_node_rel h')\ apply(auto simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def host_shadow_root_rel_def a_all_ptrs_in_heap_def a_distinct_lists_def a_shadow_root_valid_def)[1] by(auto simp add: object_ptr_kinds_eq element_ptr_kinds_eq shadow_root_ptr_kinds_eq shadow_root_eq shadow_root_eq2 tag_name_eq tag_name_eq2) qed lemma remove_preserves_type_wf: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ remove child \\<^sub>h h'" shows "type_wf h'" using CD.remove_heap_is_wellformed_preserved(1) assms unfolding heap_is_wellformed_def by auto lemma remove_preserves_known_ptrs: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ remove child \\<^sub>h h'" shows "known_ptrs h'" using CD.remove_heap_is_wellformed_preserved(2) assms unfolding heap_is_wellformed_def by auto lemma remove_heap_is_wellformed_preserved: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ remove child \\<^sub>h h'" shows "heap_is_wellformed h'" using assms by(auto simp add: remove_def elim!: bind_returns_heap_E2 intro: remove_child_heap_is_wellformed_preserved split: option.splits) lemma remove_child_removes_child: "heap_is_wellformed h \ h \ remove_child ptr' child \\<^sub>h h' \ h' \ get_child_nodes ptr \\<^sub>r children \ known_ptrs h \ type_wf h \ child \ set children" using CD.remove_child_removes_child local.heap_is_wellformed_def by blast lemma remove_child_removes_first_child: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_child_nodes ptr \\<^sub>r node_ptr # children \ h \ remove_child ptr node_ptr \\<^sub>h h' \ h' \ get_child_nodes ptr \\<^sub>r children" using CD.remove_child_removes_first_child local.heap_is_wellformed_def by blast lemma remove_removes_child: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_child_nodes ptr \\<^sub>r node_ptr # children \ h \ remove node_ptr \\<^sub>h h' \ h' \ get_child_nodes ptr \\<^sub>r children" using CD.remove_removes_child local.heap_is_wellformed_def by blast lemma remove_for_all_empty_children: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_child_nodes ptr \\<^sub>r children \ h \ forall_M remove children \\<^sub>h h' \ h' \ get_child_nodes ptr \\<^sub>r []" using CD.remove_for_all_empty_children local.heap_is_wellformed_def by blast end interpretation i_remove_child_wf2?: l_remove_child_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs known_ptr get_child_nodes get_child_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs DocumentClass.known_ptr get_parent get_parent_locs DocumentClass.type_wf get_root_node get_root_node_locs CD.a_get_owner_document get_owner_document known_ptrs get_ancestors get_ancestors_locs set_child_nodes set_child_nodes_locs remove_child remove_child_locs remove by(auto simp add: l_remove_child_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_remove_child_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma remove_child_wf2_is_l_remove_child_wf2 [instances]: "l_remove_child_wf2 type_wf known_ptr known_ptrs remove_child heap_is_wellformed get_child_nodes remove" apply(auto simp add: l_remove_child_wf2_def l_remove_child_wf2_axioms_def instances)[1] using remove_child_preserves_type_wf apply fast using remove_child_preserves_known_ptrs apply fast using remove_child_heap_is_wellformed_preserved apply (fast) using remove_preserves_type_wf apply fast using remove_preserves_known_ptrs apply fast using remove_heap_is_wellformed_preserved apply (fast) using remove_child_removes_child apply fast using remove_child_removes_first_child apply fast using remove_removes_child apply fast using remove_for_all_empty_children apply fast done subsubsection \adopt\_node\ locale l_adopt_node_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_adopt_node\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + CD: l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M _ _ _ _ _ _ _ _ _ adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M adopt_node_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma adopt_node_removes_first_child: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ adopt_node owner_document node \\<^sub>h h' \ h \ get_child_nodes ptr' \\<^sub>r node # children \ h' \ get_child_nodes ptr' \\<^sub>r children" by (smt CD.adopt_node_removes_first_child bind_returns_heap_E error_returns_heap l_adopt_node\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M.adopt_node_def local.CD.adopt_node_impl local.get_ancestors_di_pure local.l_adopt_node\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms pure_returns_heap_eq) lemma adopt_node_document_in_heap: "heap_is_wellformed h \ known_ptrs h \ type_wf h \ h \ ok (adopt_node owner_document node) \ owner_document |\| document_ptr_kinds h" by (metis (no_types, lifting) bind_returns_heap_E document_ptr_kinds_commutes is_OK_returns_heap_E is_OK_returns_result_I local.adopt_node_def local.get_ancestors_di_ptr_in_heap) end locale l_adopt_node_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes + l_get_disconnected_nodes + l_set_child_nodes_get_shadow_root + l_set_disconnected_nodes_get_shadow_root + l_set_child_nodes_get_tag_name + l_set_disconnected_nodes_get_tag_name + l_get_owner_document + l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M+ l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_root_node + l_set_disconnected_nodes_get_child_nodes + l_get_owner_document_wf + l_remove_child_wf2 + l_adopt_node_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_adopt_node\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_disconnected_document + l_get_ancestors_di\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_ancestors_di_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma adopt_node_removes_child: assumes wellformed: "heap_is_wellformed h" and adopt_node: "h \ adopt_node owner_document node_ptr \\<^sub>h h2" and children: "h2 \ get_child_nodes ptr \\<^sub>r children" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "node_ptr \ set children" proof - obtain old_document parent_opt h' where old_document: "h \ get_owner_document (cast node_ptr) \\<^sub>r old_document" and parent_opt: "h \ get_parent node_ptr \\<^sub>r parent_opt" and h': "h \ (case parent_opt of Some parent \ remove_child parent node_ptr | None \ return () ) \\<^sub>h h'" using adopt_node by(auto simp add: adopt_node_def CD.adopt_node_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_ancestors_di_pure, rotated] bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated] bind_returns_heap_E2[rotated, OF get_parent_pure, rotated] bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] split: if_splits) then have "h' \ get_child_nodes ptr \\<^sub>r children" using adopt_node apply(auto simp add: adopt_node_def CD.adopt_node_def dest!: bind_returns_heap_E3[rotated, OF old_document, rotated] bind_returns_heap_E3[rotated, OF parent_opt, rotated] elim!: bind_returns_heap_E2[rotated, OF get_ancestors_di_pure, rotated] bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated] bind_returns_heap_E4[rotated, OF h', rotated] split: if_splits)[1] apply(auto split: if_splits elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_ancestors_di_pure, rotated] bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated])[1] apply (simp add: set_disconnected_nodes_get_child_nodes children reads_writes_preserved[OF get_child_nodes_reads set_disconnected_nodes_writes]) using children by blast show ?thesis proof(insert parent_opt h', induct parent_opt) case None then show ?case using child_parent_dual wellformed known_ptrs type_wf \h' \ get_child_nodes ptr \\<^sub>r children\ returns_result_eq by fastforce next case (Some option) then show ?case using remove_child_removes_child \h' \ get_child_nodes ptr \\<^sub>r children\ known_ptrs type_wf wellformed by auto qed qed lemma adopt_node_preserves_wellformedness: assumes "heap_is_wellformed h" and "h \ adopt_node document_ptr child \\<^sub>h h'" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "heap_is_wellformed h'" proof - obtain old_document parent_opt h2 ancestors where "h \ get_ancestors_di (cast document_ptr) \\<^sub>r ancestors" and "cast child \ set ancestors" and old_document: "h \ get_owner_document (cast child) \\<^sub>r old_document" and parent_opt: "h \ get_parent child \\<^sub>r parent_opt" and h2: "h \ (case parent_opt of Some parent \ remove_child parent child | None \ return ()) \\<^sub>h h2" and h': "h2 \ (if document_ptr \ old_document then do { old_disc_nodes \ get_disconnected_nodes old_document; set_disconnected_nodes old_document (remove1 child old_disc_nodes); disc_nodes \ get_disconnected_nodes document_ptr; set_disconnected_nodes document_ptr (child # disc_nodes) } else do { return () }) \\<^sub>h h'" using assms(2) apply(auto simp add: adopt_node_def[unfolded CD.adopt_node_def] elim!: bind_returns_heap_E dest!: pure_returns_heap_eq[rotated, OF get_ancestors_di_pure])[1] apply(split if_splits) by(auto simp add: elim!: bind_returns_heap_E dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure] pure_returns_heap_eq[rotated, OF get_parent_pure]) have object_ptr_kinds_h_eq3: "object_ptr_kinds h = object_ptr_kinds h2" using h2 apply(simp split: option.splits) apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF remove_child_writes]) using remove_child_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h: "\ptrs. h \ object_ptr_kinds_M \\<^sub>r ptrs = h2 \ object_ptr_kinds_M \\<^sub>r ptrs" unfolding object_ptr_kinds_M_defs by simp then have object_ptr_kinds_eq_h: "|h \ object_ptr_kinds_M|\<^sub>r = |h2 \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq_h: "|h \ node_ptr_kinds_M|\<^sub>r = |h2 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have wellformed_h2: "heap_is_wellformed h2" using h2 remove_child_heap_is_wellformed_preserved known_ptrs type_wf by (metis (no_types, lifting) assms(1) option.case_eq_if pure_returns_heap_eq return_pure) then show "heap_is_wellformed h'" proof(cases "document_ptr = old_document") case True then show "heap_is_wellformed h'" using h' wellformed_h2 by auto next case False then obtain h3 old_disc_nodes disc_nodes_document_ptr_h3 where docs_neq: "document_ptr \ old_document" and old_disc_nodes: "h2 \ get_disconnected_nodes old_document \\<^sub>r old_disc_nodes" and h3: "h2 \ set_disconnected_nodes old_document (remove1 child old_disc_nodes) \\<^sub>h h3" and disc_nodes_document_ptr_h3: "h3 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_document_ptr_h3" and h': "h3 \ set_disconnected_nodes document_ptr (child # disc_nodes_document_ptr_h3) \\<^sub>h h'" using h' by(auto elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] ) have object_ptr_kinds_h2_eq3: "object_ptr_kinds h2 = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes h3]) using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h2: "\ptrs. h2 \ object_ptr_kinds_M \\<^sub>r ptrs = h3 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_eq_h2: "|h2 \ object_ptr_kinds_M|\<^sub>r = |h3 \ object_ptr_kinds_M|\<^sub>r" by(simp) then have node_ptr_kinds_eq_h2: "|h2 \ node_ptr_kinds_M|\<^sub>r = |h3 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast then have node_ptr_kinds_eq3_h2: "node_ptr_kinds h2 = node_ptr_kinds h3" by auto have document_ptr_kinds_eq2_h2: "|h2 \ document_ptr_kinds_M|\<^sub>r = |h3 \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_eq_h2 document_ptr_kinds_M_eq by auto then have document_ptr_kinds_eq3_h2: "document_ptr_kinds h2 = document_ptr_kinds h3" using object_ptr_kinds_eq_h2 document_ptr_kinds_M_eq by auto have children_eq_h2: "\ptr children. h2 \ get_child_nodes ptr \\<^sub>r children = h3 \ get_child_nodes ptr \\<^sub>r children" using get_child_nodes_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by (simp add: set_disconnected_nodes_get_child_nodes) then have children_eq2_h2: "\ptr. |h2 \ get_child_nodes ptr|\<^sub>r = |h3 \ get_child_nodes ptr|\<^sub>r" using select_result_eq by force have object_ptr_kinds_h3_eq3: "object_ptr_kinds h3 = object_ptr_kinds h'" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes h']) using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h3: "\ptrs. h3 \ object_ptr_kinds_M \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_eq_h3: "|h3 \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" by(simp) then have node_ptr_kinds_eq_h3: "|h3 \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast then have node_ptr_kinds_eq3_h3: "node_ptr_kinds h3 = node_ptr_kinds h'" by auto have document_ptr_kinds_eq2_h3: "|h3 \ document_ptr_kinds_M|\<^sub>r = |h' \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_eq_h3 document_ptr_kinds_M_eq by auto then have document_ptr_kinds_eq3_h3: "document_ptr_kinds h3 = document_ptr_kinds h'" using object_ptr_kinds_eq_h3 document_ptr_kinds_M_eq by auto have children_eq_h3: "\ptr children. h3 \ get_child_nodes ptr \\<^sub>r children = h' \ get_child_nodes ptr \\<^sub>r children" using get_child_nodes_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by (simp add: set_disconnected_nodes_get_child_nodes) then have children_eq2_h3: "\ptr. |h3 \ get_child_nodes ptr|\<^sub>r = |h' \ get_child_nodes ptr|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. old_document \ doc_ptr \ h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by (simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers) then have disconnected_nodes_eq2_h2: "\doc_ptr. old_document \ doc_ptr \ |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force obtain disc_nodes_old_document_h2 where disc_nodes_old_document_h2: "h2 \ get_disconnected_nodes old_document \\<^sub>r disc_nodes_old_document_h2" using old_disc_nodes by blast then have disc_nodes_old_document_h3: "h3 \ get_disconnected_nodes old_document \\<^sub>r remove1 child disc_nodes_old_document_h2" using h3 old_disc_nodes returns_result_eq set_disconnected_nodes_get_disconnected_nodes by fastforce have "distinct disc_nodes_old_document_h2" using disc_nodes_old_document_h2 local.heap_is_wellformed_disconnected_nodes_distinct wellformed_h2 by blast have "type_wf h2" proof (insert h2, induct parent_opt) case None then show ?case using type_wf by simp next case (Some option) then show ?case using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF remove_child_writes] type_wf remove_child_types_preserved by (simp add: reflp_def transp_def) qed then have "type_wf h3" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h3] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) then have "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h'] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) have disconnected_nodes_eq_h3: "\doc_ptr disc_nodes. document_ptr \ doc_ptr \ h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by (simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers) then have disconnected_nodes_eq2_h3: "\doc_ptr. document_ptr \ doc_ptr \ |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have disc_nodes_document_ptr_h2: "h2 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_document_ptr_h3" using disconnected_nodes_eq_h2 docs_neq disc_nodes_document_ptr_h3 by auto have disc_nodes_document_ptr_h': "h' \ get_disconnected_nodes document_ptr \\<^sub>r child # disc_nodes_document_ptr_h3" using h' disc_nodes_document_ptr_h3 using set_disconnected_nodes_get_disconnected_nodes by blast have document_ptr_in_heap: "document_ptr |\| document_ptr_kinds h2" using disc_nodes_document_ptr_h3 document_ptr_kinds_eq2_h2 get_disconnected_nodes_ok assms(1) unfolding heap_is_wellformed_def using disc_nodes_document_ptr_h2 get_disconnected_nodes_ptr_in_heap by blast have old_document_in_heap: "old_document |\| document_ptr_kinds h2" using disc_nodes_old_document_h3 document_ptr_kinds_eq2_h2 get_disconnected_nodes_ok assms(1) unfolding heap_is_wellformed_def using get_disconnected_nodes_ptr_in_heap old_disc_nodes by blast have "child \ set disc_nodes_old_document_h2" proof (insert parent_opt h2, induct parent_opt) case None then have "h = h2" by(auto) moreover have "CD.a_owner_document_valid h" using assms(1) by(simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) ultimately show ?case using old_document disc_nodes_old_document_h2 None(1) child_parent_dual[OF assms(1)] in_disconnected_nodes_no_parent assms(1) known_ptrs type_wf by blast next case (Some option) then show ?case apply(simp split: option.splits) using assms(1) disc_nodes_old_document_h2 old_document remove_child_in_disconnected_nodes known_ptrs by blast qed have "child \ set (remove1 child disc_nodes_old_document_h2)" using disc_nodes_old_document_h3 h3 known_ptrs wellformed_h2 \distinct disc_nodes_old_document_h2\ by auto have "child \ set disc_nodes_document_ptr_h3" proof - have "CD.a_distinct_lists h2" using heap_is_wellformed_def CD.heap_is_wellformed_def wellformed_h2 by blast then have 0: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>r) |h2 \ document_ptr_kinds_M|\<^sub>r))" by(simp add: CD.a_distinct_lists_def) show ?thesis using distinct_concat_map_E(1)[OF 0] \child \ set disc_nodes_old_document_h2\ disc_nodes_old_document_h2 disc_nodes_document_ptr_h2 by (meson \type_wf h2\ docs_neq known_ptrs local.get_owner_document_disconnected_nodes local.known_ptrs_preserved object_ptr_kinds_h_eq3 returns_result_eq wellformed_h2) qed have child_in_heap: "child |\| node_ptr_kinds h" using get_owner_document_ptr_in_heap[OF is_OK_returns_result_I[OF old_document]] node_ptr_kinds_commutes by blast have "CD.a_acyclic_heap h2" using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) have "parent_child_rel h' \ parent_child_rel h2" proof fix x assume "x \ parent_child_rel h'" then show "x \ parent_child_rel h2" using object_ptr_kinds_h2_eq3 object_ptr_kinds_h3_eq3 children_eq2_h2 children_eq2_h3 mem_Collect_eq object_ptr_kinds_M_eq_h3 select_result_eq split_cong unfolding CD.parent_child_rel_def by(simp) qed then have " CD.a_acyclic_heap h'" using \ CD.a_acyclic_heap h2\ CD.acyclic_heap_def acyclic_subset by blast moreover have " CD.a_all_ptrs_in_heap h2" using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_all_ptrs_in_heap h3" apply(auto simp add: CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq3_h2 children_eq_h2)[1] apply (metis CD.l_heap_is_wellformed_axioms \type_wf h2\ children_eq2_h2 known_ptrs l_heap_is_wellformed.heap_is_wellformed_children_in_heap local.get_child_nodes_ok local.known_ptrs_known_ptr node_ptr_kinds_eq3_h2 object_ptr_kinds_h2_eq3 object_ptr_kinds_h_eq3 returns_result_select_result wellformed_h2) by (metis (no_types, lifting) disc_nodes_old_document_h2 disc_nodes_old_document_h3 disconnected_nodes_eq2_h2 document_ptr_kinds_eq3_h2 finite_set_in select_result_I2 set_remove1_subset subsetD) then have "CD.a_all_ptrs_in_heap h'" by (smt \child \ set disc_nodes_old_document_h2\ children_eq2_h3 disc_nodes_document_ptr_h' disc_nodes_document_ptr_h2 disc_nodes_old_document_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq3_h3 finite_set_in local.CD.a_all_ptrs_in_heap_def local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq3_h2 node_ptr_kinds_eq3_h3 object_ptr_kinds_h3_eq3 select_result_I2 set_ConsD subset_code(1) wellformed_h2) moreover have "CD.a_owner_document_valid h2" using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_owner_document_valid h'" apply(simp add: CD.a_owner_document_valid_def node_ptr_kinds_eq_h2 node_ptr_kinds_eq3_h3 object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 children_eq2_h2 children_eq2_h3 ) by (metis (no_types) disc_nodes_document_ptr_h' disc_nodes_document_ptr_h2 disc_nodes_old_document_h2 disc_nodes_old_document_h3 disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_in_heap document_ptr_kinds_eq3_h2 document_ptr_kinds_eq3_h3 in_set_remove1 list.set_intros(1) list.set_intros(2) node_ptr_kinds_eq3_h2 node_ptr_kinds_eq3_h3 object_ptr_kinds_h2_eq3 object_ptr_kinds_h3_eq3 select_result_I2) have a_distinct_lists_h2: "CD.a_distinct_lists h2" using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_distinct_lists h'" apply(auto simp add: CD.a_distinct_lists_def object_ptr_kinds_eq_h3 object_ptr_kinds_eq_h2 children_eq2_h2 children_eq2_h3)[1] proof - assume 1: "distinct (concat (map (\ptr. |h' \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h')))))" and 2: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h2)))))" and 3: "(\x\fset (object_ptr_kinds h'). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h2). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" show "distinct (concat (map (\document_ptr. |h' \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h')))))" proof(rule distinct_concat_map_I) show "distinct (sorted_list_of_set (fset (document_ptr_kinds h')))" by(auto simp add: document_ptr_kinds_M_def ) next fix x assume a1: "x \ set (sorted_list_of_set (fset (document_ptr_kinds h')))" have 4: "distinct |h2 \ get_disconnected_nodes x|\<^sub>r" using a_distinct_lists_h2 "2" a1 concat_map_all_distinct document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 by fastforce then show "distinct |h' \ get_disconnected_nodes x|\<^sub>r" proof (cases "old_document \ x") case True then show ?thesis proof (cases "document_ptr \ x") case True then show ?thesis using disconnected_nodes_eq2_h2[OF \old_document \ x\] disconnected_nodes_eq2_h3[OF \document_ptr \ x\] 4 by(auto) next case False then show ?thesis using disc_nodes_document_ptr_h3 disc_nodes_document_ptr_h' 4 \child \ set disc_nodes_document_ptr_h3\ by(auto simp add: disconnected_nodes_eq2_h2[OF \old_document \ x\] ) qed next case False then show ?thesis by (metis (no_types, hide_lams) \distinct disc_nodes_old_document_h2\ disc_nodes_old_document_h3 disconnected_nodes_eq2_h3 distinct_remove1 docs_neq select_result_I2) qed next fix x y assume a0: "x \ set (sorted_list_of_set (fset (document_ptr_kinds h')))" and a1: "y \ set (sorted_list_of_set (fset (document_ptr_kinds h')))" and a2: "x \ y" moreover have 5: "set |h2 \ get_disconnected_nodes x|\<^sub>r \ set |h2 \ get_disconnected_nodes y|\<^sub>r = {}" using 2 calculation by (auto simp add: document_ptr_kinds_eq3_h2 document_ptr_kinds_eq3_h3 dest: distinct_concat_map_E(1)) ultimately show "set |h' \ get_disconnected_nodes x|\<^sub>r \ set |h' \ get_disconnected_nodes y|\<^sub>r = {}" proof(cases "old_document = x") case True have "old_document \ y" using \x \ y\ \old_document = x\ by simp have "document_ptr \ x" using docs_neq \old_document = x\ by auto show ?thesis proof(cases "document_ptr = y") case True then show ?thesis using 5 True select_result_I2[OF disc_nodes_document_ptr_h'] select_result_I2[OF disc_nodes_document_ptr_h2] select_result_I2[OF disc_nodes_old_document_h2] select_result_I2[OF disc_nodes_old_document_h3] \old_document = x\ by (metis (no_types, lifting) \child \ set (remove1 child disc_nodes_old_document_h2)\ \document_ptr \ x\ disconnected_nodes_eq2_h3 disjoint_iff_not_equal notin_set_remove1 set_ConsD) next case False then show ?thesis using 5 select_result_I2[OF disc_nodes_document_ptr_h'] select_result_I2[OF disc_nodes_document_ptr_h2] select_result_I2[OF disc_nodes_old_document_h2] select_result_I2[OF disc_nodes_old_document_h3] disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 \old_document = x\ docs_neq \old_document \ y\ by (metis (no_types, lifting) disjoint_iff_not_equal notin_set_remove1) qed next case False then show ?thesis proof(cases "old_document = y") case True then show ?thesis proof(cases "document_ptr = x") case True show ?thesis using 5 select_result_I2[OF disc_nodes_document_ptr_h'] select_result_I2[OF disc_nodes_document_ptr_h2] select_result_I2[OF disc_nodes_old_document_h2] select_result_I2[OF disc_nodes_old_document_h3] \old_document \ x\ \old_document = y\ \document_ptr = x\ apply(simp) by (metis (no_types, lifting) \child \ set (remove1 child disc_nodes_old_document_h2)\ disconnected_nodes_eq2_h3 disjoint_iff_not_equal notin_set_remove1) next case False then show ?thesis using 5 select_result_I2[OF disc_nodes_document_ptr_h'] select_result_I2[OF disc_nodes_document_ptr_h2] select_result_I2[OF disc_nodes_old_document_h2] select_result_I2[OF disc_nodes_old_document_h3] \old_document \ x\ \old_document = y\ \document_ptr \ x\ by (metis (no_types, lifting) disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disjoint_iff_not_equal docs_neq notin_set_remove1) qed next case False have "set |h2 \ get_disconnected_nodes y|\<^sub>r \ set disc_nodes_old_document_h2 = {}" by (metis DocumentMonad.ptr_kinds_M_ok DocumentMonad.ptr_kinds_M_ptr_kinds False \type_wf h2\ a1 disc_nodes_old_document_h2 document_ptr_kinds_M_def document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 l_ptr_kinds_M.ptr_kinds_ptr_kinds_M local.get_disconnected_nodes_ok local.heap_is_wellformed_one_disc_parent returns_result_select_result wellformed_h2) then show ?thesis proof(cases "document_ptr = x") case True then have "document_ptr \ y" using \x \ y\ by auto have "set |h2 \ get_disconnected_nodes y|\<^sub>r \ set disc_nodes_old_document_h2 = {}" using \set |h2 \ get_disconnected_nodes y|\<^sub>r \ set disc_nodes_old_document_h2 = {}\ by blast then show ?thesis using 5 select_result_I2[OF disc_nodes_document_ptr_h'] select_result_I2[OF disc_nodes_document_ptr_h2] select_result_I2[OF disc_nodes_old_document_h2] select_result_I2[OF disc_nodes_old_document_h3] \old_document \ x\ \old_document \ y\ \document_ptr = x\ \document_ptr \ y\ \child \ set disc_nodes_old_document_h2\ disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 \set |h2 \ get_disconnected_nodes y|\<^sub>r \ set disc_nodes_old_document_h2 = {}\ by(auto) next case False then show ?thesis proof(cases "document_ptr = y") case True have f1: "set |h2 \ get_disconnected_nodes x|\<^sub>r \ set disc_nodes_document_ptr_h3 = {}" using 2 a1 document_ptr_in_heap document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 \document_ptr \ x\ select_result_I2[OF disc_nodes_document_ptr_h3, symmetric] disconnected_nodes_eq2_h2[OF docs_neq[symmetric], symmetric] by (simp add: "5" True) moreover have f1: "set |h2 \ get_disconnected_nodes x|\<^sub>r \ set |h2 \ get_disconnected_nodes old_document|\<^sub>r = {}" using 2 a1 old_document_in_heap document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 \old_document \ x\ by (metis (no_types, lifting) a0 distinct_concat_map_E(1) document_ptr_kinds_eq3_h2 document_ptr_kinds_eq3_h3 finite_fset fmember.rep_eq set_sorted_list_of_set) ultimately show ?thesis using 5 select_result_I2[OF disc_nodes_document_ptr_h'] select_result_I2[OF disc_nodes_old_document_h2] \old_document \ x\ \document_ptr \ x\ \document_ptr = y\ \child \ set disc_nodes_old_document_h2\ disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 by auto next case False then show ?thesis using 5 select_result_I2[OF disc_nodes_old_document_h2] \old_document \ x\ \document_ptr \ x\ \document_ptr \ y\ \child \ set disc_nodes_old_document_h2\ disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 by (metis \set |h2 \ get_disconnected_nodes y|\<^sub>r \ set disc_nodes_old_document_h2 = {}\ empty_iff inf.idem) qed qed qed qed qed next fix x xa xb assume 0: "distinct (concat (map (\ptr. |h' \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h')))))" and 1: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h2)))))" and 2: "(\x\fset (object_ptr_kinds h'). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h2). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" and 3: "xa |\| object_ptr_kinds h'" and 4: "x \ set |h' \ get_child_nodes xa|\<^sub>r" and 5: "xb |\| document_ptr_kinds h'" and 6: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" then show False using \child \ set disc_nodes_old_document_h2\ disc_nodes_document_ptr_h' disc_nodes_document_ptr_h2 disc_nodes_old_document_h2 disc_nodes_old_document_h3 disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 old_document_in_heap apply(auto)[1] apply(cases "xb = old_document") proof - assume a1: "xb = old_document" assume a2: "h2 \ get_disconnected_nodes old_document \\<^sub>r disc_nodes_old_document_h2" assume a3: "h3 \ get_disconnected_nodes old_document \\<^sub>r remove1 child disc_nodes_old_document_h2" assume a4: "x \ set |h' \ get_child_nodes xa|\<^sub>r" assume "document_ptr_kinds h2 = document_ptr_kinds h'" assume a5: "(\x\fset (object_ptr_kinds h'). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h'). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" have f6: "old_document |\| document_ptr_kinds h'" using a1 \xb |\| document_ptr_kinds h'\ by blast have f7: "|h2 \ get_disconnected_nodes old_document|\<^sub>r = disc_nodes_old_document_h2" using a2 by simp have "x \ set disc_nodes_old_document_h2" using f6 a3 a1 by (metis (no_types) \type_wf h'\ \x \ set |h' \ get_disconnected_nodes xb|\<^sub>r\ disconnected_nodes_eq_h3 docs_neq get_disconnected_nodes_ok returns_result_eq returns_result_select_result set_remove1_subset subsetCE) then have "set |h' \ get_child_nodes xa|\<^sub>r \ set |h2 \ get_disconnected_nodes xb|\<^sub>r = {}" using f7 f6 a5 a4 \xa |\| object_ptr_kinds h'\ by fastforce then show ?thesis using \x \ set disc_nodes_old_document_h2\ a1 a4 f7 by blast next assume a1: "xb \ old_document" assume a2: "h2 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_document_ptr_h3" assume a3: "h2 \ get_disconnected_nodes old_document \\<^sub>r disc_nodes_old_document_h2" assume a4: "xa |\| object_ptr_kinds h'" assume a5: "h' \ get_disconnected_nodes document_ptr \\<^sub>r child # disc_nodes_document_ptr_h3" assume a6: "old_document |\| document_ptr_kinds h'" assume a7: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" assume a8: "x \ set |h' \ get_child_nodes xa|\<^sub>r" assume a9: "document_ptr_kinds h2 = document_ptr_kinds h'" assume a10: "\doc_ptr. old_document \ doc_ptr \ |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" assume a11: "\doc_ptr. document_ptr \ doc_ptr \ |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" assume a12: "(\x\fset (object_ptr_kinds h'). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h'). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" have f13: "\d. d \ set |h' \ document_ptr_kinds_M|\<^sub>r \ h2 \ ok get_disconnected_nodes d" using a9 \type_wf h2\ get_disconnected_nodes_ok by simp then have f14: "|h2 \ get_disconnected_nodes old_document|\<^sub>r = disc_nodes_old_document_h2" using a6 a3 by simp have "x \ set |h2 \ get_disconnected_nodes xb|\<^sub>r" using a12 a8 a4 \xb |\| document_ptr_kinds h'\ by (meson UN_I disjoint_iff_not_equal fmember.rep_eq) then have "x = child" using f13 a11 a10 a7 a5 a2 a1 by (metis (no_types, lifting) select_result_I2 set_ConsD) then have "child \ set disc_nodes_old_document_h2" using f14 a12 a8 a6 a4 by (metis \type_wf h'\ adopt_node_removes_child assms(1) assms(2) type_wf get_child_nodes_ok known_ptrs local.known_ptrs_known_ptr object_ptr_kinds_h2_eq3 object_ptr_kinds_h3_eq3 object_ptr_kinds_h_eq3 returns_result_select_result) then show ?thesis using \child \ set disc_nodes_old_document_h2\ by fastforce qed qed ultimately have "heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'" using \type_wf h'\ \CD.a_owner_document_valid h'\ CD.heap_is_wellformed_def by blast have shadow_root_eq_h2: "\ptr' shadow_root_ptr_opt. h2 \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt = h3 \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt" using get_shadow_root_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by(auto simp add: adopt_node_locs_def remove_child_locs_def set_child_nodes_get_shadow_root set_disconnected_nodes_get_shadow_root) then have shadow_root_eq2_h2: "\ptr'. |h2 \ get_shadow_root ptr'|\<^sub>r = |h3 \ get_shadow_root ptr'|\<^sub>r" by (meson select_result_eq) have shadow_root_eq_h3: "\ptr' shadow_root_ptr_opt. h3 \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt = h' \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt" using get_shadow_root_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: adopt_node_locs_def remove_child_locs_def set_child_nodes_get_shadow_root set_disconnected_nodes_get_shadow_root) then have shadow_root_eq2_h3: "\ptr'. |h3 \ get_shadow_root ptr'|\<^sub>r = |h' \ get_shadow_root ptr'|\<^sub>r" by (meson select_result_eq) have tag_name_eq_h2: "\ptr' tag. h2 \ get_tag_name ptr' \\<^sub>r tag = h3 \ get_tag_name ptr' \\<^sub>r tag" using get_tag_name_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by(auto simp add: adopt_node_locs_def remove_child_locs_def set_child_nodes_get_tag_name set_disconnected_nodes_get_tag_name) then have tag_name_eq2_h2: "\ptr'. |h2 \ get_tag_name ptr'|\<^sub>r = |h3 \ get_tag_name ptr'|\<^sub>r" by (meson select_result_eq) have tag_name_eq_h3: "\ptr' tag. h3 \ get_tag_name ptr' \\<^sub>r tag = h' \ get_tag_name ptr' \\<^sub>r tag" using get_tag_name_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: adopt_node_locs_def remove_child_locs_def set_child_nodes_get_tag_name set_disconnected_nodes_get_tag_name) then have tag_name_eq2_h3: "\ptr'. |h3 \ get_tag_name ptr'|\<^sub>r = |h' \ get_tag_name ptr'|\<^sub>r" by (meson select_result_eq) have object_ptr_kinds_eq_h2: "object_ptr_kinds h2 = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes h3]) unfolding adopt_node_locs_def remove_child_locs_def using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def split: if_splits) have object_ptr_kinds_eq_h3: "object_ptr_kinds h3 = object_ptr_kinds h'" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes h']) unfolding adopt_node_locs_def remove_child_locs_def using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def split: if_splits) have document_ptr_kinds_eq_h2: "document_ptr_kinds h2 = document_ptr_kinds h3" using object_ptr_kinds_eq_h2 by(auto simp add: document_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h2: "shadow_root_ptr_kinds h2 = shadow_root_ptr_kinds h3" using object_ptr_kinds_eq_h2 by(auto simp add: shadow_root_ptr_kinds_def document_ptr_kinds_def) have element_ptr_kinds_eq_h2: "element_ptr_kinds h2 = element_ptr_kinds h3" using object_ptr_kinds_eq_h2 by(auto simp add: element_ptr_kinds_def node_ptr_kinds_def) have document_ptr_kinds_eq_h3: "document_ptr_kinds h3 = document_ptr_kinds h'" using object_ptr_kinds_eq_h3 by(auto simp add: document_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h3: "shadow_root_ptr_kinds h3 = shadow_root_ptr_kinds h'" using object_ptr_kinds_eq_h3 by(auto simp add: shadow_root_ptr_kinds_def document_ptr_kinds_def) have element_ptr_kinds_eq_h3: "element_ptr_kinds h3 = element_ptr_kinds h'" using object_ptr_kinds_eq_h3 by(auto simp add: element_ptr_kinds_def node_ptr_kinds_def) have "a_host_shadow_root_rel h' = a_host_shadow_root_rel h3" and "a_host_shadow_root_rel h3 = a_host_shadow_root_rel h2" by(auto simp add: a_host_shadow_root_rel_def shadow_root_eq2_h2 shadow_root_eq2_h3 element_ptr_kinds_eq_h2 element_ptr_kinds_eq_h3) have "parent_child_rel h' = parent_child_rel h3" and "parent_child_rel h3 = parent_child_rel h2" by(auto simp add: CD.parent_child_rel_def children_eq2_h2 children_eq2_h3 object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3) have "parent_child_rel h2 \ parent_child_rel h" using h2 parent_opt proof (induct parent_opt) case None then show ?case by simp next case (Some parent) then have h2: "h \ remove_child parent child \\<^sub>h h2" by auto have child_nodes_eq_h: "\ptr children. parent \ ptr \ h \ get_child_nodes ptr \\<^sub>r children = h2 \ get_child_nodes ptr \\<^sub>r children" using get_child_nodes_reads remove_child_writes h2 apply(rule reads_writes_preserved) apply(auto simp add: remove_child_locs_def)[1] by (simp add: set_child_nodes_get_child_nodes_different_pointers) moreover obtain children where children: "h \ get_child_nodes parent \\<^sub>r children" using Some local.get_parent_child_dual by blast ultimately show ?thesis using object_ptr_kinds_eq_h h2 apply(auto simp add: CD.parent_child_rel_def split: option.splits)[1] apply(case_tac "parent = a") apply (metis (no_types, lifting) \type_wf h3\ children_eq2_h2 children_eq_h2 known_ptrs local.get_child_nodes_ok local.known_ptrs_known_ptr local.remove_child_children_subset object_ptr_kinds_h2_eq3 returns_result_select_result subset_code(1) type_wf) apply (metis (no_types, lifting) known_ptrs local.get_child_nodes_ok local.known_ptrs_known_ptr returns_result_select_result select_result_I2 type_wf) done qed have "a_host_shadow_root_rel h2 = a_host_shadow_root_rel h" using h2 proof (induct parent_opt) case None then show ?case by simp next case (Some parent) then have h2: "h \ remove_child parent child \\<^sub>h h2" by auto have "\ptr shadow_root. h \ get_shadow_root ptr \\<^sub>r shadow_root = h2 \ get_shadow_root ptr \\<^sub>r shadow_root" using get_shadow_root_reads remove_child_writes h2 apply(rule reads_writes_preserved) apply(auto simp add: remove_child_locs_def)[1] by (auto simp add: set_disconnected_nodes_get_shadow_root set_child_nodes_get_shadow_root) then show ?case apply(auto simp add: a_host_shadow_root_rel_def)[1] apply (metis (mono_tags, lifting) Collect_cong \type_wf h2\ case_prodE case_prodI l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_host_shadow_root_rel_def local.get_shadow_root_ok local.a_host_shadow_root_rel_shadow_root returns_result_select_result) by (metis (no_types, lifting) Collect_cong case_prodE case_prodI local.get_shadow_root_ok local.a_host_shadow_root_rel_def local.a_host_shadow_root_rel_shadow_root returns_result_select_result type_wf) qed have "a_ptr_disconnected_node_rel h3 = a_ptr_disconnected_node_rel h2 - {(cast old_document, cast child)}" apply(auto simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq_h2 disconnected_nodes_eq2_h2)[1] using disconnected_nodes_eq2_h2 disc_nodes_old_document_h2 disc_nodes_old_document_h3 using \distinct disc_nodes_old_document_h2\ apply (metis (no_types, lifting) \child \ set (remove1 child disc_nodes_old_document_h2)\ case_prodI in_set_remove1 mem_Collect_eq pair_imageI select_result_I2) using \child \ set (remove1 child disc_nodes_old_document_h2)\ disc_nodes_old_document_h3 apply auto[1] by (metis (no_types, lifting) case_prodI disc_nodes_old_document_h2 disc_nodes_old_document_h3 disconnected_nodes_eq2_h2 in_set_remove1 mem_Collect_eq pair_imageI select_result_I2) have "a_ptr_disconnected_node_rel h3 \ a_ptr_disconnected_node_rel h" using h2 parent_opt proof (induct parent_opt) case None then show ?case by(auto simp add: \a_ptr_disconnected_node_rel h3 = a_ptr_disconnected_node_rel h2 - {(cast old_document, cast child)}\) next case (Some parent) then have h2: "h \ remove_child parent child \\<^sub>h h2" by auto then obtain children_h h'2 disconnected_nodes_h where children_h: "h \ get_child_nodes parent \\<^sub>r children_h" and child_in_children_h: "child \ set children_h" and disconnected_nodes_h: "h \ get_disconnected_nodes old_document \\<^sub>r disconnected_nodes_h" and h'2: "h \ set_disconnected_nodes old_document (child # disconnected_nodes_h) \\<^sub>h h'2" and h': "h'2 \ set_child_nodes parent (remove1 child children_h) \\<^sub>h h2" using old_document apply(auto simp add: remove_child_def elim!: bind_returns_heap_E dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure] pure_returns_heap_eq[rotated, OF get_child_nodes_pure] pure_returns_heap_eq[rotated, OF get_disconnected_nodes_pure] split: if_splits)[1] using select_result_I2 by fastforce have "|h3 \ document_ptr_kinds_M|\<^sub>r = |h \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_eq_h object_ptr_kinds_eq_h2 by(auto simp add: document_ptr_kinds_def) have disconnected_nodes_eq_h: "\ptr disc_nodes. old_document \ ptr \ h \ get_disconnected_nodes ptr \\<^sub>r disc_nodes = h2 \ get_disconnected_nodes ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads remove_child_writes h2 apply(rule reads_writes_preserved) apply(auto simp add: remove_child_locs_def)[1] using old_document by (auto simp add:set_child_nodes_get_disconnected_nodes set_disconnected_nodes_get_disconnected_nodes_different_pointers) then have foo: "\ptr disc_nodes. old_document \ ptr \ h \ get_disconnected_nodes ptr \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes ptr \\<^sub>r disc_nodes" using disconnected_nodes_eq_h2 by simp then have foo2: "\ptr. old_document \ ptr \ |h \ get_disconnected_nodes ptr|\<^sub>r = |h3 \ get_disconnected_nodes ptr|\<^sub>r" by (meson select_result_eq) have "h'2 \ get_disconnected_nodes old_document \\<^sub>r child # disconnected_nodes_h" using h'2 using local.set_disconnected_nodes_get_disconnected_nodes by blast have "h2 \ get_disconnected_nodes old_document \\<^sub>r child # disconnected_nodes_h" using get_disconnected_nodes_reads set_child_nodes_writes h' \h'2 \ get_disconnected_nodes old_document \\<^sub>r child # disconnected_nodes_h\ apply(rule reads_writes_separate_forwards) using local.set_child_nodes_get_disconnected_nodes by blast then have "h3 \ get_disconnected_nodes old_document \\<^sub>r disconnected_nodes_h" using h3 using disc_nodes_old_document_h2 disc_nodes_old_document_h3 returns_result_eq by fastforce have "a_ptr_disconnected_node_rel h3 = a_ptr_disconnected_node_rel h" using \|h3 \ document_ptr_kinds_M|\<^sub>r = |h \ document_ptr_kinds_M|\<^sub>r\ apply(auto simp add: a_ptr_disconnected_node_rel_def )[1] apply(case_tac "old_document = aa") using disconnected_nodes_h \h3 \ get_disconnected_nodes old_document \\<^sub>r disconnected_nodes_h\ using foo2 apply(auto)[1] using disconnected_nodes_h \h3 \ get_disconnected_nodes old_document \\<^sub>r disconnected_nodes_h\ using foo2 apply(auto)[1] apply(case_tac "old_document = aa") using disconnected_nodes_h \h3 \ get_disconnected_nodes old_document \\<^sub>r disconnected_nodes_h\ using foo2 apply(auto)[1] using disconnected_nodes_h \h3 \ get_disconnected_nodes old_document \\<^sub>r disconnected_nodes_h\ using foo2 apply(auto)[1] done then show ?thesis by auto qed have "acyclic (parent_child_rel h2 \ a_host_shadow_root_rel h2 \ a_ptr_disconnected_node_rel h2)" using local.heap_is_wellformed_def wellformed_h2 by blast then have "acyclic (parent_child_rel h3 \ a_host_shadow_root_rel h3 \ a_ptr_disconnected_node_rel h3)" using \a_ptr_disconnected_node_rel h3 = a_ptr_disconnected_node_rel h2 - {(cast old_document, cast child)}\ by(auto simp add: \parent_child_rel h3 = parent_child_rel h2\ \a_host_shadow_root_rel h3 = a_host_shadow_root_rel h2\ elim!: acyclic_subset) moreover have "a_ptr_disconnected_node_rel h' = insert (cast document_ptr, cast child) (a_ptr_disconnected_node_rel h3)" using disconnected_nodes_eq2_h3[symmetric] disc_nodes_document_ptr_h3 disc_nodes_document_ptr_h' document_ptr_in_heap[unfolded document_ptr_kinds_eq_h2] apply(auto simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq_h3[symmetric])[1] apply(case_tac "document_ptr = aa") apply(auto)[1] apply(auto)[1] apply(case_tac "document_ptr = aa") apply(auto)[1] apply(auto)[1] done moreover have "(cast child, cast document_ptr) \ (parent_child_rel h \ a_host_shadow_root_rel h \ a_ptr_disconnected_node_rel h)\<^sup>*" using \h \ get_ancestors_di (cast document_ptr) \\<^sub>r ancestors\ \cast child \ set ancestors\ get_ancestors_di_parent_child_a_host_shadow_root_rel using assms(1) known_ptrs type_wf by blast moreover have "(cast child, cast document_ptr) \ (parent_child_rel h3 \ a_host_shadow_root_rel h3 \ a_ptr_disconnected_node_rel h3)\<^sup>*" proof - have "(parent_child_rel h3 \ local.a_host_shadow_root_rel h3 \ local.a_ptr_disconnected_node_rel h3) \ (parent_child_rel h \ local.a_host_shadow_root_rel h \ local.a_ptr_disconnected_node_rel h)" apply(simp add: \parent_child_rel h3 = parent_child_rel h2\ \a_host_shadow_root_rel h3 = a_host_shadow_root_rel h2\ \a_host_shadow_root_rel h2 = a_host_shadow_root_rel h\) using \local.a_ptr_disconnected_node_rel h3 \ local.a_ptr_disconnected_node_rel h\ \parent_child_rel h2 \ parent_child_rel h\ by blast then show ?thesis using calculation(3) rtrancl_mono by blast qed ultimately have "acyclic (parent_child_rel h' \ a_host_shadow_root_rel h' \ a_ptr_disconnected_node_rel h')" by(auto simp add: \parent_child_rel h' = parent_child_rel h3\ \a_host_shadow_root_rel h' = a_host_shadow_root_rel h3\) show "heap_is_wellformed h'" using \heap_is_wellformed h2\ using \heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'\ using \acyclic (parent_child_rel h' \ a_host_shadow_root_rel h' \ a_ptr_disconnected_node_rel h')\ apply(auto simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def a_all_ptrs_in_heap_def a_distinct_lists_def a_shadow_root_valid_def)[1] by(auto simp add: object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 element_ptr_kinds_eq_h2 element_ptr_kinds_eq_h3 shadow_root_ptr_kinds_eq_h2 shadow_root_ptr_kinds_eq_h3 shadow_root_eq_h2 shadow_root_eq_h3 shadow_root_eq2_h2 shadow_root_eq2_h3 tag_name_eq_h2 tag_name_eq_h3 tag_name_eq2_h2 tag_name_eq2_h3) qed qed lemma adopt_node_node_in_disconnected_nodes: assumes wellformed: "heap_is_wellformed h" and adopt_node: "h \ adopt_node owner_document node_ptr \\<^sub>h h'" and "h' \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "node_ptr \ set disc_nodes" proof - obtain old_document parent_opt h2 where old_document: "h \ get_owner_document (cast node_ptr) \\<^sub>r old_document" and parent_opt: "h \ get_parent node_ptr \\<^sub>r parent_opt" and h2: "h \ (case parent_opt of Some parent \ remove_child parent node_ptr | None \ return ()) \\<^sub>h h2" and h': "h2 \ (if owner_document \ old_document then do { old_disc_nodes \ get_disconnected_nodes old_document; set_disconnected_nodes old_document (remove1 node_ptr old_disc_nodes); disc_nodes \ get_disconnected_nodes owner_document; set_disconnected_nodes owner_document (node_ptr # disc_nodes) } else do { return () }) \\<^sub>h h'" using assms(2)[unfolded adopt_node_def CD.adopt_node_def] by(auto elim!: bind_returns_heap_E dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure] pure_returns_heap_eq[rotated, OF get_parent_pure] pure_returns_heap_eq[rotated, OF get_ancestors_di_pure] split: option.splits if_splits) show ?thesis proof (cases "owner_document = old_document") case True then show ?thesis proof (insert parent_opt h2, induct parent_opt) case None then have "h = h'" using h2 h' by(auto) then show ?case using in_disconnected_nodes_no_parent assms None old_document by blast next case (Some parent) then show ?case using remove_child_in_disconnected_nodes known_ptrs True h' assms(3) old_document by auto qed next case False then show ?thesis using assms(3) h' list.set_intros(1) select_result_I2 set_disconnected_nodes_get_disconnected_nodes apply(auto elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated])[1] proof - fix x and h'a and xb assume a1: "h' \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes" assume a2: "\h document_ptr disc_nodes h'. h \ set_disconnected_nodes document_ptr disc_nodes \\<^sub>h h' \ h' \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes" assume "h'a \ set_disconnected_nodes owner_document (node_ptr # xb) \\<^sub>h h'" then have "node_ptr # xb = disc_nodes" using a2 a1 by (meson returns_result_eq) then show ?thesis by (meson list.set_intros(1)) qed qed qed end interpretation i_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M: l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_owner_document get_parent get_parent_locs remove_child remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M adopt_node_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs set_child_nodes set_child_nodes_locs remove heap_is_wellformed parent_child_rel by(auto simp add: l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare i_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] interpretation i_adopt_node_wf?: l_adopt_node_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_owner_document get_parent get_parent_locs remove_child remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs get_ancestors_di get_ancestors_di_locs adopt_node adopt_node_locs adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M adopt_node_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs set_child_nodes set_child_nodes_locs get_host get_host_locs get_disconnected_document get_disconnected_document_locs remove heap_is_wellformed parent_child_rel by(auto simp add: l_adopt_node_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_adopt_node_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] interpretation i_adopt_node_wf2?: l_adopt_node_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs set_child_nodes set_child_nodes_locs get_shadow_root get_shadow_root_locs set_disconnected_nodes set_disconnected_nodes_locs get_tag_name get_tag_name_locs get_owner_document get_parent get_parent_locs remove_child remove_child_locs remove known_ptrs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs get_root_node get_root_node_locs get_ancestors_di get_ancestors_di_locs adopt_node adopt_node_locs adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M adopt_node_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M by(auto simp add: l_adopt_node_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_adopt_node_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma adopt_node_wf_is_l_adopt_node_wf [instances]: "l_adopt_node_wf type_wf known_ptr heap_is_wellformed parent_child_rel get_child_nodes get_disconnected_nodes known_ptrs adopt_node" apply(auto simp add: l_adopt_node_wf_def l_adopt_node_wf_axioms_def instances)[1] using adopt_node_preserves_wellformedness apply blast using adopt_node_removes_child apply blast using adopt_node_node_in_disconnected_nodes apply blast using adopt_node_removes_first_child apply blast using adopt_node_document_in_heap apply blast done subsubsection \insert\_before\ locale l_insert_before_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes + l_get_disconnected_nodes + l_set_child_nodes_get_shadow_root + l_set_disconnected_nodes_get_shadow_root + l_set_child_nodes_get_tag_name + l_set_disconnected_nodes_get_tag_name + l_set_disconnected_nodes_get_disconnected_nodes + l_set_child_nodes_get_disconnected_nodes + l_set_disconnected_nodes_get_disconnected_nodes_wf + (* l_set_disconnected_nodes_get_ancestors_si + *) l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M _ _ _ _ _ _ get_ancestors_di get_ancestors_di_locs + (* l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + *) l_get_owner_document + l_adopt_node\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_adopt_node_wf + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_adopt_node_get_shadow_root + l_get_ancestors_di_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_remove_child_wf2 begin lemma insert_before_child_preserves: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ insert_before ptr node child \\<^sub>h h'" shows "type_wf h'" and "known_ptrs h'" and "heap_is_wellformed h'" proof - obtain ancestors reference_child owner_document h2 h3 disconnected_nodes_h2 where ancestors: "h \ get_ancestors_di ptr \\<^sub>r ancestors" and node_not_in_ancestors: "cast node \ set ancestors" and reference_child: "h \ (if Some node = child then a_next_sibling node else return child) \\<^sub>r reference_child" and owner_document: "h \ get_owner_document ptr \\<^sub>r owner_document" and h2: "h \ adopt_node owner_document node \\<^sub>h h2" and disconnected_nodes_h2: "h2 \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h2" and h3: "h2 \ set_disconnected_nodes owner_document (remove1 node disconnected_nodes_h2) \\<^sub>h h3" and h': "h3 \ a_insert_node ptr node reference_child \\<^sub>h h'" using assms(4) by(auto simp add: insert_before_def a_ensure_pre_insertion_validity_def elim!: bind_returns_heap_E bind_returns_result_E bind_returns_heap_E2[rotated, OF get_parent_pure, rotated] bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated] bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] bind_returns_heap_E2[rotated, OF get_ancestors_pure, rotated] bind_returns_heap_E2[rotated, OF next_sibling_pure, rotated] bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated] split: if_splits option.splits) obtain old_document parent_opt h'2 (* ancestors *) where (* "h \ get_ancestors_di (cast owner_document) \\<^sub>r ancestors" and "cast child \ set ancestors" and *) old_document: "h \ get_owner_document (cast node) \\<^sub>r old_document" and parent_opt: "h \ get_parent node \\<^sub>r parent_opt" and h'2: "h \ (case parent_opt of Some parent \ remove_child parent node | None \ return ()) \\<^sub>h h'2" and h2': "h'2 \ (if owner_document \ old_document then do { old_disc_nodes \ get_disconnected_nodes old_document; set_disconnected_nodes old_document (remove1 node old_disc_nodes); disc_nodes \ get_disconnected_nodes owner_document; set_disconnected_nodes owner_document (node # disc_nodes) } else do { return () }) \\<^sub>h h2" using h2 apply(auto simp add: adopt_node_def[unfolded CD.adopt_node_def] elim!: bind_returns_heap_E dest!: pure_returns_heap_eq[rotated, OF get_ancestors_di_pure])[1] apply(split if_splits) by(auto simp add: elim!: bind_returns_heap_E dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure] pure_returns_heap_eq[rotated, OF get_parent_pure]) have "type_wf h2" using \type_wf h\ using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF adopt_node_writes h2] using adopt_node_types_preserved by(auto simp add: reflp_def transp_def) then have "type_wf h3" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h3] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) then show "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF insert_node_writes h'] using set_child_nodes_types_preserved by(auto simp add: reflp_def transp_def) have "object_ptr_kinds h = object_ptr_kinds h2" using adopt_node_writes h2 apply(rule writes_small_big) using adopt_node_pointers_preserved by(auto simp add: reflp_def transp_def) moreover have "\ = object_ptr_kinds h3" using set_disconnected_nodes_writes h3 apply(rule writes_small_big) using set_disconnected_nodes_pointers_preserved by(auto simp add: reflp_def transp_def) moreover have "\ = object_ptr_kinds h'" using insert_node_writes h' apply(rule writes_small_big) using set_child_nodes_pointers_preserved by(auto simp add: reflp_def transp_def) ultimately show "known_ptrs h'" using \known_ptrs h\ known_ptrs_preserved by blast have "known_ptrs h2" using \known_ptrs h\ known_ptrs_preserved \object_ptr_kinds h = object_ptr_kinds h2\ by blast then have "known_ptrs h3" using known_ptrs_preserved \object_ptr_kinds h2 = object_ptr_kinds h3\ by blast have "known_ptr ptr" by (meson get_owner_document_ptr_in_heap is_OK_returns_result_I \known_ptrs h\ l_known_ptrs.known_ptrs_known_ptr l_known_ptrs_axioms owner_document) have object_ptr_kinds_M_eq3_h: "object_ptr_kinds h = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF adopt_node_writes h2]) using adopt_node_pointers_preserved apply blast by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h: "\ptrs. h \ object_ptr_kinds_M \\<^sub>r ptrs = h2 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs ) then have object_ptr_kinds_M_eq2_h: "|h \ object_ptr_kinds_M|\<^sub>r = |h2 \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h: "|h \ node_ptr_kinds_M|\<^sub>r = |h2 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have wellformed_h2: "heap_is_wellformed h2" using adopt_node_preserves_wellformedness[OF \heap_is_wellformed h\ h2] \known_ptrs h\ \type_wf h\ . have object_ptr_kinds_M_eq3_h2: "object_ptr_kinds h2 = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes h3]) unfolding a_remove_child_locs_def using set_disconnected_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h2: "\ptrs. h2 \ object_ptr_kinds_M \\<^sub>r ptrs = h3 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_M_eq2_h2: "|h2 \ object_ptr_kinds_M|\<^sub>r = |h3 \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h2: "|h2 \ node_ptr_kinds_M|\<^sub>r = |h3 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have document_ptr_kinds_eq2_h2: "|h2 \ document_ptr_kinds_M|\<^sub>r = |h3 \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_M_eq2_h2 document_ptr_kinds_M_eq by auto have object_ptr_kinds_M_eq3_h': "object_ptr_kinds h3 = object_ptr_kinds h'" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF insert_node_writes h']) unfolding a_remove_child_locs_def using set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h3: "\ptrs. h3 \ object_ptr_kinds_M \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_M_eq2_h3: "|h3 \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h3: "|h3 \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have document_ptr_kinds_eq2_h3: "|h3 \ document_ptr_kinds_M|\<^sub>r = |h' \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_M_eq2_h3 document_ptr_kinds_M_eq by auto have shadow_root_eq_h2: "\ptr' shadow_root. h \ get_shadow_root ptr' \\<^sub>r shadow_root = h2 \ get_shadow_root ptr' \\<^sub>r shadow_root" using get_shadow_root_reads adopt_node_writes h2 apply(rule reads_writes_preserved) using local.adopt_node_get_shadow_root by blast have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. owner_document \ doc_ptr \ h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by (auto simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers) then have disconnected_nodes_eq2_h2: "\doc_ptr. doc_ptr \ owner_document \ |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have disconnected_nodes_h3: "h3 \ get_disconnected_nodes owner_document \\<^sub>r remove1 node disconnected_nodes_h2" using h3 set_disconnected_nodes_get_disconnected_nodes by blast have disconnected_nodes_eq_h3: "\doc_ptr disc_nodes. h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads insert_node_writes h' apply(rule reads_writes_preserved) using set_child_nodes_get_disconnected_nodes by fast then have disconnected_nodes_eq2_h3: "\doc_ptr. |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. h2 \ get_child_nodes ptr' \\<^sub>r children = h3 \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by (auto simp add: set_disconnected_nodes_get_child_nodes) then have children_eq2_h2: "\ptr'. |h2 \ get_child_nodes ptr'|\<^sub>r = |h3 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have children_eq_h3: "\ptr' children. ptr \ ptr' \ h3 \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads insert_node_writes h' apply(rule reads_writes_preserved) by (auto simp add: set_child_nodes_get_child_nodes_different_pointers) then have children_eq2_h3: "\ptr'. ptr \ ptr' \ |h3 \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force obtain children_h3 where children_h3: "h3 \ get_child_nodes ptr \\<^sub>r children_h3" using h' a_insert_node_def by auto have children_h': "h' \ get_child_nodes ptr \\<^sub>r insert_before_list node reference_child children_h3" using h' \type_wf h3\ \known_ptr ptr\ by(auto simp add: a_insert_node_def elim!: bind_returns_heap_E2 dest!: set_child_nodes_get_child_nodes returns_result_eq[OF children_h3]) have shadow_root_eq_h: "\ptr' shadow_root_ptr_opt. h \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt = h2 \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt" using get_shadow_root_reads adopt_node_writes h2 apply(rule reads_writes_preserved) by(auto simp add: adopt_node_locs_def CD.adopt_node_locs_def CD.remove_child_locs_def set_child_nodes_get_shadow_root set_disconnected_nodes_get_shadow_root) then have shadow_root_eq2_h: "\ptr'. |h \ get_shadow_root ptr'|\<^sub>r = |h2 \ get_shadow_root ptr'|\<^sub>r" by (meson select_result_eq) have shadow_root_eq_h2: "\ptr' shadow_root_ptr_opt. h2 \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt = h3 \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt" using get_shadow_root_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by(auto simp add: adopt_node_locs_def remove_child_locs_def set_child_nodes_get_shadow_root set_disconnected_nodes_get_shadow_root) then have shadow_root_eq2_h2: "\ptr'. |h2 \ get_shadow_root ptr'|\<^sub>r = |h3 \ get_shadow_root ptr'|\<^sub>r" by (meson select_result_eq) have shadow_root_eq_h3: "\ptr' shadow_root_ptr_opt. h3 \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt = h' \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt" using get_shadow_root_reads insert_node_writes h' apply(rule reads_writes_preserved) by(auto simp add: adopt_node_locs_def remove_child_locs_def set_child_nodes_get_shadow_root set_disconnected_nodes_get_shadow_root) then have shadow_root_eq2_h3: "\ptr'. |h3 \ get_shadow_root ptr'|\<^sub>r = |h' \ get_shadow_root ptr'|\<^sub>r" by (meson select_result_eq) have tag_name_eq_h2: "\ptr' tag. h2 \ get_tag_name ptr' \\<^sub>r tag = h3 \ get_tag_name ptr' \\<^sub>r tag" using get_tag_name_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by(auto simp add: adopt_node_locs_def remove_child_locs_def set_child_nodes_get_tag_name set_disconnected_nodes_get_tag_name) then have tag_name_eq2_h2: "\ptr'. |h2 \ get_tag_name ptr'|\<^sub>r = |h3 \ get_tag_name ptr'|\<^sub>r" by (meson select_result_eq) have tag_name_eq_h3: "\ptr' tag. h3 \ get_tag_name ptr' \\<^sub>r tag = h' \ get_tag_name ptr' \\<^sub>r tag" using get_tag_name_reads insert_node_writes h' apply(rule reads_writes_preserved) by(auto simp add: adopt_node_locs_def remove_child_locs_def set_child_nodes_get_tag_name set_disconnected_nodes_get_tag_name) then have tag_name_eq2_h3: "\ptr'. |h3 \ get_tag_name ptr'|\<^sub>r = |h' \ get_tag_name ptr'|\<^sub>r" by (meson select_result_eq) have object_ptr_kinds_eq_hx: "object_ptr_kinds h = object_ptr_kinds h'2" using h'2 apply(simp split: option.splits) apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF CD.remove_child_writes]) using CD.remove_child_pointers_preserved by (auto simp add: reflp_def transp_def) have document_ptr_kinds_eq_hx: "document_ptr_kinds h = document_ptr_kinds h'2" using object_ptr_kinds_eq_hx by(auto simp add: document_ptr_kinds_def document_ptr_kinds_def) have shadow_root_ptr_kinds_eq_hx: "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'2" using object_ptr_kinds_eq_hx by(auto simp add: shadow_root_ptr_kinds_def document_ptr_kinds_def) have element_ptr_kinds_eq_hx: "element_ptr_kinds h = element_ptr_kinds h'2" using object_ptr_kinds_eq_hx by(auto simp add: element_ptr_kinds_def node_ptr_kinds_def) have object_ptr_kinds_eq_h: "object_ptr_kinds h = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF adopt_node_writes h2]) unfolding adopt_node_locs_def CD.adopt_node_locs_def CD.remove_child_locs_def using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def split: if_splits) have document_ptr_kinds_eq_h: "document_ptr_kinds h = document_ptr_kinds h2" using object_ptr_kinds_eq_h by(auto simp add: document_ptr_kinds_def document_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h: "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h2" using object_ptr_kinds_eq_h by(auto simp add: shadow_root_ptr_kinds_def document_ptr_kinds_def) have element_ptr_kinds_eq_h: "element_ptr_kinds h = element_ptr_kinds h2" using object_ptr_kinds_eq_h by(auto simp add: element_ptr_kinds_def node_ptr_kinds_def) have object_ptr_kinds_eq_h2: "object_ptr_kinds h2 = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes h3]) unfolding adopt_node_locs_def remove_child_locs_def using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def split: if_splits) have document_ptr_kinds_eq_h2: "document_ptr_kinds h2 = document_ptr_kinds h3" using object_ptr_kinds_eq_h2 by(auto simp add: document_ptr_kinds_def document_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h2: "shadow_root_ptr_kinds h2 = shadow_root_ptr_kinds h3" using object_ptr_kinds_eq_h2 by(auto simp add: shadow_root_ptr_kinds_def document_ptr_kinds_def) have element_ptr_kinds_eq_h2: "element_ptr_kinds h2 = element_ptr_kinds h3" using object_ptr_kinds_eq_h2 by(auto simp add: element_ptr_kinds_def node_ptr_kinds_def) have object_ptr_kinds_eq_h3: "object_ptr_kinds h3 = object_ptr_kinds h'" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF insert_node_writes h']) using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def split: if_splits) have document_ptr_kinds_eq_h3: "document_ptr_kinds h3 = document_ptr_kinds h'" using object_ptr_kinds_eq_h3 by(auto simp add: document_ptr_kinds_def document_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h3: "shadow_root_ptr_kinds h3 = shadow_root_ptr_kinds h'" using object_ptr_kinds_eq_h3 by(auto simp add: shadow_root_ptr_kinds_def document_ptr_kinds_def) have element_ptr_kinds_eq_h3: "element_ptr_kinds h3 = element_ptr_kinds h'" using object_ptr_kinds_eq_h3 by(auto simp add: element_ptr_kinds_def node_ptr_kinds_def) have wellformed_h'2: "heap_is_wellformed h'2" using h'2 remove_child_heap_is_wellformed_preserved assms by (metis (no_types, lifting) assms(1) option.case_eq_if pure_returns_heap_eq return_pure) have "known_ptrs h'2" using \known_ptrs h\ known_ptrs_preserved \object_ptr_kinds h = object_ptr_kinds h'2\ by blast have "type_wf h'2" using \type_wf h\ h'2 apply(auto split: option.splits)[1] apply(drule writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF CD.remove_child_writes]) using CD.remove_child_types_preserved by(auto simp add: reflp_def transp_def ) have ptr_in_heap: "ptr |\| object_ptr_kinds h3" using children_h3 get_child_nodes_ptr_in_heap by blast have node_in_heap: "node |\| node_ptr_kinds h" using h2 adopt_node_child_in_heap by fast have child_not_in_any_children: "\p children. h2 \ get_child_nodes p \\<^sub>r children \ node \ set children" using \heap_is_wellformed h\ h2 adopt_node_removes_child \type_wf h\ \known_ptrs h\ by auto have "node \ set disconnected_nodes_h2" using disconnected_nodes_h2 h2 adopt_node_node_in_disconnected_nodes assms(1) \type_wf h\ \known_ptrs h\ by blast have node_not_in_disconnected_nodes: "\d. d |\| document_ptr_kinds h3 \ node \ set |h3 \ get_disconnected_nodes d|\<^sub>r" proof - fix d assume "d |\| document_ptr_kinds h3" show "node \ set |h3 \ get_disconnected_nodes d|\<^sub>r" proof (cases "d = owner_document") case True then show ?thesis using disconnected_nodes_h2 wellformed_h2 h3 remove_from_disconnected_nodes_removes wellformed_h2 \d |\| document_ptr_kinds h3\ disconnected_nodes_h3 by fastforce next case False then have "set |h2 \ get_disconnected_nodes d|\<^sub>r \ set |h2 \ get_disconnected_nodes owner_document|\<^sub>r = {}" using distinct_concat_map_E(1) wellformed_h2 by (metis (no_types, lifting) \d |\| document_ptr_kinds h3\ \type_wf h2\ disconnected_nodes_h2 document_ptr_kinds_M_def document_ptr_kinds_eq2_h2 l_ptr_kinds_M.ptr_kinds_ptr_kinds_M local.get_disconnected_nodes_ok local.heap_is_wellformed_one_disc_parent returns_result_select_result select_result_I2) then show ?thesis using disconnected_nodes_eq2_h2[OF False] \node \ set disconnected_nodes_h2\ disconnected_nodes_h2 by fastforce qed qed have "cast node \ ptr" using ancestors node_not_in_ancestors get_ancestors_ptr by fast have "a_host_shadow_root_rel h = a_host_shadow_root_rel h2" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h shadow_root_eq2_h) have "a_host_shadow_root_rel h2 = a_host_shadow_root_rel h3" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h2 shadow_root_eq2_h2) have "a_host_shadow_root_rel h3 = a_host_shadow_root_rel h'" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h3 shadow_root_eq2_h3) have "parent_child_rel h2 \ parent_child_rel h" proof - have "parent_child_rel h'2 \ parent_child_rel h" using h'2 parent_opt proof (induct parent_opt) case None then show ?case by simp next case (Some parent) then have h'2: "h \ remove_child parent node \\<^sub>h h'2" by auto then have "parent |\| object_ptr_kinds h" using CD.remove_child_ptr_in_heap by blast have child_nodes_eq_h: "\ptr children. parent \ ptr \ h \ get_child_nodes ptr \\<^sub>r children = h'2 \ get_child_nodes ptr \\<^sub>r children" using get_child_nodes_reads CD.remove_child_writes h'2 apply(rule reads_writes_preserved) apply(auto simp add: CD.remove_child_locs_def)[1] by (simp add: set_child_nodes_get_child_nodes_different_pointers) moreover obtain children where children: "h \ get_child_nodes parent \\<^sub>r children" using Some local.get_parent_child_dual by blast moreover obtain children_h'2 where children_h'2: "h'2 \ get_child_nodes parent \\<^sub>r children_h'2" using object_ptr_kinds_eq_hx calculation(2) \parent |\| object_ptr_kinds h\ get_child_nodes_ok by (metis \type_wf h'2\ assms(3) is_OK_returns_result_E local.known_ptrs_known_ptr) ultimately show ?thesis using object_ptr_kinds_eq_h h2 apply(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq_hx split: option.splits)[1] apply(case_tac "parent = a") using CD.remove_child_children_subset apply (metis (no_types, lifting) assms(2) assms(3) contra_subsetD h'2 select_result_I2) by (metis select_result_eq) qed moreover have "parent_child_rel h2 = parent_child_rel h'2" proof(cases "owner_document = old_document") case True then show ?thesis using h2' by simp next case False then obtain h'3 old_disc_nodes disc_nodes_document_ptr_h'3 where docs_neq: "owner_document \ old_document" and old_disc_nodes: "h'2 \ get_disconnected_nodes old_document \\<^sub>r old_disc_nodes" and h'3: "h'2 \ set_disconnected_nodes old_document (remove1 node old_disc_nodes) \\<^sub>h h'3" and disc_nodes_document_ptr_h3: "h'3 \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes_document_ptr_h'3" and h2': "h'3 \ set_disconnected_nodes owner_document (node # disc_nodes_document_ptr_h'3) \\<^sub>h h2" using h2' by(auto elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] ) have object_ptr_kinds_h'2_eq3: "object_ptr_kinds h'2 = object_ptr_kinds h'3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes h'3]) using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h'2: "\ptrs. h'2 \ object_ptr_kinds_M \\<^sub>r ptrs = h'3 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_eq_h'2: "|h'2 \ object_ptr_kinds_M|\<^sub>r = |h'3 \ object_ptr_kinds_M|\<^sub>r" by(simp) then have node_ptr_kinds_eq_h'2: "|h'2 \ node_ptr_kinds_M|\<^sub>r = |h'3 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast then have node_ptr_kinds_eq3_h'2: "node_ptr_kinds h'2 = node_ptr_kinds h'3" by auto have document_ptr_kinds_eq2_h'2: "|h'2 \ document_ptr_kinds_M|\<^sub>r = |h'3 \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_eq_h'2 document_ptr_kinds_M_eq by auto then have document_ptr_kinds_eq3_h'2: "document_ptr_kinds h'2 = document_ptr_kinds h'3" using object_ptr_kinds_eq_h'2 document_ptr_kinds_M_eq by auto have children_eq_h'2: "\ptr children. h'2 \ get_child_nodes ptr \\<^sub>r children = h'3 \ get_child_nodes ptr \\<^sub>r children" using get_child_nodes_reads set_disconnected_nodes_writes h'3 apply(rule reads_writes_preserved) by (simp add: set_disconnected_nodes_get_child_nodes) then have children_eq2_h'2: "\ptr. |h'2 \ get_child_nodes ptr|\<^sub>r = |h'3 \ get_child_nodes ptr|\<^sub>r" using select_result_eq by force have object_ptr_kinds_h'3_eq3: "object_ptr_kinds h'3 = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h2. object_ptr_kinds h = object_ptr_kinds h2", OF set_disconnected_nodes_writes h2']) using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h'3: "\ptrs. h'3 \ object_ptr_kinds_M \\<^sub>r ptrs = h2 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_eq_h'3: "|h'3 \ object_ptr_kinds_M|\<^sub>r = |h2 \ object_ptr_kinds_M|\<^sub>r" by(simp) then have node_ptr_kinds_eq_h'3: "|h'3 \ node_ptr_kinds_M|\<^sub>r = |h2 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast then have node_ptr_kinds_eq3_h'3: "node_ptr_kinds h'3 = node_ptr_kinds h2" by auto have document_ptr_kinds_eq2_h'3: "|h'3 \ document_ptr_kinds_M|\<^sub>r = |h2 \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_eq_h'3 document_ptr_kinds_M_eq by auto then have document_ptr_kinds_eq3_h'3: "document_ptr_kinds h'3 = document_ptr_kinds h2" using object_ptr_kinds_eq_h'3 document_ptr_kinds_M_eq by auto have children_eq_h'3: "\ptr children. h'3 \ get_child_nodes ptr \\<^sub>r children = h2 \ get_child_nodes ptr \\<^sub>r children" using get_child_nodes_reads set_disconnected_nodes_writes h2' apply(rule reads_writes_preserved) by (simp add: set_disconnected_nodes_get_child_nodes) then have children_eq2_h'3: "\ptr. |h'3 \ get_child_nodes ptr|\<^sub>r = |h2 \ get_child_nodes ptr|\<^sub>r" using select_result_eq by force show ?thesis by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_h'2_eq3 object_ptr_kinds_h'3_eq3 children_eq2_h'3 children_eq2_h'2) qed ultimately show ?thesis by simp qed have "parent_child_rel h2 = parent_child_rel h3" by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_M_eq3_h2 children_eq2_h2) have "parent_child_rel h' = insert (ptr, cast node) ((parent_child_rel h3))" using children_h3 children_h' ptr_in_heap apply(auto simp add: CD.parent_child_rel_def object_ptr_kinds_M_eq3_h' children_eq2_h3 insert_before_list_node_in_set)[1] apply (metis (no_types, lifting) children_eq2_h3 insert_before_list_in_set select_result_I2) by (metis (no_types, lifting) children_eq2_h3 imageI insert_before_list_in_set select_result_I2) have "a_ptr_disconnected_node_rel h3 \ a_ptr_disconnected_node_rel h" proof - have "a_ptr_disconnected_node_rel h3 = a_ptr_disconnected_node_rel h2 - {(cast owner_document, cast node)}" apply(auto simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq_h2)[1] apply(case_tac "aa = owner_document") using disconnected_nodes_h2 disconnected_nodes_h3 notin_set_remove1 apply fastforce using disconnected_nodes_eq2_h2 apply auto[1] using node_not_in_disconnected_nodes apply blast apply(case_tac "aa = owner_document") using disconnected_nodes_h2 disconnected_nodes_h3 notin_set_remove1 apply fastforce using disconnected_nodes_eq2_h2 apply auto[1] apply(case_tac "aa = owner_document") using disconnected_nodes_h2 disconnected_nodes_h3 notin_set_remove1 apply fastforce using disconnected_nodes_eq2_h2 apply auto[1] done then have "a_ptr_disconnected_node_rel h'2 \ a_ptr_disconnected_node_rel h \ {(cast old_document, cast node)}" using h'2 parent_opt proof (induct parent_opt) case None then show ?case by auto next case (Some parent) then have h'2: "h \ remove_child parent node \\<^sub>h h'2" by auto then have "parent |\| object_ptr_kinds h" using CD.remove_child_ptr_in_heap by blast obtain children_h h''2 disconnected_nodes_h where owner_document: "h \ get_owner_document (cast node) \\<^sub>r old_document" and children_h: "h \ get_child_nodes parent \\<^sub>r children_h" and child_in_children_h: "node \ set children_h" and disconnected_nodes_h: "h \ get_disconnected_nodes old_document \\<^sub>r disconnected_nodes_h" and h''2: "h \ set_disconnected_nodes old_document (node # disconnected_nodes_h) \\<^sub>h h''2" and h'2: "h''2 \ set_child_nodes parent (remove1 node children_h) \\<^sub>h h'2" using h'2 old_document apply(auto simp add: CD.remove_child_def elim!: bind_returns_heap_E dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure] pure_returns_heap_eq[rotated, OF get_child_nodes_pure] split: if_splits)[1] using pure_returns_heap_eq returns_result_eq by fastforce have disconnected_nodes_eq: "\ptr' disc_nodes. ptr' \ old_document \ h \ get_disconnected_nodes ptr' \\<^sub>r disc_nodes = h''2 \ get_disconnected_nodes ptr' \\<^sub>r disc_nodes" using local.get_disconnected_nodes_reads set_disconnected_nodes_writes h''2 apply(rule reads_writes_preserved) by (metis local.set_disconnected_nodes_get_disconnected_nodes_different_pointers) then have disconnected_nodes_eq2: "\ptr'. ptr' \ old_document \ |h \ get_disconnected_nodes ptr'|\<^sub>r = |h''2 \ get_disconnected_nodes ptr'|\<^sub>r" by (meson select_result_eq) have "h''2 \ get_disconnected_nodes old_document \\<^sub>r node # disconnected_nodes_h" using h''2 local.set_disconnected_nodes_get_disconnected_nodes by blast have disconnected_nodes_eq_h2: "\ptr' disc_nodes. h''2 \ get_disconnected_nodes ptr' \\<^sub>r disc_nodes = h'2 \ get_disconnected_nodes ptr' \\<^sub>r disc_nodes" using local.get_disconnected_nodes_reads set_child_nodes_writes h'2 apply(rule reads_writes_preserved) by (metis local.set_child_nodes_get_disconnected_nodes) then have disconnected_nodes_eq2_h2: "\ptr'. |h''2 \ get_disconnected_nodes ptr'|\<^sub>r = |h'2 \ get_disconnected_nodes ptr'|\<^sub>r" by (meson select_result_eq) show ?case apply(auto simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq_hx \\ptr'. |h''2 \ get_disconnected_nodes ptr'|\<^sub>r = |h'2 \ get_disconnected_nodes ptr'|\<^sub>r\[symmetric])[1] apply(case_tac "aa = old_document") using disconnected_nodes_h \h''2 \ get_disconnected_nodes old_document \\<^sub>r node # disconnected_nodes_h\ apply(auto)[1] apply(auto dest!: disconnected_nodes_eq2)[1] apply(case_tac "aa = old_document") using disconnected_nodes_h \h''2 \ get_disconnected_nodes old_document \\<^sub>r node # disconnected_nodes_h\ apply(auto)[1] apply(auto dest!: disconnected_nodes_eq2)[1] done qed show ?thesis proof(cases "owner_document = old_document") case True then have "a_ptr_disconnected_node_rel h'2 = a_ptr_disconnected_node_rel h2" using h2' by(auto simp add: a_ptr_disconnected_node_rel_def) then show ?thesis using \a_ptr_disconnected_node_rel h3 = a_ptr_disconnected_node_rel h2 - {(cast owner_document, cast node)}\ using True \local.a_ptr_disconnected_node_rel h'2 \ local.a_ptr_disconnected_node_rel h \ {(cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r old_document, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node)}\ by auto next case False then obtain h'3 old_disc_nodes disc_nodes_document_ptr_h'3 where docs_neq: "owner_document \ old_document" and old_disc_nodes: "h'2 \ get_disconnected_nodes old_document \\<^sub>r old_disc_nodes" and h'3: "h'2 \ set_disconnected_nodes old_document (remove1 node old_disc_nodes) \\<^sub>h h'3" and disc_nodes_document_ptr_h3: "h'3 \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes_document_ptr_h'3" and h2': "h'3 \ set_disconnected_nodes owner_document (node # disc_nodes_document_ptr_h'3) \\<^sub>h h2" using h2' by(auto elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] ) have object_ptr_kinds_h'2_eq3: "object_ptr_kinds h'2 = object_ptr_kinds h'3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes h'3]) using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h'2: "\ptrs. h'2 \ object_ptr_kinds_M \\<^sub>r ptrs = h'3 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_eq_h'2: "|h'2 \ object_ptr_kinds_M|\<^sub>r = |h'3 \ object_ptr_kinds_M|\<^sub>r" by(simp) then have node_ptr_kinds_eq_h'2: "|h'2 \ node_ptr_kinds_M|\<^sub>r = |h'3 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast then have node_ptr_kinds_eq3_h'2: "node_ptr_kinds h'2 = node_ptr_kinds h'3" by auto have document_ptr_kinds_eq2_h'2: "|h'2 \ document_ptr_kinds_M|\<^sub>r = |h'3 \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_eq_h'2 document_ptr_kinds_M_eq by auto then have document_ptr_kinds_eq3_h'2: "document_ptr_kinds h'2 = document_ptr_kinds h'3" using object_ptr_kinds_eq_h'2 document_ptr_kinds_M_eq by auto have disconnected_nodes_eq: "\ptr' disc_nodes. ptr' \ old_document \ h'2 \ get_disconnected_nodes ptr' \\<^sub>r disc_nodes = h'3 \ get_disconnected_nodes ptr' \\<^sub>r disc_nodes" using local.get_disconnected_nodes_reads set_disconnected_nodes_writes h'3 apply(rule reads_writes_preserved) by (metis local.set_disconnected_nodes_get_disconnected_nodes_different_pointers) then have disconnected_nodes_eq2: "\ptr'. ptr' \ old_document \ |h'2 \ get_disconnected_nodes ptr'|\<^sub>r = |h'3 \ get_disconnected_nodes ptr'|\<^sub>r" by (meson select_result_eq) have "h'3 \ get_disconnected_nodes old_document \\<^sub>r (remove1 node old_disc_nodes)" using h'3 local.set_disconnected_nodes_get_disconnected_nodes by blast have object_ptr_kinds_h'3_eq3: "object_ptr_kinds h'3 = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h2. object_ptr_kinds h = object_ptr_kinds h2", OF set_disconnected_nodes_writes h2']) using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_M_eq_h'3: "\ptrs. h'3 \ object_ptr_kinds_M \\<^sub>r ptrs = h2 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_eq_h'3: "|h'3 \ object_ptr_kinds_M|\<^sub>r = |h2 \ object_ptr_kinds_M|\<^sub>r" by(simp) then have node_ptr_kinds_eq_h'3: "|h'3 \ node_ptr_kinds_M|\<^sub>r = |h2 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast then have node_ptr_kinds_eq3_h'3: "node_ptr_kinds h'3 = node_ptr_kinds h2" by auto have document_ptr_kinds_eq2_h'3: "|h'3 \ document_ptr_kinds_M|\<^sub>r = |h2 \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_eq_h'3 document_ptr_kinds_M_eq by auto then have document_ptr_kinds_eq3_h'3: "document_ptr_kinds h'3 = document_ptr_kinds h2" using object_ptr_kinds_eq_h'3 document_ptr_kinds_M_eq by auto have disc_nodes_eq_h'3: "\ptr disc_nodes. h'3 \ get_child_nodes ptr \\<^sub>r disc_nodes = h2 \ get_child_nodes ptr \\<^sub>r disc_nodes" using get_child_nodes_reads set_disconnected_nodes_writes h2' apply(rule reads_writes_preserved) by (simp add: set_disconnected_nodes_get_child_nodes) then have disc_nodes_eq2_h'3: "\ptr. |h'3 \ get_child_nodes ptr|\<^sub>r = |h2 \ get_child_nodes ptr|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq: "\ptr' disc_nodes. ptr' \ owner_document \ h'3 \ get_disconnected_nodes ptr' \\<^sub>r disc_nodes = h2 \ get_disconnected_nodes ptr' \\<^sub>r disc_nodes" using local.get_disconnected_nodes_reads set_disconnected_nodes_writes h2' apply(rule reads_writes_preserved) by (metis local.set_disconnected_nodes_get_disconnected_nodes_different_pointers) then have disconnected_nodes_eq2': "\ptr'. ptr' \ owner_document \ |h'3 \ get_disconnected_nodes ptr'|\<^sub>r = |h2 \ get_disconnected_nodes ptr'|\<^sub>r" by (meson select_result_eq) have "h2 \ get_disconnected_nodes owner_document \\<^sub>r (node # disc_nodes_document_ptr_h'3)" using h2' local.set_disconnected_nodes_get_disconnected_nodes by blast have "a_ptr_disconnected_node_rel h'3 = a_ptr_disconnected_node_rel h'2 - {(cast old_document, cast node)}" apply(auto simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq2_h'2[simplified])[1] apply(case_tac "aa = old_document") using \h'3 \ get_disconnected_nodes old_document \\<^sub>r (remove1 node old_disc_nodes)\ notin_set_remove1 old_disc_nodes apply fastforce apply(auto dest!: disconnected_nodes_eq2)[1] using \h'3 \ get_disconnected_nodes old_document \\<^sub>r remove1 node old_disc_nodes\ h'3 local.remove_from_disconnected_nodes_removes old_disc_nodes wellformed_h'2 apply auto[1] defer apply(case_tac "aa = old_document") using \h'3 \ get_disconnected_nodes old_document \\<^sub>r (remove1 node old_disc_nodes)\ notin_set_remove1 old_disc_nodes apply fastforce apply(auto dest!: disconnected_nodes_eq2)[1] apply(case_tac "aa = old_document") using \h'3 \ get_disconnected_nodes old_document \\<^sub>r (remove1 node old_disc_nodes)\ notin_set_remove1 old_disc_nodes apply fastforce apply(auto dest!: disconnected_nodes_eq2)[1] done moreover have "a_ptr_disconnected_node_rel h2 = a_ptr_disconnected_node_rel h'3 \ {(cast owner_document, cast node)}" apply(auto simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq2_h'3[simplified])[1] apply(case_tac "aa = owner_document") apply(simp) apply(auto dest!: disconnected_nodes_eq2')[1] apply(case_tac "aa = owner_document") using \h2 \ get_disconnected_nodes owner_document \\<^sub>r node # disc_nodes_document_ptr_h'3\ disc_nodes_document_ptr_h3 apply auto[1] apply(auto dest!: disconnected_nodes_eq2')[1] using \node \ set disconnected_nodes_h2\ disconnected_nodes_h2 local.a_ptr_disconnected_node_rel_def local.a_ptr_disconnected_node_rel_disconnected_node apply blast defer apply(case_tac "aa = owner_document") using \h2 \ get_disconnected_nodes owner_document \\<^sub>r node # disc_nodes_document_ptr_h'3\ disc_nodes_document_ptr_h3 apply auto[1] apply(auto dest!: disconnected_nodes_eq2')[1] done ultimately show ?thesis using \a_ptr_disconnected_node_rel h3 = a_ptr_disconnected_node_rel h2 - {(cast owner_document, cast node)}\ using \a_ptr_disconnected_node_rel h'2 \ a_ptr_disconnected_node_rel h \ {(cast old_document, cast node)}\ by blast qed qed have "(cast node, ptr) \ (parent_child_rel h \ a_host_shadow_root_rel h \ a_ptr_disconnected_node_rel h)\<^sup>*" using h2 apply(auto simp add: adopt_node_def elim!: bind_returns_heap_E2 split: if_splits)[1] using ancestors assms(1) assms(2) assms(3) local.get_ancestors_di_parent_child_a_host_shadow_root_rel node_not_in_ancestors by blast then have "(cast node, ptr) \ (parent_child_rel h3 \ a_host_shadow_root_rel h3 \ a_ptr_disconnected_node_rel h3)\<^sup>*" apply(simp add: \a_host_shadow_root_rel h = a_host_shadow_root_rel h2\ \a_host_shadow_root_rel h2 = a_host_shadow_root_rel h3\) apply(simp add: \parent_child_rel h2 = parent_child_rel h3\[symmetric]) using \parent_child_rel h2 \ parent_child_rel h\ \a_ptr_disconnected_node_rel h3 \ a_ptr_disconnected_node_rel h\ by (smt Un_assoc in_rtrancl_UnI sup.orderE sup_left_commute) have "CD.a_acyclic_heap h'" proof - have "acyclic (parent_child_rel h2)" using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def) then have "acyclic (parent_child_rel h3)" by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_M_eq3_h2 children_eq2_h2) moreover have "cast node \ {x. (x, ptr) \ (parent_child_rel h3)\<^sup>*}" by (meson \(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node, ptr) \ (parent_child_rel h3 \ local.a_host_shadow_root_rel h3 \ local.a_ptr_disconnected_node_rel h3)\<^sup>*\ in_rtrancl_UnI mem_Collect_eq) ultimately show ?thesis using \parent_child_rel h' = insert (ptr, cast node) ((parent_child_rel h3))\ by(auto simp add: CD.acyclic_heap_def) qed moreover have "CD.a_all_ptrs_in_heap h2" using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) have "CD.a_all_ptrs_in_heap h'" proof - have "CD.a_all_ptrs_in_heap h3" using \CD.a_all_ptrs_in_heap h2\ apply(auto simp add: CD.a_all_ptrs_in_heap_def object_ptr_kinds_M_eq2_h2 node_ptr_kinds_eq2_h2 children_eq_h2)[1] apply (metis \known_ptrs h3\ \type_wf h3\ children_eq_h2 l_heap_is_wellformed.heap_is_wellformed_children_in_heap local.get_child_nodes_ok local.known_ptrs_known_ptr local.l_heap_is_wellformed_axioms node_ptr_kinds_commutes object_ptr_kinds_eq_h2 returns_result_select_result wellformed_h2) by (metis (mono_tags, hide_lams) disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3 document_ptr_kinds_eq_h2 finite_set_in node_ptr_kinds_commutes object_ptr_kinds_eq_h2 select_result_I2 set_remove1_subset subsetD) have "set children_h3 \ set |h' \ node_ptr_kinds_M|\<^sub>r" using children_h3 \CD.a_all_ptrs_in_heap h3\ apply(auto simp add: CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq2_h3)[1] using children_eq_h2 local.heap_is_wellformed_children_in_heap node_ptr_kinds_eq2_h2 node_ptr_kinds_eq2_h3 wellformed_h2 by auto then have "set (insert_before_list node reference_child children_h3) \ set |h' \ node_ptr_kinds_M|\<^sub>r" using node_in_heap apply(auto simp add: node_ptr_kinds_eq2_h node_ptr_kinds_eq2_h2 node_ptr_kinds_eq2_h3)[1] by (metis (no_types, hide_lams) contra_subsetD finite_set_in insert_before_list_in_set node_ptr_kinds_commutes object_ptr_kinds_M_eq3_h object_ptr_kinds_M_eq3_h' object_ptr_kinds_M_eq3_h2) then show ?thesis using \CD.a_all_ptrs_in_heap h3\ apply(auto simp add: object_ptr_kinds_M_eq3_h' CD.a_all_ptrs_in_heap_def node_ptr_kinds_def node_ptr_kinds_eq2_h3 disconnected_nodes_eq_h3)[1] apply (metis (no_types, lifting) children_eq2_h3 children_h' finite_set_in select_result_I2 subsetD) by (metis (no_types, lifting) disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 in_mono notin_fset) qed moreover have "CD.a_distinct_lists h2" using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def ) then have "CD.a_distinct_lists h3" proof(auto simp add: CD.a_distinct_lists_def object_ptr_kinds_M_eq2_h2 document_ptr_kinds_eq2_h2 children_eq2_h2 intro!: distinct_concat_map_I) fix x assume 1: "x |\| document_ptr_kinds h3" and 2: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" show "distinct |h3 \ get_disconnected_nodes x|\<^sub>r" using distinct_concat_map_E(2)[OF 2] select_result_I2[OF disconnected_nodes_h3] disconnected_nodes_eq2_h2 select_result_I2[OF disconnected_nodes_h2] 1 by (metis (full_types) distinct_remove1 finite_fset fmember.rep_eq set_sorted_list_of_set) next fix x y xa assume 1: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" and 2: "x |\| document_ptr_kinds h3" and 3: "y |\| document_ptr_kinds h3" and 4: "x \ y" and 5: "xa \ set |h3 \ get_disconnected_nodes x|\<^sub>r" and 6: "xa \ set |h3 \ get_disconnected_nodes y|\<^sub>r" show False proof (cases "x = owner_document") case True then have "y \ owner_document" using 4 by simp show ?thesis using distinct_concat_map_E(1)[OF 1] using 2 3 4 5 6 select_result_I2[OF disconnected_nodes_h3] select_result_I2[OF disconnected_nodes_h2] apply(auto simp add: True disconnected_nodes_eq2_h2[OF \y \ owner_document\])[1] by (metis (no_types, hide_lams) disconnected_nodes_eq2_h2 disjoint_iff_not_equal notin_set_remove1) next case False then show ?thesis proof (cases "y = owner_document") case True then show ?thesis using distinct_concat_map_E(1)[OF 1] using 2 3 4 5 6 select_result_I2[OF disconnected_nodes_h3] select_result_I2[OF disconnected_nodes_h2] apply(auto simp add: True disconnected_nodes_eq2_h2[OF \x \ owner_document\])[1] by (metis (no_types, hide_lams) disconnected_nodes_eq2_h2 disjoint_iff_not_equal notin_set_remove1) next case False then show ?thesis using distinct_concat_map_E(1)[OF 1, simplified, OF 2 3 4] 5 6 using disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3 disjoint_iff_not_equal finite_fset fmember.rep_eq notin_set_remove1 select_result_I2 set_sorted_list_of_set by (metis (no_types, lifting)) qed qed next fix x xa xb assume 1: "(\x\fset (object_ptr_kinds h3). set |h3 \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h3). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" and 2: "xa |\| object_ptr_kinds h3" and 3: "x \ set |h3 \ get_child_nodes xa|\<^sub>r" and 4: "xb |\| document_ptr_kinds h3" and 5: "x \ set |h3 \ get_disconnected_nodes xb|\<^sub>r" have 6: "set |h3 \ get_child_nodes xa|\<^sub>r \ set |h2 \ get_disconnected_nodes xb|\<^sub>r = {}" using 1 2 4 by (metis \type_wf h2\ children_eq2_h2 document_ptr_kinds_commutes \known_ptrs h\ local.get_child_nodes_ok local.get_disconnected_nodes_ok local.heap_is_wellformed_children_disc_nodes_different local.known_ptrs_known_ptr object_ptr_kinds_M_eq3_h object_ptr_kinds_M_eq3_h2 returns_result_select_result wellformed_h2) show False proof (cases "xb = owner_document") case True then show ?thesis using select_result_I2[OF disconnected_nodes_h3,folded select_result_I2[OF disconnected_nodes_h2]] by (metis (no_types, lifting) "3" "5" "6" disjoint_iff_not_equal notin_set_remove1) next case False show ?thesis using 2 3 4 5 6 unfolding disconnected_nodes_eq2_h2[OF False] by auto qed qed then have "CD.a_distinct_lists h'" proof(auto simp add: CD.a_distinct_lists_def document_ptr_kinds_eq2_h3 object_ptr_kinds_M_eq2_h3 disconnected_nodes_eq2_h3 intro!: distinct_concat_map_I) fix x assume 1: "distinct (concat (map (\ptr. |h3 \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h')))))" and 2: "x |\| object_ptr_kinds h'" have 3: "\p. p |\| object_ptr_kinds h' \ distinct |h3 \ get_child_nodes p|\<^sub>r" using 1 by (auto elim: distinct_concat_map_E) show "distinct |h' \ get_child_nodes x|\<^sub>r" proof(cases "ptr = x") case True show ?thesis using 3[OF 2] children_h3 children_h' by(auto simp add: True insert_before_list_distinct dest: child_not_in_any_children[unfolded children_eq_h2]) next case False show ?thesis using children_eq2_h3[OF False] 3[OF 2] by auto qed next fix x y xa assume 1: "distinct (concat (map (\ptr. |h3 \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h')))))" and 2: "x |\| object_ptr_kinds h'" and 3: "y |\| object_ptr_kinds h'" and 4: "x \ y" and 5: "xa \ set |h' \ get_child_nodes x|\<^sub>r" and 6: "xa \ set |h' \ get_child_nodes y|\<^sub>r" have 7:"set |h3 \ get_child_nodes x|\<^sub>r \ set |h3 \ get_child_nodes y|\<^sub>r = {}" using distinct_concat_map_E(1)[OF 1] 2 3 4 by auto show False proof (cases "ptr = x") case True then have "ptr \ y" using 4 by simp then show ?thesis using children_h3 children_h' child_not_in_any_children[unfolded children_eq_h2] 5 6 apply(auto simp add: True children_eq2_h3[OF \ptr \ y\])[1] by (metis (no_types, hide_lams) "3" "7" \type_wf h3\ children_eq2_h3 disjoint_iff_not_equal get_child_nodes_ok insert_before_list_in_set \known_ptrs h\ local.known_ptrs_known_ptr object_ptr_kinds_M_eq3_h object_ptr_kinds_M_eq3_h' object_ptr_kinds_M_eq3_h2 returns_result_select_result select_result_I2) next case False then show ?thesis proof (cases "ptr = y") case True then show ?thesis using children_h3 children_h' child_not_in_any_children[unfolded children_eq_h2] 5 6 apply(auto simp add: True children_eq2_h3[OF \ptr \ x\])[1] by (metis (no_types, hide_lams) "2" "4" "7" IntI \known_ptrs h3\ \type_wf h'\ children_eq_h3 empty_iff insert_before_list_in_set local.get_child_nodes_ok local.known_ptrs_known_ptr object_ptr_kinds_M_eq3_h' returns_result_select_result select_result_I2) next case False then show ?thesis using children_eq2_h3[OF \ptr \ x\] children_eq2_h3[OF \ptr \ y\] 5 6 7 by auto qed qed next fix x xa xb assume 1: " (\x\fset (object_ptr_kinds h'). set |h3 \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h'). set |h' \ get_disconnected_nodes x|\<^sub>r) = {} " and 2: "xa |\| object_ptr_kinds h'" and 3: "x \ set |h' \ get_child_nodes xa|\<^sub>r" and 4: "xb |\| document_ptr_kinds h'" and 5: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" have 6: "set |h3 \ get_child_nodes xa|\<^sub>r \ set |h' \ get_disconnected_nodes xb|\<^sub>r = {}" using 1 2 3 4 5 proof - have "\h d. \ type_wf h \ d |\| document_ptr_kinds h \ h \ ok get_disconnected_nodes d" using local.get_disconnected_nodes_ok by satx then have "h' \ ok get_disconnected_nodes xb" using "4" \type_wf h'\ by fastforce then have f1: "h3 \ get_disconnected_nodes xb \\<^sub>r |h' \ get_disconnected_nodes xb|\<^sub>r" by (simp add: disconnected_nodes_eq_h3) have "xa |\| object_ptr_kinds h3" using "2" object_ptr_kinds_M_eq3_h' by blast then show ?thesis using f1 \local.CD.a_distinct_lists h3\ CD.distinct_lists_no_parent by fastforce qed show False proof (cases "ptr = xa") case True show ?thesis using 6 node_not_in_disconnected_nodes 3 4 5 select_result_I2[OF children_h'] select_result_I2[OF children_h3] True disconnected_nodes_eq2_h3 by (metis (no_types, lifting) "2" DocumentMonad.ptr_kinds_ptr_kinds_M \CD.a_distinct_lists h3\ \type_wf h'\ disconnected_nodes_eq_h3 CD.distinct_lists_no_parent document_ptr_kinds_eq2_h3 get_disconnected_nodes_ok insert_before_list_in_set object_ptr_kinds_M_eq3_h' returns_result_select_result) next case False then show ?thesis using 1 2 3 4 5 children_eq2_h3[OF False] by fastforce qed qed moreover have "CD.a_owner_document_valid h2" using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_owner_document_valid h'" apply(auto simp add: CD.a_owner_document_valid_def object_ptr_kinds_M_eq2_h2 object_ptr_kinds_M_eq2_h3 node_ptr_kinds_eq2_h2 node_ptr_kinds_eq2_h3 document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 children_eq2_h2 )[1] apply(auto simp add: document_ptr_kinds_eq2_h2[simplified] document_ptr_kinds_eq2_h3[simplified] object_ptr_kinds_M_eq2_h2[simplified] object_ptr_kinds_M_eq2_h3[simplified] node_ptr_kinds_eq2_h2[simplified] node_ptr_kinds_eq2_h3[simplified])[1] apply(auto simp add: disconnected_nodes_eq2_h3[symmetric])[1] by (smt Core_DOM_Functions.i_insert_before.insert_before_list_in_set children_eq2_h3 children_h' children_h3 disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3 finite_set_in in_set_remove1 object_ptr_kinds_eq_h3 ptr_in_heap select_result_I2) ultimately have "heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'" by (simp add: CD.heap_is_wellformed_def) have "a_ptr_disconnected_node_rel h3 = a_ptr_disconnected_node_rel h2 - {(cast owner_document, cast node)}" apply(auto simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq_h2 disconnected_nodes_eq2_h2)[1] apply(case_tac "aa = owner_document") apply (metis (no_types, lifting) case_prodI disconnected_nodes_h2 disconnected_nodes_h3 in_set_remove1 mem_Collect_eq node_not_in_disconnected_nodes pair_imageI select_result_I2) using disconnected_nodes_eq2_h2 apply auto[1] using node_not_in_disconnected_nodes apply blast by (metis (no_types, lifting) case_prodI disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3 in_set_remove1 mem_Collect_eq pair_imageI select_result_I2) have "a_ptr_disconnected_node_rel h3 = a_ptr_disconnected_node_rel h'" by(auto simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq_h3 disconnected_nodes_eq2_h3) have "acyclic (parent_child_rel h3 \ a_host_shadow_root_rel h3 \ a_ptr_disconnected_node_rel h3)" using \heap_is_wellformed h2\ by(auto simp add: \a_ptr_disconnected_node_rel h3 = a_ptr_disconnected_node_rel h2 - {(cast owner_document, cast node)}\ heap_is_wellformed_def \parent_child_rel h2 = parent_child_rel h3\ \a_host_shadow_root_rel h2 = a_host_shadow_root_rel h3\ elim!: acyclic_subset) then have "acyclic (parent_child_rel h' \ a_host_shadow_root_rel h' \ local.a_ptr_disconnected_node_rel h')" using \(cast node, ptr) \ (parent_child_rel h3 \ a_host_shadow_root_rel h3 \ a_ptr_disconnected_node_rel h3)\<^sup>*\ by(auto simp add: \a_ptr_disconnected_node_rel h3 = a_ptr_disconnected_node_rel h'\ \a_host_shadow_root_rel h3 = a_host_shadow_root_rel h'\ \parent_child_rel h' = insert (ptr, cast node) ((parent_child_rel h3))\) then show "heap_is_wellformed h'" using \heap_is_wellformed h2\ using \heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'\ apply(auto simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def a_all_ptrs_in_heap_def a_distinct_lists_def a_shadow_root_valid_def)[1] by(auto simp add: object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 element_ptr_kinds_eq_h2 element_ptr_kinds_eq_h3 shadow_root_ptr_kinds_eq_h2 shadow_root_ptr_kinds_eq_h3 shadow_root_eq_h2 shadow_root_eq_h3 shadow_root_eq2_h2 shadow_root_eq2_h3 tag_name_eq_h2 tag_name_eq_h3 tag_name_eq2_h2 tag_name_eq2_h3) qed end interpretation i_insert_before_wf?: l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_parent get_parent_locs get_child_nodes get_child_nodes_locs set_child_nodes set_child_nodes_locs get_ancestors_di get_ancestors_di_locs adopt_node adopt_node_locs set_disconnected_nodes set_disconnected_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_owner_document insert_before insert_before_locs append_child type_wf known_ptr known_ptrs heap_is_wellformed parent_child_rel by(simp add: l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma insert_before_wf_is_l_insert_before_wf [instances]: "l_insert_before_wf Shadow_DOM.heap_is_wellformed ShadowRootClass.type_wf ShadowRootClass.known_ptr ShadowRootClass.known_ptrs Shadow_DOM.insert_before Shadow_DOM.get_child_nodes" apply(auto simp add: l_insert_before_wf_def l_insert_before_wf_axioms_def instances)[1] using insert_before_removes_child apply fast done lemma l_set_disconnected_nodes_get_disconnected_nodes_wf [instances]: "l_set_disconnected_nodes_get_disconnected_nodes_wf ShadowRootClass.type_wf ShadowRootClass.known_ptr Shadow_DOM.heap_is_wellformed Shadow_DOM.parent_child_rel Shadow_DOM.get_child_nodes get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs" apply(auto simp add: l_set_disconnected_nodes_get_disconnected_nodes_wf_def l_set_disconnected_nodes_get_disconnected_nodes_wf_axioms_def instances)[1] by (metis Diff_iff Shadow_DOM.i_heap_is_wellformed.heap_is_wellformed_disconnected_nodes_distinct Shadow_DOM.i_remove_child.set_disconnected_nodes_get_disconnected_nodes insert_iff returns_result_eq set_remove1_eq) interpretation i_insert_before_wf2?: l_insert_before_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs set_child_nodes set_child_nodes_locs get_shadow_root get_shadow_root_locs set_disconnected_nodes set_disconnected_nodes_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel get_parent get_parent_locs adopt_node adopt_node_locs get_owner_document insert_before insert_before_locs append_child known_ptrs remove_child remove_child_locs get_ancestors_di get_ancestors_di_locs adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M adopt_node_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs remove heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M by(auto simp add: l_insert_before_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_insert_before_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma insert_before_wf2_is_l_insert_before_wf2 [instances]: "l_insert_before_wf2 ShadowRootClass.type_wf ShadowRootClass.known_ptr ShadowRootClass.known_ptrs Shadow_DOM.insert_before Shadow_DOM.heap_is_wellformed" apply(auto simp add: l_insert_before_wf2_def l_insert_before_wf2_axioms_def instances)[1] using insert_before_child_preserves apply(fast, fast, fast) done subsubsection \append\_child\ locale l_append_child_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_insert_before_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma append_child_heap_is_wellformed_preserved: assumes wellformed: "heap_is_wellformed h" and append_child: "h \ append_child ptr node \\<^sub>h h'" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "heap_is_wellformed h'" and "type_wf h'" and "known_ptrs h'" using assms by(auto simp add: append_child_def intro: insert_before_child_preserves) lemma append_child_children: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_child_nodes ptr \\<^sub>r xs" assumes "h \ append_child ptr node \\<^sub>h h'" assumes "node \ set xs" shows "h' \ get_child_nodes ptr \\<^sub>r xs @ [node]" proof - obtain ancestors owner_document h2 h3 disconnected_nodes_h2 where ancestors: "h \ get_ancestors_di ptr \\<^sub>r ancestors" and node_not_in_ancestors: "cast node \ set ancestors" and owner_document: "h \ get_owner_document ptr \\<^sub>r owner_document" and h2: "h \ adopt_node owner_document node \\<^sub>h h2" and disconnected_nodes_h2: "h2 \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h2" and h3: "h2 \ set_disconnected_nodes owner_document (remove1 node disconnected_nodes_h2) \\<^sub>h h3" and h': "h3 \ a_insert_node ptr node None \\<^sub>h h'" using assms(5) by(auto simp add: append_child_def insert_before_def a_ensure_pre_insertion_validity_def elim!: bind_returns_heap_E bind_returns_result_E bind_returns_heap_E2[rotated, OF get_parent_pure, rotated] bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated] bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] bind_returns_heap_E2[rotated, OF get_ancestors_pure, rotated] bind_returns_heap_E2[rotated, OF next_sibling_pure, rotated] bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated] split: if_splits option.splits) have "\parent. |h \ get_parent node|\<^sub>r = Some parent \ parent \ ptr" using assms(1) assms(4) assms(6) by (metis (no_types, lifting) assms(2) assms(3) h2 is_OK_returns_heap_I is_OK_returns_result_E local.adopt_node_child_in_heap local.get_parent_child_dual local.get_parent_ok select_result_I2) have "h2 \ get_child_nodes ptr \\<^sub>r xs" using get_child_nodes_reads adopt_node_writes h2 assms(4) apply(rule reads_writes_separate_forwards) using \\parent. |h \ get_parent node|\<^sub>r = Some parent \ parent \ ptr\ apply(auto simp add: adopt_node_locs_def CD.adopt_node_locs_def CD.remove_child_locs_def)[1] by (meson local.set_child_nodes_get_child_nodes_different_pointers) have "h3 \ get_child_nodes ptr \\<^sub>r xs" using get_child_nodes_reads set_disconnected_nodes_writes h3 \h2 \ get_child_nodes ptr \\<^sub>r xs\ apply(rule reads_writes_separate_forwards) by(auto) have "ptr |\| object_ptr_kinds h" by (meson ancestors is_OK_returns_result_I local.get_ancestors_ptr_in_heap) then have "known_ptr ptr" using assms(3) using local.known_ptrs_known_ptr by blast have "type_wf h2" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF adopt_node_writes h2] using adopt_node_types_preserved \type_wf h\ by(auto simp add: adopt_node_locs_def remove_child_locs_def reflp_def transp_def split: if_splits) then have "type_wf h3" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h3] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) show "h' \ get_child_nodes ptr \\<^sub>r xs@[node]" using h' apply(auto simp add: a_insert_node_def dest!: bind_returns_heap_E3[rotated, OF \h3 \ get_child_nodes ptr \\<^sub>r xs\ get_child_nodes_pure, rotated])[1] using \type_wf h3\ set_child_nodes_get_child_nodes \known_ptr ptr\ by metis qed lemma append_child_for_all_on_children: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_child_nodes ptr \\<^sub>r xs" assumes "h \ forall_M (append_child ptr) nodes \\<^sub>h h'" assumes "set nodes \ set xs = {}" assumes "distinct nodes" shows "h' \ get_child_nodes ptr \\<^sub>r xs@nodes" using assms apply(induct nodes arbitrary: h xs) apply(simp) proof(auto elim!: bind_returns_heap_E)[1]fix a nodes h xs h'a assume 0: "(\h xs. heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_child_nodes ptr \\<^sub>r xs \ h \ forall_M (append_child ptr) nodes \\<^sub>h h' \ set nodes \ set xs = {} \ h' \ get_child_nodes ptr \\<^sub>r xs @ nodes)" and 1: "heap_is_wellformed h" and 2: "type_wf h" and 3: "known_ptrs h" and 4: "h \ get_child_nodes ptr \\<^sub>r xs" and 5: "h \ append_child ptr a \\<^sub>r ()" and 6: "h \ append_child ptr a \\<^sub>h h'a" and 7: "h'a \ forall_M (append_child ptr) nodes \\<^sub>h h'" and 8: "a \ set xs" and 9: "set nodes \ set xs = {}" and 10: "a \ set nodes" and 11: "distinct nodes" then have "h'a \ get_child_nodes ptr \\<^sub>r xs @ [a]" using append_child_children 6 using "1" "2" "3" "4" "8" by blast moreover have "heap_is_wellformed h'a" and "type_wf h'a" and "known_ptrs h'a" using insert_before_child_preserves 1 2 3 6 append_child_def by metis+ moreover have "set nodes \ set (xs @ [a]) = {}" using 9 10 by auto ultimately show "h' \ get_child_nodes ptr \\<^sub>r xs @ a # nodes" using 0 7 by fastforce qed lemma append_child_for_all_on_no_children: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_child_nodes ptr \\<^sub>r []" assumes "h \ forall_M (append_child ptr) nodes \\<^sub>h h'" assumes "distinct nodes" shows "h' \ get_child_nodes ptr \\<^sub>r nodes" using assms append_child_for_all_on_children by force end interpretation i_append_child_wf?: l_append_child_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs set_child_nodes set_child_nodes_locs get_shadow_root get_shadow_root_locs set_disconnected_nodes set_disconnected_nodes_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel get_parent get_parent_locs adopt_node adopt_node_locs get_owner_document insert_before insert_before_locs append_child known_ptrs remove_child remove_child_locs get_ancestors_di get_ancestors_di_locs adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M adopt_node_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs remove heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M by(auto simp add: l_append_child_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_append_child_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma append_child_wf_is_l_append_child_wf [instances]: "l_append_child_wf type_wf known_ptr known_ptrs append_child heap_is_wellformed" apply(auto simp add: l_append_child_wf_def l_append_child_wf_axioms_def instances)[1] using append_child_heap_is_wellformed_preserved by fast+ subsubsection \to\_tree\_order\ interpretation i_to_tree_order_wf?: l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf get_child_nodes get_child_nodes_locs to_tree_order known_ptrs get_parent get_parent_locs heap_is_wellformed parent_child_rel get_disconnected_nodes get_disconnected_nodes_locs apply(auto simp add: l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)[1] done declare l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma to_tree_order_wf_is_l_to_tree_order_wf [instances]: "l_to_tree_order_wf heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent get_child_nodes" apply(auto simp add: l_to_tree_order_wf_def l_to_tree_order_wf_axioms_def instances)[1] using to_tree_order_ok apply fast using to_tree_order_ptrs_in_heap apply fast using to_tree_order_parent_child_rel apply(fast, fast) using to_tree_order_child2 apply blast using to_tree_order_node_ptrs apply fast using to_tree_order_child apply fast using to_tree_order_ptr_in_result apply fast using to_tree_order_parent apply fast using to_tree_order_subset apply fast done paragraph \get\_root\_node\ interpretation i_to_tree_order_wf_get_root_node_wf?: l_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf known_ptrs heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_parent get_parent_locs get_ancestors get_ancestors_locs get_root_node get_root_node_locs to_tree_order by(auto simp add: l_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma to_tree_order_wf_get_root_node_wf_is_l_to_tree_order_wf_get_root_node_wf [instances]: "l_to_tree_order_wf_get_root_node_wf ShadowRootClass.type_wf ShadowRootClass.known_ptr ShadowRootClass.known_ptrs to_tree_order Shadow_DOM.get_root_node Shadow_DOM.heap_is_wellformed" apply(auto simp add: l_to_tree_order_wf_get_root_node_wf_def l_to_tree_order_wf_get_root_node_wf_axioms_def instances)[1] using to_tree_order_get_root_node apply fast using to_tree_order_same_root apply fast done subsubsection \to\_tree\_order\_si\ locale l_to_tree_order_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes + l_get_parent_get_host_get_disconnected_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_to_tree_order_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma to_tree_order_si_ok: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" and "ptr |\| object_ptr_kinds h" shows "h \ ok (to_tree_order_si ptr)" proof(insert assms(1) assms(4), induct rule: heap_wellformed_induct_si) case (step parent) have "known_ptr parent" using assms(2) local.known_ptrs_known_ptr step.prems by blast then show ?case using step using assms(1) assms(2) assms(3) using local.heap_is_wellformed_children_in_heap local.get_shadow_root_shadow_root_ptr_in_heap by(auto simp add: to_tree_order_si_def[of parent] intro: get_child_nodes_ok get_shadow_root_ok intro!: bind_is_OK_pure_I map_M_pure_I bind_pure_I map_M_ok_I split: option.splits) qed end interpretation i_to_tree_order_si_wf?: l_to_tree_order_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs known_ptrs get_parent get_parent_locs to_tree_order_si by(auto simp add: l_to_tree_order_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_to_tree_order_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_assigned\_nodes\ lemma forall_M_small_big: "h \ forall_M f xs \\<^sub>h h' \ P h \ (\h h' x. x \ set xs \ h \ f x \\<^sub>h h' \ P h \ P h') \ P h'" by(induct xs arbitrary: h) (auto elim!: bind_returns_heap_E) locale l_assigned_nodes_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_heap_is_wellformed + l_remove_child_wf2 + l_append_child_wf + l_remove_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma assigned_nodes_distinct: assumes "heap_is_wellformed h" assumes "h \ assigned_nodes slot \\<^sub>r nodes" shows "distinct nodes" proof - have "\ptr children. h \ get_child_nodes ptr \\<^sub>r children \ distinct children" using assms(1) local.heap_is_wellformed_children_distinct by blast then show ?thesis using assms apply(auto simp add: assigned_nodes_def elim!: bind_returns_result_E2 split: if_splits)[1] by (simp add: filter_M_distinct) qed lemma flatten_dom_preserves: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "h \ flatten_dom \\<^sub>h h'" shows "heap_is_wellformed h'" and "known_ptrs h'" and "type_wf h'" proof - obtain tups h2 element_ptrs shadow_root_ptrs where "h \ element_ptr_kinds_M \\<^sub>r element_ptrs" and tups: "h \ map_filter_M2 (\element_ptr. do { tag \ get_tag_name element_ptr; assigned_nodes \ assigned_nodes element_ptr; (if tag = ''slot'' \ assigned_nodes \ [] then return (Some (element_ptr, assigned_nodes)) else return None)}) element_ptrs \\<^sub>r tups" (is "h \ map_filter_M2 ?f element_ptrs \\<^sub>r tups") and h2: "h \ forall_M (\(slot, assigned_nodes). do { get_child_nodes (cast slot) \ forall_M remove; forall_M (append_child (cast slot)) assigned_nodes }) tups \\<^sub>h h2" and "h2 \ shadow_root_ptr_kinds_M \\<^sub>r shadow_root_ptrs" and h': "h2 \ forall_M (\shadow_root_ptr. do { host \ get_host shadow_root_ptr; get_child_nodes (cast host) \ forall_M remove; get_child_nodes (cast shadow_root_ptr) \ forall_M (append_child (cast host)); remove_shadow_root host }) shadow_root_ptrs \\<^sub>h h'" using \h \ flatten_dom \\<^sub>h h'\ apply(auto simp add: flatten_dom_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF ElementMonad.ptr_kinds_M_pure, rotated] bind_returns_heap_E2[rotated, OF ShadowRootMonad.ptr_kinds_M_pure, rotated])[1] apply(drule pure_returns_heap_eq) by(auto intro!: map_filter_M2_pure bind_pure_I) have "heap_is_wellformed h2 \ known_ptrs h2 \ type_wf h2" using h2 \heap_is_wellformed h\ \known_ptrs h\ \type_wf h\ by(auto elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated] elim!: forall_M_small_big[where P = "\h. heap_is_wellformed h \ known_ptrs h \ type_wf h", simplified] intro: remove_preserves_known_ptrs remove_heap_is_wellformed_preserved remove_preserves_type_wf append_child_preserves_known_ptrs append_child_heap_is_wellformed_preserved append_child_preserves_type_wf) then show "heap_is_wellformed h'" and "known_ptrs h'" and "type_wf h'" using h' by(auto elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_host_pure, rotated] bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated] dest!: forall_M_small_big[where P = "\h. heap_is_wellformed h \ known_ptrs h \ type_wf h", simplified] intro: remove_preserves_known_ptrs remove_heap_is_wellformed_preserved remove_preserves_type_wf append_child_preserves_known_ptrs append_child_heap_is_wellformed_preserved append_child_preserves_type_wf remove_shadow_root_preserves ) qed end interpretation i_assigned_nodes_wf?: l_assigned_nodes_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr assigned_nodes assigned_nodes_flatten flatten_dom get_child_nodes get_child_nodes_locs get_tag_name get_tag_name_locs get_root_node get_root_node_locs get_host get_host_locs find_slot assigned_slot remove insert_before insert_before_locs append_child remove_shadow_root remove_shadow_root_locs type_wf get_shadow_root get_shadow_root_locs set_shadow_root set_shadow_root_locs get_parent get_parent_locs to_tree_order heap_is_wellformed parent_child_rel get_disconnected_nodes get_disconnected_nodes_locs known_ptrs remove_child remove_child_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_document get_disconnected_document_locs by(auto simp add: l_assigned_nodes_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_assigned_nodes_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_shadow\_root\_safe\ locale l_get_shadow_root_safe_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs + l_type_wf type_wf + l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root_safe get_shadow_root_safe_locs get_shadow_root get_shadow_root_locs get_mode get_mode_locs for known_ptr :: "(_::linorder) object_ptr \ bool" and known_ptrs :: "(_) heap \ bool" and type_wf :: "(_) heap \ bool" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_shadow_root_safe :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_safe_locs :: "(_) element_ptr \ (_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" and get_child_nodes :: "(_::linorder) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" begin end subsubsection \create\_element\ locale l_create_element_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs + l_set_tag_name type_wf set_tag_name set_tag_name_locs + l_create_element_defs create_element + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs + l_new_element_get_disconnected_nodes get_disconnected_nodes get_disconnected_nodes_locs + l_set_tag_name_get_disconnected_nodes type_wf set_tag_name set_tag_name_locs get_disconnected_nodes get_disconnected_nodes_locs + l_create_element\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs set_tag_name set_tag_name_locs type_wf create_element known_ptr type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_new_element_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs + l_set_tag_name_get_child_nodes type_wf set_tag_name set_tag_name_locs known_ptr get_child_nodes get_child_nodes_locs + l_set_tag_name_get_tag_name type_wf get_tag_name get_tag_name_locs set_tag_name set_tag_name_locs + l_new_element_get_tag_name type_wf get_tag_name get_tag_name_locs + l_set_disconnected_nodes_get_child_nodes set_disconnected_nodes set_disconnected_nodes_locs get_child_nodes get_child_nodes_locs + l_set_disconnected_nodes_get_shadow_root set_disconnected_nodes set_disconnected_nodes_locs get_shadow_root get_shadow_root_locs + l_set_disconnected_nodes_get_tag_name type_wf set_disconnected_nodes set_disconnected_nodes_locs get_tag_name get_tag_name_locs + l_set_disconnected_nodes type_wf set_disconnected_nodes set_disconnected_nodes_locs + l_set_disconnected_nodes_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs + l_new_element_get_shadow_root type_wf get_shadow_root get_shadow_root_locs + l_set_tag_name_get_shadow_root type_wf set_tag_name set_tag_name_locs get_shadow_root get_shadow_root_locs + l_new_element type_wf + l_known_ptrs known_ptr known_ptrs for known_ptr :: "(_::linorder) object_ptr \ bool" and known_ptrs :: "(_) heap \ bool" and type_wf :: "(_) heap \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and set_tag_name :: "(_) element_ptr \ char list \ ((_) heap, exception, unit) prog" and set_tag_name_locs :: "(_) element_ptr \ ((_) heap, exception, unit) prog set" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" and create_element :: "(_) document_ptr \ char list \ ((_) heap, exception, (_) element_ptr) prog" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_disconnected_document :: "(_) node_ptr \ ((_) heap, exception, (_) document_ptr) prog" and get_disconnected_document_locs :: "((_) heap \ (_) heap \ bool) set" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_::linorder) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" begin lemma create_element_preserves_wellformedness: assumes "heap_is_wellformed h" and "h \ create_element document_ptr tag \\<^sub>h h'" and "type_wf h" and "known_ptrs h" shows "heap_is_wellformed h'" and "type_wf h'" and "known_ptrs h'" proof - obtain new_element_ptr h2 h3 disc_nodes_h3 where new_element_ptr: "h \ new_element \\<^sub>r new_element_ptr" and h2: "h \ new_element \\<^sub>h h2" and h3: "h2 \ set_tag_name new_element_ptr tag \\<^sub>h h3" and disc_nodes_h3: "h3 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" and h': "h3 \ set_disconnected_nodes document_ptr (cast new_element_ptr # disc_nodes_h3) \\<^sub>h h'" using assms(2) by(auto simp add: create_element_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF CD.get_disconnected_nodes_pure, rotated] ) then have "h \ create_element document_ptr tag \\<^sub>r new_element_ptr" apply(auto simp add: create_element_def intro!: bind_returns_result_I)[1] apply (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust) apply (metis is_OK_returns_heap_E is_OK_returns_result_I CD.get_disconnected_nodes_pure pure_returns_heap_eq) by (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust) have "new_element_ptr \ set |h \ element_ptr_kinds_M|\<^sub>r" using new_element_ptr ElementMonad.ptr_kinds_ptr_kinds_M h2 using new_element_ptr_not_in_heap by blast then have "cast new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r" by simp then have "cast new_element_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r" by simp have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\| {|cast new_element_ptr|}" using new_element_new_ptr h2 new_element_ptr by blast then have node_ptr_kinds_eq_h: "node_ptr_kinds h2 = node_ptr_kinds h |\| {|cast new_element_ptr|}" apply(simp add: node_ptr_kinds_def) by force then have element_ptr_kinds_eq_h: "element_ptr_kinds h2 = element_ptr_kinds h |\| {|new_element_ptr|}" apply(simp add: element_ptr_kinds_def) by force have character_data_ptr_kinds_eq_h: "character_data_ptr_kinds h2 = character_data_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: node_ptr_kinds_def character_data_ptr_kinds_def) have document_ptr_kinds_eq_h: "document_ptr_kinds h2 = document_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: document_ptr_kinds_def) have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_tag_name_writes h3]) using set_tag_name_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2" using object_ptr_kinds_eq_h2 by(auto simp add: node_ptr_kinds_def) then have element_ptr_kinds_eq_h2: "element_ptr_kinds h3 = element_ptr_kinds h2" by(simp add: element_ptr_kinds_def) have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_disconnected_nodes_writes h']) using set_disconnected_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3" using object_ptr_kinds_eq_h3 by(auto simp add: node_ptr_kinds_def) then have element_ptr_kinds_eq_h3: "element_ptr_kinds h' = element_ptr_kinds h3" by(simp add: element_ptr_kinds_def) have "known_ptr (cast new_element_ptr)" using \h \ create_element document_ptr tag \\<^sub>r new_element_ptr\ local.create_element_known_ptr by blast then have "known_ptrs h2" using known_ptrs_new_ptr object_ptr_kinds_eq_h \known_ptrs h\ h2 by blast then have "known_ptrs h3" using known_ptrs_preserved object_ptr_kinds_eq_h2 by blast then show "known_ptrs h'" using known_ptrs_preserved object_ptr_kinds_eq_h3 by blast have "document_ptr |\| document_ptr_kinds h" using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2 CD.get_disconnected_nodes_ptr_in_heap \type_wf h\ document_ptr_kinds_def by (metis is_OK_returns_result_I) have children_eq_h: "\(ptr'::(_) object_ptr) children. ptr' \ cast new_element_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads h2 get_child_nodes_new_element[rotated, OF new_element_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have children_eq2_h: "\ptr'. ptr' \ cast new_element_ptr \ |h \ get_child_nodes ptr'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have "h2 \ get_child_nodes (cast new_element_ptr) \\<^sub>r []" using new_element_ptr h2 new_element_ptr_in_heap[OF h2 new_element_ptr] new_element_is_element_ptr[OF new_element_ptr] new_element_no_child_nodes by blast have tag_name_eq_h: "\ptr' disc_nodes. ptr' \ new_element_ptr \ h \ get_tag_name ptr' \\<^sub>r disc_nodes = h2 \ get_tag_name ptr' \\<^sub>r disc_nodes" using get_tag_name_reads h2 get_tag_name_new_element[rotated, OF new_element_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by(blast)+ then have tag_name_eq2_h: "\ptr'. ptr' \ new_element_ptr \ |h \ get_tag_name ptr'|\<^sub>r = |h2 \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have "h2 \ get_tag_name new_element_ptr \\<^sub>r ''''" using new_element_ptr h2 new_element_ptr_in_heap[OF h2 new_element_ptr] new_element_is_element_ptr[OF new_element_ptr] new_element_empty_tag_name by blast have disconnected_nodes_eq_h: "\doc_ptr disc_nodes. h \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using CD.get_disconnected_nodes_reads h2 get_disconnected_nodes_new_element[OF new_element_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have disconnected_nodes_eq2_h: "\doc_ptr. |h \ get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. h2 \ get_child_nodes ptr' \\<^sub>r children = h3 \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads set_tag_name_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_tag_name_get_child_nodes) then have children_eq2_h2: "\ptr'. |h2 \ get_child_nodes ptr'|\<^sub>r = |h3 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using CD.get_disconnected_nodes_reads set_tag_name_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_tag_name_get_disconnected_nodes) then have disconnected_nodes_eq2_h2: "\doc_ptr. |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have tag_name_eq_h2: "\ptr' disc_nodes. ptr' \ new_element_ptr \ h2 \ get_tag_name ptr' \\<^sub>r disc_nodes = h3 \ get_tag_name ptr' \\<^sub>r disc_nodes" apply(rule reads_writes_preserved[OF get_tag_name_reads set_tag_name_writes h3]) by (metis local.set_tag_name_get_tag_name_different_pointers) then have tag_name_eq2_h2: "\ptr'. ptr' \ new_element_ptr \ |h2 \ get_tag_name ptr'|\<^sub>r = |h3 \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have "h2 \ get_tag_name new_element_ptr \\<^sub>r ''''" using new_element_ptr h2 new_element_ptr_in_heap[OF h2 new_element_ptr] new_element_is_element_ptr[OF new_element_ptr] new_element_empty_tag_name by blast have "type_wf h2" using \type_wf h\ new_element_types_preserved h2 by blast then have "type_wf h3" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_tag_name_writes h3] using set_tag_name_types_preserved by(auto simp add: reflp_def transp_def) then show "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h'] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) have children_eq_h3: "\ptr' children. h3 \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_disconnected_nodes_get_child_nodes) then have children_eq2_h3: "\ptr'. |h3 \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h3: "\doc_ptr disc_nodes. document_ptr \ doc_ptr \ h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using CD.get_disconnected_nodes_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers) then have disconnected_nodes_eq2_h3: "\doc_ptr. document_ptr \ doc_ptr \ |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have tag_name_eq_h2: "\ptr' disc_nodes. ptr' \ new_element_ptr \ h2 \ get_tag_name ptr' \\<^sub>r disc_nodes = h3 \ get_tag_name ptr' \\<^sub>r disc_nodes" apply(rule reads_writes_preserved[OF get_tag_name_reads set_tag_name_writes h3]) by (metis local.set_tag_name_get_tag_name_different_pointers) then have tag_name_eq2_h2: "\ptr'. ptr' \ new_element_ptr \ |h2 \ get_tag_name ptr'|\<^sub>r = |h3 \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have disc_nodes_document_ptr_h2: "h2 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" using disconnected_nodes_eq_h2 disc_nodes_h3 by auto then have disc_nodes_document_ptr_h: "h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" using disconnected_nodes_eq_h by auto then have "cast new_element_ptr \ set disc_nodes_h3" using \heap_is_wellformed h\ using \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ a_all_ptrs_in_heap_def heap_is_wellformed_def using NodeMonad.ptr_kinds_ptr_kinds_M local.heap_is_wellformed_disc_nodes_in_heap by blast have tag_name_eq_h3: "\ptr' disc_nodes. h3 \ get_tag_name ptr' \\<^sub>r disc_nodes = h' \ get_tag_name ptr' \\<^sub>r disc_nodes" apply(rule reads_writes_preserved[OF get_tag_name_reads set_disconnected_nodes_writes h']) using set_disconnected_nodes_get_tag_name by blast then have tag_name_eq2_h3: "\ptr'. |h3 \ get_tag_name ptr'|\<^sub>r = |h' \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have "acyclic (parent_child_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def) also have "parent_child_rel h = parent_child_rel h2" proof(auto simp add: CD.parent_child_rel_def)[1] fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h2" by (simp add: object_ptr_kinds_eq_h) next fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "x \ set |h2 \ get_child_nodes a|\<^sub>r" by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2_h) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h" using object_ptr_kinds_eq_h \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ by(auto) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "x \ set |h \ get_child_nodes a|\<^sub>r" by (metis (no_types, lifting) \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ children_eq2_h empty_iff empty_set image_eqI select_result_I2) qed also have "\ = parent_child_rel h3" by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2) also have "\ = parent_child_rel h'" by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3) finally have "CD.a_acyclic_heap h'" by (simp add: CD.acyclic_heap_def) have "CD.a_all_ptrs_in_heap h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_all_ptrs_in_heap h2" apply(auto simp add: CD.a_all_ptrs_in_heap_def)[1] apply (metis \known_ptrs h2\ \parent_child_rel h = parent_child_rel h2\ \type_wf h2\ assms(1) assms(3) funion_iff CD.get_child_nodes_ok local.known_ptrs_known_ptr local.parent_child_rel_child_in_heap CD.parent_child_rel_child_nodes2 node_ptr_kinds_commutes node_ptr_kinds_eq_h returns_result_select_result) by (metis (no_types, lifting) CD.get_child_nodes_ok CD.get_child_nodes_ptr_in_heap \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ assms(3) assms(4) children_eq_h disconnected_nodes_eq2_h document_ptr_kinds_eq_h finite_set_in is_OK_returns_result_I local.known_ptrs_known_ptr node_ptr_kinds_commutes returns_result_select_result subsetD) then have "CD.a_all_ptrs_in_heap h3" by (simp add: children_eq2_h2 disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 object_ptr_kinds_eq_h2) then have "CD.a_all_ptrs_in_heap h'" by (smt children_eq2_h3 disc_nodes_h3 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 element_ptr_kinds_commutes h' h2 local.CD.a_all_ptrs_in_heap_def local.set_disconnected_nodes_get_disconnected_nodes new_element_ptr new_element_ptr_in_heap node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h3 notin_fset object_ptr_kinds_eq_h3 select_result_I2 set_ConsD subset_code(1)) have "\p. p |\| object_ptr_kinds h \ cast new_element_ptr \ set |h \ get_child_nodes p|\<^sub>r" using \heap_is_wellformed h\ \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ heap_is_wellformed_children_in_heap by (meson NodeMonad.ptr_kinds_ptr_kinds_M CD.a_all_ptrs_in_heap_def assms(3) assms(4) fset_mp fset_of_list_elem CD.get_child_nodes_ok known_ptrs_known_ptr returns_result_select_result) then have "\p. p |\| object_ptr_kinds h2 \ cast new_element_ptr \ set |h2 \ get_child_nodes p|\<^sub>r" using children_eq2_h apply(auto simp add: object_ptr_kinds_eq_h)[1] using \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ apply auto[1] by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\) then have "\p. p |\| object_ptr_kinds h3 \ cast new_element_ptr \ set |h3 \ get_child_nodes p|\<^sub>r" using object_ptr_kinds_eq_h2 children_eq2_h2 by auto then have new_element_ptr_not_in_any_children: "\p. p |\| object_ptr_kinds h' \ cast new_element_ptr \ set |h' \ get_child_nodes p|\<^sub>r" using object_ptr_kinds_eq_h3 children_eq2_h3 by auto have "CD.a_distinct_lists h" using \heap_is_wellformed h\ by (simp add: CD.heap_is_wellformed_def heap_is_wellformed_def) then have "CD.a_distinct_lists h2" using \h2 \ get_child_nodes (cast new_element_ptr) \\<^sub>r []\ apply(auto simp add: CD.a_distinct_lists_def object_ptr_kinds_eq_h document_ptr_kinds_eq_h disconnected_nodes_eq2_h intro!: distinct_concat_map_I)[1] - apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert) + apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert_remove) apply(case_tac "x=cast new_element_ptr") apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply (metis IntI assms(1) assms(3) assms(4) empty_iff CD.get_child_nodes_ok local.heap_is_wellformed_one_parent local.known_ptrs_known_ptr returns_result_select_result) apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] by (metis \ CD.a_distinct_lists h\ \type_wf h2\ disconnected_nodes_eq_h document_ptr_kinds_eq_h CD.distinct_lists_no_parent get_disconnected_nodes_ok returns_result_select_result) then have " CD.a_distinct_lists h3" by(auto simp add: CD.a_distinct_lists_def disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 children_eq2_h2 object_ptr_kinds_eq_h2) then have " CD.a_distinct_lists h'" proof(auto simp add: CD.a_distinct_lists_def disconnected_nodes_eq2_h3 children_eq2_h3 object_ptr_kinds_eq_h3 document_ptr_kinds_eq_h3 intro!: distinct_concat_map_I)[1] fix x assume "distinct (concat (map (\document_ptr. |h3 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" and "x |\| document_ptr_kinds h3" then show "distinct |h' \ get_disconnected_nodes x|\<^sub>r" using document_ptr_kinds_eq_h3 disconnected_nodes_eq_h3 h' set_disconnected_nodes_get_disconnected_nodes by (metis (no_types, lifting) \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set disc_nodes_h3\ \ CD.a_distinct_lists h3\ \type_wf h'\ disc_nodes_h3 distinct.simps(2) CD.distinct_lists_disconnected_nodes get_disconnected_nodes_ok returns_result_eq returns_result_select_result) next fix x y xa assume "distinct (concat (map (\document_ptr. |h3 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" and "x |\| document_ptr_kinds h3" and "y |\| document_ptr_kinds h3" and "x \ y" and "xa \ set |h' \ get_disconnected_nodes x|\<^sub>r" and "xa \ set |h' \ get_disconnected_nodes y|\<^sub>r" moreover have "set |h3 \ get_disconnected_nodes x|\<^sub>r \ set |h3 \ get_disconnected_nodes y|\<^sub>r = {}" using calculation by(auto dest: distinct_concat_map_E(1)) ultimately show "False" apply(-) apply(cases "x = document_ptr") apply(smt NodeMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ \CD.a_all_ptrs_in_heap h\ disc_nodes_h3 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 finite_set_in h' set_disconnected_nodes_get_disconnected_nodes CD.a_all_ptrs_in_heap_def select_result_I2 set_ConsD subsetD) by (smt NodeMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ \CD.a_all_ptrs_in_heap h\ disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 finite_set_in h' l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes CD.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms select_result_I2 set_ConsD subsetD) next fix x xa xb assume 2: "(\x\fset (object_ptr_kinds h3). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h3). set |h3 \ get_disconnected_nodes x|\<^sub>r) = {}" and 3: "xa |\| object_ptr_kinds h3" and 4: "x \ set |h' \ get_child_nodes xa|\<^sub>r" and 5: "xb |\| document_ptr_kinds h3" and 6: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" show "False" using disc_nodes_document_ptr_h disconnected_nodes_eq2_h3 apply - apply(cases "xb = document_ptr") apply (metis (no_types, hide_lams) "3" "4" "6" \\p. p |\| object_ptr_kinds h3 \ cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h3 \ get_child_nodes p|\<^sub>r\ \ CD.a_distinct_lists h3\ children_eq2_h3 disc_nodes_h3 CD.distinct_lists_no_parent h' select_result_I2 set_ConsD set_disconnected_nodes_get_disconnected_nodes) by (metis "3" "4" "5" "6" \ CD.a_distinct_lists h3\ \type_wf h3\ children_eq2_h3 CD.distinct_lists_no_parent get_disconnected_nodes_ok returns_result_select_result) qed have "CD.a_owner_document_valid h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_owner_document_valid h'" using disc_nodes_h3 \document_ptr |\| document_ptr_kinds h\ apply(auto simp add: CD.a_owner_document_valid_def)[1] apply(auto simp add: object_ptr_kinds_eq_h object_ptr_kinds_eq_h3 )[1] apply(auto simp add: object_ptr_kinds_eq_h2)[1] apply(auto simp add: document_ptr_kinds_eq_h document_ptr_kinds_eq_h3 )[1] apply(auto simp add: document_ptr_kinds_eq_h2)[1] apply(auto simp add: node_ptr_kinds_eq_h node_ptr_kinds_eq_h3 )[1] apply(auto simp add: node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h )[1] apply(auto simp add: children_eq2_h2[symmetric] children_eq2_h3[symmetric] disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3)[1] apply (metis (no_types, lifting) document_ptr_kinds_eq_h h' list.set_intros(1) local.set_disconnected_nodes_get_disconnected_nodes select_result_I2) apply(simp add: object_ptr_kinds_eq_h) by(metis (no_types, lifting) NodeMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ children_eq2_h children_eq2_h2 children_eq2_h3 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h finite_set_in h' l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes list.set_intros(2) local.l_set_disconnected_nodes_get_disconnected_nodes_axioms node_ptr_kinds_commutes select_result_I2) have "CD.a_heap_is_wellformed h'" using \CD.a_acyclic_heap h'\ \CD.a_all_ptrs_in_heap h'\ \CD.a_distinct_lists h'\ \CD.a_owner_document_valid h'\ by(simp add: CD.a_heap_is_wellformed_def) have shadow_root_ptr_kinds_eq_h: "shadow_root_ptr_kinds h2 = shadow_root_ptr_kinds h" using document_ptr_kinds_eq_h by(auto simp add: shadow_root_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h2: "shadow_root_ptr_kinds h3 = shadow_root_ptr_kinds h2" using document_ptr_kinds_eq_h2 by(auto simp add: shadow_root_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h3: "shadow_root_ptr_kinds h' = shadow_root_ptr_kinds h3" using document_ptr_kinds_eq_h3 by(auto simp add: shadow_root_ptr_kinds_def) have shadow_root_eq_h: "\element_ptr shadow_root_opt. element_ptr \ new_element_ptr \ h \ get_shadow_root element_ptr \\<^sub>r shadow_root_opt = h2 \ get_shadow_root element_ptr \\<^sub>r shadow_root_opt" proof - fix element_ptr shadow_root_opt assume "element_ptr \ new_element_ptr " have "\P \ get_shadow_root_locs element_ptr. P h h2" using get_shadow_root_new_element new_element_ptr h2 using \element_ptr \ new_element_ptr\ by blast then have "preserved (get_shadow_root element_ptr) h h2" using get_shadow_root_new_element[rotated, OF new_element_ptr h2] using get_shadow_root_reads by(simp add: reads_def) then show "h \ get_shadow_root element_ptr \\<^sub>r shadow_root_opt = h2 \ get_shadow_root element_ptr \\<^sub>r shadow_root_opt" by (simp add: preserved_def) qed have shadow_root_none: "h2 \ get_shadow_root (new_element_ptr) \\<^sub>r None" using new_element_ptr h2 new_element_ptr_in_heap[OF h2 new_element_ptr] new_element_is_element_ptr[OF new_element_ptr] new_element_no_shadow_root by blast have shadow_root_eq_h2: "\ptr' children. h2 \ get_shadow_root ptr' \\<^sub>r children = h3 \ get_shadow_root ptr' \\<^sub>r children" using get_shadow_root_reads set_tag_name_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_tag_name_get_shadow_root) have shadow_root_eq_h3: "\ptr' children. h3 \ get_shadow_root ptr' \\<^sub>r children = h' \ get_shadow_root ptr' \\<^sub>r children" using get_shadow_root_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) using set_disconnected_nodes_get_shadow_root by(auto simp add: set_disconnected_nodes_get_shadow_root) have "a_all_ptrs_in_heap h" by (simp add: assms(1) local.a_all_ptrs_in_heap_def local.get_shadow_root_shadow_root_ptr_in_heap) then have "a_all_ptrs_in_heap h2" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h)[1] using returns_result_eq shadow_root_eq_h shadow_root_none by fastforce then have "a_all_ptrs_in_heap h3" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h2)[1] using shadow_root_eq_h2 by blast then have "a_all_ptrs_in_heap h'" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h3)[1] by (simp add: shadow_root_eq_h3) have "a_distinct_lists h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_distinct_lists h2" apply(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h)[1] apply(auto simp add: distinct_insort intro!: distinct_concat_map_I split: option.splits)[1] apply(case_tac "x = new_element_ptr") using shadow_root_none apply auto[1] using shadow_root_eq_h by (smt Diff_empty Diff_insert0 ElementMonad.ptr_kinds_M_ptr_kinds ElementMonad.ptr_kinds_ptr_kinds_M assms(1) assms(3) finite_set_in h2 insort_split local.get_shadow_root_ok local.shadow_root_same_host new_element_ptr new_element_ptr_not_in_heap option.distinct(1) returns_result_select_result select_result_I2 shadow_root_none) then have "a_distinct_lists h3" by(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h2 select_result_eq[OF shadow_root_eq_h2]) then have "a_distinct_lists h'" by(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h3 select_result_eq[OF shadow_root_eq_h3]) have "a_shadow_root_valid h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_shadow_root_valid h2" proof (unfold a_shadow_root_valid_def; safe) fix shadow_root_ptr assume "\shadow_root_ptr\fset (shadow_root_ptr_kinds h). \host\fset (element_ptr_kinds h). |h \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types \ |h \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr" assume "shadow_root_ptr \ fset (shadow_root_ptr_kinds h2)" obtain previous_host where "previous_host \ fset (element_ptr_kinds h)" and "|h \ get_tag_name previous_host|\<^sub>r \ safe_shadow_root_element_types" and "|h \ get_shadow_root previous_host|\<^sub>r = Some shadow_root_ptr" by (metis \local.a_shadow_root_valid h\ \shadow_root_ptr \ fset (shadow_root_ptr_kinds h2)\ local.a_shadow_root_valid_def shadow_root_ptr_kinds_eq_h) moreover have "previous_host \ new_element_ptr" using calculation(1) h2 new_element_ptr new_element_ptr_not_in_heap by auto ultimately have "|h2 \ get_tag_name previous_host|\<^sub>r \ safe_shadow_root_element_types" and "|h2 \ get_shadow_root previous_host|\<^sub>r = Some shadow_root_ptr" using shadow_root_eq_h apply (simp add: tag_name_eq2_h) by (metis \previous_host \ new_element_ptr\ \|h \ get_shadow_root previous_host|\<^sub>r = Some shadow_root_ptr\ select_result_eq shadow_root_eq_h) then show "\host\fset (element_ptr_kinds h2). |h2 \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types \ |h2 \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr" by (meson \previous_host \ fset (element_ptr_kinds h)\ \previous_host \ new_element_ptr\ assms(3) local.get_shadow_root_ok local.get_shadow_root_ptr_in_heap notin_fset returns_result_select_result shadow_root_eq_h) qed then have "a_shadow_root_valid h3" proof (unfold a_shadow_root_valid_def; safe) fix shadow_root_ptr assume "\shadow_root_ptr\fset (shadow_root_ptr_kinds h2). \host\fset (element_ptr_kinds h2). |h2 \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types \ |h2 \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr" assume "shadow_root_ptr \ fset (shadow_root_ptr_kinds h3)" obtain previous_host where "previous_host \ fset (element_ptr_kinds h2)" and "|h2 \ get_tag_name previous_host|\<^sub>r \ safe_shadow_root_element_types" and "|h2 \ get_shadow_root previous_host|\<^sub>r = Some shadow_root_ptr" by (metis \local.a_shadow_root_valid h2\ \shadow_root_ptr \ fset (shadow_root_ptr_kinds h3)\ local.a_shadow_root_valid_def shadow_root_ptr_kinds_eq_h2) moreover have "previous_host \ new_element_ptr" using calculation(1) h3 new_element_ptr new_element_ptr_not_in_heap using calculation(3) shadow_root_none by auto ultimately have "|h2 \ get_tag_name previous_host|\<^sub>r \ safe_shadow_root_element_types" and "|h2 \ get_shadow_root previous_host|\<^sub>r = Some shadow_root_ptr" using shadow_root_eq_h2 apply (simp add: tag_name_eq2_h2) by (metis \previous_host \ new_element_ptr\ \|h2 \ get_shadow_root previous_host|\<^sub>r = Some shadow_root_ptr\ select_result_eq shadow_root_eq_h) then show "\host\fset (element_ptr_kinds h3). |h3 \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types \ |h3 \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr" by (smt \previous_host \ fset (element_ptr_kinds h2)\ \previous_host \ new_element_ptr\ \type_wf h2\ \type_wf h3\ element_ptr_kinds_eq_h2 finite_set_in local.get_shadow_root_ok returns_result_eq returns_result_select_result shadow_root_eq_h2 tag_name_eq2_h2) qed then have "a_shadow_root_valid h'" apply(auto simp add: a_shadow_root_valid_def element_ptr_kinds_eq_h3 shadow_root_eq_h3 shadow_root_ptr_kinds_eq_h3 tag_name_eq2_h3)[1] by (smt \type_wf h3\ finite_set_in local.get_shadow_root_ok returns_result_select_result select_result_I2 shadow_root_eq_h3) have "a_host_shadow_root_rel h = a_host_shadow_root_rel h2" apply(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h shadow_root_eq_h)[1] apply (smt assms(3) case_prod_conv h2 image_iff local.get_shadow_root_ok mem_Collect_eq new_element_ptr new_element_ptr_not_in_heap returns_result_select_result select_result_I2 shadow_root_eq_h) using shadow_root_none apply auto[1] apply (metis (no_types, lifting) Collect_cong assms(3) case_prodE case_prodI h2 host_shadow_root_rel_def i_get_parent_get_host_get_disconnected_document_wf.a_host_shadow_root_rel_shadow_root local.a_host_shadow_root_rel_def local.get_shadow_root_impl local.get_shadow_root_ok new_element_ptr new_element_ptr_not_in_heap returns_result_select_result select_result_I2 shadow_root_eq_h) done have "a_host_shadow_root_rel h2 = a_host_shadow_root_rel h3" apply(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h2 shadow_root_eq_h2)[1] apply (smt Collect_cong \type_wf h2\ case_prodE case_prodI element_ptr_kinds_eq_h2 host_shadow_root_rel_def i_get_root_node_si_wf.a_host_shadow_root_rel_shadow_root local.a_host_shadow_root_rel_def local.get_shadow_root_impl local.get_shadow_root_ok returns_result_select_result shadow_root_eq_h2) by (metis (no_types, lifting) Collect_cong \type_wf h3\ case_prodI2 case_prod_conv element_ptr_kinds_eq_h2 host_shadow_root_rel_def i_get_root_node_si_wf.a_host_shadow_root_rel_shadow_root local.a_host_shadow_root_rel_def local.get_shadow_root_impl local.get_shadow_root_ok returns_result_select_result shadow_root_eq_h2) have "a_host_shadow_root_rel h3 = a_host_shadow_root_rel h'" apply(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h2 shadow_root_eq_h2)[1] apply (smt Collect_cong Shadow_DOM.a_host_shadow_root_rel_def \type_wf h3\ case_prodD case_prodI2 element_ptr_kinds_eq_h2 i_get_root_node_si_wf.a_host_shadow_root_rel_shadow_root local.get_shadow_root_impl local.get_shadow_root_ok returns_result_select_result shadow_root_eq_h3) apply (smt Collect_cong \type_wf h'\ case_prodE case_prodI element_ptr_kinds_eq_h2 host_shadow_root_rel_def i_get_root_node_si_wf.a_host_shadow_root_rel_shadow_root l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_host_shadow_root_rel_def local.get_shadow_root_impl local.get_shadow_root_ok returns_result_select_result shadow_root_eq_h3) done have "a_ptr_disconnected_node_rel h = a_ptr_disconnected_node_rel h2" by(simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq_h disconnected_nodes_eq2_h) have "a_ptr_disconnected_node_rel h2 = a_ptr_disconnected_node_rel h3" by(simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq_h2 disconnected_nodes_eq2_h2) have "h' \ get_disconnected_nodes document_ptr \\<^sub>r cast new_element_ptr # disc_nodes_h3" using h' local.set_disconnected_nodes_get_disconnected_nodes by auto have " document_ptr |\| document_ptr_kinds h3" by (simp add: \document_ptr |\| document_ptr_kinds h\ document_ptr_kinds_eq_h document_ptr_kinds_eq_h2) have "cast new_element_ptr \ set |h' \ get_disconnected_nodes document_ptr|\<^sub>r" using \h' \ get_disconnected_nodes document_ptr \\<^sub>r cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr # disc_nodes_h3\ by auto have "a_ptr_disconnected_node_rel h' = {(cast document_ptr, cast new_element_ptr)} \ a_ptr_disconnected_node_rel h3" apply(auto simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq_h3 disconnected_nodes_eq2_h3)[1] apply(case_tac "aa = document_ptr") using disc_nodes_h3 h' \h' \ get_disconnected_nodes document_ptr \\<^sub>r cast new_element_ptr # disc_nodes_h3\ apply(auto)[1] using disconnected_nodes_eq2_h3 apply auto[1] using \h' \ get_disconnected_nodes document_ptr \\<^sub>r cast new_element_ptr # disc_nodes_h3\ using \cast new_element_ptr \ set |h' \ get_disconnected_nodes document_ptr|\<^sub>r\ using \document_ptr |\| document_ptr_kinds h3\ apply auto[1] apply(case_tac "document_ptr = aa") using \h' \ get_disconnected_nodes document_ptr \\<^sub>r cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr # disc_nodes_h3\ disc_nodes_h3 apply auto[1] using disconnected_nodes_eq_h3[THEN select_result_eq, simplified] by auto have "acyclic (parent_child_rel h \ a_host_shadow_root_rel h \ a_ptr_disconnected_node_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) have "parent_child_rel h \ a_host_shadow_root_rel h \ a_ptr_disconnected_node_rel h = parent_child_rel h2 \ a_host_shadow_root_rel h2 \ a_ptr_disconnected_node_rel h2" using \local.a_host_shadow_root_rel h = local.a_host_shadow_root_rel h2\ \local.a_ptr_disconnected_node_rel h = local.a_ptr_disconnected_node_rel h2\ \parent_child_rel h = parent_child_rel h2\ by auto have "parent_child_rel h2 \ a_host_shadow_root_rel h2 \ a_ptr_disconnected_node_rel h2 = parent_child_rel h3 \ a_host_shadow_root_rel h3 \ a_ptr_disconnected_node_rel h3" using \local.a_host_shadow_root_rel h2 = local.a_host_shadow_root_rel h3\ \local.a_ptr_disconnected_node_rel h2 = local.a_ptr_disconnected_node_rel h3\ \parent_child_rel h2 = parent_child_rel h3\ by auto have "parent_child_rel h' \ a_host_shadow_root_rel h' \ a_ptr_disconnected_node_rel h' = {(cast document_ptr, cast new_element_ptr)} \ parent_child_rel h3 \ a_host_shadow_root_rel h3 \ a_ptr_disconnected_node_rel h3" by (simp add: \local.a_host_shadow_root_rel h3 = local.a_host_shadow_root_rel h'\ \local.a_ptr_disconnected_node_rel h' = {(cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr, cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr)} \ local.a_ptr_disconnected_node_rel h3\ \parent_child_rel h3 = parent_child_rel h'\) have "\a b. (a, b) \ parent_child_rel h3 \ a \ cast new_element_ptr" using CD.parent_child_rel_parent_in_heap \parent_child_rel h = parent_child_rel h2\ \parent_child_rel h2 = parent_child_rel h3\ element_ptr_kinds_commutes h2 new_element_ptr new_element_ptr_not_in_heap node_ptr_kinds_commutes by blast moreover have "\a b. (a, b) \ a_host_shadow_root_rel h3 \ a \ cast new_element_ptr" using shadow_root_eq_h2 shadow_root_none by(auto simp add: a_host_shadow_root_rel_def) moreover have "\a b. (a, b) \ a_ptr_disconnected_node_rel h3 \ a \ cast new_element_ptr" by(auto simp add: a_ptr_disconnected_node_rel_def) moreover have "cast new_element_ptr \ {x. (x, cast document_ptr) \ (parent_child_rel h3 \ a_host_shadow_root_rel h3 \ a_ptr_disconnected_node_rel h3)\<^sup>*}" by (smt Un_iff \\b a. (a, b) \ local.a_host_shadow_root_rel h3 \ a \ cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr\ \\b a. (a, b) \ local.a_ptr_disconnected_node_rel h3 \ a \ cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr\ \\b a. (a, b) \ parent_child_rel h3 \ a \ cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr\ cast_document_ptr_not_node_ptr(1) converse_rtranclE mem_Collect_eq) moreover have "acyclic (parent_child_rel h3 \ a_host_shadow_root_rel h3 \ a_ptr_disconnected_node_rel h3)" using \acyclic (parent_child_rel h \ local.a_host_shadow_root_rel h \ local.a_ptr_disconnected_node_rel h)\ \parent_child_rel h \ local.a_host_shadow_root_rel h \ local.a_ptr_disconnected_node_rel h = parent_child_rel h2 \ local.a_host_shadow_root_rel h2 \ local.a_ptr_disconnected_node_rel h2\ \parent_child_rel h2 \ local.a_host_shadow_root_rel h2 \ local.a_ptr_disconnected_node_rel h2 = parent_child_rel h3 \ local.a_host_shadow_root_rel h3 \ local.a_ptr_disconnected_node_rel h3\ by auto ultimately have "acyclic (parent_child_rel h' \ a_host_shadow_root_rel h' \ a_ptr_disconnected_node_rel h')" by(simp add: \parent_child_rel h' \ a_host_shadow_root_rel h' \ a_ptr_disconnected_node_rel h' = {(cast document_ptr, cast new_element_ptr)} \ parent_child_rel h3 \ a_host_shadow_root_rel h3 \ a_ptr_disconnected_node_rel h3\) show " heap_is_wellformed h' " using \acyclic (parent_child_rel h' \ local.a_host_shadow_root_rel h' \ local.a_ptr_disconnected_node_rel h')\ by(simp add: heap_is_wellformed_def CD.heap_is_wellformed_impl \local.CD.a_heap_is_wellformed h'\ \local.a_all_ptrs_in_heap h'\ \local.a_distinct_lists h'\ \local.a_shadow_root_valid h'\) qed end interpretation i_create_element_wf?: l_create_element_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr known_ptrs type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed parent_child_rel set_tag_name set_tag_name_locs set_disconnected_nodes set_disconnected_nodes_locs create_element get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs DocumentClass.known_ptr DocumentClass.type_wf by(auto simp add: l_create_element_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_create_element_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \create\_character\_data\ locale l_create_character_data_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs + l_create_character_data\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs set_val set_val_locs create_character_data known_ptr type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_new_character_data_get_disconnected_nodes get_disconnected_nodes get_disconnected_nodes_locs + l_set_val_get_disconnected_nodes type_wf set_val set_val_locs get_disconnected_nodes get_disconnected_nodes_locs + l_new_character_data_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs + l_set_val_get_child_nodes type_wf set_val set_val_locs known_ptr get_child_nodes get_child_nodes_locs + l_set_disconnected_nodes_get_child_nodes set_disconnected_nodes set_disconnected_nodes_locs get_child_nodes get_child_nodes_locs + l_set_disconnected_nodes type_wf set_disconnected_nodes set_disconnected_nodes_locs + l_set_disconnected_nodes_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs + l_set_val_get_shadow_root type_wf set_val set_val_locs get_shadow_root get_shadow_root_locs + l_set_disconnected_nodes_get_shadow_root set_disconnected_nodes set_disconnected_nodes_locs get_shadow_root get_shadow_root_locs + l_new_character_data_get_tag_name get_tag_name get_tag_name_locs + l_set_val_get_tag_name type_wf set_val set_val_locs get_tag_name get_tag_name_locs + l_get_tag_name type_wf get_tag_name get_tag_name_locs + l_set_disconnected_nodes_get_tag_name type_wf set_disconnected_nodes set_disconnected_nodes_locs get_tag_name get_tag_name_locs + l_new_character_data type_wf + l_known_ptrs known_ptr known_ptrs for known_ptr :: "(_::linorder) object_ptr \ bool" and known_ptrs :: "(_) heap \ bool" and type_wf :: "(_) heap \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and set_tag_name :: "(_) element_ptr \ char list \ ((_) heap, exception, unit) prog" and set_tag_name_locs :: "(_) element_ptr \ ((_) heap, exception, unit) prog set" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" and create_element :: "(_) document_ptr \ char list \ ((_) heap, exception, (_) element_ptr) prog" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_disconnected_document :: "(_) node_ptr \ ((_) heap, exception, (_) document_ptr) prog" and get_disconnected_document_locs :: "((_) heap \ (_) heap \ bool) set" and set_val :: "(_) character_data_ptr \ char list \ ((_) heap, exception, unit) prog" and set_val_locs :: "(_) character_data_ptr \ ((_) heap, exception, unit) prog set" and create_character_data :: "(_) document_ptr \ char list \ ((_) heap, exception, (_) character_data_ptr) prog" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_::linorder) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" begin lemma create_character_data_preserves_wellformedness: assumes "heap_is_wellformed h" and "h \ create_character_data document_ptr text \\<^sub>h h'" and "type_wf h" and "known_ptrs h" shows "heap_is_wellformed h'" and "type_wf h'" and "known_ptrs h'" proof - obtain new_character_data_ptr h2 h3 disc_nodes_h3 where new_character_data_ptr: "h \ new_character_data \\<^sub>r new_character_data_ptr" and h2: "h \ new_character_data \\<^sub>h h2" and h3: "h2 \ set_val new_character_data_ptr text \\<^sub>h h3" and disc_nodes_h3: "h3 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" and h': "h3 \ set_disconnected_nodes document_ptr (cast new_character_data_ptr # disc_nodes_h3) \\<^sub>h h'" using assms(2) by(auto simp add: CD.create_character_data_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF CD.get_disconnected_nodes_pure, rotated] ) then have "h \ create_character_data document_ptr text \\<^sub>r new_character_data_ptr" apply(auto simp add: CD.create_character_data_def intro!: bind_returns_result_I)[1] apply (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust) apply (metis is_OK_returns_heap_E is_OK_returns_result_I local.CD.get_disconnected_nodes_pure pure_returns_heap_eq) by (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust) have "new_character_data_ptr \ set |h \ character_data_ptr_kinds_M|\<^sub>r" using new_character_data_ptr CharacterDataMonad.ptr_kinds_ptr_kinds_M h2 using new_character_data_ptr_not_in_heap by blast then have "cast new_character_data_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r" by simp then have "cast new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r" by simp have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\| {|cast new_character_data_ptr|}" using new_character_data_new_ptr h2 new_character_data_ptr by blast then have node_ptr_kinds_eq_h: "node_ptr_kinds h2 = node_ptr_kinds h |\| {|cast new_character_data_ptr|}" apply(simp add: node_ptr_kinds_def) by force then have character_data_ptr_kinds_eq_h: "character_data_ptr_kinds h2 = character_data_ptr_kinds h |\| {|new_character_data_ptr|}" apply(simp add: character_data_ptr_kinds_def) by force have element_ptr_kinds_eq_h: "element_ptr_kinds h2 = element_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: node_ptr_kinds_def element_ptr_kinds_def) have document_ptr_kinds_eq_h: "document_ptr_kinds h2 = document_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: document_ptr_kinds_def) have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF CD.set_val_writes h3]) using CD.set_val_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2" using object_ptr_kinds_eq_h2 by(auto simp add: node_ptr_kinds_def) have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_disconnected_nodes_writes h']) using set_disconnected_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3" using object_ptr_kinds_eq_h3 by(auto simp add: node_ptr_kinds_def) have "known_ptr (cast new_character_data_ptr)" using \h \ create_character_data document_ptr text \\<^sub>r new_character_data_ptr\ local.create_character_data_known_ptr by blast then have "known_ptrs h2" using known_ptrs_new_ptr object_ptr_kinds_eq_h \known_ptrs h\ h2 by blast then have "known_ptrs h3" using known_ptrs_preserved object_ptr_kinds_eq_h2 by blast then show "known_ptrs h'" using known_ptrs_preserved object_ptr_kinds_eq_h3 by blast have "document_ptr |\| document_ptr_kinds h" using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2 CD.get_disconnected_nodes_ptr_in_heap \type_wf h\ document_ptr_kinds_def by (metis is_OK_returns_result_I) have children_eq_h: "\(ptr'::(_) object_ptr) children. ptr' \ cast new_character_data_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads h2 get_child_nodes_new_character_data[rotated, OF new_character_data_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have children_eq2_h: "\ptr'. ptr' \ cast new_character_data_ptr \ |h \ get_child_nodes ptr'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\| {|cast new_character_data_ptr|}" using new_character_data_new_ptr h2 new_character_data_ptr by blast then have node_ptr_kinds_eq_h: "node_ptr_kinds h2 = node_ptr_kinds h |\| {|cast new_character_data_ptr|}" apply(simp add: node_ptr_kinds_def) by force then have character_data_ptr_kinds_eq_h: "character_data_ptr_kinds h2 = character_data_ptr_kinds h |\| {|new_character_data_ptr|}" apply(simp add: character_data_ptr_kinds_def) by force have element_ptr_kinds_eq_h: "element_ptr_kinds h2 = element_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: node_ptr_kinds_def element_ptr_kinds_def) have document_ptr_kinds_eq_h: "document_ptr_kinds h2 = document_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: document_ptr_kinds_def) have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF CD.set_val_writes h3]) using CD.set_val_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2" using object_ptr_kinds_eq_h2 by(auto simp add: node_ptr_kinds_def) then have character_data_ptr_kinds_eq_h2: "character_data_ptr_kinds h3 = character_data_ptr_kinds h2" by(simp add: character_data_ptr_kinds_def) have element_ptr_kinds_eq_h2: "element_ptr_kinds h3 = element_ptr_kinds h2" using node_ptr_kinds_eq_h2 by(simp add: element_ptr_kinds_def) have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_disconnected_nodes_writes h']) using set_disconnected_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3" using object_ptr_kinds_eq_h3 by(auto simp add: node_ptr_kinds_def) then have character_data_ptr_kinds_eq_h3: "character_data_ptr_kinds h' = character_data_ptr_kinds h3" by(simp add: character_data_ptr_kinds_def) have element_ptr_kinds_eq_h3: "element_ptr_kinds h' = element_ptr_kinds h3" using node_ptr_kinds_eq_h3 by(simp add: element_ptr_kinds_def) have "document_ptr |\| document_ptr_kinds h" using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2 CD.get_disconnected_nodes_ptr_in_heap \type_wf h\ document_ptr_kinds_def by (metis is_OK_returns_result_I) have children_eq_h: "\(ptr'::(_) object_ptr) children. ptr' \ cast new_character_data_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads h2 get_child_nodes_new_character_data[rotated, OF new_character_data_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have children_eq2_h: "\ptr'. ptr' \ cast new_character_data_ptr \ |h \ get_child_nodes ptr'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have "h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []" using new_character_data_ptr h2 new_character_data_ptr_in_heap[OF h2 new_character_data_ptr] new_character_data_is_character_data_ptr[OF new_character_data_ptr] new_character_data_no_child_nodes by blast have disconnected_nodes_eq_h: "\doc_ptr disc_nodes. h \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using CD.get_disconnected_nodes_reads h2 get_disconnected_nodes_new_character_data[OF new_character_data_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have disconnected_nodes_eq2_h: "\doc_ptr. |h \ get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have tag_name_eq_h: "\ptr' disc_nodes. h \ get_tag_name ptr' \\<^sub>r disc_nodes = h2 \ get_tag_name ptr' \\<^sub>r disc_nodes" using get_tag_name_reads h2 get_tag_name_new_character_data[OF new_character_data_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have tag_name_eq2_h: "\ptr'. |h \ get_tag_name ptr'|\<^sub>r = |h2 \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. h2 \ get_child_nodes ptr' \\<^sub>r children = h3 \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads CD.set_val_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_val_get_child_nodes) then have children_eq2_h2: "\ptr'. |h2 \ get_child_nodes ptr'|\<^sub>r = |h3 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using CD.get_disconnected_nodes_reads CD.set_val_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_val_get_disconnected_nodes) then have disconnected_nodes_eq2_h2: "\doc_ptr. |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have tag_name_eq_h2: "\ptr' disc_nodes. h2 \ get_tag_name ptr' \\<^sub>r disc_nodes = h3 \ get_tag_name ptr' \\<^sub>r disc_nodes" using get_tag_name_reads CD.set_val_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_val_get_tag_name) then have tag_name_eq2_h2: "\ptr'. |h2 \ get_tag_name ptr'|\<^sub>r = |h3 \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have "type_wf h2" using \type_wf h\ new_character_data_types_preserved h2 by blast then have "type_wf h3" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF CD.set_val_writes h3] using set_val_types_preserved by(auto simp add: reflp_def transp_def) then show "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h'] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) have children_eq_h3: "\ptr' children. h3 \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_disconnected_nodes_get_child_nodes) then have children_eq2_h3: " \ptr'. |h3 \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h3: "\doc_ptr disc_nodes. document_ptr \ doc_ptr \ h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using CD.get_disconnected_nodes_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers) then have disconnected_nodes_eq2_h3: "\doc_ptr. document_ptr \ doc_ptr \ |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have tag_name_eq_h3: "\ptr' disc_nodes. h3 \ get_tag_name ptr' \\<^sub>r disc_nodes = h' \ get_tag_name ptr' \\<^sub>r disc_nodes" using get_tag_name_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_disconnected_nodes_get_tag_name) then have tag_name_eq2_h3: "\ptr'. |h3 \ get_tag_name ptr'|\<^sub>r = |h' \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have disc_nodes_document_ptr_h2: "h2 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" using disconnected_nodes_eq_h2 disc_nodes_h3 by auto then have disc_nodes_document_ptr_h: "h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" using disconnected_nodes_eq_h by auto then have "cast new_character_data_ptr \ set disc_nodes_h3" using \heap_is_wellformed h\ using \cast new_character_data_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ a_all_ptrs_in_heap_def heap_is_wellformed_def using NodeMonad.ptr_kinds_ptr_kinds_M local.heap_is_wellformed_disc_nodes_in_heap by blast have "acyclic (parent_child_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def) also have "parent_child_rel h = parent_child_rel h2" proof(auto simp add: CD.parent_child_rel_def)[1] fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h2" by (simp add: object_ptr_kinds_eq_h) next fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "x \ set |h2 \ get_child_nodes a|\<^sub>r" by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2_h) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h" using object_ptr_kinds_eq_h \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ by(auto) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "x \ set |h \ get_child_nodes a|\<^sub>r" by (metis (no_types, lifting) \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ children_eq2_h empty_iff empty_set image_eqI select_result_I2) qed also have "\ = parent_child_rel h3" by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2) also have "\ = parent_child_rel h'" by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3) finally have "CD.a_acyclic_heap h'" by (simp add: CD.acyclic_heap_def) have "CD.a_all_ptrs_in_heap h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_all_ptrs_in_heap h2" apply(auto simp add: CD.a_all_ptrs_in_heap_def)[1] using node_ptr_kinds_eq_h \cast new_character_data_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ apply (metis (no_types, lifting) NodeMonad.ptr_kinds_ptr_kinds_M \parent_child_rel h = parent_child_rel h2\ children_eq2_h finite_set_in finsert_iff funion_finsert_right CD.parent_child_rel_child CD.parent_child_rel_parent_in_heap node_ptr_kinds_commutes object_ptr_kinds_eq_h select_result_I2 subsetD sup_bot.right_neutral) by (metis (no_types, lifting) CD.get_child_nodes_ok CD.get_child_nodes_ptr_in_heap \h2 \ get_child_nodes (cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_character_data_ptr) \\<^sub>r []\ assms(3) assms(4) children_eq_h disconnected_nodes_eq2_h document_ptr_kinds_eq_h finite_set_in is_OK_returns_result_I local.known_ptrs_known_ptr node_ptr_kinds_commutes returns_result_select_result subset_code(1)) then have "CD.a_all_ptrs_in_heap h3" by (simp add: children_eq2_h2 disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 object_ptr_kinds_eq_h2) then have "CD.a_all_ptrs_in_heap h'" by (smt character_data_ptr_kinds_commutes character_data_ptr_kinds_eq_h2 children_eq2_h3 disc_nodes_h3 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 h' h2 local.CD.a_all_ptrs_in_heap_def local.set_disconnected_nodes_get_disconnected_nodes new_character_data_ptr new_character_data_ptr_in_heap node_ptr_kinds_eq_h3 notin_fset object_ptr_kinds_eq_h3 select_result_I2 set_ConsD subset_code(1)) have "\p. p |\| object_ptr_kinds h \ cast new_character_data_ptr \ set |h \ get_child_nodes p|\<^sub>r" using \heap_is_wellformed h\ \cast new_character_data_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ heap_is_wellformed_children_in_heap by (meson NodeMonad.ptr_kinds_ptr_kinds_M CD.a_all_ptrs_in_heap_def assms(3) assms(4) fset_mp fset_of_list_elem CD.get_child_nodes_ok known_ptrs_known_ptr returns_result_select_result) then have "\p. p |\| object_ptr_kinds h2 \ cast new_character_data_ptr \ set |h2 \ get_child_nodes p|\<^sub>r" using children_eq2_h apply(auto simp add: object_ptr_kinds_eq_h)[1] using \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ apply auto[1] by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\) then have "\p. p |\| object_ptr_kinds h3 \ cast new_character_data_ptr \ set |h3 \ get_child_nodes p|\<^sub>r" using object_ptr_kinds_eq_h2 children_eq2_h2 by auto then have new_character_data_ptr_not_in_any_children: "\p. p |\| object_ptr_kinds h' \ cast new_character_data_ptr \ set |h' \ get_child_nodes p|\<^sub>r" using object_ptr_kinds_eq_h3 children_eq2_h3 by auto have "CD.a_distinct_lists h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_distinct_lists h2" using \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ apply(auto simp add: CD.a_distinct_lists_def object_ptr_kinds_eq_h document_ptr_kinds_eq_h disconnected_nodes_eq2_h intro!: distinct_concat_map_I)[1] - apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert) + apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert_remove) apply(case_tac "x=cast new_character_data_ptr") apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply (metis IntI assms(1) assms(3) assms(4) empty_iff CD.get_child_nodes_ok local.heap_is_wellformed_one_parent local.known_ptrs_known_ptr returns_result_select_result) apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] thm children_eq2_h using \CD.a_distinct_lists h\ \type_wf h2\ disconnected_nodes_eq_h document_ptr_kinds_eq_h CD.distinct_lists_no_parent get_disconnected_nodes_ok returns_result_select_result by metis then have "CD.a_distinct_lists h3" by(auto simp add: CD.a_distinct_lists_def disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 children_eq2_h2 object_ptr_kinds_eq_h2)[1] then have "CD.a_distinct_lists h'" proof(auto simp add: CD.a_distinct_lists_def disconnected_nodes_eq2_h3 children_eq2_h3 object_ptr_kinds_eq_h3 document_ptr_kinds_eq_h3 intro!: distinct_concat_map_I)[1] fix x assume "distinct (concat (map (\document_ptr. |h3 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" and "x |\| document_ptr_kinds h3" then show "distinct |h' \ get_disconnected_nodes x|\<^sub>r" using document_ptr_kinds_eq_h3 disconnected_nodes_eq_h3 h' set_disconnected_nodes_get_disconnected_nodes by (metis (no_types, hide_lams) \cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_character_data_ptr \ set disc_nodes_h3\ \type_wf h2\ assms(1) disc_nodes_document_ptr_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disconnected_nodes_eq_h distinct.simps(2) document_ptr_kinds_eq_h2 local.get_disconnected_nodes_ok local.heap_is_wellformed_disconnected_nodes_distinct returns_result_select_result select_result_I2) next fix x y xa assume "distinct (concat (map (\document_ptr. |h3 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" and "x |\| document_ptr_kinds h3" and "y |\| document_ptr_kinds h3" and "x \ y" and "xa \ set |h' \ get_disconnected_nodes x|\<^sub>r" and "xa \ set |h' \ get_disconnected_nodes y|\<^sub>r" moreover have "set |h3 \ get_disconnected_nodes x|\<^sub>r \ set |h3 \ get_disconnected_nodes y|\<^sub>r = {}" using calculation by(auto dest: distinct_concat_map_E(1)) ultimately show "False" using NodeMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_character_data_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ by (smt local.CD.a_all_ptrs_in_heap_def \CD.a_all_ptrs_in_heap h\ disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 finite_set_in h' l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms select_result_I2 set_ConsD subsetD) next fix x xa xb assume 2: "(\x\fset (object_ptr_kinds h3). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h3). set |h3 \ get_disconnected_nodes x|\<^sub>r) = {}" and 3: "xa |\| object_ptr_kinds h3" and 4: "x \ set |h' \ get_child_nodes xa|\<^sub>r" and 5: "xb |\| document_ptr_kinds h3" and 6: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" show "False" using disc_nodes_document_ptr_h disconnected_nodes_eq2_h3 apply(cases "document_ptr = xb") apply (metis (no_types, lifting) "3" "4" "5" "6" CD.distinct_lists_no_parent \local.CD.a_distinct_lists h2\ \type_wf h'\ children_eq2_h2 children_eq2_h3 disc_nodes_document_ptr_h2 document_ptr_kinds_eq_h3 h' local.get_disconnected_nodes_ok local.set_disconnected_nodes_get_disconnected_nodes new_character_data_ptr_not_in_any_children object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 returns_result_eq returns_result_select_result set_ConsD) by (metis "3" "4" "5" "6" CD.distinct_lists_no_parent \local.CD.a_distinct_lists h3\ \type_wf h3\ children_eq2_h3 local.get_disconnected_nodes_ok returns_result_select_result) qed have "CD.a_owner_document_valid h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_owner_document_valid h'" using disc_nodes_h3 \document_ptr |\| document_ptr_kinds h\ apply(simp add: CD.a_owner_document_valid_def) apply(simp add: object_ptr_kinds_eq_h object_ptr_kinds_eq_h3 ) apply(simp add: object_ptr_kinds_eq_h2) apply(simp add: document_ptr_kinds_eq_h document_ptr_kinds_eq_h3 ) apply(simp add: document_ptr_kinds_eq_h2) apply(simp add: node_ptr_kinds_eq_h node_ptr_kinds_eq_h3 ) apply(simp add: node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h ) apply(auto simp add: children_eq2_h2[symmetric] children_eq2_h3[symmetric] disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3)[1] apply (metis (no_types, lifting) document_ptr_kinds_eq_h h' list.set_intros(1) local.set_disconnected_nodes_get_disconnected_nodes select_result_I2) apply(simp add: object_ptr_kinds_eq_h) by (metis (mono_tags, lifting) \cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2_h disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h finite_set_in h' l_ptr_kinds_M.ptr_kinds_ptr_kinds_M l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes list.set_intros(2) local.l_set_disconnected_nodes_get_disconnected_nodes_axioms object_ptr_kinds_M_def select_result_I2) have shadow_root_ptr_kinds_eq_h: "shadow_root_ptr_kinds h2 = shadow_root_ptr_kinds h" using document_ptr_kinds_eq_h by(auto simp add: shadow_root_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h2: "shadow_root_ptr_kinds h3 = shadow_root_ptr_kinds h2" using document_ptr_kinds_eq_h2 by(auto simp add: shadow_root_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h3: "shadow_root_ptr_kinds h' = shadow_root_ptr_kinds h3" using document_ptr_kinds_eq_h3 by(auto simp add: shadow_root_ptr_kinds_def) have shadow_root_eq_h: "\character_data_ptr shadow_root_opt. h \ get_shadow_root character_data_ptr \\<^sub>r shadow_root_opt = h2 \ get_shadow_root character_data_ptr \\<^sub>r shadow_root_opt" using get_shadow_root_reads h2 get_shadow_root_new_character_data[rotated, OF h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] using local.get_shadow_root_locs_impl new_character_data_ptr apply blast using local.get_shadow_root_locs_impl new_character_data_ptr by blast have shadow_root_eq_h2: "\ptr' children. h2 \ get_shadow_root ptr' \\<^sub>r children = h3 \ get_shadow_root ptr' \\<^sub>r children" using get_shadow_root_reads set_val_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_val_get_shadow_root) have shadow_root_eq_h3: "\ptr' children. h3 \ get_shadow_root ptr' \\<^sub>r children = h' \ get_shadow_root ptr' \\<^sub>r children" using get_shadow_root_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) using set_disconnected_nodes_get_shadow_root by(auto simp add: set_disconnected_nodes_get_shadow_root) have "a_all_ptrs_in_heap h" by (simp add: assms(1) local.a_all_ptrs_in_heap_def local.get_shadow_root_shadow_root_ptr_in_heap) then have "a_all_ptrs_in_heap h2" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h)[1] using returns_result_eq shadow_root_eq_h by fastforce then have "a_all_ptrs_in_heap h3" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h2)[1] using shadow_root_eq_h2 by blast then have "a_all_ptrs_in_heap h'" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h3)[1] by (simp add: shadow_root_eq_h3) have "a_distinct_lists h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_distinct_lists h2" apply(auto simp add: a_distinct_lists_def character_data_ptr_kinds_eq_h)[1] apply(auto simp add: distinct_insort intro!: distinct_concat_map_I split: option.splits)[1] by (metis \type_wf h2\ assms(1) assms(3) local.get_shadow_root_ok local.shadow_root_same_host returns_result_select_result shadow_root_eq_h) then have "a_distinct_lists h3" by(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h2 select_result_eq[OF shadow_root_eq_h2]) then have "a_distinct_lists h'" by(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h3 select_result_eq[OF shadow_root_eq_h3]) have "a_shadow_root_valid h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_shadow_root_valid h2" by(auto simp add: a_shadow_root_valid_def shadow_root_ptr_kinds_eq_h element_ptr_kinds_eq_h select_result_eq[OF shadow_root_eq_h] tag_name_eq2_h) then have "a_shadow_root_valid h3" by(auto simp add: a_shadow_root_valid_def shadow_root_ptr_kinds_eq_h2 element_ptr_kinds_eq_h2 select_result_eq[OF shadow_root_eq_h2] tag_name_eq2_h2) then have "a_shadow_root_valid h'" by(auto simp add: a_shadow_root_valid_def shadow_root_ptr_kinds_eq_h3 element_ptr_kinds_eq_h3 select_result_eq[OF shadow_root_eq_h3] tag_name_eq2_h3) have "a_host_shadow_root_rel h = a_host_shadow_root_rel h2" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h select_result_eq[OF shadow_root_eq_h]) have "a_host_shadow_root_rel h2 = a_host_shadow_root_rel h3" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h2 select_result_eq[OF shadow_root_eq_h2]) have "a_host_shadow_root_rel h3 = a_host_shadow_root_rel h'" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h3 select_result_eq[OF shadow_root_eq_h3]) have "a_ptr_disconnected_node_rel h = a_ptr_disconnected_node_rel h2" by(simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq_h disconnected_nodes_eq2_h) have "a_ptr_disconnected_node_rel h2 = a_ptr_disconnected_node_rel h3" by(simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq_h2 disconnected_nodes_eq2_h2) have "h' \ get_disconnected_nodes document_ptr \\<^sub>r cast new_character_data_ptr # disc_nodes_h3" using h' local.set_disconnected_nodes_get_disconnected_nodes by auto have " document_ptr |\| document_ptr_kinds h3" by (simp add: \document_ptr |\| document_ptr_kinds h\ document_ptr_kinds_eq_h document_ptr_kinds_eq_h2) have "cast new_character_data_ptr \ set |h' \ get_disconnected_nodes document_ptr|\<^sub>r" using \h' \ get_disconnected_nodes document_ptr \\<^sub>r cast new_character_data_ptr # disc_nodes_h3\ by auto have "a_ptr_disconnected_node_rel h' = {(cast document_ptr, cast new_character_data_ptr)} \ a_ptr_disconnected_node_rel h3" apply(auto simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq_h3 disconnected_nodes_eq2_h3)[1] apply(case_tac "aa = document_ptr") using disc_nodes_h3 h' \h' \ get_disconnected_nodes document_ptr \\<^sub>r cast new_character_data_ptr # disc_nodes_h3\ apply(auto)[1] using disconnected_nodes_eq2_h3 apply auto[1] using \h' \ get_disconnected_nodes document_ptr \\<^sub>r cast new_character_data_ptr # disc_nodes_h3\ using \cast new_character_data_ptr \ set |h' \ get_disconnected_nodes document_ptr|\<^sub>r\ using \document_ptr |\| document_ptr_kinds h3\ apply auto[1] apply(case_tac "document_ptr = aa") using \h' \ get_disconnected_nodes document_ptr \\<^sub>r cast new_character_data_ptr # disc_nodes_h3\ disc_nodes_h3 apply auto[1] using disconnected_nodes_eq_h3[THEN select_result_eq, simplified] by auto have "acyclic (parent_child_rel h \ a_host_shadow_root_rel h \ a_ptr_disconnected_node_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) have "parent_child_rel h \ a_host_shadow_root_rel h \ a_ptr_disconnected_node_rel h = parent_child_rel h2 \ a_host_shadow_root_rel h2 \ a_ptr_disconnected_node_rel h2" using \local.a_host_shadow_root_rel h = local.a_host_shadow_root_rel h2\ \local.a_ptr_disconnected_node_rel h = local.a_ptr_disconnected_node_rel h2\ \parent_child_rel h = parent_child_rel h2\ by auto have "parent_child_rel h2 \ a_host_shadow_root_rel h2 \ a_ptr_disconnected_node_rel h2 = parent_child_rel h3 \ a_host_shadow_root_rel h3 \ a_ptr_disconnected_node_rel h3" using \local.a_host_shadow_root_rel h2 = local.a_host_shadow_root_rel h3\ \local.a_ptr_disconnected_node_rel h2 = local.a_ptr_disconnected_node_rel h3\ \parent_child_rel h2 = parent_child_rel h3\ by auto have "parent_child_rel h' \ a_host_shadow_root_rel h' \ a_ptr_disconnected_node_rel h' = {(cast document_ptr, cast new_character_data_ptr)} \ parent_child_rel h3 \ a_host_shadow_root_rel h3 \ a_ptr_disconnected_node_rel h3" by (simp add: \local.a_host_shadow_root_rel h3 = local.a_host_shadow_root_rel h'\ \local.a_ptr_disconnected_node_rel h' = {(cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr, cast new_character_data_ptr)} \ local.a_ptr_disconnected_node_rel h3\ \parent_child_rel h3 = parent_child_rel h'\) have "\a b. (a, b) \ parent_child_rel h3 \ a \ cast new_character_data_ptr" using CD.parent_child_rel_parent_in_heap \parent_child_rel h = parent_child_rel h2\ \parent_child_rel h2 = parent_child_rel h3\ character_data_ptr_kinds_commutes h2 new_character_data_ptr new_character_data_ptr_not_in_heap node_ptr_kinds_commutes by blast moreover have "\a b. (a, b) \ a_host_shadow_root_rel h3 \ a \ cast new_character_data_ptr" using shadow_root_eq_h2 by(auto simp add: a_host_shadow_root_rel_def) moreover have "\a b. (a, b) \ a_ptr_disconnected_node_rel h3 \ a \ cast new_character_data_ptr" by(auto simp add: a_ptr_disconnected_node_rel_def) moreover have "cast new_character_data_ptr \ {x. (x, cast document_ptr) \ (parent_child_rel h3 \ a_host_shadow_root_rel h3 \ a_ptr_disconnected_node_rel h3)\<^sup>*}" by (smt Un_iff calculation(1) calculation(2) calculation(3) cast_document_ptr_not_node_ptr(2) converse_rtranclE mem_Collect_eq) moreover have "acyclic (parent_child_rel h3 \ a_host_shadow_root_rel h3 \ a_ptr_disconnected_node_rel h3)" using \acyclic (parent_child_rel h \ local.a_host_shadow_root_rel h \ local.a_ptr_disconnected_node_rel h)\ \parent_child_rel h \ local.a_host_shadow_root_rel h \ local.a_ptr_disconnected_node_rel h = parent_child_rel h2 \ local.a_host_shadow_root_rel h2 \ local.a_ptr_disconnected_node_rel h2\ \parent_child_rel h2 \ local.a_host_shadow_root_rel h2 \ local.a_ptr_disconnected_node_rel h2 = parent_child_rel h3 \ local.a_host_shadow_root_rel h3 \ local.a_ptr_disconnected_node_rel h3\ by auto ultimately have "acyclic (parent_child_rel h' \ a_host_shadow_root_rel h' \ a_ptr_disconnected_node_rel h')" by(simp add: \parent_child_rel h' \ a_host_shadow_root_rel h' \ a_ptr_disconnected_node_rel h' = {(cast document_ptr, cast new_character_data_ptr)} \ parent_child_rel h3 \ a_host_shadow_root_rel h3 \ a_ptr_disconnected_node_rel h3\) have "CD.a_heap_is_wellformed h'" apply(simp add: CD.a_heap_is_wellformed_def) by (simp add: \local.CD.a_acyclic_heap h'\ \local.CD.a_all_ptrs_in_heap h'\ \local.CD.a_distinct_lists h'\ \local.CD.a_owner_document_valid h'\) show " heap_is_wellformed h' " using \acyclic (parent_child_rel h' \ local.a_host_shadow_root_rel h' \ local.a_ptr_disconnected_node_rel h')\ by(simp add: heap_is_wellformed_def CD.heap_is_wellformed_impl \local.CD.a_heap_is_wellformed h'\ \local.a_all_ptrs_in_heap h'\ \local.a_distinct_lists h'\ \local.a_shadow_root_valid h'\) qed end subsubsection \create\_document\ locale l_create_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs + l_new_document_get_disconnected_nodes get_disconnected_nodes get_disconnected_nodes_locs + l_create_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M create_document + l_new_document_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs + l_get_tag_name type_wf get_tag_name get_tag_name_locs + l_new_document_get_tag_name get_tag_name get_tag_name_locs + l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs + l_new_document type_wf + l_known_ptrs known_ptr known_ptrs for known_ptr :: "(_::linorder) object_ptr \ bool" and type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_disconnected_document :: "(_) node_ptr \ ((_) heap, exception, (_) document_ptr) prog" and get_disconnected_document_locs :: "((_) heap \ (_) heap \ bool) set" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and set_val :: "(_) character_data_ptr \ char list \ ((_) heap, exception, unit) prog" and set_val_locs :: "(_) character_data_ptr \ ((_) heap, exception, unit) prog set" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" and create_document :: "((_) heap, exception, (_) document_ptr) prog" and known_ptrs :: "(_) heap \ bool" begin lemma create_document_preserves_wellformedness: assumes "heap_is_wellformed h" and "h \ create_document \\<^sub>h h'" and "type_wf h" and "known_ptrs h" shows "heap_is_wellformed h'" proof - obtain new_document_ptr where new_document_ptr: "h \ new_document \\<^sub>r new_document_ptr" and h': "h \ new_document \\<^sub>h h'" using assms(2) apply(simp add: create_document_def) using new_document_ok by blast have "new_document_ptr \ set |h \ document_ptr_kinds_M|\<^sub>r" using new_document_ptr DocumentMonad.ptr_kinds_ptr_kinds_M using new_document_ptr_not_in_heap h' by blast then have "cast new_document_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r" by simp have "new_document_ptr |\| document_ptr_kinds h" using new_document_ptr DocumentMonad.ptr_kinds_ptr_kinds_M using new_document_ptr_not_in_heap h' by blast then have "cast new_document_ptr |\| object_ptr_kinds h" by simp have object_ptr_kinds_eq: "object_ptr_kinds h' = object_ptr_kinds h |\| {|cast new_document_ptr|}" using new_document_new_ptr h' new_document_ptr by blast then have node_ptr_kinds_eq: "node_ptr_kinds h' = node_ptr_kinds h" apply(simp add: node_ptr_kinds_def) by force then have character_data_ptr_kinds_eq_h: "character_data_ptr_kinds h' = character_data_ptr_kinds h" by(simp add: character_data_ptr_kinds_def) have element_ptr_kinds_eq_h: "element_ptr_kinds h' = element_ptr_kinds h" using object_ptr_kinds_eq by(auto simp add: node_ptr_kinds_def element_ptr_kinds_def) have document_ptr_kinds_eq_h: "document_ptr_kinds h' = document_ptr_kinds h |\| {|new_document_ptr|}" using object_ptr_kinds_eq apply(auto simp add: document_ptr_kinds_def)[1] by (metis (no_types, lifting) document_ptr_kinds_commutes document_ptr_kinds_def finsertI1 fset.map_comp) have children_eq: "\(ptr'::(_) object_ptr) children. ptr' \ cast new_document_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads h' get_child_nodes_new_document[rotated, OF new_document_ptr h'] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have children_eq2: "\ptr'. ptr' \ cast new_document_ptr \ |h \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have "h' \ get_child_nodes (cast new_document_ptr) \\<^sub>r []" using new_document_ptr h' new_document_ptr_in_heap[OF h' new_document_ptr] new_document_is_document_ptr[OF new_document_ptr] new_document_no_child_nodes by blast have disconnected_nodes_eq_h: "\doc_ptr disc_nodes. doc_ptr \ new_document_ptr \ h \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using CD.get_disconnected_nodes_reads h' get_disconnected_nodes_new_document_different_pointers new_document_ptr apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by (metis(full_types) \\thesis. (\new_document_ptr. \h \ new_document \\<^sub>r new_document_ptr; h \ new_document \\<^sub>h h'\ \ thesis) \ thesis\ local.get_disconnected_nodes_new_document_different_pointers new_document_ptr)+ then have disconnected_nodes_eq2_h: "\doc_ptr. doc_ptr \ new_document_ptr \ |h \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have "h' \ get_disconnected_nodes new_document_ptr \\<^sub>r []" using h' local.new_document_no_disconnected_nodes new_document_ptr by blast have "type_wf h'" using \type_wf h\ new_document_types_preserved h' by blast have "acyclic (parent_child_rel h)" using \heap_is_wellformed h\ by (auto simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def) also have "parent_child_rel h = parent_child_rel h'" proof(auto simp add: CD.parent_child_rel_def)[1] fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h'" by (simp add: object_ptr_kinds_eq) next fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "x \ set |h' \ get_child_nodes a|\<^sub>r" by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast new_document_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2) next fix a x assume 0: "a |\| object_ptr_kinds h'" and 1: "x \ set |h' \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h" using object_ptr_kinds_eq \h' \ get_child_nodes (cast new_document_ptr) \\<^sub>r []\ by(auto) next fix a x assume 0: "a |\| object_ptr_kinds h'" and 1: "x \ set |h' \ get_child_nodes a|\<^sub>r" then show "x \ set |h \ get_child_nodes a|\<^sub>r" by (metis (no_types, lifting) \h' \ get_child_nodes (cast new_document_ptr) \\<^sub>r []\ children_eq2 empty_iff empty_set image_eqI select_result_I2) qed finally have "CD.a_acyclic_heap h'" by (simp add: CD.acyclic_heap_def) have "CD.a_all_ptrs_in_heap h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def ) then have "CD.a_all_ptrs_in_heap h'" apply(auto simp add: CD.a_all_ptrs_in_heap_def)[1] using ObjectMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ \parent_child_rel h = parent_child_rel h'\ assms(1) children_eq fset_of_list_elem local.heap_is_wellformed_children_in_heap CD.parent_child_rel_child CD.parent_child_rel_parent_in_heap node_ptr_kinds_eq apply (metis (no_types, lifting) \h' \ get_child_nodes (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr) \\<^sub>r []\ children_eq2 finite_set_in finsert_iff funion_finsert_right object_ptr_kinds_eq select_result_I2 subsetD sup_bot.right_neutral) by (metis (no_types, lifting) \h' \ get_disconnected_nodes new_document_ptr \\<^sub>r []\ \type_wf h'\ assms(1) disconnected_nodes_eq_h empty_iff empty_set local.get_disconnected_nodes_ok local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq returns_result_select_result select_result_I2) have "CD.a_distinct_lists h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_distinct_lists h'" using \h' \ get_disconnected_nodes new_document_ptr \\<^sub>r []\ \h' \ get_child_nodes (cast new_document_ptr) \\<^sub>r []\ apply(auto simp add: children_eq2[symmetric] CD.a_distinct_lists_def insort_split object_ptr_kinds_eq document_ptr_kinds_eq_h disconnected_nodes_eq2_h intro!: distinct_concat_map_I)[1] - apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert) + apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert_remove) apply(auto simp add: dest: distinct_concat_map_E)[1] apply(auto simp add: dest: distinct_concat_map_E)[1] using \new_document_ptr |\| document_ptr_kinds h\ apply(auto simp add: distinct_insort dest: distinct_concat_map_E)[1] apply (metis assms(1) assms(3) disconnected_nodes_eq2_h get_disconnected_nodes_ok local.heap_is_wellformed_disconnected_nodes_distinct returns_result_select_result) proof - fix x :: "(_) document_ptr" and y :: "(_) document_ptr" and xa :: "(_) node_ptr" assume a1: "x \ y" assume a2: "x |\| document_ptr_kinds h" assume a3: "x \ new_document_ptr" assume a4: "y |\| document_ptr_kinds h" assume a5: "y \ new_document_ptr" assume a6: "distinct (concat (map (\document_ptr. |h \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h)))))" assume a7: "xa \ set |h' \ get_disconnected_nodes x|\<^sub>r" assume a8: "xa \ set |h' \ get_disconnected_nodes y|\<^sub>r" have f9: "xa \ set |h \ get_disconnected_nodes x|\<^sub>r" using a7 a3 disconnected_nodes_eq2_h by presburger have f10: "xa \ set |h \ get_disconnected_nodes y|\<^sub>r" using a8 a5 disconnected_nodes_eq2_h by presburger have f11: "y \ set (sorted_list_of_set (fset (document_ptr_kinds h)))" using a4 by simp have "x \ set (sorted_list_of_set (fset (document_ptr_kinds h)))" using a2 by simp then show False using f11 f10 f9 a6 a1 by (meson disjoint_iff_not_equal distinct_concat_map_E(1)) next fix x xa xb assume 0: "h' \ get_disconnected_nodes new_document_ptr \\<^sub>r []" and 1: "h' \ get_child_nodes (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr) \\<^sub>r []" and 2: "distinct (concat (map (\ptr. |h \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h)))))" and 3: "distinct (concat (map (\document_ptr. |h \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h)))))" and 4: "(\x\fset (object_ptr_kinds h). set |h \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h). set |h \ get_disconnected_nodes x|\<^sub>r) = {}" and 5: "x \ set |h \ get_child_nodes xa|\<^sub>r" and 6: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" and 7: "xa |\| object_ptr_kinds h" and 8: "xa \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr" and 9: "xb |\| document_ptr_kinds h" and 10: "xb \ new_document_ptr" then show "False" by (metis \CD.a_distinct_lists h\ assms(3) disconnected_nodes_eq2_h CD.distinct_lists_no_parent get_disconnected_nodes_ok returns_result_select_result) qed have "CD.a_owner_document_valid h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_owner_document_valid h'" apply(auto simp add: CD.a_owner_document_valid_def)[1] by (metis \cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr |\| object_ptr_kinds h\ children_eq2 disconnected_nodes_eq2_h document_ptr_kinds_commutes finite_set_in funion_iff node_ptr_kinds_eq object_ptr_kinds_eq) have shadow_root_eq_h: "\character_data_ptr shadow_root_opt. h \ get_shadow_root character_data_ptr \\<^sub>r shadow_root_opt = h' \ get_shadow_root character_data_ptr \\<^sub>r shadow_root_opt" using get_shadow_root_reads assms(2) get_shadow_root_new_document[rotated, OF h'] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] using local.get_shadow_root_locs_impl new_document_ptr apply blast using local.get_shadow_root_locs_impl new_document_ptr by blast have "a_all_ptrs_in_heap h" by (simp add: assms(1) local.a_all_ptrs_in_heap_def local.get_shadow_root_shadow_root_ptr_in_heap) then have "a_all_ptrs_in_heap h'" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_def document_ptr_kinds_eq_h)[1] using shadow_root_eq_h by fastforce have "a_distinct_lists h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_distinct_lists h'" apply(auto simp add: a_distinct_lists_def character_data_ptr_kinds_eq_h)[1] apply(auto simp add: distinct_insort intro!: distinct_concat_map_I split: option.splits)[1] by (metis \type_wf h'\ assms(1) assms(3) local.get_shadow_root_ok local.shadow_root_same_host returns_result_select_result shadow_root_eq_h) have tag_name_eq_h: "\ptr' disc_nodes. h \ get_tag_name ptr' \\<^sub>r disc_nodes = h' \ get_tag_name ptr' \\<^sub>r disc_nodes" using get_tag_name_reads h' get_tag_name_new_document[OF new_document_ptr h'] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ have "a_shadow_root_valid h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_shadow_root_valid h'" using new_document_is_document_ptr[OF new_document_ptr] by(auto simp add: a_shadow_root_valid_def element_ptr_kinds_eq_h document_ptr_kinds_eq_h shadow_root_ptr_kinds_def select_result_eq[OF shadow_root_eq_h] select_result_eq[OF tag_name_eq_h] is_shadow_root_ptr_kind\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def is_document_ptr\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def split: option.splits) have "a_host_shadow_root_rel h = a_host_shadow_root_rel h'" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h select_result_eq[OF shadow_root_eq_h]) have "a_ptr_disconnected_node_rel h = a_ptr_disconnected_node_rel h'" apply(auto simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq_h disconnected_nodes_eq2_h)[1] using \new_document_ptr |\| document_ptr_kinds h\ disconnected_nodes_eq2_h apply fastforce using new_document_disconnected_nodes[OF h' new_document_ptr] apply(simp add: CD.get_disconnected_nodes_impl CD.a_get_disconnected_nodes_def) using \new_document_ptr |\| document_ptr_kinds h\ disconnected_nodes_eq2_h apply fastforce done have "acyclic (parent_child_rel h \ a_host_shadow_root_rel h \ a_ptr_disconnected_node_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) moreover have "parent_child_rel h \ a_host_shadow_root_rel h \ a_ptr_disconnected_node_rel h = parent_child_rel h' \ a_host_shadow_root_rel h' \ a_ptr_disconnected_node_rel h'" by (simp add: \local.a_host_shadow_root_rel h = local.a_host_shadow_root_rel h'\ \local.a_ptr_disconnected_node_rel h = local.a_ptr_disconnected_node_rel h'\ \parent_child_rel h = parent_child_rel h'\) ultimately have "acyclic (parent_child_rel h' \ a_host_shadow_root_rel h' \ a_ptr_disconnected_node_rel h')" by simp have "CD.a_heap_is_wellformed h'" apply(simp add: CD.a_heap_is_wellformed_def) by (simp add: \local.CD.a_acyclic_heap h'\ \local.CD.a_all_ptrs_in_heap h'\ \local.CD.a_distinct_lists h'\ \local.CD.a_owner_document_valid h'\) show "heap_is_wellformed h'" using CD.heap_is_wellformed_impl \acyclic (parent_child_rel h' \ local.a_host_shadow_root_rel h' \ local.a_ptr_disconnected_node_rel h')\ \local.CD.a_heap_is_wellformed h'\ \local.a_all_ptrs_in_heap h'\ \local.a_distinct_lists h'\ \local.a_shadow_root_valid h'\ local.heap_is_wellformed_def by auto qed end interpretation l_create_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf DocumentClass.type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs heap_is_wellformed parent_child_rel set_val set_val_locs set_disconnected_nodes set_disconnected_nodes_locs create_document known_ptrs by(auto simp add: l_create_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) subsubsection \attach\_shadow\_root\ locale l_attach_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs + l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr set_shadow_root set_shadow_root_locs set_mode set_mode_locs attach_shadow_root type_wf get_tag_name get_tag_name_locs get_shadow_root get_shadow_root_locs + l_new_shadow_root_get_disconnected_nodes get_disconnected_nodes get_disconnected_nodes_locs + l_set_mode_get_disconnected_nodes type_wf set_mode set_mode_locs get_disconnected_nodes get_disconnected_nodes_locs + l_new_shadow_root_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs + l_new_shadow_root_get_tag_name type_wf get_tag_name get_tag_name_locs + l_set_mode_get_child_nodes type_wf set_mode set_mode_locs known_ptr get_child_nodes get_child_nodes_locs + l_set_shadow_root_get_child_nodes type_wf set_shadow_root set_shadow_root_locs known_ptr get_child_nodes get_child_nodes_locs + l_set_shadow_root type_wf set_shadow_root set_shadow_root_locs + l_set_shadow_root_get_disconnected_nodes set_shadow_root set_shadow_root_locs get_disconnected_nodes get_disconnected_nodes_locs + l_set_mode_get_shadow_root type_wf set_mode set_mode_locs get_shadow_root get_shadow_root_locs + l_set_shadow_root_get_shadow_root type_wf set_shadow_root set_shadow_root_locs get_shadow_root get_shadow_root_locs + l_new_character_data_get_tag_name get_tag_name get_tag_name_locs + l_set_mode_get_tag_name type_wf set_mode set_mode_locs get_tag_name get_tag_name_locs + l_get_tag_name type_wf get_tag_name get_tag_name_locs + l_set_shadow_root_get_tag_name set_shadow_root set_shadow_root_locs get_tag_name get_tag_name_locs + l_new_shadow_root type_wf + l_known_ptrs known_ptr known_ptrs for known_ptr :: "(_::linorder) object_ptr \ bool" and known_ptrs :: "(_) heap \ bool" and type_wf :: "(_) heap \ bool" and get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and set_tag_name :: "(_) element_ptr \ char list \ ((_) heap, exception, unit) prog" and set_tag_name_locs :: "(_) element_ptr \ ((_) heap, exception, unit) prog set" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" and create_element :: "(_) document_ptr \ char list \ ((_) heap, exception, (_) element_ptr) prog" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_disconnected_document :: "(_) node_ptr \ ((_) heap, exception, (_) document_ptr) prog" and get_disconnected_document_locs :: "((_) heap \ (_) heap \ bool) set" and set_val :: "(_) character_data_ptr \ char list \ ((_) heap, exception, unit) prog" and set_val_locs :: "(_) character_data_ptr \ ((_) heap, exception, unit) prog set" and create_character_data :: "(_) document_ptr \ char list \ ((_) heap, exception, (_) character_data_ptr) prog" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_::linorder) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and set_shadow_root :: "(_) element_ptr \ (_) shadow_root_ptr option \ (_, unit) dom_prog" and set_shadow_root_locs :: "(_) element_ptr \ (_, unit) dom_prog set" and set_mode :: "(_) shadow_root_ptr \ shadow_root_mode \ (_, unit) dom_prog" and set_mode_locs :: "(_) shadow_root_ptr \ (_, unit) dom_prog set" and attach_shadow_root :: "(_) element_ptr \ shadow_root_mode \ (_, (_) shadow_root_ptr) dom_prog" begin lemma attach_shadow_root_child_preserves: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ attach_shadow_root element_ptr new_mode \\<^sub>h h'" shows "type_wf h'" and "known_ptrs h'" and "heap_is_wellformed h'" proof - obtain h2 h3 new_shadow_root_ptr element_tag_name where element_tag_name: "h \ get_tag_name element_ptr \\<^sub>r element_tag_name" and "element_tag_name \ safe_shadow_root_element_types" and prev_shadow_root: "h \ get_shadow_root element_ptr \\<^sub>r None" and h2: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h2" and new_shadow_root_ptr: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr" and h3: "h2 \ set_mode new_shadow_root_ptr new_mode \\<^sub>h h3" and h': "h3 \ set_shadow_root element_ptr (Some new_shadow_root_ptr) \\<^sub>h h'" using assms(4) by(auto simp add: attach_shadow_root_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_tag_name_pure, rotated] bind_returns_heap_E2[rotated, OF get_shadow_root_pure, rotated] split: if_splits) have "h \ attach_shadow_root element_ptr new_mode \\<^sub>r new_shadow_root_ptr" thm bind_pure_returns_result_I[OF get_tag_name_pure] apply(unfold attach_shadow_root_def)[1] using element_tag_name apply(rule bind_pure_returns_result_I[OF get_tag_name_pure]) apply(rule bind_pure_returns_result_I) using \element_tag_name \ safe_shadow_root_element_types\ apply(simp) using \element_tag_name \ safe_shadow_root_element_types\ apply(simp) using prev_shadow_root apply(rule bind_pure_returns_result_I[OF get_shadow_root_pure]) apply(rule bind_pure_returns_result_I) apply(simp) apply(simp) using h2 new_shadow_root_ptr h3 h' by(auto intro!: bind_returns_result_I intro: is_OK_returns_result_E[OF is_OK_returns_heap_I[OF h3]] is_OK_returns_result_E[OF is_OK_returns_heap_I[OF h']]) have "new_shadow_root_ptr \ set |h \ shadow_root_ptr_kinds_M|\<^sub>r" using new_shadow_root_ptr ShadowRootMonad.ptr_kinds_ptr_kinds_M h2 using new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ptr_not_in_heap by blast then have "cast new_shadow_root_ptr \ set |h \ document_ptr_kinds_M|\<^sub>r" by simp then have "cast new_shadow_root_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r" by simp have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\| {|cast new_shadow_root_ptr|}" using new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_new_ptr h2 new_shadow_root_ptr by blast then have document_ptr_kinds_eq_h: "document_ptr_kinds h2 = document_ptr_kinds h |\| {|cast new_shadow_root_ptr|}" apply(simp add: document_ptr_kinds_def) by force then have shadow_root_ptr_kinds_eq_h: "shadow_root_ptr_kinds h2 = shadow_root_ptr_kinds h |\| {|new_shadow_root_ptr|}" apply(simp add: shadow_root_ptr_kinds_def) by force have element_ptr_kinds_eq_h: "element_ptr_kinds h2 = element_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: node_ptr_kinds_def element_ptr_kinds_def) have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_mode_writes h3]) using set_mode_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2" by (auto simp add: document_ptr_kinds_def) then have shadow_root_ptr_kinds_eq_h2: "shadow_root_ptr_kinds h3 = shadow_root_ptr_kinds h2" by (auto simp add: shadow_root_ptr_kinds_def) have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2" using object_ptr_kinds_eq_h2 by(auto simp add: node_ptr_kinds_def) have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_shadow_root_writes h']) using set_shadow_root_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3" by (auto simp add: document_ptr_kinds_def) then have shadow_root_ptr_kinds_eq_h3: "shadow_root_ptr_kinds h' = shadow_root_ptr_kinds h3" by (auto simp add: shadow_root_ptr_kinds_def) have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3" using object_ptr_kinds_eq_h3 by(auto simp add: node_ptr_kinds_def) have "known_ptr (cast new_shadow_root_ptr)" using \h \ attach_shadow_root element_ptr new_mode \\<^sub>r new_shadow_root_ptr\ create_shadow_root_known_ptr by blast then have "known_ptrs h2" using known_ptrs_new_ptr object_ptr_kinds_eq_h \known_ptrs h\ h2 by blast then have "known_ptrs h3" using known_ptrs_preserved object_ptr_kinds_eq_h2 by blast then show "known_ptrs h'" using known_ptrs_preserved object_ptr_kinds_eq_h3 by blast have "element_ptr |\| element_ptr_kinds h" by (meson \h \ attach_shadow_root element_ptr new_mode \\<^sub>r new_shadow_root_ptr\ is_OK_returns_result_I local.attach_shadow_root_element_ptr_in_heap) have children_eq_h: "\(ptr'::(_) object_ptr) children. ptr' \ cast new_shadow_root_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads h2 get_child_nodes_new_shadow_root[rotated, OF new_shadow_root_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have children_eq2_h: "\ptr'. ptr' \ cast new_shadow_root_ptr \ |h \ get_child_nodes ptr'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\| {|cast new_shadow_root_ptr|}" using new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_new_ptr h2 new_shadow_root_ptr object_ptr_kinds_eq_h by blast then have node_ptr_kinds_eq_h: "node_ptr_kinds h2 = node_ptr_kinds h" apply(simp add: node_ptr_kinds_def) by force then have character_data_ptr_kinds_eq_h: "character_data_ptr_kinds h2 = character_data_ptr_kinds h" apply(simp add: character_data_ptr_kinds_def) done have element_ptr_kinds_eq_h: "element_ptr_kinds h2 = element_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: node_ptr_kinds_def element_ptr_kinds_def) have document_ptr_kinds_eq_h: "document_ptr_kinds h2 = document_ptr_kinds h |\| {|cast new_shadow_root_ptr|}" using object_ptr_kinds_eq_h apply(auto simp add: document_ptr_kinds_def)[1] by (metis (full_types) document_ptr_kinds_def document_ptr_kinds_eq_h finsert_fsubset fset.map_comp funion_upper2) have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_mode_writes h3]) using set_mode_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2" using object_ptr_kinds_eq_h2 by(auto simp add: node_ptr_kinds_def) then have character_data_ptr_kinds_eq_h2: "character_data_ptr_kinds h3 = character_data_ptr_kinds h2" by(simp add: character_data_ptr_kinds_def) have element_ptr_kinds_eq_h2: "element_ptr_kinds h3 = element_ptr_kinds h2" using node_ptr_kinds_eq_h2 by(simp add: element_ptr_kinds_def) have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_shadow_root_writes h']) using set_shadow_root_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3" using object_ptr_kinds_eq_h3 by(auto simp add: node_ptr_kinds_def) then have character_data_ptr_kinds_eq_h3: "character_data_ptr_kinds h' = character_data_ptr_kinds h3" by(simp add: character_data_ptr_kinds_def) have element_ptr_kinds_eq_h3: "element_ptr_kinds h' = element_ptr_kinds h3" using node_ptr_kinds_eq_h3 by(simp add: element_ptr_kinds_def) have children_eq_h: "\(ptr'::(_) object_ptr) children. ptr' \ cast new_shadow_root_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads h2 get_child_nodes_new_shadow_root[rotated, OF new_shadow_root_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have children_eq2_h: "\ptr'. ptr' \ cast new_shadow_root_ptr \ |h \ get_child_nodes ptr'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have "h2 \ get_child_nodes (cast new_shadow_root_ptr) \\<^sub>r []" using h2 local.new_shadow_root_no_child_nodes new_shadow_root_ptr by auto have disconnected_nodes_eq_h: "\doc_ptr disc_nodes. doc_ptr \ cast new_shadow_root_ptr \ h \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads h2 get_disconnected_nodes_new_shadow_root_different_pointers[rotated, OF new_shadow_root_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by (metis (no_types, lifting))+ then have disconnected_nodes_eq2_h: "\doc_ptr. doc_ptr \ cast new_shadow_root_ptr \ |h \ get_disconnected_nodes doc_ptr|\<^sub>r = |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have "h2 \ get_disconnected_nodes (cast new_shadow_root_ptr) \\<^sub>r []" using h2 new_shadow_root_no_disconnected_nodes new_shadow_root_ptr by auto have tag_name_eq_h: "\ptr' disc_nodes. h \ get_tag_name ptr' \\<^sub>r disc_nodes = h2 \ get_tag_name ptr' \\<^sub>r disc_nodes" using get_tag_name_reads h2 get_tag_name_new_shadow_root[OF new_shadow_root_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have tag_name_eq2_h: "\ptr'. |h \ get_tag_name ptr'|\<^sub>r = |h2 \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. h2 \ get_child_nodes ptr' \\<^sub>r children = h3 \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads set_mode_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_mode_get_child_nodes) then have children_eq2_h2: "\ptr'. |h2 \ get_child_nodes ptr'|\<^sub>r = |h3 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_mode_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_mode_get_disconnected_nodes) then have disconnected_nodes_eq2_h2: "\doc_ptr. |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have tag_name_eq_h2: "\ptr' disc_nodes. h2 \ get_tag_name ptr' \\<^sub>r disc_nodes = h3 \ get_tag_name ptr' \\<^sub>r disc_nodes" using get_tag_name_reads set_mode_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_mode_get_tag_name) then have tag_name_eq2_h2: "\ptr'. |h2 \ get_tag_name ptr'|\<^sub>r = |h3 \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have "type_wf h2" using \type_wf h\ new_shadow_root_types_preserved h2 by blast then have "type_wf h3" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_mode_writes h3] using set_mode_types_preserved by(auto simp add: reflp_def transp_def) then show "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_shadow_root_writes h'] using set_shadow_root_types_preserved by(auto simp add: reflp_def transp_def) have children_eq_h3: "\ptr' children. h3 \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads set_shadow_root_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_shadow_root_get_child_nodes) then have children_eq2_h3: " \ptr'. |h3 \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h3: "\doc_ptr disc_nodes. h3 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads set_shadow_root_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_shadow_root_get_disconnected_nodes) then have disconnected_nodes_eq2_h3: "\doc_ptr. |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have tag_name_eq_h3: "\ptr' disc_nodes. h3 \ get_tag_name ptr' \\<^sub>r disc_nodes = h' \ get_tag_name ptr' \\<^sub>r disc_nodes" using get_tag_name_reads set_shadow_root_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_shadow_root_get_tag_name) then have tag_name_eq2_h3: "\ptr'. |h3 \ get_tag_name ptr'|\<^sub>r = |h' \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have "acyclic (parent_child_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def) also have "parent_child_rel h = parent_child_rel h2" proof(auto simp add: CD.parent_child_rel_def)[1] fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h2" by (simp add: object_ptr_kinds_eq_h) next fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>r" then show "x \ set |h2 \ get_child_nodes a|\<^sub>r" by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast new_shadow_root_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2_h) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h" using object_ptr_kinds_eq_h \h2 \ get_child_nodes (cast new_shadow_root_ptr) \\<^sub>r []\ by(auto) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "x \ set |h \ get_child_nodes a|\<^sub>r" by (metis (no_types, lifting) \h2 \ get_child_nodes (cast new_shadow_root_ptr) \\<^sub>r []\ children_eq2_h empty_iff empty_set image_eqI select_result_I2) qed also have "\ = parent_child_rel h3" by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2) also have "\ = parent_child_rel h'" by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3) finally have "CD.a_acyclic_heap h'" by (simp add: CD.acyclic_heap_def) have "CD.a_all_ptrs_in_heap h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_all_ptrs_in_heap h2" apply(auto simp add: CD.a_all_ptrs_in_heap_def)[1] using node_ptr_kinds_eq_h \h2 \ get_child_nodes (cast new_shadow_root_ptr) \\<^sub>r []\ apply (metis (no_types, lifting) CD.get_child_nodes_ok CD.l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms \known_ptrs h2\ \parent_child_rel h = parent_child_rel h2\ \type_wf h2\ assms(1) assms(2) l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_child local.known_ptrs_known_ptr local.parent_child_rel_child_in_heap node_ptr_kinds_commutes returns_result_select_result) by (metis (no_types, hide_lams) \h2 \ get_disconnected_nodes (cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_shadow_root_ptr) \\<^sub>r []\ \type_wf h2\ disconnected_nodes_eq_h empty_iff finite_set_in is_OK_returns_result_E is_OK_returns_result_I local.get_disconnected_nodes_ok local.get_disconnected_nodes_ptr_in_heap node_ptr_kinds_eq_h select_result_I2 set_empty subset_code(1)) then have "CD.a_all_ptrs_in_heap h3" by (simp add: children_eq2_h2 disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 object_ptr_kinds_eq_h2) then have "CD.a_all_ptrs_in_heap h'" by (simp add: children_eq2_h3 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h3 object_ptr_kinds_eq_h3) have "cast new_shadow_root_ptr |\| document_ptr_kinds h" using h2 new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ptr_not_in_heap new_shadow_root_ptr shadow_root_ptr_kinds_commutes by blast have "CD.a_distinct_lists h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_distinct_lists h2" using \h2 \ get_disconnected_nodes (cast new_shadow_root_ptr) \\<^sub>r []\ \h2 \ get_child_nodes (cast new_shadow_root_ptr) \\<^sub>r []\ apply(auto simp add: children_eq2_h[symmetric] CD.a_distinct_lists_def insort_split object_ptr_kinds_eq_h document_ptr_kinds_eq_h disconnected_nodes_eq2_h intro!: distinct_concat_map_I)[1] - apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert) + apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert_remove) apply(auto simp add: dest: distinct_concat_map_E)[1] apply(auto simp add: dest: distinct_concat_map_E)[1] using \cast new_shadow_root_ptr |\| document_ptr_kinds h\ apply(auto simp add: distinct_insort dest: distinct_concat_map_E)[1] apply (metis (no_types) DocumentMonad.ptr_kinds_M_ptr_kinds DocumentMonad.ptr_kinds_ptr_kinds_M concat_map_all_distinct disconnected_nodes_eq2_h select_result_I2) proof - fix x :: "(_) document_ptr" and y :: "(_) document_ptr" and xa :: "(_) node_ptr" assume a1: "x \ y" assume a2: "x |\| document_ptr_kinds h" assume a3: "x \ cast new_shadow_root_ptr" assume a4: "y |\| document_ptr_kinds h" assume a5: "y \ cast new_shadow_root_ptr" assume a6: "distinct (concat (map (\document_ptr. |h \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h)))))" assume a7: "xa \ set |h2 \ get_disconnected_nodes x|\<^sub>r" assume a8: "xa \ set |h2 \ get_disconnected_nodes y|\<^sub>r" have f9: "xa \ set |h \ get_disconnected_nodes x|\<^sub>r" using a7 a3 disconnected_nodes_eq2_h by (simp add: disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3) have f10: "xa \ set |h \ get_disconnected_nodes y|\<^sub>r" using a8 a5 disconnected_nodes_eq2_h by (simp add: disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3) have f11: "y \ set (sorted_list_of_set (fset (document_ptr_kinds h)))" using a4 by simp have "x \ set (sorted_list_of_set (fset (document_ptr_kinds h)))" using a2 by simp then show False using f11 f10 f9 a6 a1 by (meson disjoint_iff_not_equal distinct_concat_map_E(1)) next fix x xa xb assume 0: "h2 \ get_disconnected_nodes (cast new_shadow_root_ptr) \\<^sub>r []" and 1: "h2 \ get_child_nodes (cast new_shadow_root_ptr) \\<^sub>r []" and 2: "distinct (concat (map (\ptr. |h \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h)))))" and 3: "distinct (concat (map (\document_ptr. |h \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h)))))" and 4: "(\x\fset (object_ptr_kinds h). set |h \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h). set |h \ get_disconnected_nodes x|\<^sub>r) = {}" and 5: "x \ set |h \ get_child_nodes xa|\<^sub>r" and 6: "x \ set |h2 \ get_disconnected_nodes xb|\<^sub>r" and 7: "xa |\| object_ptr_kinds h" and 8: "xa \ cast new_shadow_root_ptr" and 9: "xb |\| document_ptr_kinds h" and 10: "xb \ cast new_shadow_root_ptr" then show "False" by (metis CD.distinct_lists_no_parent \local.CD.a_distinct_lists h\ assms(2) disconnected_nodes_eq2_h local.get_disconnected_nodes_ok returns_result_select_result) qed then have "CD.a_distinct_lists h3" by(auto simp add: CD.a_distinct_lists_def disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 children_eq2_h2 object_ptr_kinds_eq_h2)[1] then have "CD.a_distinct_lists h'" by(auto simp add: CD.a_distinct_lists_def disconnected_nodes_eq2_h3 children_eq2_h3 object_ptr_kinds_eq_h3 document_ptr_kinds_eq_h3 intro!: distinct_concat_map_I) have "CD.a_owner_document_valid h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_owner_document_valid h'" (* using disc_nodes_h3 \document_ptr |\| document_ptr_kinds h\ *) apply(simp add: CD.a_owner_document_valid_def) apply(simp add: object_ptr_kinds_eq_h object_ptr_kinds_eq_h3 ) apply(simp add: object_ptr_kinds_eq_h2) apply(simp add: document_ptr_kinds_eq_h document_ptr_kinds_eq_h3 ) apply(simp add: document_ptr_kinds_eq_h2) apply(simp add: node_ptr_kinds_eq_h node_ptr_kinds_eq_h3 ) apply(simp add: node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h ) apply(auto simp add: children_eq2_h2[symmetric] children_eq2_h3[symmetric] disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3)[1] by (metis CD.get_child_nodes_ok CD.get_child_nodes_ptr_in_heap \cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_shadow_root_ptr |\| document_ptr_kinds h\ assms(2) assms(3) children_eq2_h children_eq_h disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_commutes finite_set_in is_OK_returns_result_I local.known_ptrs_known_ptr returns_result_select_result) have shadow_root_eq_h: "\character_data_ptr shadow_root_opt. h \ get_shadow_root character_data_ptr \\<^sub>r shadow_root_opt = h2 \ get_shadow_root character_data_ptr \\<^sub>r shadow_root_opt" using get_shadow_root_reads h2 get_shadow_root_new_shadow_root[rotated, OF h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] using local.get_shadow_root_locs_impl new_shadow_root_ptr apply blast using local.get_shadow_root_locs_impl new_shadow_root_ptr by blast have shadow_root_eq_h2: "\ptr' children. h2 \ get_shadow_root ptr' \\<^sub>r children = h3 \ get_shadow_root ptr' \\<^sub>r children" using get_shadow_root_reads set_mode_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_mode_get_shadow_root) have shadow_root_eq_h3: "\ptr' children. element_ptr \ ptr' \ h3 \ get_shadow_root ptr' \\<^sub>r children = h' \ get_shadow_root ptr' \\<^sub>r children" using get_shadow_root_reads set_shadow_root_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_shadow_root_get_shadow_root_different_pointers) have shadow_root_h3: "h' \ get_shadow_root element_ptr \\<^sub>r Some new_shadow_root_ptr" using \type_wf h3\ h' local.set_shadow_root_get_shadow_root by blast have "a_all_ptrs_in_heap h" by (simp add: assms(1) local.a_all_ptrs_in_heap_def local.get_shadow_root_shadow_root_ptr_in_heap) then have "a_all_ptrs_in_heap h2" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h)[1] using returns_result_eq shadow_root_eq_h by fastforce then have "a_all_ptrs_in_heap h3" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h2)[1] using shadow_root_eq_h2 by blast then have "a_all_ptrs_in_heap h'" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h3)[1] apply(case_tac "shadow_root_ptr = new_shadow_root_ptr") using h2 new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ptr_in_heap new_shadow_root_ptr shadow_root_ptr_kinds_eq_h2 apply blast using \type_wf h3\ h' local.set_shadow_root_get_shadow_root returns_result_eq shadow_root_eq_h3 apply fastforce done have "a_distinct_lists h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_distinct_lists h2" apply(auto simp add: a_distinct_lists_def character_data_ptr_kinds_eq_h)[1] apply(auto simp add: distinct_insort intro!: distinct_concat_map_I split: option.splits)[1] by (metis \type_wf h2\ assms(1) assms(2) local.get_shadow_root_ok local.shadow_root_same_host returns_result_select_result shadow_root_eq_h) then have "a_distinct_lists h3" by(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h2 select_result_eq[OF shadow_root_eq_h2]) then have "a_distinct_lists h'" apply(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h3 select_result_eq[OF shadow_root_eq_h3])[1] apply(auto simp add: distinct_insort intro!: distinct_concat_map_I split: option.splits)[1] by (smt \type_wf h3\ assms(1) assms(2) h' h2 local.get_shadow_root_ok local.get_shadow_root_shadow_root_ptr_in_heap local.set_shadow_root_get_shadow_root local.shadow_root_same_host new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ptr_not_in_heap new_shadow_root_ptr returns_result_select_result select_result_I2 shadow_root_eq_h shadow_root_eq_h2 shadow_root_eq_h3) have "a_shadow_root_valid h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_shadow_root_valid h'" proof(unfold a_shadow_root_valid_def; safe) fix shadow_root_ptr assume "\shadow_root_ptr\fset (shadow_root_ptr_kinds h). \host\fset (element_ptr_kinds h). |h \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types \ |h \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr" assume "a_shadow_root_valid h" assume "shadow_root_ptr \ fset (shadow_root_ptr_kinds h')" show "\host\fset (element_ptr_kinds h'). |h' \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types \ |h' \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr" proof (cases "shadow_root_ptr = new_shadow_root_ptr") case True have "element_ptr \ fset (element_ptr_kinds h')" by (simp add: \element_ptr |\| element_ptr_kinds h\ element_ptr_kinds_eq_h element_ptr_kinds_eq_h2 element_ptr_kinds_eq_h3) moreover have "|h' \ get_tag_name element_ptr|\<^sub>r \ safe_shadow_root_element_types" by (smt \\thesis. (\h2 h3 new_shadow_root_ptr element_tag_name. \h \ get_tag_name element_ptr \\<^sub>r element_tag_name; element_tag_name \ safe_shadow_root_element_types; h \ get_shadow_root element_ptr \\<^sub>r None; h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h2; h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr; h2 \ set_mode new_shadow_root_ptr new_mode \\<^sub>h h3; h3 \ set_shadow_root element_ptr (Some new_shadow_root_ptr) \\<^sub>h h'\ \ thesis) \ thesis\ select_result_I2 tag_name_eq2_h tag_name_eq2_h2 tag_name_eq2_h3) moreover have "|h' \ get_shadow_root element_ptr|\<^sub>r = Some shadow_root_ptr" using shadow_root_h3 by (simp add: True) ultimately show ?thesis by meson next case False then obtain host where host: "host \ fset (element_ptr_kinds h)" and "|h \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types" and "|h \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr" using \shadow_root_ptr \ fset (shadow_root_ptr_kinds h')\ using \\shadow_root_ptr\fset (shadow_root_ptr_kinds h). \host\fset (element_ptr_kinds h). |h \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types \ |h \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr\ apply(simp add: shadow_root_ptr_kinds_eq_h3 shadow_root_ptr_kinds_eq_h2 shadow_root_ptr_kinds_eq_h) by (meson finite_set_in) moreover have "host \ element_ptr" using calculation(3) prev_shadow_root by auto ultimately show ?thesis using element_ptr_kinds_eq_h3 element_ptr_kinds_eq_h2 element_ptr_kinds_eq_h by (smt \type_wf h'\ assms(2) finite_set_in local.get_shadow_root_ok returns_result_eq returns_result_select_result shadow_root_eq_h shadow_root_eq_h2 shadow_root_eq_h3 tag_name_eq2_h tag_name_eq2_h2 tag_name_eq2_h3) qed qed have "a_host_shadow_root_rel h = a_host_shadow_root_rel h2" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h select_result_eq[OF shadow_root_eq_h]) have "a_host_shadow_root_rel h2 = a_host_shadow_root_rel h3" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h2 select_result_eq[OF shadow_root_eq_h2]) have "a_host_shadow_root_rel h' = {(cast element_ptr, cast new_shadow_root_ptr)} \ a_host_shadow_root_rel h3" apply(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h3 )[1] apply(case_tac "element_ptr \ aa") using select_result_eq[OF shadow_root_eq_h3] apply (simp add: image_iff) using select_result_eq[OF shadow_root_eq_h3] apply (metis (no_types, lifting) \local.a_host_shadow_root_rel h = local.a_host_shadow_root_rel h2\ \local.a_host_shadow_root_rel h2 = local.a_host_shadow_root_rel h3\ \type_wf h3\ host_shadow_root_rel_def i_get_parent_get_host_get_disconnected_document_wf.a_host_shadow_root_rel_shadow_root local.get_shadow_root_impl local.get_shadow_root_ok option.distinct(1) prev_shadow_root returns_result_select_result) apply (metis (mono_tags, lifting) \\ptr'. (\x. element_ptr \ ptr') \ |h3 \ get_shadow_root ptr'|\<^sub>r = |h' \ get_shadow_root ptr'|\<^sub>r\ case_prod_conv image_iff is_OK_returns_result_I mem_Collect_eq option.inject returns_result_eq returns_result_select_result shadow_root_h3) using element_ptr_kinds_eq_h3 local.get_shadow_root_ptr_in_heap shadow_root_h3 apply fastforce apply (smt Shadow_DOM.a_host_shadow_root_rel_def \\ptr'. (\x. element_ptr \ ptr') \ |h3 \ get_shadow_root ptr'|\<^sub>r = |h' \ get_shadow_root ptr'|\<^sub>r\ \type_wf h3\ case_prodE case_prodI i_get_root_node_si_wf.a_host_shadow_root_rel_shadow_root image_iff local.get_shadow_root_impl local.get_shadow_root_ok mem_Collect_eq option.distinct(1) prev_shadow_root returns_result_eq returns_result_select_result shadow_root_eq_h shadow_root_eq_h2) done have "a_ptr_disconnected_node_rel h = a_ptr_disconnected_node_rel h2" apply(auto simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq_h)[1] apply (metis (no_types, lifting) \cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_shadow_root_ptr |\| document_ptr_kinds h\ case_prodI disconnected_nodes_eq2_h mem_Collect_eq pair_imageI) using \h2 \ get_disconnected_nodes (cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_shadow_root_ptr) \\<^sub>r []\ apply auto[1] apply(case_tac "cast new_shadow_root_ptr \ aa") apply (simp add: disconnected_nodes_eq2_h image_iff) using \cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_shadow_root_ptr |\| document_ptr_kinds h\ apply blast done have "a_ptr_disconnected_node_rel h2 = a_ptr_disconnected_node_rel h3" by(simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq_h2 disconnected_nodes_eq2_h2) have "a_ptr_disconnected_node_rel h3 = a_ptr_disconnected_node_rel h'" by(simp add: a_ptr_disconnected_node_rel_def document_ptr_kinds_eq_h3 disconnected_nodes_eq2_h3) have "acyclic (parent_child_rel h \ a_host_shadow_root_rel h \ a_ptr_disconnected_node_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) have "parent_child_rel h \ a_host_shadow_root_rel h \ a_ptr_disconnected_node_rel h = parent_child_rel h2 \ a_host_shadow_root_rel h2 \ a_ptr_disconnected_node_rel h2" using \local.a_host_shadow_root_rel h = local.a_host_shadow_root_rel h2\ \local.a_ptr_disconnected_node_rel h = local.a_ptr_disconnected_node_rel h2\ \parent_child_rel h = parent_child_rel h2\ by auto have "parent_child_rel h2 \ a_host_shadow_root_rel h2 \ a_ptr_disconnected_node_rel h2 = parent_child_rel h3 \ a_host_shadow_root_rel h3 \ a_ptr_disconnected_node_rel h3" using \local.a_host_shadow_root_rel h2 = local.a_host_shadow_root_rel h3\ \local.a_ptr_disconnected_node_rel h2 = local.a_ptr_disconnected_node_rel h3\ \parent_child_rel h2 = parent_child_rel h3\ by auto have "parent_child_rel h' \ a_host_shadow_root_rel h' \ a_ptr_disconnected_node_rel h' = {(cast element_ptr, cast new_shadow_root_ptr)} \ parent_child_rel h3 \ a_host_shadow_root_rel h3 \ a_ptr_disconnected_node_rel h3" by (simp add: \local.a_host_shadow_root_rel h' = {(cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r element_ptr, cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_shadow_root_ptr)} \ local.a_host_shadow_root_rel h3\ \local.a_ptr_disconnected_node_rel h3 = local.a_ptr_disconnected_node_rel h'\ \parent_child_rel h3 = parent_child_rel h'\) have "\a b. (a, b) \ parent_child_rel h3 \ a \ cast new_shadow_root_ptr" using CD.parent_child_rel_parent_in_heap \cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_shadow_root_ptr |\| document_ptr_kinds h\ \parent_child_rel h = parent_child_rel h2\ \parent_child_rel h2 = parent_child_rel h3\ document_ptr_kinds_commutes by blast moreover have "\a b. (a, b) \ a_host_shadow_root_rel h3 \ a \ cast new_shadow_root_ptr" using shadow_root_eq_h2 by(auto simp add: a_host_shadow_root_rel_def) moreover have "\a b. (a, b) \ a_ptr_disconnected_node_rel h3 \ a \ cast new_shadow_root_ptr" using \h2 \ get_disconnected_nodes (cast new_shadow_root_ptr) \\<^sub>r []\ by(auto simp add: a_ptr_disconnected_node_rel_def disconnected_nodes_eq_h2) moreover have "cast new_shadow_root_ptr \ {x. (x, cast element_ptr) \ (parent_child_rel h3 \ a_host_shadow_root_rel h3 \ a_ptr_disconnected_node_rel h3)\<^sup>*}" by (smt Un_iff calculation(1) calculation(2) calculation(3) cast_document_ptr_not_node_ptr(2) converse_rtranclE mem_Collect_eq) moreover have "acyclic (parent_child_rel h3 \ a_host_shadow_root_rel h3 \ a_ptr_disconnected_node_rel h3)" using \acyclic (parent_child_rel h \ local.a_host_shadow_root_rel h \ local.a_ptr_disconnected_node_rel h)\ \parent_child_rel h \ local.a_host_shadow_root_rel h \ local.a_ptr_disconnected_node_rel h = parent_child_rel h2 \ local.a_host_shadow_root_rel h2 \ local.a_ptr_disconnected_node_rel h2\ \parent_child_rel h2 \ local.a_host_shadow_root_rel h2 \ local.a_ptr_disconnected_node_rel h2 = parent_child_rel h3 \ local.a_host_shadow_root_rel h3 \ local.a_ptr_disconnected_node_rel h3\ by auto ultimately have "acyclic (parent_child_rel h' \ a_host_shadow_root_rel h' \ a_ptr_disconnected_node_rel h')" by(simp add: \parent_child_rel h' \ a_host_shadow_root_rel h' \ a_ptr_disconnected_node_rel h' = {(cast element_ptr, cast new_shadow_root_ptr)} \ parent_child_rel h3 \ a_host_shadow_root_rel h3 \ a_ptr_disconnected_node_rel h3\) have "CD.a_heap_is_wellformed h'" apply(simp add: CD.a_heap_is_wellformed_def) by (simp add: \local.CD.a_acyclic_heap h'\ \local.CD.a_all_ptrs_in_heap h'\ \local.CD.a_distinct_lists h'\ \local.CD.a_owner_document_valid h'\) show "heap_is_wellformed h' " using \acyclic (parent_child_rel h' \ local.a_host_shadow_root_rel h' \ local.a_ptr_disconnected_node_rel h')\ by(simp add: heap_is_wellformed_def CD.heap_is_wellformed_impl \local.CD.a_heap_is_wellformed h'\ \local.a_all_ptrs_in_heap h'\ \local.a_distinct_lists h'\ \local.a_shadow_root_valid h'\) qed end interpretation l_attach_shadow_root_wf?: l_attach_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr known_ptrs type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed parent_child_rel set_tag_name set_tag_name_locs set_disconnected_nodes set_disconnected_nodes_locs create_element get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs set_val set_val_locs create_character_data DocumentClass.known_ptr DocumentClass.type_wf set_shadow_root set_shadow_root_locs set_mode set_mode_locs attach_shadow_root by(auto simp add: l_attach_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_attach_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] end diff --git a/thys/Skip_Lists/Misc.thy b/thys/Skip_Lists/Misc.thy --- a/thys/Skip_Lists/Misc.thy +++ b/thys/Skip_Lists/Misc.thy @@ -1,192 +1,192 @@ (* File: Misc.thy Authors: Max W. Haslbeck, Manuel Eberl *) section \Auxiliary material\ theory Misc imports "HOL-Analysis.Analysis" begin text \Based on @{term sorted_list_of_set} and @{term the_inv_into} we construct a bijection between a finite set A of type 'a::linorder and a set of natural numbers @{term "{..< card A}"}\ lemma bij_betw_mono_on_the_inv_into: fixes A::"'a::linorder set" and B::"'b::linorder set" assumes b: "bij_betw f A B" and m: "mono_on f A" shows "mono_on (the_inv_into A f) B" proof (rule ccontr) assume "\ mono_on (the_inv_into A f) B" then have "\r s. r \ B \ s \ B \ r \ s \ \ the_inv_into A f s \ the_inv_into A f r" unfolding mono_on_def by blast then obtain r s where rs: "r \ B" "s \ B" "r \ s" "the_inv_into A f s < the_inv_into A f r" by fastforce have f: "f (the_inv_into A f b) = b" if "b \ B" for b using that assms f_the_inv_into_f_bij_betw by metis have "the_inv_into A f s \ A" "the_inv_into A f r \ A" using rs assms by (auto simp add: bij_betw_def the_inv_into_into) then have "f (the_inv_into A f s) \ f (the_inv_into A f r)" using rs by (intro mono_onD[OF m]) (auto) then have "r = s" using rs f by simp then show False using rs by auto qed lemma rev_removeAll_removeAll_rev: "rev (removeAll x xs) = removeAll x (rev xs)" by (simp add: removeAll_filter_not_eq rev_filter) lemma sorted_list_of_set_Min_Cons: assumes "finite A" "A \ {}" shows "sorted_list_of_set A = Min A # sorted_list_of_set (A - {Min A})" proof - have *: "A = insert (Min A) A" using assms Min_in by (auto) then have "sorted_list_of_set A = insort (Min A) (sorted_list_of_set (A - {Min A}))" - using assms by (subst *, intro sorted_list_of_set_insert) auto + using assms by (subst *, intro sorted_list_of_set_insert_remove) auto also have "\ = Min A # sorted_list_of_set (A - {Min A})" using assms by (intro insort_is_Cons) (auto) finally show ?thesis by simp qed lemma sorted_list_of_set_filter: assumes "finite A" shows "sorted_list_of_set ({x\A. P x}) = filter P (sorted_list_of_set A)" using assms proof (induction "sorted_list_of_set A" arbitrary: A) case (Cons x xs) have x: "x \ A" using Cons sorted_list_of_set list.set_intros(1) by metis have "sorted_list_of_set A = Min A # sorted_list_of_set (A - {Min A})" using Cons by (intro sorted_list_of_set_Min_Cons) auto then have 1: "x = Min A" "xs = sorted_list_of_set (A - {x})" using Cons by auto { assume Px: "P x" have 2: "sorted_list_of_set {x \ A. P x} = Min {x \ A. P x} # sorted_list_of_set ({x \ A. P x} - {Min {x \ A. P x}})" using Px Cons 1 sorted_list_of_set_eq_Nil_iff by (intro sorted_list_of_set_Min_Cons) fastforce+ also have 3: "Min {x \ A. P x} = x" using Cons 1 Px x by (auto intro!: Min_eqI) also have 4: "{x \ A. P x} - {x} = {y \ A - {x}. P y}" by blast also have 5: "sorted_list_of_set {y \ A - {x}. P y} = filter P (sorted_list_of_set (A - {x}))" using 1 Cons by (intro Cons) (auto) also have "\ = filter P xs" using 1 by simp also have "filter P (sorted_list_of_set A) = x # filter P xs" using Px by (simp flip: \x # xs = sorted_list_of_set A\) finally have ?case by auto } moreover { assume Px: "\ P x" then have "{x \ A. P x} = {y \ A - {x}. P y}" by blast also have "sorted_list_of_set \ = filter P (sorted_list_of_set (A - {x}))" using 1 Cons by (intro Cons) auto also have "filter P (sorted_list_of_set (A - {x})) = filter P (sorted_list_of_set A)" using 1 Px by (simp flip: \x # xs = sorted_list_of_set A\) finally have ?case by simp } ultimately show ?case by blast qed (use sorted_list_of_set_eq_Nil_iff in fastforce) lemma sorted_list_of_set_Max_snoc: assumes "finite A" "A \ {}" shows "sorted_list_of_set A = sorted_list_of_set (A - {Max A}) @ [Max A]" proof - have *: "A = insert (Max A) A" using assms Max_in by (auto) then have "sorted_list_of_set A = insort (Max A) (sorted_list_of_set (A - {Max A}))" - using assms by (subst *, intro sorted_list_of_set_insert) auto + using assms by (subst *, intro sorted_list_of_set_insert_remove) auto also have "\ = sorted_list_of_set (A - {Max A}) @ [Max A]" using assms by (intro sorted_insort_is_snoc) (auto) finally show ?thesis by simp qed lemma sorted_list_of_set_image: assumes "mono_on g A" "inj_on g A" shows "(sorted_list_of_set (g ` A)) = map g (sorted_list_of_set A)" proof (cases "finite A") case True then show ?thesis using assms proof (induction "sorted_list_of_set A" arbitrary: A) case Nil then show ?case using sorted_list_of_set_eq_Nil_iff by fastforce next case (Cons x xs A) have not_empty_A: "A \ {}" using Cons sorted_list_of_set_eq_Nil_iff by auto have *: "Min (g ` A) = g (Min A)" proof - have "g (Min A) \ g a" if "a \ A" for a using that Cons Min_in Min_le not_empty_A by (auto intro!: mono_onD[of g]) then show ?thesis using Cons not_empty_A by (intro Min_eqI) auto qed have "g ` A \ {}" "finite (g ` A)" using Cons by auto then have "(sorted_list_of_set (g ` A)) = Min (g ` A) # sorted_list_of_set ((g ` A) - {Min (g ` A)})" by (auto simp add: sorted_list_of_set_Min_Cons) also have "(g ` A) - {Min (g ` A)} = g ` (A - {Min A})" using Cons Min_in not_empty_A * by (subst inj_on_image_set_diff[of _ A]) auto also have "sorted_list_of_set (g ` (A - {Min A})) = map g (sorted_list_of_set (A - {Min A}))" using not_empty_A Cons mono_on_subset[of _ A "A - {Min A}"] inj_on_subset[of _ A "A - {Min A}"] by (intro Cons) (auto simp add: sorted_list_of_set_Min_Cons) finally show ?case using Cons not_empty_A * by (auto simp add: sorted_list_of_set_Min_Cons) qed next case False then show ?thesis using assms by (simp add: finite_image_iff) qed lemma sorted_list_of_set_length: "length (sorted_list_of_set A) = card A" using distinct_card sorted_list_of_set[of A] by (cases "finite A") fastforce+ lemma sorted_list_of_set_bij_betw: assumes "finite A" shows "bij_betw (\n. sorted_list_of_set A ! n) {..n. xs ! n) {.. mono_on (\n. sorted_list_of_set A ! n) {.. 'a \ nat" where "bij_mono_map_set_to_nat A = (\x. if x \ A then the_inv_into {..The Cauchy--Binet formula\ theory Cauchy_Binet imports Diagonal_To_Smith SNF_Missing_Lemmas begin subsection \Previous missing results about @{text "pick"} and @{text "insert"}\ lemma pick_insert: assumes a_notin_I: "a \ I" and i2: "i < card I" and a_def: "pick (insert a I) a' = a" (*Alternative: index (insort a (sorted_list_of_set I)) a = a'*) and ia': "i < a'" (*Case 1*) and a'_card: "a' < card I + 1" shows "pick (insert a I) i = pick I i" proof - have finI: "finite I" using i2 using card.infinite by force have "pick (insert a I) i = sorted_list_of_set (insert a I) ! i" proof (rule sorted_list_of_set_eq_pick[symmetric]) have "finite (insert a I)" using card.infinite i2 by force thus "i < length (sorted_list_of_set (insert a I))" by (metis a_notin_I card_insert_disjoint distinct_card finite_insert i2 less_Suc_eq sorted_list_of_set(1) sorted_list_of_set(3)) qed also have "... = insort a (sorted_list_of_set I) ! i" - using sorted_list_of_set.insert - by (metis a_notin_I card.infinite i2 not_less0) + by (simp add: a_notin_I finI) also have "... = (sorted_list_of_set I) ! i" proof (rule insort_nth[OF]) show "sorted (sorted_list_of_set I)" by auto show "a \ set (sorted_list_of_set I)" using a_notin_I by (metis card.infinite i2 not_less_zero set_sorted_list_of_set) have "index (sorted_list_of_set (insert a I)) a = a'" using pick_index a_def using a'_card a_notin_I finI by auto hence "index (insort a (sorted_list_of_set I)) a = a'" by (simp add: a_notin_I finI) thus "i < index (insort a (sorted_list_of_set I)) a" using ia' by auto show "sorted_list_of_set I \ []" using finI i2 by fastforce qed also have "... = pick I i" proof (rule sorted_list_of_set_eq_pick) have "finite I" using card.infinite i2 by fastforce thus "i < length (sorted_list_of_set I)" by (metis distinct_card distinct_sorted_list_of_set i2 set_sorted_list_of_set) qed finally show ?thesis . qed lemma pick_insert2: assumes a_notin_I: "a \ I" and i2: "i < card I" and a_def: "pick (insert a I) a' = a" (*Alternative: index (sorted_list_of_set (insert a I)) a = a'*) and ia': "i \ a'" (*Case 2*) and a'_card: "a' < card I + 1" shows "pick (insert a I) i < pick I i" proof (cases "i = 0") case True then show ?thesis by (auto, metis (mono_tags, lifting) DL_Missing_Sublist.pick.simps(1) Least_le a_def a_notin_I dual_order.order_iff_strict i2 ia' insertCI le_zero_eq not_less_Least pick_in_set_le) next case False hence i0: "i = Suc (i - 1)" using a'_card ia' by auto have finI: "finite I" using i2 card.infinite by force have index_a'1: "index (sorted_list_of_set (insert a I)) a = a'" using pick_index using a'_card a_def a_notin_I finI by auto hence index_a': "index (insort a (sorted_list_of_set I)) a = a'" by (simp add: a_notin_I finI) have i1_length: "i - 1 < length (sorted_list_of_set I)" using i2 by (metis distinct_card distinct_sorted_list_of_set finI less_imp_diff_less set_sorted_list_of_set) have 1: "pick (insert a I) i = sorted_list_of_set (insert a I) ! i" proof (rule sorted_list_of_set_eq_pick[symmetric]) have "finite (insert a I)" using card.infinite i2 by force thus "i < length (sorted_list_of_set (insert a I))" by (metis a_notin_I card_insert_disjoint distinct_card finite_insert i2 less_Suc_eq sorted_list_of_set(1) sorted_list_of_set(3)) qed also have 2: "... = insort a (sorted_list_of_set I) ! i" - using sorted_list_of_set.insert - by (metis a_notin_I card.infinite i2 not_less0) + by (simp add: a_notin_I finI) also have "... = insort a (sorted_list_of_set I) ! Suc (i-1)" using i0 by auto also have "... < pick I i" proof (cases "i = a'") case True have "(sorted_list_of_set I) ! i > a" by (smt "1" Suc_less_eq True a_def a_notin_I distinct_card distinct_sorted_list_of_set finI i2 ia' index_a' insort_nth2 length_insort lessI list.size(3) nat_less_le not_less_zero - pick_in_set_le set_sorted_list_of_set sorted_list_of_set(2) sorted_list_of_set.insert + pick_in_set_le set_sorted_list_of_set sorted_list_of_set(2) + sorted_list_of_set_insert sorted_list_of_set_eq_pick sorted_wrt_nth_less) moreover have "a = insort a (sorted_list_of_set I) ! i" using True 1 2 a_def by auto ultimately show ?thesis using 1 2 by (metis distinct_card finI i0 i2 set_sorted_list_of_set sorted_list_of_set(3) sorted_list_of_set_eq_pick) next case False have "insort a (sorted_list_of_set I) ! Suc (i-1) = (sorted_list_of_set I) ! (i-1)" by (rule insort_nth2, insert i1_length False ia' index_a', auto simp add: a_notin_I finI) also have "... = pick I (i-1)" by (rule sorted_list_of_set_eq_pick[OF i1_length]) also have "... < pick I i" using i0 i2 pick_mono_le by auto finally show ?thesis . qed finally show ?thesis . qed lemma pick_insert3: assumes a_notin_I: "a \ I" and i2: "i < card I" and a_def: "pick (insert a I) a' = a" (*Alternative: index (sorted_list_of_set (insert a I)) a = a'.*) and ia': "i \ a'" (*Case 2*) and a'_card: "a' < card I + 1" shows "pick (insert a I) (Suc i) = pick I i" proof (cases "i = 0") case True have a_LEAST: "a < (LEAST aa. aa\I)" using True a_def a_notin_I i2 ia' pick_insert2 by fastforce have Least_rw: "(LEAST aa. aa = a \ aa \ I) = a" by (rule Least_equality, insert a_notin_I, auto, metis a_LEAST le_less_trans nat_le_linear not_less_Least) let ?P = "\aa. (aa = a \ aa \ I) \ (LEAST aa. aa = a \ aa \ I) < aa" let ?Q = "\aa. aa \ I" have "?P = ?Q" unfolding Least_rw fun_eq_iff by (auto, metis a_LEAST le_less_trans not_le not_less_Least) thus ?thesis using True by auto next case False have finI: "finite I" using i2 card.infinite by force have index_a'1: "index (sorted_list_of_set (insert a I)) a = a'" using pick_index using a'_card a_def a_notin_I finI by auto hence index_a': "index (insort a (sorted_list_of_set I)) a = a'" by (simp add: a_notin_I finI) have i1_length: "i < length (sorted_list_of_set I)" using i2 by (metis distinct_card distinct_sorted_list_of_set finI set_sorted_list_of_set) have 1: "pick (insert a I) (Suc i) = sorted_list_of_set (insert a I) ! (Suc i)" proof (rule sorted_list_of_set_eq_pick[symmetric]) have "finite (insert a I)" using card.infinite i2 by force thus "Suc i < length (sorted_list_of_set (insert a I))" by (metis Suc_mono a_notin_I card_insert_disjoint distinct_card distinct_sorted_list_of_set finI i2 set_sorted_list_of_set) qed also have 2: "... = insort a (sorted_list_of_set I) ! Suc i" - using sorted_list_of_set.insert - by (metis a_notin_I card.infinite i2 not_less0) + by (simp add: a_notin_I finI) also have "... = pick I i" proof (cases "i = a'") case True show ?thesis by (metis True a_notin_I finI i1_length index_a' insort_nth2 le_refl list.size(3) not_less0 set_sorted_list_of_set sorted_list_of_set(2) sorted_list_of_set_eq_pick) next case False have "insort a (sorted_list_of_set I) ! Suc i = (sorted_list_of_set I) ! i" by (rule insort_nth2, insert i1_length False ia' index_a', auto simp add: a_notin_I finI) also have "... = pick I i" by (rule sorted_list_of_set_eq_pick[OF i1_length]) finally show ?thesis . qed finally show ?thesis . qed lemma pick_insert_index: assumes Ik: "card I = k" and a_notin_I: "a \ I" and ik: "i < k" and a_def: "pick (insert a I) a' = a" and a'k: "a' < card I + 1" shows "pick (insert a I) (insert_index a' i) = pick I i" proof (cases "i I" by (simp add: Ik ik pick_in_set_le) show "pick (insert a I) i < pick I i" by (rule pick_insert2[OF a_notin_I _ a_def _ a'k], insert False, auto simp add: Ik ik) fix y assume y: "y \ I \ pick (insert a I) i < y" let ?xs = "sorted_list_of_set (insert a I)" have "y \ set ?xs" using y by (metis fin_aI insertI2 set_sorted_list_of_set y) from this obtain j where xs_j_y: "?xs ! j = y" and j: "j < length ?xs" using in_set_conv_nth by metis have ij: "i pick (insert a I) j" by (metis Ik Suc_lessI card.infinite distinct_card distinct_sorted_list_of_set eq_iff finite_insert ij ik j less_imp_le_nat not_less_zero pick_mono_le set_sorted_list_of_set) also have "... = ?xs ! j" by (rule sorted_list_of_set_eq_pick[symmetric, OF j]) also have "... = y" by (rule xs_j_y) finally show "pick I i \ y" . qed finally show ?thesis unfolding insert_index_def using False by auto qed subsection\Start of the proof\ definition "strict_from_inj n f = (\i. if i\{0.. nat" assumes "inj_on f {0.. {0.. {0.. f ` {0.. f ` {0.. strict_from_inj n f ` {0..)|f \. f \ {0.. {0.. (\i. i \ {0.. f i = i) \ \ permutes {0.. {0.. {0.. (\i. i \ {0.. f i = i)} \ {\. \ permutes {0.. carrier_mat n m" and B: "B \ carrier_mat m n" shows "det (A*B) = det (mat\<^sub>r n n (\i. finsum_vec TYPE('a::comm_ring_1) n (\k. B $$ (k, i) \\<^sub>v Matrix.col A k) {0..T \ carrier_mat m n" using A by auto have BT: "B\<^sup>T \ carrier_mat n m" using B by auto let ?f = "(\i. finsum_vec TYPE('a) n (\k. B\<^sup>T $$ (i, k) \\<^sub>v Matrix.row A\<^sup>T k) {0..k\{0..T $$ (i, k) \\<^sub>v row A\<^sup>T k) $ j)" by (rule index_finsum_vec[OF _ j_n], auto simp add: A) also have "... = (\k\{0..\<^sub>v col A k) $ j)" proof (rule sum.cong, auto) fix x assume x: "xT x = col A x" by (rule row_transpose, insert A x, auto) have B_rw: "B\<^sup>T $$ (i,x) = B $$ (x, i)" by (rule index_transpose_mat, insert x i B, auto) have "(B\<^sup>T $$ (i, x) \\<^sub>v Matrix.row A\<^sup>T x) $v j = B\<^sup>T $$ (i, x) * Matrix.row A\<^sup>T x $v j" by (rule index_smult_vec, insert A j_n, auto) also have "... = B $$ (x, i) * col A x $v j" unfolding row_rw B_rw by simp also have "... = (B $$ (x, i) \\<^sub>v col A x) $v j" by (rule index_smult_vec[symmetric], insert A j_n, auto) finally show " (B\<^sup>T $$ (i, x) \\<^sub>v Matrix.row A\<^sup>T x) $v j = (B $$ (x, i) \\<^sub>v col A x) $v j" . qed also have "... = ?g i $v j" by (rule index_finsum_vec[symmetric, OF _ j_n], auto simp add: A) also have "... = ?rhs $$ (i, j)" by (rule index_mat[symmetric], insert i j, auto) finally show "?lhs $$ (i, j) = ?rhs $$ (i, j)" . qed have "det (A*B) = det (B\<^sup>T*A\<^sup>T)" using det_transpose by (metis A B Matrix.transpose_mult mult_carrier_mat) also have "... = det (mat\<^sub>r n n (\i. finsum_vec TYPE('a) n (\k. B\<^sup>T $$ (i, k) \\<^sub>v Matrix.row A\<^sup>T k) {0..r n n (\i. finsum_vec TYPE('a) n (\k. B $$ (k, i) \\<^sub>v Matrix.col A k) {0.. carrier_mat n m" and B: "B \ carrier_mat m n" shows "det (A*B) = (\f | (\i\{0.. {0.. (\i. i \ {0.. f i = i). (\i = 0..r n n (\i. col A (f i))))" proof - let ?V="{0..r n n (\i. B $$ (f i, i) \\<^sub>v col A (f i))) = (prod (\i. B $$ (f i, i)) {0..r n n (\i. col A (f i)))" if f: "(\i\{0.. {0.. (\i. i \ {0.. f i = i)" for f by (rule det_rows_mul, insert A col_dim, auto) have "det (A*B) = det (mat\<^sub>r n n (\i. finsum_vec TYPE('a::comm_ring_1) n (\k. B $$ (k, i) \\<^sub>v Matrix.col A k) ?U))" by (rule det_mul_finsum_alt[OF A B]) also have "... = sum ?g ?F" by (rule det_linear_rows_sum[OF fm], auto simp add: A) also have "... = (\f\?F. prod (\i. B $$ (f i, i)) {0..r n n (\i. col A (f i))))" using det_rw by auto finally show ?thesis . qed lemma det_cols_mul': assumes A: "A \ carrier_mat n m" and B: "B \ carrier_mat m n" shows "det (A*B) = (\f | (\i\{0.. {0.. (\i. i \ {0.. f i = i). (\i = 0..r n n (\i. row B (f i))))" proof - let ?F="{f. (\i\{0.. {0.. (\i. i \ {0.. f i = i)}" have t: "A * B = (B\<^sup>T*A\<^sup>T)\<^sup>T" using transpose_mult[OF A B] transpose_transpose by metis have "det (B\<^sup>T*A\<^sup>T) = (\f\?F. (\i = 0..T $$ (f i, i)) * det (mat\<^sub>r n n (\i. col B\<^sup>T (f i))))" by (rule det_cols_mul, auto simp add: A B) also have "... = (\f \?F. (\i = 0..r n n (\i. row B (f i))))" proof (rule sum.cong, rule refl) fix f assume f: "f \ ?F" have "(\i = 0..T $$ (f i, i)) = (\i = 0.. {0..T $$ (f x, x) = A $$ (x, f x)" by (rule index_transpose_mat(1), insert f A x, auto) qed moreover have "det (mat\<^sub>r n n (\i. col B\<^sup>T (f i))) = det (mat\<^sub>r n n (\i. row B (f i)))" proof - have row_eq_colT: "row B (f i) $v j = col B\<^sup>T (f i) $v j" if i: "i < n" and j: "j < n" for i j proof - have fi_m: "f i < m" using f i by auto have "col B\<^sup>T (f i) $v j = B\<^sup>T $$(j, f i)" by (rule index_col, insert B fi_m j, auto) also have "... = B $$ (f i, j)" using B fi_m j by auto also have "... = row B (f i) $v j" by (rule index_row[symmetric], insert B fi_m j, auto) finally show ?thesis .. qed show ?thesis by (rule arg_cong[of _ _ det], rule eq_matI, insert row_eq_colT, auto) qed ultimately show "(\i = 0..T $$ (f i, i)) * det (mat\<^sub>r n n (\i. col B\<^sup>T (f i))) = (\i = 0..r n n (\i. row B (f i)))" by simp qed finally show ?thesis by (metis (no_types, lifting) A B det_transpose transpose_mult mult_carrier_mat) qed (*We need a more general version of this lemma*) lemma assumes F: "F= {f. f \ {0.. {0.. (\i. i \ {0.. f i = i)}" and p: " \ permutes {0..f\F. (\i = 0.. i))) = (\f\F. (\i = 0.. \ = g \ \" have "f x = g x" for x proof (cases "x \ {0.. F" show "xa \ \ \ F" unfolding o_def F using F PiE p xa by (auto, smt F atLeastLessThan_iff mem_Collect_eq p permutes_def xa) show "\x\F. xa = x \ \" proof (rule bexI[of _ "xa \ Hilbert_Choice.inv \"]) show "xa = xa \ Hilbert_Choice.inv \ \ \" using p by auto show "xa \ Hilbert_Choice.inv \ \ F" unfolding o_def F using F PiE p xa by (auto, smt atLeastLessThan_iff permutes_def permutes_less(3)) qed qed have prod_rw: "(\i = 0..i = 0.. i), \ i))" if "f\F" for f using prod.permute[OF p] by auto let ?g = "\f. (\i = 0.. i))" have "(\f\F. (\i = 0..f\F. (\i = 0.. i), \ i)))" using prod_rw by auto also have "... = (\f\(?h`F). \i = 0.. i))" using sum.reindex[OF inj_on_F, of ?g] unfolding hF by auto also have "... = (\f\F. \i = 0.. i))" unfolding hF by auto finally show ?thesis .. qed lemma detAB_Znm_aux: assumes F: "F= {f. f \ {0.. {0.. (\i. i \ {0.. f i = i)}" shows"(\\ | \ permutes {0..f\F. prod (\i. B $$ (f i, i)) {0.. * (\i = 0.. i, f i))))) = (\\ | \ permutes {0..f\F. (\i = 0.. i)) * (signof \ * (\i = 0..\ | \ permutes {0..f\F. prod (\i. B $$ (f i, i)) {0.. * (\i = 0.. i, f i))))) = (\\ | \ permutes {0..f\F. signof \ * (\i = 0.. i, f i)))" by (smt mult.left_commute prod.cong prod.distrib sum.cong) also have "... = (\\ | \ permutes {0..f\F. signof (Hilbert_Choice.inv \) * (\i = 0.. i, f i)))" by (rule sum_permutations_inverse) also have "... = (\\ | \ permutes {0..f\F. signof (Hilbert_Choice.inv \) * (\i = 0.. i), (\ i)) * A $$ (Hilbert_Choice.inv \ (\ i), f (\ i))))" proof (rule sum.cong) fix x assume x: "x \ {\. \ permutes {0..i = 0..i = 0.. F" for f using prod.permute[OF p] by auto then show "(\f\F. signof ?inv_x * (\i = 0..f\F. signof ?inv_x * (\i = 0..\ | \ permutes {0..f\F. signof \ * (\i = 0.. i), (\ i)) * A $$ (i, f (\ i))))" by (rule sum.cong, auto, rule sum.cong, auto) (metis (no_types, lifting) finite_atLeastLessThan signof_inv) also have "... = (\\ | \ permutes {0..f\F. signof \ * (\i = 0.. i)) * A $$ (i, f i)))" proof (rule sum.cong) fix \ assume p: "\ \ {\. \ permutes {0.. permutes {0.. ?inv_pi = g \ ?inv_pi" have "f x = g x" for x proof (cases "x \ {0.. F" show "xa \ ?inv_pi \ F" unfolding o_def F using F PiE p xa by (auto, smt atLeastLessThan_iff permutes_def permutes_less(3)) show "\x\F. xa = x \ ?inv_pi" proof (rule bexI[of _ "xa \ \"]) show "xa = xa \ \ \ Hilbert_Choice.inv \ " using p by auto show "xa \ \ \ F" unfolding o_def F using F PiE p xa by (auto, smt atLeastLessThan_iff permutes_def permutes_less(3)) qed qed let ?g = "\f. signof \ * (\i = 0.. i), \ i) * A $$ (i, f (\ i)))" show "(\f\F. signof \ * (\i = 0.. i), \ i) * A $$ (i, f (\ i)))) = (\f\F. signof \ * (\i = 0.. i) * A $$ (i, f i)))" using sum.reindex[OF inj_on_F, of "?g"] p unfolding hF unfolding o_def by auto qed (simp) also have "... = (\\ | \ permutes {0..f\F. (\i = 0.. i)) * (signof \ * (\i = 0.. carrier_mat n m" and B: "B \ carrier_mat m n" shows "det (A*B) = (\(f, \)\Z n m. signof \ * (\i = 0.. i)))" proof - let ?V="{0.. {0.. {0.. (\i. i \ {0.. f i = i)}" by auto have det_rw: "det (mat\<^sub>r n n (\i. B $$ (f i, i) \\<^sub>v col A (f i))) = (prod (\i. B $$ (f i, i)) {0..r n n (\i. col A (f i)))" if f: "(\i\{0.. {0.. (\i. i \ {0.. f i = i)" for f by (rule det_rows_mul, insert A col_dim, auto) have det_rw2: "det (mat\<^sub>r n n (\i. col A (f i))) = (\\ | \ permutes {0.. * (\i = 0.. i, f i)))" if f: "f \ ?F" for f proof (unfold Determinant.det_def, auto, rule sum.cong, auto) fix x assume x: "x permutes {0..i = 0..i = 0.. {0..i = 0..i = 0..r n n (\i. finsum_vec TYPE('a::comm_ring_1) n (\k. B $$ (k, i) \\<^sub>v Matrix.col A k) {0..f\?F. prod (\i. B $$ (f i, i)) {0..r n n (\i. col A (f i))))" using det_rw by auto also have "... = (\f\?F. prod (\i. B $$ (f i, i)) {0..\ | \ permutes {0.. * (\i = 0.. i, f (i)))))" by (rule sum.cong, auto simp add: det_rw2) also have "... = (\f\?F. \\ | \ permutes {0..i. B $$ (f i, i)) {0.. * (\i = 0.. i, f (i)))))" by (simp add: mult_hom.hom_sum) also have "... = (\\ | \ permutes {0..f\?F.prod (\i. B $$ (f i, i)) {0.. * (\i = 0.. i, f i))))" by (rule VS_Connect.class_semiring.finsum_finsum_swap, insert finite_permutations finite_bounded_functions[OF fin_m fin_n], auto) thm detAB_Znm_aux also have "... = (\\ | \ permutes {0..f\?F. (\i = 0.. i)) * (signof \ * (\i = 0..f\?F.\\ | \ permutes {0..i = 0.. i)) * (signof \ * (\i = 0..f\?F.\\ | \ permutes {0.. * (\i = 0.. i)))" unfolding prod.distrib by (rule sum.cong, auto, rule sum.cong, auto) also have "... = sum (\(f,\). (signof \) * (prod (\i. A$$(i,f i) * B $$ (f i, \ i)) {0.. carrier_mat n m" and B: "B \ carrier_mat m n" begin private definition "Z_inj = ({f. f \ {0.. {0.. (\i. i \ {0.. f i = i) \ inj_on f {0.. {\. \ permutes {0.. {0.. {0.. (\i. i \ {0.. f i = i) \ \ inj_on f {0.. {\. \ permutes {0.. {0.. {0.. (\i. i \ {0.. f i = i) \ strict_mono_on f {0.. {\. \ permutes {0.. {0.. {0.. (\i. i \ {0.. f i = i) \ \ strict_mono_on f {0.. {\. \ permutes {0.. = (signof \) * (prod (\i. A$$(i,f i) * B $$ (f i, \ i)) {0.. {0.. {0.. (\i. i \ {0.. f i = i) \ inj_on f {0.. (f`{0.. {\. \ permutes {0.. {0.. {0.. (\i. i \ {0.. f i = i) \ strict_mono_on f {0.. {0.. {0.. (\i. i \ {0.. f i = i) \ inj_on f {0.. {0.. {0.. (\i. i \ {0.. f i = i) \ \ inj_on f {0.. {0.. {0.. (\i. i \ {0.. f i = i)}" text\The Cauchy--Binet formula is proven in \url{https://core.ac.uk/download/pdf/82475020.pdf} In that work, they define @{text "\ \ inv \ \ \"}. I had problems following this proof in Isabelle, since I was demanded to show that such permutations commute, which is false. It is a notation problem of the @{text "\"} operator, the author means @{text "\ \ \ \ inv \"} using the Isabelle notation (i.e., @{text "\ x = \ ((inv \) x)"}). \ lemma step_weight: fixes \ \ defines "\ \ \ \ Hilbert_Choice.inv \" assumes f_inj: "f \ F_inj" and gF: "g \ F" and pi: "\ permutes {0.. permutes {0..x \ {0.. x)" shows "weight f \ = (signof \) * (\i = 0.. i))) * (signof \) * (\i = 0.. i))" proof - let ?A = "(\i = 0.. i))) " let ?B = "(\i = 0.. i))" have sigma: "\ permutes {0.._def by (rule permutes_compose[OF permutes_inv[OF phi] pi]) have A_rw: "?A = (\i = 0..i = 0.. i), \ (\ i)))" by (rule prod.permute[unfolded o_def, OF phi]) also have "... = (\i = 0.. i))" using fg_phi unfolding \_def unfolding o_def unfolding permutes_inverses(2)[OF phi] by auto finally have B_rw: "?B = (\i = 0.. i))" . have "(signof \) * ?A * (signof \) * ?B = (signof \) * (signof \) * ?A * ?B" by auto also have "... = signof (\ \ \) * ?A * ?B" unfolding signof_compose[OF phi sigma] by simp also have "... = signof \ * ?A * ?B" by (metis (no_types, lifting) \_def mult.commute o_inv_o_cancel permutes_inj phi sigma signof_compose) also have "... = signof \ * (\i = 0..i = 0.. i))" using A_rw B_rw by auto also have "... = signof \ * (\i = 0.. i))" by auto also have "... = weight f \" unfolding weight_def by simp finally show ?thesis .. qed lemma Z_good_fun_alt_sum: fixes g defines "Z_good_fun \ {f. f \ {0.. {0.. (\i. i \ {0.. f i = i) \ inj_on f {0.. (f`{0.. F_inj" shows "(\f\Z_good_fun. P f)= (\\\{\. \ permutes {0.. \))" proof - let ?f = "\\. g \ \" let ?P = "{\. \ permutes {0.. i < n" hence "xa i = i" unfolding permutes_def by auto thus "g (xa i) = i" using g i_ge_n unfolding F_inj_def by auto next fix xa assume "xa permutes {0.. xa) {0.. xb assume "\ permutes {0.. (\x. g (\ x)) ` {0.. {0.. {0..i. \ i < n \ x i = i" and inj_on_x: "inj_on x {0.. = "\i. if i x i = g j) else i" show "x \ (\) g ` {\. \ permutes {0..], rule conjI) have "?\ i = i" if i: "i \ {0..!j. ?\ j = i" for i proof (cases "i x a = g j) = i" proof (rule theI2) show "i < n \ x a = g i" using xa_gi True by auto fix xa assume "xa < n \ x a = g xa" thus "xa = i" by (metis (mono_tags, lifting) F_inj_def True atLeast0LessThan g inj_onD lessThan_iff mem_Collect_eq xa_gi) thus "xa = i" . qed thus ta: "?\ a = i" using a by auto fix j assume tj: "?\ j = i" show "j = a" proof (cases "j x j = g ja) = i" using tj True by auto have "?P (THE ja. ?P ja)" proof (rule theI) show "b < n \ x j = g b" using xj_gb b by auto fix xa assume "xa < n \ x j = g xa" thus "xa = b" by (metis (mono_tags, lifting) F_inj_def b atLeast0LessThan g inj_onD lessThan_iff mem_Collect_eq xj_gb) qed hence "x j = g i" unfolding the_ji by auto hence "x j = x a" using xa_gi by auto then show ?thesis using inj_on_x a True unfolding inj_on_def by auto next case False then show ?thesis using tj True by auto qed qed next case False note i_ge_n = False show ?thesis proof (rule ex1I[of _ i]) show "?\ i = i" using False by simp fix j assume tj: "?\ j = i" show "j = i" proof (cases "j x j = g ja) < n" proof (rule theI2) show "a < n \ x j = g a" using xj_ga a by auto fix xa assume a1: "xa < n \ x j = g xa" thus "xa = a" using F_inj_def a atLeast0LessThan g inj_on_eq_iff xj_ga by fastforce show "xa < n" by (simp add: a1) qed then show ?thesis using tj i_ge_n by auto next case False then show ?thesis using tj by auto qed qed qed ultimately show "?\ permutes {0.. ?\" proof - have "x xa = g (THE j. j < n \ x xa = g j)" if xa: "xa < n" for xa proof - obtain c where c: "c < n" and xxa_gc: "x xa = g c" by (metis (mono_tags, hide_lams) atLeast0LessThan imageE imageI lessThan_iff xa xg) show ?thesis proof (rule theI2) show c1: "c < n \ x xa = g c" using c xxa_gc by auto fix xb assume c2: "xb < n \ x xa = g xb" thus "xb = c" by (metis (mono_tags, lifting) F_inj_def c1 atLeast0LessThan g inj_onD lessThan_iff mem_Collect_eq) show "x xa = g xb" using c1 c2 by simp qed qed moreover have "x xa = g xa" if xa: "\ xa < n" for xa using g x1 x2 xa unfolding F_inj_def by simp ultimately show ?thesis unfolding o_def fun_eq_iff by auto qed qed qed have inj: "inj_on ?f ?P" proof (rule inj_onI) fix x y assume x: "x \ ?P" and y: "y \ ?P" and gx_gy: "g \ x = g \ y" have "x i = y i" for i proof (cases "i {0.. {0..f\Z_good_fun. P f) = (\f\?f`?P. P f)" using fP by simp also have "... = sum (P \ (\) g) {\. \ permutes {0..\ | \ permutes {0.. \))" by auto finally show ?thesis . qed lemma F_injI: assumes "f \ {0.. {0..i. i \ {0.. f i = i)" and "inj_on f {0.. F_inj" using assms unfolding F_inj_def by simp lemma F_inj_composition_permutation: assumes phi: "\ permutes {0.. F_inj" shows "g \ \ \ F_inj" proof (rule F_injI) show "g \ \ \ {0.. {0..i. i \ {0.. (g \ \) i = i" using g phi unfolding permutes_def F_inj_def by simp show "inj_on (g \ \) {0.. F_strict" shows "f \ F_inj" using f strict_mono_on_imp_inj_on unfolding F_strict_def F_inj_def by auto lemma one_step: assumes g1: "g \ F_strict" shows "det (submatrix A UNIV (g`{0..(x, y) \ Z_good g. weight x y)" (is "?lhs = ?rhs") proof - define Z_good_fun where "Z_good_fun = {f. f \ {0.. {0.. (\i. i \ {0.. f i = i) \ inj_on f {0.. (f`{0.. F_inj" by (rule F_strict_imp_F_inj[OF g1]) have detA: "(\\\{\. \ permutes {0.. * (\i = 0.. i)))) = det (submatrix A UNIV (g`{0.. j \ g ` {0.. g ` {0.. j \ g ` {0.. carrier_mat n n" unfolding submatrix_def card_J using A by auto have "det (submatrix A UNIV (g`{0..p | p permutes {0..i = 0..\\{\. \ permutes {0.. * (\i = 0.. i))))" proof (rule sum.cong) fix x assume x: "x \ {\. \ permutes {0..i = 0..i = 0.. {0.. (g ` {0..i = 0..i = 0..\ \ ?Perm. signof (\ \ ?inv \) * (\i = 0.. \ ?inv \) i))) = (\\ \ ?Perm. signof (\) * (\i = 0.. i)))" if phi: "\ permutes {0.. proof - let ?h="\\. \ \ ?inv \" let ?g = "\\. signof (\) * (\i = 0.. i))" have "?h`?Perm = ?Perm" proof - have "\ \ ?inv \ permutes {0.. permutes {0.. using permutes_compose permutes_inv phi that by blast moreover have "x \ (\\. \ \ ?inv \) ` ?Perm" if "x permutes {0.. \ permutes {0.. \ \ ?inv \" using phi by auto ultimately show ?thesis unfolding image_def by auto qed ultimately show ?thesis by auto qed hence "(\\ \ ?Perm. ?g \) = (\\ \ ?h`?Perm. ?g \)" by simp also have "... = sum (?g \ ?h) ?Perm" proof (rule sum.reindex) show "inj_on (\\. \ \ ?inv \) {\. \ permutes {0..\ \ ?Perm. signof (\ \ ?inv \) * (\i = 0.. \ ?inv \) i)))" unfolding o_def by auto finally show ?thesis by simp qed have detB: "det (submatrix B (g`{0..\ \ ?Perm. signof \ * (\i = 0.. i)))" proof - have "{i. i < dim_row B \ i \ g ` {0.. g ` {0.. j \ g ` {0.. carrier_mat n n" unfolding submatrix_def using card_I B by auto have "det (submatrix B (g`{0..p \ ?Perm. signof p * (\i=0..\ \ ?Perm. signof \ * (\i = 0.. i)))" proof (rule sum.cong, rule refl) fix x assume x: "x \ {\. \ permutes {0..i=0..i=0.. {0.. (g ` {0..i = 0..i = 0..f\Z_good_fun. \\\?Perm. weight f \)" unfolding Z_good_def sum.cartesian_product Z_good_fun_def by blast also have "... = (\\\{\. \ permutes {0.. \))" unfolding Z_good_fun_def by (rule Z_good_fun_alt_sum[OF g]) also have "... = (\\\{\. \ permutes {0..\\{\. \ permutes {0.. * (\i = 0.. i))) * signof (\ \ ?inv \) * (\i = 0.. \ ?inv \) i)))" proof (rule sum.cong, simp, rule sum.cong, simp) fix \ \ assume phi: "\ \ ?Perm" and pi: "\ \ ?Perm" show "weight (g \ \) \ = signof \ * (\i = 0.. i))) * signof (\ \ ?inv \) * (\i = 0.. \ ?inv \) i))" proof (rule step_weight) show "g \ \ \ F_inj" by (rule F_inj_composition_permutation[OF _ g], insert phi, auto) show "g \ F" using g unfolding F_def F_inj_def by simp qed (insert phi pi, auto) qed also have "... = (\\\{\. \ permutes {0.. * (\i = 0.. i))) * (\\ | \ permutes {0.. \ ?inv \) * (\i = 0.. \ ?inv \) i))))" by (metis (mono_tags, lifting) Groups.mult_ac(1) semiring_0_class.sum_distrib_left sum.cong) also have "... = (\\ \ ?Perm. signof \ * (\i = 0.. i))) * (\\ \ ?Perm. signof \ * (\i = 0.. i))))" using detB_rw by auto also have "... = (\\ \ ?Perm. signof \ * (\i = 0.. i)))) * (\\ \ ?Perm. signof \ * (\i = 0.. i)))" by (simp add: semiring_0_class.sum_distrib_right) also have "... = ?lhs" unfolding detA detB .. finally show ?thesis .. qed lemma gather_by_strictness: "sum (\g. sum (\(f,\). weight f \) (Z_good g)) F_strict = sum (\g. det (submatrix A UNIV (g`{0.. F_strict" show "(\(x, y)\Z_good f. weight x y) = det (submatrix A UNIV (f ` {0..i\{0.. {0.. (\i. i \ {0.. f i = i)} = ?B" by auto have "?A\?B" by auto moreover have "finite ?B" using B finite_bounded_functions[OF finM finN] by auto ultimately show "finite ?A" using rev_finite_subset by blast show "finite {\. \ permutes {0..i\{0.. {0.. (\i. i \ {0.. f i = i)} = ?B" by auto have "?A\?B" by auto moreover have "finite ?B" using B finite_bounded_functions[OF finM finN] by auto ultimately show "finite ?A" using rev_finite_subset by blast show "finite {\. \ permutes {0..i\{0.. {0.. (\i. i \ {0.. f i = i)} = ?B" by auto have "?A\?B" by auto moreover have "finite ?B" using B finite_bounded_functions[OF finM finN] by auto ultimately show "finite ?A" using rev_finite_subset by blast show "finite {\. \ permutes {0..i\{0.. {0.. (\i. i \ {0.. f i = i)} = ?B" by auto have "?A\?B" by auto moreover have "finite ?B" using B finite_bounded_functions[OF finM finN] by auto ultimately show "finite F_inj" unfolding F_inj_def using rev_finite_subset by blast qed lemma finite_F_strict[simp]: "finite F_strict" proof - have finN: "finite {0..i\{0.. {0.. (\i. i \ {0.. f i = i)} = ?B" by auto have "?A\?B" by auto moreover have "finite ?B" using B finite_bounded_functions[OF finM finN] by auto ultimately show "finite F_strict" unfolding F_strict_def using rev_finite_subset by blast qed lemma nth_strict_mono: fixes f::"nat \ nat" assumes strictf: "strict_mono f" and i: "i ?I. a < f i} = i" using i proof (induct i) case 0 then show ?case by (auto simp add: strict_mono_less strictf) next case (Suc i) have i: "i < n" using Suc.prems by auto let ?J'="{a \ f ` {0.. f i" and 2: "f xa < f (Suc i)" show "f xa < f i" using 1 2 not_less_less_Suc_eq strict_mono_less strictf by fastforce next fix xa assume "f xa < f i" thus "f xa < f (Suc i)" using less_SucI strict_mono_less strictf by blast next show "f i \ f ` {0..card (f ` {0.. i) also have "... = pick ?I (card {a \ ?I. a < f i})" unfolding card_eq by simp also have "... = f i" by (rule pick_card_in_set, simp add: i) finally show ?thesis .. qed lemma nth_strict_mono_on: fixes f::"nat \ nat" assumes strictf: "strict_mono_on f {0.. ?I. a < f i} = i" using i proof (induct i) case 0 then show ?case by (auto, metis (no_types, lifting) atLeast0LessThan lessThan_iff less_Suc_eq not_less0 not_less_eq strict_mono_on_def strictf) next case (Suc i) have i: "i < n" using Suc.prems by auto let ?J'="{a \ f ` {0.. f i" and 2: "f xa < f (Suc i)" and 3: "xa < n" show "f xa < f i" by (metis (full_types) 1 2 3 antisym_conv3 atLeast0LessThan i lessThan_iff less_SucE order.asym strict_mono_onD strictf) next fix xa assume "f xa < f i" and "xa < n" thus "f xa < f (Suc i)" using less_SucI strictf by (metis (no_types, lifting) Suc.prems atLeast0LessThan lessI lessThan_iff less_trans strict_mono_onD) next show "f i \ f ` {0..card (f ` {0.. i) also have "... = pick ?I (card {a \ ?I. a < f i})" unfolding card_eq by simp also have "... = f i" by (rule pick_card_in_set, simp add: i) finally show ?thesis .. qed lemma strict_fun_eq: assumes f: "f \ F_strict" and g: "g \ F_strict" and fg: "f`{0.. F_inj" shows "strict_from_inj n f \ F" proof - { fix x assume x: "x < n" have inj_on: "inj_on f {0.. a \ f ` {0.. a \ f ` {0.. F_strict" if xa: "xa \ F_inj" for xa proof - have "strict_mono_on (strict_from_inj n xa) {0.. F_inj" shows "strict_from_inj n f ` {0.. a \ f ` {0.. a \ f ` {0.. f ` {0..card (f ` {0.. xa) finally show "strict_from_inj n f xa \ f ` {0.. strict_from_inj n f ` {0.. F_strict" shows "Z_good g = {x \ F_inj. strict_from_inj n x = g} \ {\. \ permutes {0.. {0.. {0.. (\i. i \ {0.. f i = i) \ inj_on f {0.. (f`{0.. F_inj. strict_from_inj n x = g}" proof (auto) fix f assume f: "f \ Z_good_fun" thus f_inj: "f \ F_inj" unfolding F_inj_def Z_good_fun_def by auto show "strict_from_inj n f = g" proof (rule strict_fun_eq[OF _ g]) show "strict_from_inj n f ` {0.. F_strict" using F_strict_def f_inj strict_from_inj_F_strict by blast qed next fix f assume f_inj: "f \ F_inj" and g_strict_f: "g = strict_from_inj n f" have "f xa \ g ` {0.. f ` {0.. Z_good_fun" using f_inj g_strict_f unfolding Z_good_fun_def F_inj_def by auto qed thus ?thesis unfolding Z_good_fun_def Z_good_def by simp qed lemma weight_0: "(\(f, \) \ Z_not_inj. weight f \) = 0" proof - let ?F="{f. (\i\{0.. {0.. (\i. i \ {0.. f i = i)}" let ?Perm = "{\. \ permutes {0..(f, \)\Z_not_inj. weight f \) = (\f \ F_not_inj. (\i = 0..r n n (\i. row B (f i))))" proof - have dim_row_rw: "dim_row (mat\<^sub>r n n (\i. col A (f i))) = n" for f by auto have dim_row_rw2: "dim_row (mat\<^sub>r n n (\i. Matrix.row B (f i))) = n" for f by auto have prod_rw: "(\i = 0.. i)) = (\i = 0.. i)" if f: "f \ F_not_inj" and pi: "\ \ ?Perm" for f \ proof (rule prod.cong, rule refl) fix x assume x: "x \ {0.. x < dim_col B" using x pi B by auto ultimately show "B $$ (f x, \ x) = Matrix.row B (f x) $v \ x" by (rule index_row[symmetric]) qed have sum_rw: "(\\ | \ permutes {0.. * (\i = 0.. i))) = det (mat\<^sub>r n n (\i. row B (f i)))" if f: "f \ F_not_inj" for f unfolding Determinant.det_def using dim_row_rw2 prod_rw f by auto have "(\(f, \)\Z_not_inj. weight f \) = (\f\F_not_inj.\\ \ ?Perm. weight f \)" unfolding Z_not_inj_def unfolding sum.cartesian_product unfolding F_not_inj_def by simp also have "... = (\f\F_not_inj. \\ | \ permutes {0.. * (\i = 0.. i)))" unfolding weight_def by simp also have "... = (\f\F_not_inj. (\i = 0..\ | \ permutes {0.. * (\i = 0.. i))))" by (rule sum.cong, rule refl, auto) (metis (no_types, lifting) mult.left_commute mult_hom.hom_sum sum.cong) also have "... = (\f \ F_not_inj. (\i = 0..r n n (\i. row B (f i))))" using sum_rw by auto finally show ?thesis by auto qed also have "... = 0" by (rule sum.neutral, insert det_not_inj_on[of _ n B], auto simp add: F_not_inj_def) finally show ?thesis . qed subsection \Final theorem\ lemma Cauchy_Binet1: shows "det (A*B) = sum (\f. det (submatrix A UNIV (f`{0..(f, \) \ Z_not_inj. weight f \) = 0" by (rule weight_0) let ?f = "strict_from_inj n" have sum_rw: "sum g F_inj = (\y \ F_strict. sum g {x \ F_inj. ?f x = y})" for g by (rule sum.group[symmetric], insert strict_from_inj_F_strict, auto) have Z_Union: "Z_inj \ Z_not_inj = Z n m" unfolding Z_def Z_not_inj_def Z_inj_def by auto have Z_Inter: "Z_inj \ Z_not_inj = {}" unfolding Z_def Z_not_inj_def Z_inj_def by auto have "det (A*B) = (\(f, \)\Z n m. weight f \)" using detAB_Znm[OF A B] unfolding weight_def by auto also have "... = (\(f, \)\Z_inj. weight f \) + (\(f, \)\Z_not_inj. weight f \)" by (metis Z_Inter Z_Union finite_Un finite_Znm sum.union_disjoint) also have "... = (\(f, \)\Z_inj. weight f \)" using sum0 by force also have "... = (\f \ F_inj. \\\{\. \ permutes {0..)" unfolding Z_inj_def unfolding F_inj_def sum.cartesian_product .. also have "... = (\y\F_strict. \f\{x \ F_inj. strict_from_inj n x = y}. sum (weight f) {\. \ permutes {0..y\F_strict. \(f,\)\({x \ F_inj. strict_from_inj n x = y} \ {\. \ permutes {0..)" unfolding F_inj_def sum.cartesian_product .. also have "... = sum (\g. sum (\(f,\). weight f \) (Z_good g)) F_strict" using Z_good_alt by auto also have "... = ?rhs" unfolding gather_by_strictness by simp finally show ?thesis . qed lemma Cauchy_Binet: "det (A*B) = (\I\{I. I\{0.. card I=n}. det (submatrix A UNIV I) * det (submatrix B I UNIV))" proof - let ?f="(\I. (\i. if i ?setI" and J: "J \ ?setI" and fI_fJ: "?f I = ?f J" have "x \ J" if x: "x \ I" for x by (metis (mono_tags) fI_fJ I J distinct_card in_set_conv_nth mem_Collect_eq sorted_list_of_set(1) sorted_list_of_set(3) subset_eq_atLeast0_lessThan_finite x) moreover have "x \ I" if x: "x \ J" for x by (metis (mono_tags) fI_fJ I J distinct_card in_set_conv_nth mem_Collect_eq sorted_list_of_set(1) sorted_list_of_set(3) subset_eq_atLeast0_lessThan_finite x) ultimately show "I = J" by auto qed have rw: "?f I ` {0.. ?setI" for I proof - have "sorted_list_of_set I ! xa \ I" if "xa < n" for xa by (metis (mono_tags, lifting) I distinct_card distinct_sorted_list_of_set mem_Collect_eq nth_mem set_sorted_list_of_set subset_eq_atLeast0_lessThan_finite that) moreover have "\xa\{0..I" for x by (metis (full_types) x I atLeast0LessThan distinct_card in_set_conv_nth mem_Collect_eq lessThan_iff sorted_list_of_set(1) sorted_list_of_set(3) subset_eq_atLeast0_lessThan_finite) ultimately show ?thesis unfolding image_def by auto qed have f_setI: "?f` ?setI = F_strict" proof - have "sorted_list_of_set I ! xa < m" if I: "I \ {0..xa < card I\ atLeast0LessThan distinct_card finite_atLeastLessThan lessThan_iff pick_in_set_le rev_finite_subset sorted_list_of_set(1) sorted_list_of_set(3) sorted_list_of_set_eq_pick subsetCE) moreover have "strict_mono_on (\i. if i < card I then sorted_list_of_set I ! i else i) {0.. {0..I \ {0.. atLeastLessThan_iff distinct_card finite_atLeastLessThan pick_mono_le rev_finite_subset sorted_list_of_set(1) sorted_list_of_set(3) sorted_list_of_set_eq_pick strict_mono_on_def) moreover have "x \ ?f ` {I. I \ {0.. card I = n}" if x1: "x \ {0.. {0..i. \ i < n \ x i = i" and s: "strict_mono_on x {0..i. if i < n then sorted_list_of_set (x ` {0..f. det (submatrix A UNIV (f ` {0.. ?f) {I. I \ {0.. card I = n}" unfolding Cauchy_Binet1 f_setI[symmetric] by (rule sum.reindex[OF inj_on]) also have "... = (\I\{I. I\{0.. card I=n}.det(submatrix A UNIV I)*det(submatrix B I UNIV))" by (rule sum.cong, insert rw, auto) finally show ?thesis . qed end end diff --git a/thys/Smith_Normal_Form/SNF_Missing_Lemmas.thy b/thys/Smith_Normal_Form/SNF_Missing_Lemmas.thy --- a/thys/Smith_Normal_Form/SNF_Missing_Lemmas.thy +++ b/thys/Smith_Normal_Form/SNF_Missing_Lemmas.thy @@ -1,1168 +1,1168 @@ (* Author: Jose Divasón Email: jose.divason@unirioja.es *) section \Missing results\ theory SNF_Missing_Lemmas imports Hermite.Hermite Mod_Type_Connect Jordan_Normal_Form.DL_Rank_Submatrix "List-Index.List_Index" begin text \This theory presents some missing lemmas that are required for the Smith normal form development. Some of them could be added to different AFP entries, such as the Jordan Normal Form AFP entry by Ren\'e Thiemann and Akihisa Yamada. However, not all the lemmas can be added directly, since some imports are required.\ hide_const (open) C hide_const (open) measure subsection \Miscellaneous lemmas\ lemma sum_two_rw: "(\i = 0..<2. f i) = (\i \ {0,1::nat}. f i)" by (rule sum.cong, auto) lemma sum_common_left: fixes f::"'a \ 'b::comm_ring_1" assumes "finite A" shows "sum (\i. c * f i) A = c * sum f A" by (simp add: mult_hom.hom_sum) lemma prod3_intro: assumes "fst A = a" and "fst (snd A) = b" and "snd (snd A) = c" shows "A = (a,b,c)" using assms by auto subsection \Transfer rules for the HMA\_Connect file of the Perron-Frobenius development\ hide_const (open) HMA_M HMA_I to_hma\<^sub>m from_hma\<^sub>m hide_fact (open) from_hma\<^sub>m_def from_hma_to_hma\<^sub>m HMA_M_def HMA_I_def dim_row_transfer_rule dim_col_transfer_rule context includes lifting_syntax begin lemma HMA_invertible_matrix[transfer_rule]: "((HMA_Connect.HMA_M :: _ \ 'a :: comm_ring_1 ^ 'n ^ 'n \ _) ===> (=)) invertible_mat invertible" proof (intro rel_funI, goal_cases) case (1 x y) note rel_xy[transfer_rule] = "1" have eq_dim: "dim_col x = dim_row x" using HMA_Connect.dim_col_transfer_rule HMA_Connect.dim_row_transfer_rule rel_xy by fastforce moreover have "\A'. y ** A' = Finite_Cartesian_Product.mat 1 \ A' ** y = Finite_Cartesian_Product.mat 1" if xB: "x * B = 1\<^sub>m (dim_row x)" and Bx: "B * x = 1\<^sub>m (dim_row B)" for B proof - let ?A' = "HMA_Connect.to_hma\<^sub>m B:: 'a :: comm_ring_1 ^ 'n ^ 'n" have rel_BA[transfer_rule]: "HMA_M B ?A'" by (metis (no_types, lifting) Bx HMA_M_def eq_dim carrier_mat_triv dim_col_mat(1) from_hma\<^sub>m_def from_hma_to_hma\<^sub>m index_mult_mat(3) index_one_mat(3) rel_xy xB) have [simp]: "dim_row B = CARD('n)" using dim_row_transfer_rule rel_BA by blast have [simp]: "dim_row x = CARD('n)" using dim_row_transfer_rule rel_xy by blast have "y ** ?A' = Finite_Cartesian_Product.mat 1" using xB by (transfer, simp) moreover have "?A' ** y = Finite_Cartesian_Product.mat 1" using Bx by (transfer, simp) ultimately show ?thesis by blast qed moreover have "\B. x * B = 1\<^sub>m (dim_row x) \ B * x = 1\<^sub>m (dim_row B)" if yA: "y ** A' = Finite_Cartesian_Product.mat 1" and Ay: "A' ** y = Finite_Cartesian_Product.mat 1" for A' proof - let ?B = "(from_hma\<^sub>m A')" have [simp]: "dim_row x = CARD('n)" using dim_row_transfer_rule rel_xy by blast have [transfer_rule]: "HMA_M ?B A'" by (simp add: HMA_M_def) hence [simp]: "dim_row ?B = CARD('n)" using dim_row_transfer_rule by auto have "x * ?B = 1\<^sub>m (dim_row x)" using yA by (transfer', auto) moreover have "?B * x = 1\<^sub>m (dim_row ?B)" using Ay by (transfer', auto) ultimately show ?thesis by auto qed ultimately show ?case unfolding invertible_mat_def invertible_def inverts_mat_def by auto qed end subsection \Lemmas obtained from HOL Analysis using local type definitions\ thm Cartesian_Space.invertible_mult (*In HOL Analysis*) thm invertible_iff_is_unit (*In HOL Analysis*) thm det_non_zero_imp_unit (*In JNF, but only for fields*) thm mat_mult_left_right_inverse (*In JNF, but only for fields*) lemma invertible_mat_zero: assumes A: "A \ carrier_mat 0 0" shows "invertible_mat A" using A unfolding invertible_mat_def inverts_mat_def one_mat_def times_mat_def scalar_prod_def Matrix.row_def col_def carrier_mat_def by (auto, metis (no_types, lifting) cong_mat not_less_zero) lemma invertible_mult_JNF: fixes A::"'a::comm_ring_1 mat" assumes A: "A\carrier_mat n n" and B: "B\carrier_mat n n" and inv_A: "invertible_mat A" and inv_B: "invertible_mat B" shows "invertible_mat (A*B)" proof (cases "n = 0") case True then show ?thesis using assms by (simp add: invertible_mat_zero) next case False then show ?thesis using invertible_mult[where ?'a="'a::comm_ring_1", where ?'b="'n::finite", where ?'c="'n::finite", where ?'d="'n::finite", untransferred, cancel_card_constraint, OF assms] by auto qed lemma invertible_iff_is_unit_JNF: assumes A: "A \ carrier_mat n n" shows "invertible_mat A \ (Determinant.det A) dvd 1" proof (cases "n=0") case True then show ?thesis using det_dim_zero invertible_mat_zero A by auto next case False then show ?thesis using invertible_iff_is_unit[untransferred, cancel_card_constraint] A by auto qed subsection \Lemmas about matrices, submatrices and determinants\ (*This is a generalization of thm mat_mult_left_right_inverse*) thm mat_mult_left_right_inverse lemma mat_mult_left_right_inverse: fixes A :: "'a::comm_ring_1 mat" assumes A: "A \ carrier_mat n n" and B: "B \ carrier_mat n n" and AB: "A * B = 1\<^sub>m n" shows "B * A = 1\<^sub>m n" proof - have "Determinant.det (A * B) = Determinant.det (1\<^sub>m n)" using AB by auto hence "Determinant.det A * Determinant.det B = 1" using Determinant.det_mult[OF A B] det_one by auto hence det_A: "(Determinant.det A) dvd 1" and det_B: "(Determinant.det B) dvd 1" using dvd_triv_left dvd_triv_right by metis+ hence inv_A: "invertible_mat A" and inv_B: "invertible_mat B" using A B invertible_iff_is_unit_JNF by blast+ obtain B' where inv_BB': "inverts_mat B B'" and inv_B'B: "inverts_mat B' B" using inv_B unfolding invertible_mat_def by auto have B'_carrier: "B' \ carrier_mat n n" by (metis B inv_B'B inv_BB' carrier_matD(1) carrier_matD(2) carrier_mat_triv index_mult_mat(3) index_one_mat(3) inverts_mat_def) have "B * A * B = B" using A AB B by auto hence "B * A * (B * B') = B * B'" by (smt A AB B B'_carrier assoc_mult_mat carrier_matD(1) inv_BB' inverts_mat_def one_carrier_mat) thus ?thesis by (metis A B carrier_matD(1) carrier_matD(2) index_mult_mat(3) inv_BB' inverts_mat_def right_mult_one_mat') qed context comm_ring_1 begin lemma col_submatrix_UNIV: assumes "j < card {i. i < dim_col A \ i \ J}" shows "col (submatrix A UNIV J) j = col A (pick J j)" proof (rule eq_vecI) show dim_eq:"dim_vec (col (submatrix A UNIV J) j) = dim_vec (col A (pick J j))" by (simp add: dim_submatrix(1)) fix i assume "i < dim_vec (col A (pick J j))" show "col (submatrix A UNIV J) j $v i = col A (pick J j) $v i" by (smt Collect_cong assms col_def dim_col dim_eq dim_submatrix(1) eq_vecI index_vec pick_UNIV submatrix_index) qed lemma submatrix_split2: "submatrix A I J = submatrix (submatrix A I UNIV) UNIV J" (is "?lhs = ?rhs") proof (rule eq_matI) show dr: "dim_row ?lhs = dim_row ?rhs" by (simp add: dim_submatrix(1)) show dc: "dim_col ?lhs = dim_col ?rhs" by (simp add: dim_submatrix(2)) fix i j assume i: "i < dim_row ?rhs" and j: "j < dim_col ?rhs" have "?rhs $$ (i, j) = (submatrix A I UNIV) $$ (pick UNIV i, pick J j)" proof (rule submatrix_index) show "i < card {i. i < dim_row (submatrix A I UNIV) \ i \ UNIV}" by (metis (full_types) dim_submatrix(1) i) show "j < card {j. j < dim_col (submatrix A I UNIV) \ j \ J}" by (metis (full_types) dim_submatrix(2) j) qed also have "... = A $$ (pick I (pick UNIV i), pick UNIV (pick J j))" proof (rule submatrix_index) show "pick UNIV i < card {i. i < dim_row A \ i \ I}" by (metis (full_types) dr dim_submatrix(1) i pick_UNIV) show "pick J j < card {j. j < dim_col A \ j \ UNIV}" by (metis (full_types) dim_submatrix(2) j pick_le) qed also have "... = ?lhs $$ (i,j)" proof (unfold pick_UNIV, rule submatrix_index[symmetric]) show "i < card {i. i < dim_row A \ i \ I}" by (metis (full_types) dim_submatrix(1) dr i) show "j < card {j. j < dim_col A \ j \ J}" by (metis (full_types) dim_submatrix(2) dc j) qed finally show "?lhs $$ (i, j) = ?rhs $$ (i, j)" .. qed lemma submatrix_mult: "submatrix (A*B) I J = submatrix A I UNIV * submatrix B UNIV J" (is "?lhs = ?rhs") proof (rule eq_matI) show "dim_row ?lhs = dim_row ?rhs" unfolding submatrix_def by auto show "dim_col ?lhs = dim_col ?rhs" unfolding submatrix_def by auto fix i j assume i: "i < dim_row ?rhs" and j: "j < dim_col ?rhs" have i1: "i < card {i. i < dim_row (A * B) \ i \ I}" by (metis (full_types) dim_submatrix(1) i index_mult_mat(2)) have j1: "j < card {j. j < dim_col (A * B) \ j \ J}" by (metis dim_submatrix(2) index_mult_mat(3) j) have pi: "pick I i < dim_row A" using i1 pick_le by auto have pj: "pick J j < dim_col B" using j1 pick_le by auto have row_rw: "Matrix.row (submatrix A I UNIV) i = Matrix.row A (pick I i)" using i1 row_submatrix_UNIV by auto have col_rw: "col (submatrix B UNIV J) j = col B (pick J j)" using j1 col_submatrix_UNIV by auto have "?lhs $$ (i,j) = (A*B) $$ (pick I i, pick J j)" by (rule submatrix_index[OF i1 j1]) also have "... = Matrix.row A (pick I i) \ col B (pick J j)" by (rule index_mult_mat(1)[OF pi pj]) also have "... = Matrix.row (submatrix A I UNIV) i \ col (submatrix B UNIV J) j" using row_rw col_rw by simp also have "... = (?rhs) $$ (i,j)" by (rule index_mult_mat[symmetric], insert i j, auto) finally show "?lhs $$ (i, j) = ?rhs $$ (i, j)" . qed lemma det_singleton: assumes A: "A \ carrier_mat 1 1" shows "det A = A $$ (0,0)" using A unfolding carrier_mat_def Determinant.det_def by auto lemma submatrix_singleton_index: assumes A: "A \ carrier_mat n m" and an: "a < n" and bm: "b < m" shows "submatrix A {a} {b} $$ (0,0) = A $$ (a,b)" proof - have a: "{i. i = a \ i < dim_row A} = {a}" using an A unfolding carrier_mat_def by auto have b: "{i. i = b \ i < dim_col A} = {b}" using bm A unfolding carrier_mat_def by auto have "submatrix A {a} {b} $$ (0,0) = A $$ (pick {a} 0,pick {b} 0)" by (rule submatrix_index, insert a b, auto) moreover have "pick {a} 0 = a" by (auto, metis (full_types) LeastI) moreover have "pick {b} 0 = b" by (auto, metis (full_types) LeastI) ultimately show ?thesis by simp qed end lemma det_not_inj_on: assumes not_inj_on: "\ inj_on f {0..r n n (\i. Matrix.row B (f i))) = 0" proof - obtain i j where i: "ij" using not_inj_on unfolding inj_on_def by auto show ?thesis proof (rule det_identical_rows[OF _ ij i j]) let ?B="(mat\<^sub>r n n (\i. row B (f i)))" show "row ?B i = row ?B j" proof (rule eq_vecI, auto) fix ia assume ia: "ia < n" have "row ?B i $ ia = ?B $$ (i, ia)" by (rule index_row(1), insert i ia, auto) also have "... = ?B $$ (j, ia)" by (simp add: fi_fj i ia j) also have "... = row ?B j $ ia" by (rule index_row(1)[symmetric], insert j ia, auto) finally show "row ?B i $ ia = row (mat\<^sub>r n n (\i. row B (f i))) j $ ia" by simp qed show "mat\<^sub>r n n (\i. Matrix.row B (f i)) \ carrier_mat n n" by auto qed qed lemma mat_row_transpose: "(mat\<^sub>r nr nc f)\<^sup>T = mat nc nr (\(i,j). vec_index (f j) i)" by (rule eq_matI, auto) lemma obtain_inverse_matrix: assumes A: "A \ carrier_mat n n" and i: "invertible_mat A" obtains B where "inverts_mat A B" and "inverts_mat B A" and "B \ carrier_mat n n" proof - have "(\B. inverts_mat A B \ inverts_mat B A)" using i unfolding invertible_mat_def by auto from this obtain B where AB: "inverts_mat A B" and BA: "inverts_mat B A" by auto moreover have "B \ carrier_mat n n" using A AB BA unfolding carrier_mat_def inverts_mat_def by (auto, metis index_mult_mat(3) index_one_mat(3))+ ultimately show ?thesis using that by blast qed lemma invertible_mat_smult_mat: fixes A :: "'a::comm_ring_1 mat" assumes inv_A: "invertible_mat A" and k: "k dvd 1" shows "invertible_mat (k \\<^sub>m A)" proof - obtain n where A: "A \ carrier_mat n n" using inv_A unfolding invertible_mat_def by auto have det_dvd_1: "Determinant.det A dvd 1" using inv_A invertible_iff_is_unit_JNF[OF A] by auto have "Determinant.det (k \\<^sub>m A) = k ^ dim_col A * Determinant.det A" by simp also have "... dvd 1" by (rule unit_prod, insert k det_dvd_1 dvd_power_same, force+) finally show ?thesis using invertible_iff_is_unit_JNF by (metis A smult_carrier_mat) qed lemma invertible_mat_one[simp]: "invertible_mat (1\<^sub>m n)" unfolding invertible_mat_def using inverts_mat_def by fastforce lemma four_block_mat_dim0: assumes A: "A \ carrier_mat n n" and B: "B \ carrier_mat n 0" and C: "C \ carrier_mat 0 n" and D: "D \ carrier_mat 0 0" shows "four_block_mat A B C D = A" unfolding four_block_mat_def using assms by auto lemma det_four_block_mat_lower_right_id: assumes A: "A \ carrier_mat m m" and B: "B = 0\<^sub>m m (n-m)" and C: "C = 0\<^sub>m (n-m) m" and D: "D = 1\<^sub>m (n-m)" and "n>m" shows "Determinant.det (four_block_mat A B C D) = Determinant.det A" using assms proof (induct n arbitrary: A B C D) case 0 then show ?case by auto next case (Suc n) let ?block = "(four_block_mat A B C D)" let ?B = "Matrix.mat m (n-m) (\(i,j). 0)" let ?C = "Matrix.mat (n-m) m (\(i,j). 0)" let ?D = "1\<^sub>m (n-m)" have mat_eq: "(mat_delete ?block n n) = four_block_mat A ?B ?C ?D" (is "?lhs = ?rhs") proof (rule eq_matI) fix i j assume i: "i < dim_row (four_block_mat A ?B ?C ?D)" and j: "j < dim_col (four_block_mat A ?B ?C ?D)" let ?f = " (if i < dim_row A then if j < dim_col A then A $$ (i, j) else B $$ (i, j - dim_col A) else if j < dim_col A then C $$ (i - dim_row A, j) else D $$ (i - dim_row A, j - dim_col A))" let ?g = "(if i < dim_row A then if j < dim_col A then A $$ (i, j) else ?B $$ (i, j - dim_col A) else if j < dim_col A then ?C $$ (i - dim_row A, j) else ?D $$ (i - dim_row A, j - dim_col A))" have "(mat_delete ?block n n) $$ (i,j) = ?block $$ (i,j)" using i j Suc.prems unfolding mat_delete_def by auto also have "... = ?f" by (rule index_mat_four_block, insert Suc.prems i j, auto) also have "... = ?g" using i j Suc.prems by auto also have "... = four_block_mat A ?B ?C ?D $$ (i,j)" by (rule index_mat_four_block[symmetric], insert Suc.prems i j, auto) finally show "?lhs $$ (i,j) = ?rhs $$ (i,j)" . qed (insert Suc.prems, auto) have nn_1: "?block $$ (n, n) = 1" using Suc.prems by auto have rw0: "(\i {..ii carrier_mat 1 n" and B: "B \ carrier_mat m n" and m0: "m \ 0" and r: "Matrix.row A 0 = Matrix.row B 0" shows "Matrix.row (A * V) 0 = Matrix.row (B * V) 0" proof (rule eq_vecI) show "dim_vec (Matrix.row (A * V) 0) = dim_vec (Matrix.row (B * V) 0)" using A B r by auto fix i assume i: "i < dim_vec (Matrix.row (B * V) 0)" have "Matrix.row (A * V) 0 $v i = (A * V) $$ (0,i)" by (rule index_row, insert i A, auto) also have "... = Matrix.row A 0 \ col V i" by (rule index_mult_mat, insert A i, auto) also have "... = Matrix.row B 0 \ col V i" using r by auto also have "... = (B * V) $$ (0,i)" by (rule index_mult_mat[symmetric], insert m0 B i, auto) also have "... = Matrix.row (B * V) 0 $v i" by (rule index_row[symmetric], insert i B m0, auto) finally show "Matrix.row (A * V) 0 $v i = Matrix.row (B * V) 0 $v i" . qed lemma smult_mat_mat_one_element: assumes A: "A \ carrier_mat 1 1" and B: "B \ carrier_mat 1 n" shows "A * B = A $$ (0,0) \\<^sub>m B" proof (rule eq_matI) fix i j assume i: "i < dim_row (A $$ (0, 0) \\<^sub>m B)" and j: "j < dim_col (A $$ (0, 0) \\<^sub>m B)" have i0: "i = 0" using A B i by auto have "(A * B) $$ (i, j) = Matrix.row A i \ col B j" by (rule index_mult_mat, insert i j A B, auto) also have "... = Matrix.row A i $v 0 * col B j $v 0" unfolding scalar_prod_def using B by auto also have "... = A$$(i,i) * B$$(i,j)" using A i i0 j by auto also have "... = (A $$ (i, i) \\<^sub>m B) $$ (i, j)" unfolding i by (rule index_smult_mat[symmetric], insert i j B, auto) finally show "(A * B) $$ (i, j) = (A $$ (0, 0) \\<^sub>m B) $$ (i, j)" using i0 by simp qed (insert A B, auto) lemma determinant_one_element: assumes A: "A \ carrier_mat 1 1" shows "Determinant.det A = A $$ (0,0)" proof - have "Determinant.det A = prod_list (diag_mat A)" by (rule det_upper_triangular[OF _ A], insert A, unfold upper_triangular_def, auto) also have "... = A $$ (0,0)" using A unfolding diag_mat_def by auto finally show ?thesis . qed lemma invertible_mat_transpose: assumes inv_A: "invertible_mat (A::'a::comm_ring_1 mat)" shows "invertible_mat A\<^sup>T" proof - obtain n where A: "A \ carrier_mat n n" using inv_A unfolding invertible_mat_def square_mat.simps by auto hence At: "A\<^sup>T \ carrier_mat n n" by simp have "Determinant.det A\<^sup>T = Determinant.det A" by (metis Determinant.det_def Determinant.det_transpose carrier_matI index_transpose_mat(2) index_transpose_mat(3)) also have "... dvd 1" using invertible_iff_is_unit_JNF[OF A] inv_A by simp finally show ?thesis using invertible_iff_is_unit_JNF[OF At] by auto qed lemma dvd_elements_mult_matrix_left: assumes A: "(A::'a::comm_ring_1 mat) \ carrier_mat m n" and P: "P \ carrier_mat m m" and x: "(\i j. i j x dvd A$$(i,j))" shows "(\i j. i j x dvd (P*A)$$(i,j))" proof - have "x dvd (P * A) $$ (i, j)" if i: "i < m" and j: "j < n" for i j proof - have "(P * A) $$ (i, j) = (\ia = 0..ia = 0.. carrier_mat m n" and Q: "Q \ carrier_mat n n" and x: "(\i j. i j x dvd A$$(i,j))" shows "(\i j. i j x dvd (A*Q)$$(i,j))" proof - have "x dvd (A*Q) $$ (i, j)" if i: "i < m" and j: "j < n" for i j proof - have "(A*Q) $$ (i, j) = (\ia = 0..ia = 0.. carrier_mat m n" and P: "P \ carrier_mat m m" and Q: "Q \ carrier_mat n n" and x: "(\i j. i j x dvd A$$(i,j))" shows "(\i j. i j x dvd (P*A*Q)$$(i,j))" using dvd_elements_mult_matrix_left[OF A P x] by (meson P A Q dvd_elements_mult_matrix_right mult_carrier_mat) definition append_cols :: "'a :: zero mat \ 'a mat \ 'a mat" (infixr "@\<^sub>c" 65)where "A @\<^sub>c B = four_block_mat A B (0\<^sub>m 0 (dim_col A)) (0\<^sub>m 0 (dim_col B))" lemma append_cols_carrier[simp,intro]: "A \ carrier_mat n a \ B \ carrier_mat n b \ (A @\<^sub>c B) \ carrier_mat n (a+b)" unfolding append_cols_def by auto lemma append_cols_mult_left: assumes A: "A \ carrier_mat n a" and B: "B \ carrier_mat n b" and P: "P \ carrier_mat n n" shows "P * (A @\<^sub>c B) = (P*A) @\<^sub>c (P*B)" proof - let ?P = "four_block_mat P (0\<^sub>m n 0) (0\<^sub>m 0 n) (0\<^sub>m 0 0)" have "P = ?P" by (rule eq_matI, auto) hence "P * (A @\<^sub>c B) = ?P * (A @\<^sub>c B)" by simp also have "?P * (A @\<^sub>c B) = four_block_mat (P * A + 0\<^sub>m n 0 * 0\<^sub>m 0 (dim_col A)) (P * B + 0\<^sub>m n 0 * 0\<^sub>m 0 (dim_col B)) (0\<^sub>m 0 n * A + 0\<^sub>m 0 0 * 0\<^sub>m 0 (dim_col A)) (0\<^sub>m 0 n * B + 0\<^sub>m 0 0 * 0\<^sub>m 0 (dim_col B))" unfolding append_cols_def by (rule mult_four_block_mat, insert A B P, auto) also have "... = four_block_mat (P * A) (P * B) (0\<^sub>m 0 (dim_col (P*A))) (0\<^sub>m 0 (dim_col (P*B)))" by (rule cong_four_block_mat, insert P, auto) also have "... = (P*A) @\<^sub>c (P*B)" unfolding append_cols_def by auto finally show ?thesis . qed lemma append_cols_mult_right_id: assumes A: "(A::'a::semiring_1 mat) \ carrier_mat n 1" and B: "B \ carrier_mat n (m-1)" and C: "C = four_block_mat (1\<^sub>m 1) (0\<^sub>m 1 (m - 1)) (0\<^sub>m (m - 1) 1) D" and D: "D \ carrier_mat (m-1) (m-1)" shows "(A @\<^sub>c B) * C = A @\<^sub>c (B * D)" proof - let ?C = "four_block_mat (1\<^sub>m 1) (0\<^sub>m 1 (m - 1)) (0\<^sub>m (m - 1) 1) D" have "(A @\<^sub>c B) * C = (A @\<^sub>c B) * ?C" unfolding C by auto also have "... = four_block_mat A B (0\<^sub>m 0 (dim_col A)) (0\<^sub>m 0 (dim_col B)) * ?C" unfolding append_cols_def by auto also have "... = four_block_mat (A * 1\<^sub>m 1 + B * 0\<^sub>m (m - 1) 1) (A * 0\<^sub>m 1 (m - 1) + B * D) (0\<^sub>m 0 (dim_col A) * 1\<^sub>m 1 + 0\<^sub>m 0 (dim_col B) * 0\<^sub>m (m - 1) 1) (0\<^sub>m 0 (dim_col A) * 0\<^sub>m 1 (m - 1) + 0\<^sub>m 0 (dim_col B) * D)" by (rule mult_four_block_mat, insert assms, auto) also have "... = four_block_mat A (B * D) (0\<^sub>m 0 (dim_col A)) (0\<^sub>m 0 (dim_col (B*D)))" by (rule cong_four_block_mat, insert assms, auto) also have "... = A @\<^sub>c (B * D)" unfolding append_cols_def by auto finally show ?thesis . qed lemma append_cols_mult_right_id2: assumes A: "(A::'a::semiring_1 mat) \ carrier_mat n a" and B: "B \ carrier_mat n b" and C: "C = four_block_mat D (0\<^sub>m a b) (0\<^sub>m b a) (1\<^sub>m b)" and D: "D \ carrier_mat a a" shows "(A @\<^sub>c B) * C = (A * D) @\<^sub>c B" proof - let ?C = "four_block_mat D (0\<^sub>m a b) (0\<^sub>m b a) (1\<^sub>m b)" have "(A @\<^sub>c B) * C = (A @\<^sub>c B) * ?C" unfolding C by auto also have "... = four_block_mat A B (0\<^sub>m 0 a) (0\<^sub>m 0 b) * ?C" unfolding append_cols_def using A B by auto also have "... = four_block_mat (A * D + B * 0\<^sub>m b a) (A * 0\<^sub>m a b + B * 1\<^sub>m b) (0\<^sub>m 0 a * D + 0\<^sub>m 0 b * 0\<^sub>m b a) (0\<^sub>m 0 a * 0\<^sub>m a b + 0\<^sub>m 0 b * 1\<^sub>m b)" by (rule mult_four_block_mat, insert A B C D, auto) also have "... = four_block_mat (A * D) B (0\<^sub>m 0 (dim_col (A*D))) (0\<^sub>m 0 (dim_col B))" by (rule cong_four_block_mat, insert assms, auto) also have "... = (A * D) @\<^sub>c B" unfolding append_cols_def by auto finally show ?thesis . qed lemma append_cols_nth: assumes A: "A \ carrier_mat n a" and B: "B \ carrier_mat n b" and i: "ic B) $$ (i, j) = (if j < dim_col A then A $$(i,j) else B$$(i,j-a))" (is "?lhs = ?rhs") proof - let ?C = "(0\<^sub>m 0 (dim_col A))" let ?D = "(0\<^sub>m 0 (dim_col B))" have i2: "i < dim_row A + dim_row ?D" using i A by auto have j2: "j < dim_col A + dim_col (0\<^sub>m 0 (dim_col B))" using j B A by auto have "(A @\<^sub>c B) $$ (i, j) = four_block_mat A B ?C ?D $$ (i, j)" unfolding append_cols_def by auto also have "... = (if i < dim_row A then if j < dim_col A then A $$ (i, j) else B $$ (i, j - dim_col A) else if j < dim_col A then ?C $$ (i - dim_row A, j) else 0\<^sub>m 0 (dim_col B) $$ (i - dim_row A, j - dim_col A))" by (rule index_mat_four_block(1)[OF i2 j2]) also have "... = ?rhs" using i A by auto finally show ?thesis . qed lemma append_cols_split: assumes d: "dim_col A > 0" shows "A = mat_of_cols (dim_row A) [col A 0] @\<^sub>c mat_of_cols (dim_row A) (map (col A) [1..c ?A2") proof (rule eq_matI) fix i j assume i: "i < dim_row (?A1 @\<^sub>c ?A2)" and j: "j < dim_col (?A1 @\<^sub>c ?A2)" have "(?A1 @\<^sub>c ?A2) $$ (i, j) = (if j < dim_col ?A1 then ?A1 $$(i,j) else ?A2$$(i,j-(dim_col ?A1)))" by (rule append_cols_nth, insert i j, auto simp add: append_cols_def) also have "... = A $$ (i,j)" proof (cases "j< dim_col ?A1") case True then show ?thesis by (metis One_nat_def Suc_eq_plus1 add.right_neutral append_cols_def col_def i index_mat_four_block(2) index_vec index_zero_mat(2) less_one list.size(3) list.size(4) mat_of_cols_Cons_index_0 mat_of_cols_carrier(2) mat_of_cols_carrier(3)) next case False then show ?thesis by (metis (no_types, lifting) Suc_eq_plus1 Suc_less_eq Suc_pred add_diff_cancel_right' append_cols_def diff_zero i index_col index_mat_four_block(2) index_mat_four_block(3) index_zero_mat(2) index_zero_mat(3) j length_map length_upt linordered_semidom_class.add_diff_inverse list.size(3) list.size(4) mat_of_cols_carrier(2) mat_of_cols_carrier(3) mat_of_cols_index nth_map_upt plus_1_eq_Suc upt_0) qed finally show "A $$ (i, j) = (?A1 @\<^sub>c ?A2) $$ (i, j)" .. qed (auto simp add: append_cols_def d) lemma append_rows_nth: assumes A: "A \ carrier_mat a n" and B: "B \ carrier_mat b n" and i: "ir B) $$ (i, j) = (if i < dim_row A then A $$(i,j) else B$$(i-a,j))" (is "?lhs = ?rhs") proof - let ?C = "(0\<^sub>m (dim_row A) 0)" let ?D = "(0\<^sub>m (dim_row B) 0)" have i2: "i < dim_row A + dim_row ?D" using i j A B by auto have j2: "j < dim_col A + dim_col ?D" using i j A B by auto have "(A @\<^sub>r B) $$ (i, j) = four_block_mat A ?C B ?D $$ (i, j)" unfolding append_rows_def by auto also have "... = (if i < dim_row A then if j < dim_col A then A $$ (i, j) else ?C $$ (i, j - dim_col A) else if j < dim_col A then B $$ (i - dim_row A, j) else ?D $$ (i - dim_row A, j - dim_col A))" by (rule index_mat_four_block(1)[OF i2 j2]) also have "... = ?rhs" using i A j B by auto finally show ?thesis . qed lemma append_rows_split: assumes k: "k\dim_row A" shows "A = (mat_of_rows (dim_col A) [Matrix.row A i. i \ [0..r (mat_of_rows (dim_col A) [Matrix.row A i. i \ [k..r ?A2") proof (rule eq_matI) have "(?A1 @\<^sub>r ?A2) \ carrier_mat (k + (dim_row A-k)) (dim_col A)" by (rule carrier_append_rows, insert k, auto) hence A1_A2: "(?A1 @\<^sub>r ?A2) \ carrier_mat (dim_row A) (dim_col A)" using k by simp thus "dim_row A = dim_row (?A1 @\<^sub>r ?A2)" and "dim_col A = dim_col (?A1 @\<^sub>r ?A2)" by auto fix i j assume i: "i < dim_row (?A1 @\<^sub>r ?A2)" and j: "j < dim_col (?A1 @\<^sub>r ?A2)" have "(?A1 @\<^sub>r ?A2) $$ (i, j) = (if i < dim_row ?A1 then ?A1 $$(i,j) else ?A2$$(i-(dim_row ?A1),j))" by (rule append_rows_nth, insert k i j, auto simp add: append_rows_def) also have "... = A $$ (i,j)" proof (cases "ir ?A2) $$ (i,j)" by simp qed lemma transpose_mat_append_rows: assumes A: "A \ carrier_mat a n" and B: "B \ carrier_mat b n" shows "(A @\<^sub>r B)\<^sup>T = A\<^sup>T @\<^sub>c B\<^sup>T" by (smt append_cols_def append_rows_def A B carrier_matD(1) index_transpose_mat(3) transpose_four_block_mat zero_carrier_mat zero_transpose_mat) lemma transpose_mat_append_cols: assumes A: "A \ carrier_mat n a" and B: "B \ carrier_mat n b" shows "(A @\<^sub>c B)\<^sup>T = A\<^sup>T @\<^sub>r B\<^sup>T" by (metis Matrix.transpose_transpose A B carrier_matD(1) carrier_mat_triv index_transpose_mat(3) transpose_mat_append_rows) lemma append_rows_mult_right: assumes A: "(A::'a::comm_semiring_1 mat) \ carrier_mat a n" and B: "B \ carrier_mat b n" and Q: "Q\ carrier_mat n n" shows "(A @\<^sub>r B) * Q = (A * Q) @\<^sub>r (B*Q)" proof - have "transpose_mat ((A @\<^sub>r B) * Q) = Q\<^sup>T * (A @\<^sub>r B)\<^sup>T" by (rule transpose_mult, insert A B Q, auto) also have "... = Q\<^sup>T * (A\<^sup>T @\<^sub>c B\<^sup>T)" using transpose_mat_append_rows assms by metis also have "... = Q\<^sup>T * A\<^sup>T @\<^sub>c Q\<^sup>T * B\<^sup>T" using append_cols_mult_left assms by (metis transpose_carrier_mat) also have "transpose_mat ... = (A * Q) @\<^sub>r (B*Q)" by (smt A B Matrix.transpose_mult Matrix.transpose_transpose append_cols_def append_rows_def Q carrier_mat_triv index_mult_mat(2) index_transpose_mat(2) transpose_four_block_mat zero_carrier_mat zero_transpose_mat) finally show ?thesis by simp qed lemma append_rows_mult_left_id: assumes A: "(A::'a::comm_semiring_1 mat) \ carrier_mat 1 n" and B: "B \ carrier_mat (m-1) n" and C: "C = four_block_mat (1\<^sub>m 1) (0\<^sub>m 1 (m - 1)) (0\<^sub>m (m - 1) 1) D" and D: "D \ carrier_mat (m-1) (m-1)" shows "C * (A @\<^sub>r B) = A @\<^sub>r (D * B)" proof - have "transpose_mat (C * (A @\<^sub>r B)) = (A @\<^sub>r B)\<^sup>T * C\<^sup>T" by (metis (no_types, lifting) B C D Matrix.transpose_mult append_rows_def A carrier_matD carrier_mat_triv index_mat_four_block(2,3) index_zero_mat(2) one_carrier_mat) also have "... = (A\<^sup>T @\<^sub>c B\<^sup>T) * C\<^sup>T" using transpose_mat_append_rows[OF A B] by auto also have "... = A\<^sup>T @\<^sub>c (B\<^sup>T * D\<^sup>T)" by (rule append_cols_mult_right_id, insert A B C D, auto) also have "transpose_mat ... = A @\<^sub>r (D * B)" by (smt B D Matrix.transpose_mult Matrix.transpose_transpose append_cols_def append_rows_def A carrier_matD(2) carrier_mat_triv index_mult_mat(3) index_transpose_mat(3) transpose_four_block_mat zero_carrier_mat zero_transpose_mat) finally show ?thesis by auto qed lemma append_rows_mult_left_id2: assumes A: "(A::'a::comm_semiring_1 mat) \ carrier_mat a n" and B: "B \ carrier_mat b n" and C: "C = four_block_mat D (0\<^sub>m a b) (0\<^sub>m b a) (1\<^sub>m b)" and D: "D \ carrier_mat a a" shows "C * (A @\<^sub>r B) = (D * A) @\<^sub>r B" proof - have "(C * (A @\<^sub>r B))\<^sup>T = (A @\<^sub>r B)\<^sup>T * C\<^sup>T" by (rule transpose_mult, insert assms, auto) also have "... = (A\<^sup>T @\<^sub>c B\<^sup>T) * C\<^sup>T" by (metis A B transpose_mat_append_rows) also have "... = (A\<^sup>T * D\<^sup>T @\<^sub>c B\<^sup>T)" by (rule append_cols_mult_right_id2, insert assms, auto) also have "...\<^sup>T = (D * A) @\<^sub>r B" by (metis A B D transpose_mult transpose_transpose mult_carrier_mat transpose_mat_append_rows) finally show ?thesis by simp qed lemma four_block_mat_preserves_column: assumes A: "(A::'a::semiring_1 mat) \ carrier_mat n m" and B: "B = four_block_mat (1\<^sub>m 1) (0\<^sub>m 1 (m - 1)) (0\<^sub>m (m - 1) 1) C" and C: "C \ carrier_mat (m-1) (m-1)" and i: "ic ?A2" by (rule append_cols_split[of A, unfolded n2], insert m A, auto) hence "A * B = (?A1 @\<^sub>c ?A2) * B" by simp also have "... = ?A1 @\<^sub>c (?A2 * C)" by (rule append_cols_mult_right_id[OF _ _ B C], insert A, auto) also have "... $$ (i,0) = ?A1 $$ (i,0)" using append_cols_nth by (simp add: append_cols_def i) also have "... = A $$ (i,0)" by (metis A i carrier_matD(1) col_def index_vec mat_of_cols_Cons_index_0) finally show ?thesis . qed definition "lower_triangular A = (\i j. i < j \ i < dim_row A \ j < dim_col A \ A $$ (i,j) = 0)" lemma lower_triangular_index: assumes "lower_triangular A" "i carrier_mat n n" shows "A * (k \\<^sub>m (1\<^sub>m n)) = (k \\<^sub>m (1\<^sub>m n)) * A" proof - have "(\ia = 0..ia = 0..ia \ ({0..ia \ ({0..ia \ ({0..ia \ ({0.. carrier_mat 2 2" shows "Determinant.det A = A$$(0,0) * A $$ (1,1) - A$$(0,1)*A$$(1,0)" proof - let ?A = "(Mod_Type_Connect.to_hma\<^sub>m A)::'a^2^2" have [transfer_rule]: "Mod_Type_Connect.HMA_M A ?A" unfolding Mod_Type_Connect.HMA_M_def using from_hma_to_hma\<^sub>m A by auto have [transfer_rule]: "Mod_Type_Connect.HMA_I 0 0" unfolding Mod_Type_Connect.HMA_I_def by (simp add: to_nat_0) have [transfer_rule]: "Mod_Type_Connect.HMA_I 1 1" unfolding Mod_Type_Connect.HMA_I_def by (simp add: to_nat_1) have "Determinant.det A = Determinants.det ?A" by (transfer, simp) also have "... = ?A $h 1 $h 1 * ?A $h 2 $h 2 - ?A $h 1 $h 2 * ?A $h 2 $h 1" unfolding det_2 by simp also have "... = ?A $h 0 $h 0 * ?A $h 1 $h 1 - ?A $h 0 $h 1 * ?A $h 1 $h 0" by (smt Groups.mult_ac(2) exhaust_2 semiring_norm(160)) also have "... = A$$(0,0) * A $$ (1,1) - A$$(0,1)*A$$(1,0)" unfolding index_hma_def[symmetric] by (transfer, auto) finally show ?thesis . qed lemma mat_diag_smult: "mat_diag n (\ x. (k::'a::comm_ring_1)) = (k \\<^sub>m 1\<^sub>m n)" proof - have "mat_diag n (\ x. k) = mat_diag n (\ x. k * 1)" by auto also have "... = mat_diag n (\ x. k) * mat_diag n (\ x. 1)" using mat_diag_diag by (simp add: mat_diag_def) also have "... = mat_diag n (\ x. k) * (1\<^sub>m n)" by auto thm mat_diag_mult_left also have "... = Matrix.mat n n (\(i, j). k * (1\<^sub>m n) $$ (i, j))" by (rule mat_diag_mult_left, auto) also have "... = (k \\<^sub>m 1\<^sub>m n)" unfolding smult_mat_def by auto finally show ?thesis . qed lemma invertible_mat_four_block_mat_lower_right: assumes A: "(A::'a::comm_ring_1 mat) \ carrier_mat n n" and inv_A: "invertible_mat A" shows "invertible_mat (four_block_mat (1\<^sub>m 1) (0\<^sub>m 1 n) (0\<^sub>m n 1) A)" proof - let ?I = "(four_block_mat (1\<^sub>m 1) (0\<^sub>m 1 n) (0\<^sub>m n 1) A)" have "Determinant.det ?I = Determinant.det (1\<^sub>m 1) * Determinant.det A" by (rule det_four_block_mat_lower_left_zero_col, insert assms, auto) also have "... = Determinant.det A" by auto finally have "Determinant.det ?I = Determinant.det A" . thus ?thesis by (metis (no_types, lifting) assms carrier_matD(1) carrier_matD(2) carrier_mat_triv index_mat_four_block(2) index_mat_four_block(3) index_one_mat(2) index_one_mat(3) invertible_iff_is_unit_JNF) qed lemma invertible_mat_four_block_mat_lower_right_id: assumes A: "(A::'a::comm_ring_1 mat) \ carrier_mat m m" and B: "B = 0\<^sub>m m (n-m)" and C: "C = 0\<^sub>m (n-m) m" and D: "D = 1\<^sub>m (n-m)" and "n>m" and inv_A: "invertible_mat A" shows "invertible_mat (four_block_mat A B C D)" proof - have "Determinant.det (four_block_mat A B C D) = Determinant.det A" by (rule det_four_block_mat_lower_right_id, insert assms, auto) thus ?thesis using inv_A by (metis (no_types, lifting) assms(1) assms(4) carrier_matD(1) carrier_matD(2) carrier_mat_triv index_mat_four_block(2) index_mat_four_block(3) index_one_mat(2) index_one_mat(3) invertible_iff_is_unit_JNF) qed lemma split_block4_decreases_dim_row: assumes E: "(A,B,C,D) = split_block E 1 1" and E1: "dim_row E > 1" and E2: "dim_col E > 1" shows "dim_row D < dim_row E" proof - have "D \ carrier_mat (1 + (dim_row E - 2)) (1 + (dim_col E - 2))" by (rule split_block(4)[OF E[symmetric]], insert E1 E2, auto) hence "D \ carrier_mat (dim_row E - 1) (dim_col E - 1)" using E1 E2 by auto thus ?thesis using E1 by auto qed lemma inv_P'PAQQ': assumes A: "A \ carrier_mat n n" and P: "P \ carrier_mat n n" and inv_P: "inverts_mat P' P" and inv_Q: "inverts_mat Q Q'" and Q: "Q \ carrier_mat n n" and P': "P' \ carrier_mat n n" and Q': "Q' \ carrier_mat n n" shows "(P'*(P*A*Q)*Q') = A" proof - have "(P'*(P*A*Q)*Q') = (P'*(P*A*Q*Q'))" by (smt P P' Q Q' assoc_mult_mat carrier_matD(1) carrier_matD(2) carrier_mat_triv index_mult_mat(2) index_mult_mat(3)) also have "... = ((P'*P)*A*(Q*Q'))" by (smt A P P' Q Q' assoc_mult_mat carrier_matD(1) carrier_matD(2) carrier_mat_triv index_mult_mat(3) inv_Q inverts_mat_def right_mult_one_mat') finally show ?thesis by (metis P' Q A inv_P inv_Q carrier_matD(1) inverts_mat_def left_mult_one_mat right_mult_one_mat) qed lemma assumes "U \ carrier_mat 2 2" and "V \ carrier_mat 2 2" and "A = U * V" shows mat_mult2_00: "A $$ (0,0) = U $$ (0,0)*V $$ (0,0) + U $$ (0,1)*V $$ (1,0)" and mat_mult2_01: "A $$ (0,1) = U $$ (0,0)*V $$ (0,1) + U $$ (0,1)*V $$ (1,1)" and mat_mult2_10: "A $$ (1,0) = U $$ (1,0)*V $$ (0,0) + U $$ (1,1)*V $$ (1,0)" and mat_mult2_11: "A $$ (1,1) = U $$ (1,0)*V $$ (0,1) + U $$ (1,1)*V $$ (1,1)" using assms unfolding times_mat_def Matrix.row_def col_def scalar_prod_def using sum_two_rw by auto subsection\Lemmas about @{text "sorted lists"}, @{text "insort"} and @{text "pick"}\ lemma sorted_distinct_imp_sorted_wrt: assumes "sorted xs" and "distinct xs" shows "sorted_wrt (<) xs" using assms by (induct xs, insert le_neq_trans, auto) lemma sorted_map_strict: assumes "strict_mono_on g {0.. g ` {0..x\set (map g [0.. g n" using sg unfolding strict_mono_on_def by (simp add: less_imp_le) qed finally show ?case . qed lemma sorted_nth_strict_mono: "sorted xs \ distinct xs \i < j \ j < length xs \ xs!i < xs!j" by (simp add: less_le nth_eq_iff_index_eq sorted_iff_nth_mono_less) lemma sorted_list_of_set_0_LEAST: assumes finI: "finite I" and I: "I \ {}" shows "sorted_list_of_set I ! 0 = (LEAST n. n\I)" proof (rule Least_equality[symmetric]) show "sorted_list_of_set I ! 0 \ I" by (metis I Max_in finI gr_zeroI in_set_conv_nth not_less_zero set_sorted_list_of_set) fix y assume "y \ I" thus "sorted_list_of_set I ! 0 \ y" by (metis eq_iff finI in_set_conv_nth neq0_conv sorted_iff_nth_mono_less sorted_list_of_set(1) sorted_sorted_list_of_set) qed lemma sorted_list_of_set_eq_pick: assumes i: "i < length (sorted_list_of_set I)" shows "sorted_list_of_set I ! i = pick I i" proof - have finI: "finite I" proof (rule ccontr) assume "infinite I" - hence "length (sorted_list_of_set I) = 0" using sorted_list_of_set.infinite by auto + hence "length (sorted_list_of_set I) = 0" by auto thus False using i by simp qed show ?thesis using i proof (induct i) case 0 have I: "I \ {}" using "0.prems" sorted_list_of_set_empty by blast show ?case unfolding pick.simps by (rule sorted_list_of_set_0_LEAST[OF finI I]) next case (Suc i) note x_less = Suc.prems show ?case proof (unfold pick.simps, rule Least_equality[symmetric], rule conjI) show 1: "pick I i < sorted_list_of_set I ! Suc i" by (metis Suc.hyps Suc.prems Suc_lessD distinct_sorted_list_of_set find_first_unique lessI nat_less_le sorted_sorted_list_of_set sorted_wrt_nth_less) show "sorted_list_of_set I ! Suc i \ I" using Suc.prems finI nth_mem set_sorted_list_of_set by blast have rw: "sorted_list_of_set I ! i = pick I i" using Suc.hyps Suc_lessD x_less by blast have sorted_less: "sorted_list_of_set I ! i < sorted_list_of_set I ! Suc i" by (simp add: 1 rw) fix y assume y: "y \ I \ pick I i < y" show "sorted_list_of_set I ! Suc i \ y" by (smt antisym_conv finI in_set_conv_nth less_Suc_eq less_Suc_eq_le nat_neq_iff rw sorted_iff_nth_mono_less sorted_list_of_set(1) sorted_sorted_list_of_set x_less y) qed qed qed text\$b$ is the position where we add, $a$ the element to be added and $i$ the position that is checked\ lemma insort_nth': assumes "\j set xs" and "i < length xs + 1" and "i < b" and "xs \ []" and "b < length xs" shows "insort a xs ! i = xs ! i" using assms proof (induct xs arbitrary: a b i) case Nil then show ?case by auto next case (Cons x xs) note less = Cons.prems(1) note sorted = Cons.prems(2) note a_notin = Cons.prems(3) note i_length = Cons.prems(4) note i_b = Cons.prems(5) note b_length = Cons.prems(7) show ?case proof (cases "a \ x") case True have "insort a (x # xs) ! i = (a # x # xs) ! i" using True by simp also have "... = (x # xs) ! i" using Cons.prems(1) Cons.prems(5) True by force finally show ?thesis . next case False note x_less_a = False have "insort a (x # xs) ! i = (x # insort a xs) ! i" using False by simp also have "... = (x # xs) ! i" proof (cases "i = 0") case True then show ?thesis by auto next case False have "(x # insort a xs) ! i = (insort a xs) ! (i-1)" by (simp add: False nth_Cons') also have "... = xs ! (i-1)" proof (rule Cons.hyps) show "sorted xs" using sorted by simp show "a \ set xs" using a_notin by simp show "i - 1 < length xs + 1" using i_length False by auto show "xs \ []" using i_b b_length by force show "i - 1 < b - 1" by (simp add: False diff_less_mono i_b leI) show "b - 1 < length xs" using b_length i_b by auto show "\j set xs" and "i < index (insort a xs) a" and "xs \ []" shows "insort a xs ! i = xs ! i" using assms proof (induct xs arbitrary: a i) case Nil then show ?case by auto next case (Cons x xs) note sorted = Cons.prems(1) note a_notin = Cons.prems(2) note i_index = Cons.prems(3) show ?case proof (cases "a \ x") case True have "insort a (x # xs) ! i = (a # x # xs) ! i" using True by simp also have "... = (x # xs) ! i" using Cons.prems(1) Cons.prems(3) True by force finally show ?thesis . next case False note x_less_a = False show ?thesis proof (cases "xs = []") case True have "x \ a" using False by auto then show ?thesis using True i_index False by auto next case False note xs_not_empty = False have "insort a (x # xs) ! i = (x # insort a xs) ! i" using x_less_a by simp also have "... = (x # xs) ! i" proof (cases "i = 0") case True then show ?thesis by auto next case False note i0 = False have "(x # insort a xs) ! i = (insort a xs) ! (i-1)" by (simp add: False nth_Cons') also have "... = xs ! (i-1)" proof (rule Cons.hyps[OF _ _ _ xs_not_empty]) show "sorted xs" using sorted by simp show "a \ set xs" using a_notin by simp have "index (insort a (x # xs)) a = index ((x # insort a xs)) a" using x_less_a by auto also have "... = index (insort a xs) a + 1" unfolding index_Cons using x_less_a by simp finally show "i - 1 < index (insort a xs) a" using False i_index by linarith qed also have "... = (x # xs) ! i" by (simp add: False nth_Cons') finally show ?thesis . qed finally show ?thesis . qed qed qed lemma insort_nth2: assumes "sorted xs" and "a \ set xs" and "i < length xs" and "i \ index (insort a xs) a" and "xs \ []" shows "insort a xs ! (Suc i) = xs ! i" using assms proof (induct xs arbitrary: a i) case Nil then show ?case by auto next case (Cons x xs) note sorted = Cons.prems(1) note a_notin = Cons.prems(2) note i_length = Cons.prems(3) note index_i = Cons.prems(4) show ?case proof (cases "a \ x") case True have "insort a (x # xs) ! (Suc i) = (a # x # xs) ! (Suc i)" using True by simp also have "... = (x # xs) ! i" using Cons.prems(1) Cons.prems(5) True by force finally show ?thesis . next case False note x_less_a = False have "insort a (x # xs) ! (Suc i) = (x # insort a xs) ! (Suc i)" using False by simp also have "... = (x # xs) ! i" proof (cases "i = 0") case True then show ?thesis using index_i linear x_less_a by fastforce next case False note i0 = False show ?thesis proof - have Suc_i: "Suc (i - 1) = i" using i0 by auto have "(x # insort a xs) ! (Suc i) = (insort a xs) ! i" by (simp add: nth_Cons') also have "... = (insort a xs) ! Suc (i - 1)" using Suc_i by simp also have "... = xs ! (i - 1)" proof (rule Cons.hyps) show "sorted xs" using sorted by simp show "a \ set xs" using a_notin by simp show "i - 1 < length xs" using i_length using Suc_i by auto thus "xs \ []" by auto have "index (insort a (x # xs)) a = index ((x # insort a xs)) a" using x_less_a by simp also have "... = index (insort a xs) a + 1" unfolding index_Cons using x_less_a by simp finally show "index (insort a xs) a \ i - 1" using index_i i0 by auto qed also have "... = (x # xs) ! i" using Suc_i by auto finally show ?thesis . qed qed finally show ?thesis . qed qed lemma pick_index: assumes a: "a \ I" and a'_card: "a' < card I" shows "(pick I a' = a) = (index (sorted_list_of_set I) a = a')" proof - have finI: "finite I" using a'_card card.infinite by force have length_I: "length (sorted_list_of_set I) = card I" by (metis a'_card card.infinite distinct_card distinct_sorted_list_of_set not_less_zero set_sorted_list_of_set) let ?i = "index (sorted_list_of_set I) a" have "(sorted_list_of_set I) ! a' = pick I a'" by (rule sorted_list_of_set_eq_pick, auto simp add: finI a'_card length_I) moreover have "(sorted_list_of_set I) ! ?i = a" by (rule nth_index, simp add: a finI) ultimately show ?thesis by (metis a'_card distinct_sorted_list_of_set index_nth_id length_I) qed end diff --git a/thys/UTP/toolkit/Sequence.thy b/thys/UTP/toolkit/Sequence.thy --- a/thys/UTP/toolkit/Sequence.thy +++ b/thys/UTP/toolkit/Sequence.thy @@ -1,287 +1,287 @@ (******************************************************************************) (* Project: Isabelle/UTP Toolkit *) (* File: Sequence.thy *) (* Authors: Simon Foster and Frank Zeyda *) (* Emails: simon.foster@york.ac.uk and frank.zeyda@york.ac.uk *) (******************************************************************************) section \ Infinite Sequences \ theory Sequence imports HOL.Real List_Extra "HOL-Library.Sublist" "HOL-Library.Nat_Bijection" begin typedef 'a seq = "UNIV :: (nat \ 'a) set" by (auto) setup_lifting type_definition_seq definition ssubstr :: "nat \ nat \ 'a seq \ 'a list" where "ssubstr i j xs = map (Rep_seq xs) [i ..< j]" lift_definition nth_seq :: "'a seq \ nat \ 'a" (infixl "!\<^sub>s" 100) is "\ f i. f i" . abbreviation sinit :: "nat \ 'a seq \ 'a list" where "sinit i xs \ ssubstr 0 i xs" lemma sinit_len [simp]: "length (sinit i xs) = i" by (simp add: ssubstr_def) lemma sinit_0 [simp]: "sinit 0 xs = []" by (simp add: ssubstr_def) lemma prefix_upt_0 [intro]: "i \ j \ prefix [0.. j \ prefix (sinit i xs) (sinit j xs)" by (simp add: map_mono_prefix prefix_upt_0 ssubstr_def) lemma sinit_strict_prefix: "i < j \ strict_prefix (sinit i xs) (sinit j xs)" by (metis sinit_len sinit_prefix le_less nat_neq_iff prefix_order.dual_order.strict_iff_order) lemma nth_sinit: "i < n \ sinit n xs ! i = xs !\<^sub>s i" apply (auto simp add: ssubstr_def) apply (transfer, auto) done lemma sinit_append_split: assumes "i < j" shows "sinit j xs = sinit i xs @ ssubstr i j xs" proof - have "[0.. lexord R" "(sinit j ys, sinit j xs) \ lexord R" shows False proof - have sinit_xs: "sinit j xs = sinit i xs @ ssubstr i j xs" by (metis assms(2) sinit_append_split) have sinit_ys: "sinit j ys = sinit i ys @ ssubstr i j ys" by (metis assms(2) sinit_append_split) from sinit_xs sinit_ys assms(4) have "(sinit i ys, sinit i xs) \ lexord R \ (sinit i ys = sinit i xs \ (ssubstr i j ys, ssubstr i j xs) \ lexord R)" by (auto dest: lexord_append) with assms lexord_asymmetric show False by (force) qed lemma sinit_linear_asym_lemma2: assumes "asym R" "(sinit i xs, sinit i ys) \ lexord R" "(sinit j ys, sinit j xs) \ lexord R" shows False proof (cases i j rule: linorder_cases) case less with assms show ?thesis by (auto dest: sinit_linear_asym_lemma1) next case equal with assms show ?thesis by (simp add: lexord_asymmetric) next case greater with assms show ?thesis by (auto dest: sinit_linear_asym_lemma1) qed lemma range_ext: assumes "\i :: nat. \x\{0.. x" by (metis lessI) with assms show "f(x) = g(x)" by (auto) qed lemma sinit_ext: "(\ i. sinit i xs = sinit i ys) \ xs = ys" by (simp add: ssubstr_def, transfer, auto intro: range_ext) definition seq_lexord :: "'a rel \ ('a seq) rel" where "seq_lexord R = {(xs, ys). (\ i. (sinit i xs, sinit i ys) \ lexord R)}" lemma seq_lexord_irreflexive: "\x. (x, x) \ R \ (xs, xs) \ seq_lexord R" by (auto dest: lexord_irreflexive simp add: irrefl_def seq_lexord_def) lemma seq_lexord_irrefl: "irrefl R \ irrefl (seq_lexord R)" by (simp add: irrefl_def seq_lexord_irreflexive) lemma seq_lexord_transitive: assumes "trans R" shows "trans (seq_lexord R)" unfolding seq_lexord_def proof (rule transI, clarify) fix xs ys zs :: "'a seq" and m n :: nat assume las: "(sinit m xs, sinit m ys) \ lexord R" "(sinit n ys, sinit n zs) \ lexord R" hence inz: "m > 0" using gr0I by force from las(1) obtain i where sinitm: "(sinit m xs!i, sinit m ys!i) \ R" "i < m" "\ j R" "j < n" "\ ki. (sinit i xs, sinit i zs) \ lexord R" proof (cases "i \ j") case True note lt = this with sinitm sinitn have "(sinit n xs!i, sinit n zs!i) \ R" by (metis assms le_eq_less_or_eq le_less_trans nth_sinit transD) moreover from lt sinitm sinitn have "\ j lexord R" using sinitm(2) sinitn(2) lt apply (rule_tac lexord_intro_elems) apply (auto) apply (metis less_le_trans less_trans nth_sinit) done thus ?thesis by auto next case False then have ge: "i > j" by auto with assms sinitm sinitn have "(sinit n xs!j, sinit n zs!j) \ R" by (metis less_trans nth_sinit) moreover from ge sinitm sinitn have "\ k lexord R" using sinitm(2) sinitn(2) ge apply (rule_tac lexord_intro_elems) apply (auto) apply (metis less_trans nth_sinit) done thus ?thesis by auto qed qed lemma seq_lexord_trans: "\ (xs, ys) \ seq_lexord R; (ys, zs) \ seq_lexord R; trans R \ \ (xs, zs) \ seq_lexord R" by (meson seq_lexord_transitive transE) lemma seq_lexord_antisym: "\ asym R; (a, b) \ seq_lexord R \ \ (b, a) \ seq_lexord R" by (auto dest: sinit_linear_asym_lemma2 simp add: seq_lexord_def) lemma seq_lexord_asym: assumes "asym R" shows "asym (seq_lexord R)" by (meson assms asym.simps seq_lexord_antisym seq_lexord_irrefl) lemma seq_lexord_total: assumes "total R" shows "total (seq_lexord R)" using assms by (auto simp add: total_on_def seq_lexord_def, meson lexord_linear sinit_ext) lemma seq_lexord_linear: assumes "(\ a b. (a,b)\ R \ a = b \ (b,a) \ R)" shows "(x,y) \ seq_lexord R \ x = y \ (y,x) \ seq_lexord R" proof - have "total R" using assms total_on_def by blast hence "total (seq_lexord R)" using seq_lexord_total by blast thus ?thesis by (auto simp add: total_on_def) qed instantiation seq :: (ord) ord begin definition less_seq :: "'a seq \ 'a seq \ bool" where "less_seq xs ys \ (xs, ys) \ seq_lexord {(xs, ys). xs < ys}" definition less_eq_seq :: "'a seq \ 'a seq \ bool" where "less_eq_seq xs ys = (xs = ys \ xs < ys)" instance .. end instance seq :: (order) order proof fix xs :: "'a seq" show "xs \ xs" by (simp add: less_eq_seq_def) next fix xs ys zs :: "'a seq" assume "xs \ ys" and "ys \ zs" then show "xs \ zs" by (force dest: seq_lexord_trans simp add: less_eq_seq_def less_seq_def trans_def) next fix xs ys :: "'a seq" assume "xs \ ys" and "ys \ xs" then show "xs = ys" apply (auto simp add: less_eq_seq_def less_seq_def) apply (rule seq_lexord_irreflexive [THEN notE]) defer apply (rule seq_lexord_trans) apply (auto intro: transI) done next fix xs ys :: "'a seq" show "xs < ys \ xs \ ys \ \ ys \ xs" apply (auto simp add: less_seq_def less_eq_seq_def) defer apply (rule seq_lexord_irreflexive [THEN notE]) apply auto apply (rule seq_lexord_irreflexive [THEN notE]) defer apply (rule seq_lexord_trans) apply (auto intro: transI) apply (simp add: seq_lexord_irreflexive) done qed instance seq :: (linorder) linorder proof fix xs ys :: "'a seq" have "(xs, ys) \ seq_lexord {(u, v). u < v} \ xs = ys \ (ys, xs) \ seq_lexord {(u, v). u < v}" by (rule seq_lexord_linear) auto then show "xs \ ys \ ys \ xs" by (auto simp add: less_eq_seq_def less_seq_def) qed lemma seq_lexord_mono [mono]: "(\ x y. f x y \ g x y) \ (xs, ys) \ seq_lexord {(x, y). f x y} \ (xs, ys) \ seq_lexord {(x, y). g x y}" apply (auto simp add: seq_lexord_def) apply (metis case_prodD case_prodI lexord_take_index_conv mem_Collect_eq) done fun insort_rel :: "'a rel \ 'a \ 'a list \ 'a list" where "insort_rel R x [] = [x]" | "insort_rel R x (y # ys) = (if (x = y \ (x,y) \ R) then x # y # ys else y # insort_rel R x ys)" inductive sorted_rel :: "'a rel \ 'a list \ bool" where Nil_rel [iff]: "sorted_rel R []" | Cons_rel: "\ y \ set xs. (x = y \ (x, y) \ R) \ sorted_rel R xs \ sorted_rel R (x # xs)" definition list_of_set :: "'a rel \ 'a set \ 'a list" where -"list_of_set R = folding.F (insort_rel R) []" +"list_of_set R = folding_on.F (insort_rel R) []" lift_definition seq_inj :: "'a seq seq \ 'a seq" is "\ f i. f (fst (prod_decode i)) (snd (prod_decode i))" . lift_definition seq_proj :: "'a seq \ 'a seq seq" is "\ f i j. f (prod_encode (i, j))" . lemma seq_inj_inverse: "seq_proj (seq_inj x) = x" by (transfer, simp) lemma seq_proj_inverse: "seq_inj (seq_proj x) = x" by (transfer, simp) lemma seq_inj: "inj seq_inj" by (metis injI seq_inj_inverse) lemma seq_inj_surj: "bij seq_inj" apply (rule bijI) apply (auto simp add: seq_inj) apply (metis rangeI seq_proj_inverse) done end \ No newline at end of file