diff --git a/thys/CakeML_Codegen/Rewriting/Rewriting_Nterm.thy b/thys/CakeML_Codegen/Rewriting/Rewriting_Nterm.thy --- a/thys/CakeML_Codegen/Rewriting/Rewriting_Nterm.thy +++ b/thys/CakeML_Codegen/Rewriting/Rewriting_Nterm.thy @@ -1,773 +1,786 @@ section \Higher-order term rewriting using explicit bound variable names\ theory Rewriting_Nterm imports Rewriting_Term Higher_Order_Terms.Term_to_Nterm "../Terms/Strong_Term" begin subsection \Definitions\ type_synonym nrule = "term \ nterm" abbreviation nrule :: "nrule \ bool" where "nrule \ basic_rule" fun (in constants) not_shadowing :: "nrule \ bool" where "not_shadowing (lhs, rhs) \ \ shadows_consts lhs \ \ shadows_consts rhs" locale nrules = constants C_info "heads_of rs" for C_info and rs :: "nrule fset" + assumes all_rules: "fBall rs nrule" assumes arity: "arity_compatibles rs" assumes fmap: "is_fmap rs" assumes patterns: "pattern_compatibles rs" assumes nonempty: "rs \ {||}" assumes not_shadows: "fBall rs not_shadowing" assumes welldefined_rs: "fBall rs (\(_, rhs). welldefined rhs)" subsection \Matching and rewriting\ inductive nrewrite :: "nrule fset \ nterm \ nterm \ bool" ("_/ \\<^sub>n/ _ \/ _" [50,0,50] 50) for rs where step: "r |\| rs \ r \ t \ u \ rs \\<^sub>n t \ u" | beta: "rs \\<^sub>n ((\\<^sub>n x. t) $\<^sub>n t') \ subst t (fmap_of_list [(x, t')])" | "fun": "rs \\<^sub>n t \ t' \ rs \\<^sub>n t $\<^sub>n u \ t' $\<^sub>n u" | arg: "rs \\<^sub>n u \ u' \ rs \\<^sub>n t $\<^sub>n u \ t $\<^sub>n u'" global_interpretation nrewrite: rewriting "nrewrite rs" for rs by standard (auto intro: nrewrite.intros simp: app_nterm_def)+ abbreviation nrewrite_rt :: "nrule fset \ nterm \ nterm \ bool" ("_/ \\<^sub>n/ _ \*/ _" [50,0,50] 50) where "nrewrite_rt rs \ (nrewrite rs)\<^sup>*\<^sup>*" lemma (in nrules) nrewrite_closed: assumes "rs \\<^sub>n t \ t'" "closed t" shows "closed t'" using assms proof induction case (step r t u) obtain lhs rhs where "r = (lhs, rhs)" by force with step have "nrule (lhs, rhs)" using all_rules by blast hence "frees rhs |\| frees lhs" by simp have "(lhs, rhs) \ t \ u" using step unfolding \r = _\ by simp show ?case apply (rule rewrite_step_closed) apply fact+ done next case beta show ?case apply simp apply (subst closed_except_def) apply (subst subst_frees) using beta unfolding closed_except_def by auto qed (auto simp: closed_except_def) corollary (in nrules) nrewrite_rt_closed: assumes "rs \\<^sub>n t \* t'" "closed t" shows "closed t'" using assms by induction (auto intro: nrewrite_closed) subsection \Translation from @{typ term} to @{typ nterm}\ context begin private lemma term_to_nterm_all_vars0: assumes "wellformed' (length \) t" shows "\T. all_frees (fst (run_state (term_to_nterm \ t) x)) |\| fset_of_list \ |\| frees t |\| T \ fBall T (\y. y > x)" using assms proof (induction \ t arbitrary: x rule: term_to_nterm_induct) case (bound \ i) hence "\ ! i |\| fset_of_list \" by (simp add: fset_of_list_elem) with bound show ?case by (auto simp: State_Monad.return_def) next case (abs \ t) obtain T where "all_frees (fst (run_state (term_to_nterm (next x # \) t) (next x))) |\| fset_of_list (next x # \) |\| frees t |\| T" and "fBall T ((<) (next x))" apply atomize_elim apply (rule abs(1)) using abs by auto show ?case apply (simp add: split_beta create_alt_def) apply (rule exI[where x = "finsert (next x) T"]) apply (intro conjI) subgoal by auto subgoal using \all_frees (fst (run_state _ (next x))) |\| _\ by simp subgoal apply simp apply (rule conjI) apply (rule next_ge) using \fBall T ((<) (next x))\ by (metis fBallE fBallI fresh_class.next_ge order.strict_trans) done next case (app \ t\<^sub>1 t\<^sub>2 x\<^sub>1) obtain t\<^sub>1' x\<^sub>2 where "run_state (term_to_nterm \ t\<^sub>1) x\<^sub>1 = (t\<^sub>1', x\<^sub>2)" by fastforce moreover obtain T\<^sub>1 where "all_frees (fst (run_state (term_to_nterm \ t\<^sub>1) x\<^sub>1)) |\| fset_of_list \ |\| frees t\<^sub>1 |\| T\<^sub>1" and "fBall T\<^sub>1 ((<) x\<^sub>1)" apply atomize_elim apply (rule app(1)) using app by auto ultimately have "all_frees t\<^sub>1' |\| fset_of_list \ |\| frees t\<^sub>1 |\| T\<^sub>1" by simp obtain T\<^sub>2 where "all_frees (fst (run_state (term_to_nterm \ t\<^sub>2) x\<^sub>2)) |\| fset_of_list \ |\| frees t\<^sub>2 |\| T\<^sub>2" and "fBall T\<^sub>2 ((<) x\<^sub>2)" apply atomize_elim apply (rule app(2)) using app by auto moreover obtain t\<^sub>2' x' where "run_state (term_to_nterm \ t\<^sub>2) x\<^sub>2 = (t\<^sub>2', x')" by fastforce ultimately have "all_frees t\<^sub>2' |\| fset_of_list \ |\| frees t\<^sub>2 |\| T\<^sub>2" by simp have "x\<^sub>1 \ x\<^sub>2" apply (rule state_io_relD[OF term_to_nterm_mono]) apply fact done show ?case apply simp unfolding \run_state (term_to_nterm \ t\<^sub>1) x\<^sub>1 = _\ apply simp unfolding \run_state (term_to_nterm \ t\<^sub>2) x\<^sub>2 = _\ apply simp apply (rule exI[where x = "T\<^sub>1 |\| T\<^sub>2"]) apply (intro conjI) subgoal using \all_frees t\<^sub>1' |\| _\ by auto subgoal using \all_frees t\<^sub>2' |\| _\ by auto subgoal apply auto using \fBall T\<^sub>1 ((<) x\<^sub>1)\ apply auto[] using \fBall T\<^sub>2 ((<) x\<^sub>2)\ \x\<^sub>1 \ x\<^sub>2\ - by (meson fBallE less_le_not_le order_trans) + by auto done qed auto lemma term_to_nterm_all_vars: assumes "wellformed t" "fdisjnt (frees t) S" shows "fdisjnt (all_frees (fresh_frun (term_to_nterm [] t) (T |\| S))) S" proof - let ?\ = "[]" let ?x = "fresh_fNext (T |\| S)" from assms have "wellformed' (length ?\) t" by simp then obtain F where "all_frees (fst (run_state (term_to_nterm ?\ t) ?x)) |\| fset_of_list ?\ |\| frees t |\| F" and "fBall F (\y. y > ?x)" by (metis term_to_nterm_all_vars0) have "fdisjnt F (T |\| S)" if "S \ {||}" apply (rule fdisjnt_ge_max) apply (rule fBall_pred_weaken[OF _ \fBall F (\y. y > ?x)\]) apply (rule less_trans) apply (rule fNext_ge_max) using that by auto show ?thesis apply (rule fdisjnt_subset_left) apply (subst fresh_frun_def) apply (subst fresh_fNext_def[symmetric]) apply fact apply simp apply (rule fdisjnt_union_left) apply fact using \_ \ fdisjnt F (T |\| S)\ by (auto simp: fdisjnt_alt_def) qed end fun translate_rule :: "name fset \ rule \ nrule" where "translate_rule S (lhs, rhs) = (lhs, fresh_frun (term_to_nterm [] rhs) (frees lhs |\| S))" lemma translate_rule_alt_def: "translate_rule S = (\(lhs, rhs). (lhs, fresh_frun (term_to_nterm [] rhs) (frees lhs |\| S)))" by auto definition compile' where "compile' C_info rs = translate_rule (pre_constants.all_consts C_info (heads_of rs)) |`| rs" context rules begin definition compile :: "nrule fset" where "compile = translate_rule all_consts |`| rs" lemma compile'_compile_eq[simp]: "compile' C_info rs = compile" unfolding compile'_def compile_def .. lemma compile_heads: "heads_of compile = heads_of rs" unfolding compile_def translate_rule_alt_def head_def[abs_def] by force lemma compile_rules: "nrules C_info compile" proof have "fBall compile (\(lhs, rhs). nrule (lhs, rhs))" proof safe fix lhs rhs assume "(lhs, rhs) |\| compile" then obtain rhs' where "(lhs, rhs') |\| rs" and rhs: "rhs = fresh_frun (term_to_nterm [] rhs') (frees lhs |\| all_consts)" unfolding compile_def by force then have rule: "rule (lhs, rhs')" using all_rules by blast show "nrule (lhs, rhs)" proof from rule show "linear lhs" "is_const (fst (strip_comb lhs))" "\ is_const lhs" by auto have "frees rhs |\| frees rhs'" unfolding rhs using rule by (metis rule.simps term_to_nterm_vars) also have "frees rhs' |\| frees lhs" using rule by auto finally show "frees rhs |\| frees lhs" . qed qed thus "fBall compile nrule" by fast next show "arity_compatibles compile" proof safe fix lhs\<^sub>1 rhs\<^sub>1 lhs\<^sub>2 rhs\<^sub>2 assume "(lhs\<^sub>1, rhs\<^sub>1) |\| compile" "(lhs\<^sub>2, rhs\<^sub>2) |\| compile" then obtain rhs\<^sub>1' rhs\<^sub>2' where "(lhs\<^sub>1, rhs\<^sub>1') |\| rs" "(lhs\<^sub>2, rhs\<^sub>2') |\| rs" unfolding compile_def by force thus "arity_compatible lhs\<^sub>1 lhs\<^sub>2" using arity by (blast dest: fpairwiseD) qed next have "is_fmap rs" using fmap by simp thus "is_fmap compile" unfolding compile_def translate_rule_alt_def by (rule is_fmap_image) next have "pattern_compatibles rs" using patterns by simp thus "pattern_compatibles compile" unfolding compile_def translate_rule_alt_def by (auto dest: fpairwiseD) next show "fdisjnt (heads_of compile) C" using disjnt by (simp add: compile_heads) next have "fBall compile not_shadowing" proof safe fix lhs rhs assume "(lhs, rhs) |\| compile" then obtain rhs' where "rhs = fresh_frun (term_to_nterm [] rhs') (frees lhs |\| all_consts)" and "(lhs, rhs') |\| rs" unfolding compile_def translate_rule_alt_def by auto hence "rule (lhs, rhs')" "\ shadows_consts lhs" using all_rules not_shadows by blast+ moreover hence "wellformed rhs'" "frees rhs' |\| frees lhs" "fdisjnt all_consts (frees lhs)" unfolding shadows_consts_def by simp+ moreover have "\ shadows_consts rhs" apply (subst shadows_consts_def) apply simp unfolding \rhs = _\ apply (rule fdisjnt_swap) apply (rule term_to_nterm_all_vars) apply fact apply (rule fdisjnt_subset_left) apply fact apply (rule fdisjnt_swap) apply fact done ultimately show "not_shadowing (lhs, rhs)" unfolding compile_heads by simp qed thus "fBall compile (constants.not_shadowing C_info (heads_of compile))" unfolding compile_heads . have "fBall compile (\(_, rhs). welldefined rhs)" unfolding compile_heads unfolding compile_def ball_simps apply (rule fBall_pred_weaken[OF _ welldefined_rs]) subgoal for x apply (cases x) apply simp apply (subst fresh_frun_def) apply (subst term_to_nterm_consts[THEN pred_state_run_state]) by auto done thus "fBall compile (\(_, rhs). consts rhs |\| pre_constants.all_consts C_info (heads_of compile))" unfolding compile_heads . next show "compile \ {||}" using nonempty unfolding compile_def by auto next show "distinct all_constructors" by (fact distinct_ctr) qed sublocale rules_as_nrules: nrules C_info compile by (fact compile_rules) end subsection \Correctness of translation\ theorem (in rules) compile_correct: assumes "compile \\<^sub>n u \ u'" "closed u" shows "rs \ nterm_to_term' u \ nterm_to_term' u'" using assms proof (induction u u') case (step r u u') moreover obtain pat rhs' where "r = (pat, rhs')" by force ultimately obtain nenv where "match pat u = Some nenv" "u' = subst rhs' nenv" by auto then obtain env where "nrelated.P_env [] env nenv" "match pat (nterm_to_term [] u) = Some env" by (metis nrelated.match_rel option.exhaust option.rel_distinct(1) option.rel_inject(2)) have "closed_env nenv" using step \match pat u = Some nenv\ by (intro closed.match) from step obtain rhs where "rhs' = fresh_frun (term_to_nterm [] rhs) (frees pat |\| all_consts)" "(pat, rhs) |\| rs" unfolding \r = _\ compile_def by auto with assms have "rule (pat, rhs)" using all_rules by blast hence "rhs = nterm_to_term [] rhs'" unfolding \rhs' = _\ by (simp add: term_to_nterm_nterm_to_term fresh_frun_def) have "compile \\<^sub>n u \ u'" using step by (auto intro: nrewrite.step) hence "closed u'" by (rule rules_as_nrules.nrewrite_closed) fact show ?case proof (rule rewrite.step) show "(pat, rhs) \ nterm_to_term [] u \ nterm_to_term [] u'" apply (subst nterm_to_term_eq_closed) apply fact apply simp apply (rule exI[where x = env]) apply (rule conjI) apply fact unfolding \rhs = _\ apply (subst nrelated_subst) apply fact apply fact unfolding fdisjnt_alt_def apply simp unfolding \u' = subst rhs' nenv\ by simp qed fact next case beta show ?case apply simp apply (subst subst_single_eq[symmetric, simplified]) apply (subst nterm_to_term_subst_replace_bound[where n = 0]) subgoal using beta by (simp add: closed_except_def) subgoal by simp subgoal by simp subgoal by simp (rule rewrite.beta) done qed (auto intro: rewrite.intros simp: closed_except_def) subsection \Completeness of translation\ context rules begin context notes [simp] = closed_except_def fdisjnt_alt_def begin private lemma compile_complete0: assumes "rs \ t \ t'" "closed t" "wellformed t" obtains u' where "compile \\<^sub>n fst (run_state (term_to_nterm [] t) s) \ u'" "u' \\<^sub>\ fst (run_state (term_to_nterm [] t') s')" apply atomize_elim using assms proof (induction t t' arbitrary: s s') case (step r t t') let ?t\<^sub>n = "fst (run_state (term_to_nterm [] t) s)" let ?t\<^sub>n' = "fst (run_state (term_to_nterm [] t') s')" from step have "closed t" "closed ?t\<^sub>n" using term_to_nterm_vars0[where \ = "[]"] by force+ from step have "nterm_to_term' ?t\<^sub>n = t" by (auto intro!: term_to_nterm_nterm_to_term0) obtain pat rhs' where "r = (pat, rhs')" by fastforce with step obtain env' where "match pat t = Some env'" "t' = subst rhs' env'" by fastforce with \_ = t\ have "rel_option (nrelated.P_env []) (match pat t) (match pat (?t\<^sub>n))" by (metis nrelated.match_rel) with step \_ = Some env'\ obtain env where "nrelated.P_env [] env' env" "match pat ?t\<^sub>n = Some env" by (metis (no_types, lifting) option_rel_Some1) with \closed ?t\<^sub>n\ have "closed_env env" by (blast intro: closed.match) from step obtain rhs where "rhs = fresh_frun (term_to_nterm [] rhs') (frees pat |\| all_consts)" "(pat, rhs) |\| compile" unfolding \r = _\ compile_def by force with step have "rule (pat, rhs')" unfolding \r = _\ using all_rules by fast hence "nterm_to_term' rhs = rhs'" unfolding \rhs = _\ by (simp add: fresh_frun_def term_to_nterm_nterm_to_term) obtain u' where "subst rhs env = u'" by simp have "t' = nterm_to_term' u'" unfolding \t' = _\ unfolding \_ = rhs'\[symmetric] apply (subst nrelated_subst) apply fact+ using \_ = u'\ by simp+ have "compile \\<^sub>n ?t\<^sub>n \ u'" apply (rule nrewrite.step) apply fact apply simp apply (intro conjI exI) apply fact+ done with \closed ?t\<^sub>n\ have "closed u'" by (blast intro:rules_as_nrules.nrewrite_closed) with \t' = nterm_to_term' _\ have "u' \\<^sub>\ ?t\<^sub>n'" by (force intro: nterm_to_term_term_to_nterm[where \ = "[]" and \' = "[]",simplified]) show ?case apply (intro conjI exI) apply (rule nrewrite.step) apply fact apply simp apply (intro conjI exI) apply fact+ done next case (beta t t') let ?name = "next s" let ?t\<^sub>n = "fst (run_state (term_to_nterm [?name] t) (?name))" let ?t\<^sub>n' = "fst (run_state (term_to_nterm [] t') (snd (run_state (term_to_nterm [?name] t) (?name))))" from beta have "closed t" "closed (t [t']\<^sub>\)" "closed (\ t $ t')" "closed t'" using replace_bound_frees by fastforce+ moreover from beta have "wellformed' (Suc 0) t" "wellformed t'" by simp+ ultimately have "t = nterm_to_term [?name] ?t\<^sub>n" and "t' = nterm_to_term' ?t\<^sub>n'" and *:"frees ?t\<^sub>n = {|?name|} \ frees ?t\<^sub>n = fempty" and "closed ?t\<^sub>n'" using term_to_nterm_vars0[where \ = "[?name]"] using term_to_nterm_vars0[where \ = "[]"] by (force simp: term_to_nterm_nterm_to_term0)+ hence **:"t [t']\<^sub>\ = nterm_to_term' (subst_single ?t\<^sub>n ?name ?t\<^sub>n')" by (auto simp: nterm_to_term_subst_replace_bound[where n = 0]) from \closed ?t\<^sub>n'\ have "closed_env (fmap_of_list [(?name, ?t\<^sub>n')])" by auto show ?case apply (rule exI) apply (auto simp: split_beta create_alt_def) apply (rule nrewrite.beta) apply (subst subst_single_eq[symmetric]) apply (subst **) apply (rule nterm_to_term_term_to_nterm[where \ = "[]" and \' = "[]", simplified]) apply (subst subst_single_eq) apply (subst subst_frees[OF \closed_env _\]) using * by force next case ("fun" t t' u) hence "closed t" "closed u" "closed (t $ u)" and wellform:"wellformed t" "wellformed u" by fastforce+ from "fun" obtain u' where "compile \\<^sub>n fst (run_state (term_to_nterm [] t) s) \ u'" "u' \\<^sub>\ fst (run_state (term_to_nterm [] t') s')" by force show ?case apply (rule exI) apply (auto simp: split_beta create_alt_def) apply (rule nrewrite.fun) apply fact apply rule apply fact apply (subst term_to_nterm_alpha_equiv[of "[]" "[]", simplified]) using \closed u\ \wellformed u\ by auto next case (arg u u' t) hence "closed t" "closed u" "closed (t $ u)" and wellform:"wellformed t" "wellformed u" by fastforce+ from arg obtain t' where "compile \\<^sub>n fst (run_state (term_to_nterm [] u) (snd (run_state (term_to_nterm [] t) s))) \ t'" "t' \\<^sub>\ fst (run_state (term_to_nterm [] u') (snd (run_state (term_to_nterm [] t) s')))" by force show ?case apply (rule exI) apply (auto simp: split_beta create_alt_def) apply rule apply fact apply rule apply (subst term_to_nterm_alpha_equiv[of "[]" "[]", simplified]) using \closed t\ \wellformed t\ apply force+ by fact qed lemma compile_complete: assumes "rs \ t \ t'" "closed t" "wellformed t" obtains u' where "compile \\<^sub>n term_to_nterm' t \ u'" "u' \\<^sub>\ term_to_nterm' t'" unfolding term_to_nterm'_def using assms by (metis compile_complete0) end end subsection \Splitting into constants\ type_synonym crules = "(term list \ nterm) fset" type_synonym crule_set = "(name \ crules) fset" abbreviation arity_compatibles :: "(term list \ 'a) fset \ bool" where "arity_compatibles \ fpairwise (\(pats\<^sub>1, _) (pats\<^sub>2, _). length pats\<^sub>1 = length pats\<^sub>2)" lemma arity_compatible_length: assumes "arity_compatibles rs" "(pats, rhs) |\| rs" shows "length pats = arity rs" proof - have "fBall rs (\(pats\<^sub>1, _). fBall rs (\(pats\<^sub>2, _). length pats\<^sub>1 = length pats\<^sub>2))" using assms unfolding fpairwise_alt_def by blast hence "fBall rs (\x. fBall rs (\y. (length \ fst) x = (length \ fst) y))" by force hence "(length \ fst) (pats, rhs) = arity rs" using assms(2) unfolding arity_def fthe_elem'_eq by (rule singleton_fset_holds) thus ?thesis by simp qed locale pre_crules = constants C_info "fst |`| rs" for C_info and rs :: "crule_set" locale crules = pre_crules + assumes fmap: "is_fmap rs" assumes nonempty: "rs \ {||}" assumes inner: "fBall rs (\(_, crs). arity_compatibles crs \ is_fmap crs \ patterns_compatibles crs \ crs \ {||} \ fBall crs (\(pats, rhs). linears pats \ pats \ [] \ fdisjnt (freess pats) all_consts \ \ shadows_consts rhs \ frees rhs |\| freess pats \ welldefined rhs))" lemma (in pre_crules) crulesI: assumes "\name crs. (name, crs) |\| rs \ arity_compatibles crs" assumes "\name crs. (name, crs) |\| rs \ is_fmap crs" assumes "\name crs. (name, crs) |\| rs \ patterns_compatibles crs" assumes "\name crs. (name, crs) |\| rs \ crs \ {||}" assumes "\name crs pats rhs. (name, crs) |\| rs \ (pats, rhs) |\| crs \ linears pats" assumes "\name crs pats rhs. (name, crs) |\| rs \ (pats, rhs) |\| crs \ pats \ []" assumes "\name crs pats rhs. (name, crs) |\| rs \ (pats, rhs) |\| crs \ fdisjnt (freess pats) all_consts" assumes "\name crs pats rhs. (name, crs) |\| rs \ (pats, rhs) |\| crs \ \ shadows_consts rhs" assumes "\name crs pats rhs. (name, crs) |\| rs \ (pats, rhs) |\| crs \ frees rhs |\| freess pats" assumes "\name crs pats rhs. (name, crs) |\| rs \ (pats, rhs) |\| crs \ welldefined rhs" assumes "is_fmap rs" "rs \ {||}" shows "crules C_info rs" -using assms unfolding crules_axioms_def crules_def -by (auto simp: prod_fBallI intro: pre_crules_axioms) +proof unfold_locales + show "is_fmap rs" + using assms(11) . +next + show "rs \ {||}" + using assms(12) . +next + show "fBall rs (\(_, crs). Rewriting_Nterm.arity_compatibles crs \ is_fmap crs \ + patterns_compatibles crs \ crs \ {||} \ + fBall crs (\(pats, rhs). linears pats \ pats \ [] \ fdisjnt (freess pats) all_consts \ + \ shadows_consts rhs \ frees rhs |\| freess pats \ consts rhs |\| all_consts))" + using assms(1-10) + by (intro prod_BallI conjI) metis+ +qed lemmas crulesI[intro!] = pre_crules.crulesI[unfolded pre_crules_def] definition "consts_of" :: "nrule fset \ crule_set" where "consts_of = fgroup_by split_rule" lemma consts_of_heads: "fst |`| consts_of rs = heads_of rs" unfolding consts_of_def by (simp add: split_rule_fst comp_def) lemma (in nrules) consts_rules: "crules C_info (consts_of rs)" proof have "is_fmap rs" using fmap by simp thus "is_fmap (consts_of rs)" unfolding consts_of_def by auto show "consts_of rs \ {||}" using nonempty unfolding consts_of_def by (meson fgroup_by_nonempty) show "constants C_info (fst |`| consts_of rs)" proof show "fdisjnt (fst |`| consts_of rs) C" using disjnt by (auto simp: consts_of_heads) next show "distinct all_constructors" by (fact distinct_ctr) qed fix name crs assume crs: "(name, crs) |\| consts_of rs" thus "crs \ {||}" unfolding consts_of_def by (meson femptyE fgroup_by_nonempty_inner) show "arity_compatibles crs" "patterns_compatibles crs" proof safe fix pats\<^sub>1 rhs\<^sub>1 pats\<^sub>2 rhs\<^sub>2 assume "(pats\<^sub>1, rhs\<^sub>1) |\| crs" "(pats\<^sub>2, rhs\<^sub>2) |\| crs" with crs obtain lhs\<^sub>1 lhs\<^sub>2 where rs: "(lhs\<^sub>1, rhs\<^sub>1) |\| rs" "(lhs\<^sub>2, rhs\<^sub>2) |\| rs" and split: "split_rule (lhs\<^sub>1, rhs\<^sub>1) = (name, (pats\<^sub>1, rhs\<^sub>1))" "split_rule (lhs\<^sub>2, rhs\<^sub>2) = (name, (pats\<^sub>2, rhs\<^sub>2))" unfolding consts_of_def by (force simp: split_beta) hence arity: "arity_compatible lhs\<^sub>1 lhs\<^sub>2" using arity by (force dest: fpairwiseD) from rs have const: "is_const (fst (strip_comb lhs\<^sub>1))" "is_const (fst (strip_comb lhs\<^sub>2))" using all_rules by force+ have "name = const_name (fst (strip_comb lhs\<^sub>1))" "name = const_name (fst (strip_comb lhs\<^sub>2))" using split by (auto simp: split_beta) with const have "fst (strip_comb lhs\<^sub>1) = Const name" "fst (strip_comb lhs\<^sub>2) = Const name" apply (fold const_term_def) subgoal by simp subgoal by fastforce done hence fst: "fst (strip_comb lhs\<^sub>1) = fst (strip_comb lhs\<^sub>2)" by simp with arity have "length (snd (strip_comb lhs\<^sub>1)) = length (snd (strip_comb lhs\<^sub>2))" unfolding arity_compatible_def by (simp add: split_beta) with split show "length pats\<^sub>1 = length pats\<^sub>2" by (auto simp: split_beta) have "pattern_compatible lhs\<^sub>1 lhs\<^sub>2" using rs patterns by (auto dest: fpairwiseD) moreover have "lhs\<^sub>1 = name $$ pats\<^sub>1" using split(1) const(1) by (auto simp: split_beta) moreover have "lhs\<^sub>2 = name $$ pats\<^sub>2" using split(2) const(2) by (auto simp: split_beta) ultimately have "pattern_compatible (name $$ pats\<^sub>1) (name $$ pats\<^sub>2)" by simp thus "patterns_compatible pats\<^sub>1 pats\<^sub>2" using \length pats\<^sub>1 = _\ by (auto dest: pattern_compatible_combD) qed show "is_fmap crs" proof fix pats rhs\<^sub>1 rhs\<^sub>2 assume "(pats, rhs\<^sub>1) |\| crs" "(pats, rhs\<^sub>2) |\| crs" with crs obtain lhs\<^sub>1 lhs\<^sub>2 where rs: "(lhs\<^sub>1, rhs\<^sub>1) |\| rs" "(lhs\<^sub>2, rhs\<^sub>2) |\| rs" and split: "split_rule (lhs\<^sub>1, rhs\<^sub>1) = (name, (pats, rhs\<^sub>1))" "split_rule (lhs\<^sub>2, rhs\<^sub>2) = (name, (pats, rhs\<^sub>2))" unfolding consts_of_def by (force simp: split_beta) have "lhs\<^sub>1 = lhs\<^sub>2" proof (rule ccontr) assume "lhs\<^sub>1 \ lhs\<^sub>2" then consider (fst) "fst (strip_comb lhs\<^sub>1) \ fst (strip_comb lhs\<^sub>2)" | (snd) "snd (strip_comb lhs\<^sub>1) \ snd (strip_comb lhs\<^sub>2)" by (metis list_strip_comb) thus False proof cases case fst moreover have "is_const (fst (strip_comb lhs\<^sub>1))" "is_const (fst (strip_comb lhs\<^sub>2))" using rs all_rules by force+ ultimately show ?thesis using split const_name_simps by (fastforce simp: split_beta) next case snd with split show ?thesis by (auto simp: split_beta) qed qed with rs show "rhs\<^sub>1 = rhs\<^sub>2" using \is_fmap rs\ by (auto dest: is_fmapD) qed fix pats rhs assume "(pats, rhs) |\| crs" then obtain lhs where "(lhs, rhs) |\| rs" "pats = snd (strip_comb lhs)" using crs unfolding consts_of_def by (force simp: split_beta) hence "nrule (lhs, rhs)" using all_rules by blast hence "linear lhs" "frees rhs |\| frees lhs" by auto thus "linears pats" unfolding \pats = _\ by (intro linears_strip_comb) have "\ is_const lhs" "is_const (fst (strip_comb lhs))" using \nrule _\ by auto thus "pats \ []" unfolding \pats = _\ using \linear lhs\ apply (cases lhs) apply (fold app_term_def) by (auto split: prod.splits) from \nrule (lhs, rhs)\ have "frees (fst (strip_comb lhs)) = {||}" by (cases "fst (strip_comb lhs)") (auto simp: is_const_def) hence "frees lhs = freess (snd (strip_comb lhs))" by (subst frees_strip_comb) auto thus "frees rhs |\| freess pats" unfolding \pats = _\ using \frees rhs |\| frees lhs\ by simp have "\ shadows_consts rhs" using \(lhs, rhs) |\| rs\ not_shadows by force thus "\ pre_constants.shadows_consts C_info (fst |`| consts_of rs) rhs" by (simp add: consts_of_heads) have "fdisjnt all_consts (frees lhs)" using \(lhs, rhs) |\| rs\ not_shadows by (force simp: shadows_consts_def) moreover have "freess pats |\| frees lhs" unfolding \pats = _\ \frees lhs = _\ by simp ultimately have "fdisjnt (freess pats) all_consts" by (metis fdisjnt_subset_right fdisjnt_swap) thus "fdisjnt (freess pats) (pre_constants.all_consts C_info (fst |`| consts_of rs))" by (simp add: consts_of_heads) show "pre_constants.welldefined C_info (fst |`| consts_of rs) rhs" using welldefined_rs \(lhs, rhs) |\| rs\ by (force simp: consts_of_heads) qed sublocale nrules \ nrules_as_crules?: crules C_info "consts_of rs" by (fact consts_rules) subsection \Computability\ export_code translate_rule consts_of arity nterm_to_term checking Scala -end \ No newline at end of file +end + \ No newline at end of file diff --git a/thys/CakeML_Codegen/Rewriting/Rewriting_Pterm_Elim.thy b/thys/CakeML_Codegen/Rewriting/Rewriting_Pterm_Elim.thy --- a/thys/CakeML_Codegen/Rewriting/Rewriting_Pterm_Elim.thy +++ b/thys/CakeML_Codegen/Rewriting/Rewriting_Pterm_Elim.thy @@ -1,1733 +1,1750 @@ section \Higher-order term rewriting with explicit pattern matching\ theory Rewriting_Pterm_Elim imports Rewriting_Nterm "../Terms/Pterm" begin subsection \Intermediate rule sets\ type_synonym irules = "(term list \ pterm) fset" type_synonym irule_set = "(name \ irules) fset" locale pre_irules = constants C_info "fst |`| rs" for C_info and rs :: "irule_set" locale irules = pre_irules + assumes fmap: "is_fmap rs" assumes nonempty: "rs \ {||}" assumes inner: "fBall rs (\(_, irs). arity_compatibles irs \ is_fmap irs \ patterns_compatibles irs \ irs \ {||} \ fBall irs (\(pats, rhs). linears pats \ abs_ish pats rhs \ closed_except rhs (freess pats) \ fdisjnt (freess pats) all_consts \ wellformed rhs \ \ shadows_consts rhs \ welldefined rhs))" lemma (in pre_irules) irulesI: assumes "\name irs. (name, irs) |\| rs \ arity_compatibles irs" assumes "\name irs. (name, irs) |\| rs \ is_fmap irs" assumes "\name irs. (name, irs) |\| rs \ patterns_compatibles irs" assumes "\name irs. (name, irs) |\| rs \ irs \ {||}" assumes "\name irs pats rhs. (name, irs) |\| rs \ (pats, rhs) |\| irs \ linears pats" assumes "\name irs pats rhs. (name, irs) |\| rs \ (pats, rhs) |\| irs \ abs_ish pats rhs" assumes "\name irs pats rhs. (name, irs) |\| rs \ (pats, rhs) |\| irs \ fdisjnt (freess pats) all_consts" assumes "\name irs pats rhs. (name, irs) |\| rs \ (pats, rhs) |\| irs \ closed_except rhs (freess pats)" assumes "\name irs pats rhs. (name, irs) |\| rs \ (pats, rhs) |\| irs \ wellformed rhs" assumes "\name irs pats rhs. (name, irs) |\| rs \ (pats, rhs) |\| irs \ \ shadows_consts rhs" assumes "\name irs pats rhs. (name, irs) |\| rs \ (pats, rhs) |\| irs \ welldefined rhs" assumes "is_fmap rs" "rs \ {||}" shows "irules C_info rs" -using assms unfolding irules_axioms_def irules_def -by (auto simp: prod_fBallI intro: pre_irules_axioms) +proof unfold_locales + show "is_fmap rs" + using assms(12) . +next + show "rs \ {||}" + using assms(13) . +next + show "fBall rs (\(_, irs). Rewriting_Nterm.arity_compatibles irs \ is_fmap irs \ + patterns_compatibles irs \ irs \ {||} \ + fBall irs (\(pats, rhs). linears pats \ abs_ish pats rhs \ closed_except rhs (freess pats) \ + fdisjnt (freess pats) all_consts \ pre_strong_term_class.wellformed rhs \ + \ shadows_consts rhs \ consts rhs |\| all_consts))" + using assms(1-11) + by (intro prod_BallI conjI) metis+ +qed lemmas irulesI[intro!] = pre_irules.irulesI[unfolded pre_irules_def] subsubsection \Translation from @{typ nterm} to @{typ pterm}\ fun nterm_to_pterm :: "nterm \ pterm" where "nterm_to_pterm (Nvar s) = Pvar s" | "nterm_to_pterm (Nconst s) = Pconst s" | "nterm_to_pterm (t\<^sub>1 $\<^sub>n t\<^sub>2) = nterm_to_pterm t\<^sub>1 $\<^sub>p nterm_to_pterm t\<^sub>2" | "nterm_to_pterm (\\<^sub>n x. t) = (\\<^sub>p x. nterm_to_pterm t)" lemma nterm_to_pterm_inj: "nterm_to_pterm x = nterm_to_pterm y \ x = y" by (induction y arbitrary: x) (auto elim: nterm_to_pterm.elims) lemma nterm_to_pterm: assumes "no_abs t" shows "nterm_to_pterm t = convert_term t" using assms apply induction apply auto by (auto simp: free_nterm_def free_pterm_def const_nterm_def const_pterm_def app_nterm_def app_pterm_def) lemma nterm_to_pterm_frees[simp]: "frees (nterm_to_pterm t) = frees t" by (induct t) auto lemma closed_nterm_to_pterm[intro]: "closed_except (nterm_to_pterm t) (frees t)" unfolding closed_except_def by simp lemma (in constants) shadows_nterm_to_pterm[simp]: "shadows_consts (nterm_to_pterm t) = shadows_consts t" by (induct t) (auto simp: shadows_consts_def fdisjnt_alt_def) lemma wellformed_nterm_to_pterm[intro]: "wellformed (nterm_to_pterm t)" by (induct t) auto lemma consts_nterm_to_pterm[simp]: "consts (nterm_to_pterm t) = consts t" by (induct t) auto subsubsection \Translation from @{typ crule_set} to @{typ irule_set}\ definition translate_crules :: "crules \ irules" where "translate_crules = fimage (map_prod id nterm_to_pterm)" definition compile :: "crule_set \ irule_set" where "compile = fimage (map_prod id translate_crules)" lemma compile_heads: "fst |`| compile rs = fst |`| rs" unfolding compile_def by simp lemma (in crules) compile_rules: "irules C_info (compile rs)" proof have "is_fmap rs" using fmap by simp thus "is_fmap (compile rs)" unfolding compile_def map_prod_def id_apply by (rule is_fmap_image) show "compile rs \ {||}" using nonempty unfolding compile_def by auto show "constants C_info (fst |`| compile rs)" proof show "fdisjnt (fst |`| compile rs) C" using disjnt unfolding compile_def by force next show "distinct all_constructors" by (fact distinct_ctr) qed fix name irs assume irs: "(name, irs) |\| compile rs" then obtain irs' where "(name, irs') |\| rs" "irs = translate_crules irs'" unfolding compile_def by force hence "arity_compatibles irs'" - using inner by (blast dest: fpairwiseD) + using inner + by (metis (no_types, lifting) case_prodD) thus "arity_compatibles irs" unfolding \irs = translate_crules irs'\ translate_crules_def by (force dest: fpairwiseD) have "patterns_compatibles irs'" using \(name, irs') |\| rs\ inner by (blast dest: fpairwiseD) thus "patterns_compatibles irs" unfolding \irs = _\ translate_crules_def by (auto dest: fpairwiseD) have "is_fmap irs'" using \(name, irs') |\| rs\ inner by auto thus "is_fmap irs" unfolding \irs = translate_crules irs'\ translate_crules_def map_prod_def id_apply by (rule is_fmap_image) have "irs' \ {||}" using \(name, irs') |\| rs\ inner by auto thus "irs \ {||}" unfolding \irs = translate_crules irs'\ translate_crules_def by simp fix pats rhs assume "(pats, rhs) |\| irs" then obtain rhs' where "(pats, rhs') |\| irs'" "rhs = nterm_to_pterm rhs'" unfolding \irs = translate_crules irs'\ translate_crules_def by force hence "linears pats" "pats \ []" "frees rhs' |\| freess pats" "\ shadows_consts rhs'" using fbspec[OF inner \(name, irs') |\| rs\] by blast+ show "linears pats" by fact show "closed_except rhs (freess pats)" unfolding \rhs = nterm_to_pterm rhs'\ using \frees rhs' |\| freess pats\ by (metis dual_order.trans closed_nterm_to_pterm closed_except_def) show "wellformed rhs" unfolding \rhs = nterm_to_pterm rhs'\ by auto have "fdisjnt (freess pats) all_consts" using \(pats, rhs') |\| irs'\ \(name, irs') |\| rs\ inner - by blast + by auto thus "fdisjnt (freess pats) (pre_constants.all_consts C_info (fst |`| compile rs))" unfolding compile_def by simp have "\ shadows_consts rhs" unfolding \rhs = _\ using \\ shadows_consts _\ by simp thus "\ pre_constants.shadows_consts C_info (fst |`| compile rs) rhs" unfolding compile_heads . show "abs_ish pats rhs" using \pats \ []\ unfolding abs_ish_def by simp have "welldefined rhs'" using fbspec[OF inner \(name, irs') |\| rs\, simplified] using \(pats, rhs') |\| irs'\ by blast thus "pre_constants.welldefined C_info (fst |`| compile rs) rhs" unfolding compile_def \rhs = _\ by simp qed sublocale crules \ crules_as_irules: irules C_info "compile rs" by (fact compile_rules) subsubsection \Transformation of @{typ irule_set}\ definition transform_irules :: "irules \ irules" where "transform_irules rs = ( if arity rs = 0 then rs else map_prod id Pabs |`| fgroup_by (\(pats, rhs). (butlast pats, (last pats, rhs))) rs)" lemma arity_compatibles_transform_irules: assumes "arity_compatibles rs" shows "arity_compatibles (transform_irules rs)" proof (cases "arity rs = 0") case True thus ?thesis unfolding transform_irules_def using assms by simp next case False let ?rs' = "transform_irules rs" let ?f = "\(pats, rhs). (butlast pats, (last pats, rhs))" let ?grp = "fgroup_by ?f rs" have rs': "?rs' = map_prod id Pabs |`| ?grp" using False unfolding transform_irules_def by simp show ?thesis proof safe fix pats\<^sub>1 rhs\<^sub>1 pats\<^sub>2 rhs\<^sub>2 assume "(pats\<^sub>1, rhs\<^sub>1) |\| ?rs'" "(pats\<^sub>2, rhs\<^sub>2) |\| ?rs'" then obtain rhs\<^sub>1' rhs\<^sub>2' where "(pats\<^sub>1, rhs\<^sub>1') |\| ?grp" "(pats\<^sub>2, rhs\<^sub>2') |\| ?grp" unfolding rs' by auto then obtain pats\<^sub>1' pats\<^sub>2' x y \ \dummies\ where "fst (?f (pats\<^sub>1', x)) = pats\<^sub>1" "(pats\<^sub>1', x) |\| rs" and "fst (?f (pats\<^sub>2', y)) = pats\<^sub>2" "(pats\<^sub>2', y) |\| rs" by (fastforce simp: split_beta elim: fgroup_byE2) hence "pats\<^sub>1 = butlast pats\<^sub>1'" "pats\<^sub>2 = butlast pats\<^sub>2'" "length pats\<^sub>1' = length pats\<^sub>2'" using assms by (force dest: fpairwiseD)+ thus "length pats\<^sub>1 = length pats\<^sub>2" by auto qed qed lemma arity_transform_irules: assumes "arity_compatibles rs" "rs \ {||}" shows "arity (transform_irules rs) = (if arity rs = 0 then 0 else arity rs - 1)" proof (cases "arity rs = 0") case True thus ?thesis unfolding transform_irules_def by simp next case False let ?f = "\(pats, rhs). (butlast pats, (last pats, rhs))" let ?grp = "fgroup_by ?f rs" let ?rs' = "map_prod id Pabs |`| ?grp" have "arity ?rs' = arity rs - 1" proof (rule arityI) show "fBall ?rs' (\(pats, _). length pats = arity rs - 1)" proof (rule prod_fBallI) fix pats rhs assume "(pats, rhs) |\| ?rs'" then obtain cs where "(pats, cs) |\| ?grp" "rhs = Pabs cs" by force then obtain pats' x \ \dummy\ where "pats = butlast pats'" "(pats', x) |\| rs" by (fastforce simp: split_beta elim: fgroup_byE2) hence "length pats' = arity rs" using assms by (metis arity_compatible_length) thus "length pats = arity rs - 1" unfolding \pats = butlast pats'\ using False by simp qed next show "?rs' \ {||}" using assms by (simp add: fgroup_by_nonempty) qed with False show ?thesis unfolding transform_irules_def by simp qed definition transform_irule_set :: "irule_set \ irule_set" where "transform_irule_set = fimage (map_prod id transform_irules)" lemma transform_irule_set_heads: "fst |`| transform_irule_set rs = fst |`| rs" unfolding transform_irule_set_def by simp lemma (in irules) rules_transform: "irules C_info (transform_irule_set rs)" proof have "is_fmap rs" using fmap by simp thus "is_fmap (transform_irule_set rs)" unfolding transform_irule_set_def map_prod_def id_apply by (rule is_fmap_image) show "transform_irule_set rs \ {||}" using nonempty unfolding transform_irule_set_def by auto show "constants C_info (fst |`| transform_irule_set rs)" proof show "fdisjnt (fst |`| transform_irule_set rs) C" using disjnt unfolding transform_irule_set_def by force next show "distinct all_constructors" by (fact distinct_ctr) qed fix name irs assume irs: "(name, irs) |\| transform_irule_set rs" then obtain irs' where "(name, irs') |\| rs" "irs = transform_irules irs'" unfolding transform_irule_set_def by force hence "arity_compatibles irs'" - using inner by (blast dest: fpairwiseD) + using inner + by (metis (no_types, lifting) case_prodD) thus "arity_compatibles irs" unfolding \irs = transform_irules irs'\ by (rule arity_compatibles_transform_irules) have "irs' \ {||}" using \(name, irs') |\| rs\ inner by blast thus "irs \ {||}" unfolding \irs = transform_irules irs'\ transform_irules_def by (simp add: fgroup_by_nonempty) let ?f = "\(pats, rhs). (butlast pats, (last pats, rhs))" let ?grp = "fgroup_by ?f irs'" have "patterns_compatibles irs'" using \(name, irs') |\| rs\ inner by (blast dest: fpairwiseD) show "patterns_compatibles irs" proof (cases "arity irs' = 0") case True thus ?thesis unfolding \irs = transform_irules irs'\ transform_irules_def using \patterns_compatibles irs'\ by simp next case False hence irs': "irs = map_prod id Pabs |`| ?grp" unfolding \irs = transform_irules irs'\ transform_irules_def by simp show ?thesis proof safe fix pats\<^sub>1 rhs\<^sub>1 pats\<^sub>2 rhs\<^sub>2 assume "(pats\<^sub>1, rhs\<^sub>1) |\| irs" "(pats\<^sub>2, rhs\<^sub>2) |\| irs" with irs' obtain cs\<^sub>1 cs\<^sub>2 where "(pats\<^sub>1, cs\<^sub>1) |\| ?grp" "(pats\<^sub>2, cs\<^sub>2) |\| ?grp" by force then obtain pats\<^sub>1' pats\<^sub>2' and x y \ \dummies\ where "(pats\<^sub>1', x) |\| irs'" "(pats\<^sub>2', y) |\| irs'" and "pats\<^sub>1 = butlast pats\<^sub>1'" "pats\<^sub>2 = butlast pats\<^sub>2'" unfolding irs' by (fastforce elim: fgroup_byE2) hence "patterns_compatible pats\<^sub>1' pats\<^sub>2'" using \patterns_compatibles irs'\ by (auto dest: fpairwiseD) thus "patterns_compatible pats\<^sub>1 pats\<^sub>2" unfolding \pats\<^sub>1 = _\ \pats\<^sub>2 = _\ by auto qed qed have "is_fmap irs'" using \(name, irs') |\| rs\ inner by blast show "is_fmap irs" proof (cases "arity irs' = 0") case True thus ?thesis unfolding \irs = transform_irules irs'\ transform_irules_def using \is_fmap irs'\ by simp next case False hence irs': "irs = map_prod id Pabs |`| ?grp" unfolding \irs = transform_irules irs'\ transform_irules_def by simp show ?thesis proof fix pats rhs\<^sub>1 rhs\<^sub>2 assume "(pats, rhs\<^sub>1) |\| irs" "(pats, rhs\<^sub>2) |\| irs" with irs' obtain cs\<^sub>1 cs\<^sub>2 where "(pats, cs\<^sub>1) |\| ?grp" "rhs\<^sub>1 = Pabs cs\<^sub>1" and "(pats, cs\<^sub>2) |\| ?grp" "rhs\<^sub>2 = Pabs cs\<^sub>2" by force moreover have "is_fmap ?grp" by auto ultimately show "rhs\<^sub>1 = rhs\<^sub>2" by (auto dest: is_fmapD) qed qed fix pats rhs assume "(pats, rhs) |\| irs" show "linears pats" proof (cases "arity irs' = 0") case True thus ?thesis using \(pats, rhs) |\| irs\ \(name, irs') |\| rs\ inner unfolding \irs = transform_irules irs'\ transform_irules_def - by (smt fBallE split_conv) + by auto next case False hence irs': "irs = map_prod id Pabs |`| ?grp" unfolding \irs = transform_irules irs'\ transform_irules_def by simp then obtain cs where "(pats, cs) |\| ?grp" using \(pats, rhs) |\| irs\ by force then obtain pats' x \ \dummy\ where "fst (?f (pats', x)) = pats" "(pats', x) |\| irs'" by (fastforce simp: split_beta elim: fgroup_byE2) hence "pats = butlast pats'" by simp moreover have "linears pats'" using \(pats', x) |\| irs'\ \(name, irs') |\| rs\ inner - by blast + by auto ultimately show ?thesis by auto qed have "fdisjnt (freess pats) all_consts" proof (cases "arity irs' = 0") case True thus ?thesis using \(pats, rhs) |\| irs\ \(name, irs') |\| rs\ inner unfolding \irs = transform_irules irs'\ transform_irules_def - by (smt fBallE split_conv) + by auto next case False hence irs': "irs = map_prod id Pabs |`| ?grp" unfolding \irs = transform_irules irs'\ transform_irules_def by simp then obtain cs where "(pats, cs) |\| ?grp" using \(pats, rhs) |\| irs\ by force then obtain pats' x \ \dummy\ where "fst (?f (pats', x)) = pats" "(pats', x) |\| irs'" by (fastforce simp: split_beta elim: fgroup_byE2) hence "pats = butlast pats'" by simp moreover have "fdisjnt (freess pats') all_consts" - using \(pats', x) |\| irs'\ \(name, irs') |\| rs\ inner by blast + using \(pats', x) |\| irs'\ \(name, irs') |\| rs\ inner by auto ultimately show ?thesis by (metis subsetI in_set_butlastD freess_subset fdisjnt_subset_left) qed thus "fdisjnt (freess pats) (pre_constants.all_consts C_info (fst |`| transform_irule_set rs))" unfolding transform_irule_set_def by simp show "closed_except rhs (freess pats)" proof (cases "arity irs' = 0") case True thus ?thesis using \(pats, rhs) |\| irs\ \(name, irs') |\| rs\ inner unfolding \irs = transform_irules irs'\ transform_irules_def - by (smt fBallE split_conv) + by auto next case False hence irs': "irs = map_prod id Pabs |`| ?grp" unfolding \irs = transform_irules irs'\ transform_irules_def by simp then obtain cs where "(pats, cs) |\| ?grp" "rhs = Pabs cs" using \(pats, rhs) |\| irs\ by force show ?thesis unfolding \rhs = Pabs cs\ closed_except_simps proof safe fix pat t assume "(pat, t) |\| cs" then obtain pats' where "(pats', t) |\| irs'" "?f (pats', t) = (pats, (pat, t))" using \(pats, cs) |\| ?grp\ by auto hence "closed_except t (freess pats')" - using \(name, irs') |\| rs\ inner by blast + using \(name, irs') |\| rs\ inner by auto have "pats' \ []" using \arity_compatibles irs'\ \(pats', t) |\| irs'\ False by (metis list.size(3) arity_compatible_length) hence "pats' = pats @ [pat]" using \?f (pats', t) = (pats, (pat, t))\ by (fastforce simp: split_beta snoc_eq_iff_butlast) hence "freess pats |\| frees pat = freess pats'" unfolding freess_def by auto thus "closed_except t (freess pats |\| frees pat)" using \closed_except t (freess pats')\ by simp qed qed show "wellformed rhs" proof (cases "arity irs' = 0") case True thus ?thesis using \(pats, rhs) |\| irs\ \(name, irs') |\| rs\ inner unfolding \irs = transform_irules irs'\ transform_irules_def - by (smt fBallE split_conv) + by auto next case False hence irs': "irs = map_prod id Pabs |`| ?grp" unfolding \irs = transform_irules irs'\ transform_irules_def by simp then obtain cs where "(pats, cs) |\| ?grp" "rhs = Pabs cs" using \(pats, rhs) |\| irs\ by force show ?thesis unfolding \rhs = Pabs cs\ proof (rule wellformed_PabsI) show "cs \ {||}" using \(pats, cs) |\| ?grp\ \irs' \ {||}\ by (meson femptyE fgroup_by_nonempty_inner) next show "is_fmap cs" proof fix pat t\<^sub>1 t\<^sub>2 assume "(pat, t\<^sub>1) |\| cs" "(pat, t\<^sub>2) |\| cs" then obtain pats\<^sub>1' pats\<^sub>2' where "(pats\<^sub>1', t\<^sub>1) |\| irs'" "?f (pats\<^sub>1', t\<^sub>1) = (pats, (pat, t\<^sub>1))" and "(pats\<^sub>2', t\<^sub>2) |\| irs'" "?f (pats\<^sub>2', t\<^sub>2) = (pats, (pat, t\<^sub>2))" using \(pats, cs) |\| ?grp\ by force moreover hence "pats\<^sub>1' \ []" "pats\<^sub>2' \ []" using \arity_compatibles irs'\ False unfolding prod.case by (metis list.size(3) arity_compatible_length)+ ultimately have "pats\<^sub>1' = pats @ [pat]" "pats\<^sub>2' = pats @ [pat]" unfolding split_beta fst_conv snd_conv by (metis prod.inject snoc_eq_iff_butlast)+ with \is_fmap irs'\ show "t\<^sub>1 = t\<^sub>2" using \(pats\<^sub>1', t\<^sub>1) |\| irs'\ \(pats\<^sub>2', t\<^sub>2) |\| irs'\ by (blast dest: is_fmapD) qed next show "pattern_compatibles cs" proof safe fix pat\<^sub>1 rhs\<^sub>1 pat\<^sub>2 rhs\<^sub>2 assume "(pat\<^sub>1, rhs\<^sub>1) |\| cs" "(pat\<^sub>2, rhs\<^sub>2) |\| cs" then obtain pats\<^sub>1' pats\<^sub>2' where "(pats\<^sub>1', rhs\<^sub>1) |\| irs'" "?f (pats\<^sub>1', rhs\<^sub>1) = (pats, (pat\<^sub>1, rhs\<^sub>1))" and "(pats\<^sub>2', rhs\<^sub>2) |\| irs'" "?f (pats\<^sub>2', rhs\<^sub>2) = (pats, (pat\<^sub>2, rhs\<^sub>2))" using \(pats, cs) |\| ?grp\ by force moreover hence "pats\<^sub>1' \ []" "pats\<^sub>2' \ []" using \arity_compatibles irs'\ False unfolding prod.case by (metis list.size(3) arity_compatible_length)+ ultimately have "pats\<^sub>1' = pats @ [pat\<^sub>1]" "pats\<^sub>2' = pats @ [pat\<^sub>2]" unfolding split_beta fst_conv snd_conv by (metis prod.inject snoc_eq_iff_butlast)+ moreover have "patterns_compatible pats\<^sub>1' pats\<^sub>2'" using \(pats\<^sub>1', rhs\<^sub>1) |\| irs'\ \(pats\<^sub>2', rhs\<^sub>2) |\| irs'\ \patterns_compatibles irs'\ by (auto dest: fpairwiseD) ultimately show "pattern_compatible pat\<^sub>1 pat\<^sub>2" by (auto elim: rev_accum_rel_snoc_eqE) qed next fix pat t assume "(pat, t) |\| cs" then obtain pats' where "(pats', t) |\| irs'" "pat = last pats'" using \(pats, cs) |\| ?grp\ by auto moreover hence "pats' \ []" using \arity_compatibles irs'\ False by (metis list.size(3) arity_compatible_length) ultimately have "pat \ set pats'" by auto moreover have "linears pats'" - using \(pats', t) |\| irs'\ \(name, irs') |\| rs\ inner by blast + using \(pats', t) |\| irs'\ \(name, irs') |\| rs\ inner by auto ultimately show "linear pat" by (metis linears_linear) show "wellformed t" - using \(pats', t) |\| irs'\ \(name, irs') |\| rs\ inner by blast + using \(pats', t) |\| irs'\ \(name, irs') |\| rs\ inner by auto qed qed have "\ shadows_consts rhs" proof (cases "arity irs' = 0") case True thus ?thesis using \(pats, rhs) |\| irs\ \(name, irs') |\| rs\ inner unfolding \irs = transform_irules irs'\ transform_irules_def - by (smt fBallE split_conv) + by auto next case False hence irs': "irs = map_prod id Pabs |`| ?grp" unfolding \irs = transform_irules irs'\ transform_irules_def by simp then obtain cs where "(pats, cs) |\| ?grp" "rhs = Pabs cs" using \(pats, rhs) |\| irs\ by force show ?thesis unfolding \rhs = _\ proof assume "shadows_consts (Pabs cs)" then obtain pat t where "(pat, t) |\| cs" "shadows_consts t \ shadows_consts pat" by force then obtain pats' where "(pats', t) |\| irs'" "pat = last pats'" using \(pats, cs) |\| ?grp\ by auto moreover hence "pats' \ []" using \arity_compatibles irs'\ False by (metis list.size(3) arity_compatible_length) ultimately have "pat \ set pats'" by auto show False using \shadows_consts t \ shadows_consts pat\ proof assume "shadows_consts t" thus False - using \(name, irs') |\| rs\ \(pats', t) |\| irs'\ inner by blast + using \(name, irs') |\| rs\ \(pats', t) |\| irs'\ inner by auto next assume "shadows_consts pat" have "fdisjnt (freess pats') all_consts" - using \(name, irs') |\| rs\ \(pats', t) |\| irs'\ inner by blast + using \(name, irs') |\| rs\ \(pats', t) |\| irs'\ inner by auto have "fdisjnt (frees pat) all_consts" apply (rule fdisjnt_subset_left) apply (subst freess_single[symmetric]) apply (rule freess_subset) apply simp apply fact+ done thus False using \shadows_consts pat\ unfolding shadows_consts_def fdisjnt_alt_def by auto qed qed qed thus "\ pre_constants.shadows_consts C_info (fst |`| transform_irule_set rs) rhs" by (simp add: transform_irule_set_heads) show "abs_ish pats rhs" proof (cases "arity irs' = 0") case True thus ?thesis using \(pats, rhs) |\| irs\ \(name, irs') |\| rs\ inner unfolding \irs = transform_irules irs'\ transform_irules_def - by (smt fBallE split_conv) + by auto next case False hence irs': "irs = map_prod id Pabs |`| ?grp" unfolding \irs = transform_irules irs'\ transform_irules_def by simp then obtain cs where "(pats, cs) |\| ?grp" "rhs = Pabs cs" using \(pats, rhs) |\| irs\ by force thus ?thesis unfolding abs_ish_def by (simp add: is_abs_def term_cases_def) qed have "welldefined rhs" proof (cases "arity irs' = 0") case True hence \(pats, rhs) |\| irs'\ using \(pats, rhs) |\| irs\ \(name, irs') |\| rs\ inner unfolding \irs = transform_irules irs'\ transform_irules_def by (smt fBallE split_conv) thus ?thesis unfolding transform_irule_set_def using fbspec[OF inner \(name, irs') |\| rs\, simplified] by force next case False hence irs': "irs = map_prod id Pabs |`| ?grp" unfolding \irs = transform_irules irs'\ transform_irules_def by simp then obtain cs where "(pats, cs) |\| ?grp" "rhs = Pabs cs" using \(pats, rhs) |\| irs\ by force show ?thesis unfolding \rhs = _\ apply simp apply (rule ffUnion_least) unfolding ball_simps apply rule apply (rename_tac x, case_tac x, hypsubst_thin) apply simp subgoal premises prems for pat t proof - from prems obtain pats' where "(pats', t) |\| irs'" using \(pats, cs) |\| ?grp\ by auto hence "welldefined t" using fbspec[OF inner \(name, irs') |\| rs\, simplified] by blast thus ?thesis unfolding transform_irule_set_def by simp qed done qed thus "pre_constants.welldefined C_info (fst |`| transform_irule_set rs) rhs" unfolding transform_irule_set_heads . qed subsubsection \Matching and rewriting\ definition irewrite_step :: "name \ term list \ pterm \ pterm \ pterm option" where "irewrite_step name pats rhs t = map_option (subst rhs) (match (name $$ pats) t)" abbreviation irewrite_step' :: "name \ term list \ pterm \ pterm \ pterm \ bool" ("_, _, _ \\<^sub>i/ _ \/ _" [50,0,50] 50) where "name, pats, rhs \\<^sub>i t \ u \ irewrite_step name pats rhs t = Some u" lemma irewrite_stepI: assumes "match (name $$ pats) t = Some env" "subst rhs env = u" shows "name, pats, rhs \\<^sub>i t \ u" using assms unfolding irewrite_step_def by simp inductive irewrite :: "irule_set \ pterm \ pterm \ bool" ("_/ \\<^sub>i/ _ \/ _" [50,0,50] 50) for irs where step: "\ (name, rs) |\| irs; (pats, rhs) |\| rs; name, pats, rhs \\<^sub>i t \ t' \ \ irs \\<^sub>i t \ t'" | beta: "\ c |\| cs; c \ t \ t' \ \ irs \\<^sub>i Pabs cs $\<^sub>p t \ t'" | "fun": "irs \\<^sub>i t \ t' \ irs \\<^sub>i t $\<^sub>p u \ t' $\<^sub>p u" | arg: "irs \\<^sub>i u \ u' \ irs \\<^sub>i t $\<^sub>p u \ t $\<^sub>p u'" global_interpretation irewrite: rewriting "irewrite rs" for rs by standard (auto intro: irewrite.intros simp: app_pterm_def)+ abbreviation irewrite_rt :: "irule_set \ pterm \ pterm \ bool" ("_/ \\<^sub>i/ _ \*/ _" [50,0,50] 50) where "irewrite_rt rs \ (irewrite rs)\<^sup>*\<^sup>*" lemma (in irules) irewrite_closed: assumes "rs \\<^sub>i t \ u" "closed t" shows "closed u" using assms proof induction case (step name rs pats rhs t t') then obtain env where "match (name $$ pats) t = Some env" "t' = subst rhs env" unfolding irewrite_step_def by auto hence "closed_env env" using step by (auto intro: closed.match) show ?case unfolding \t' = _\ apply (subst closed_except_def) apply (subst subst_frees) apply fact apply (subst match_dom) apply fact apply (subst frees_list_comb) apply simp apply (subst closed_except_def[symmetric]) using inner step by blast next case (beta c cs t t') then obtain pat rhs where "c = (pat, rhs)" by (cases c) auto with beta obtain env where "match pat t = Some env" "t' = subst rhs env" by auto moreover have "closed t" using beta unfolding closed_except_def by simp ultimately have "closed_env env" using beta by (auto intro: closed.match) show ?case unfolding \t' = subst rhs env\ apply (subst closed_except_def) apply (subst subst_frees) apply fact apply (subst match_dom) apply fact apply simp apply (subst closed_except_def[symmetric]) using inner beta \c = _\ by (auto simp: closed_except_simps) qed (auto simp: closed_except_def) corollary (in irules) irewrite_rt_closed: assumes "rs \\<^sub>i t \* u" "closed t" shows "closed u" using assms by induction (auto intro: irewrite_closed) subsubsection \Correctness of translation\ abbreviation irelated :: "nterm \ pterm \ bool" ("_ \\<^sub>i _" [0,50] 50) where "n \\<^sub>i p \ nterm_to_pterm n = p" global_interpretation irelated: term_struct_rel_strong irelated by standard (auto simp: app_pterm_def app_nterm_def const_pterm_def const_nterm_def elim: nterm_to_pterm.elims) lemma irelated_vars: "t \\<^sub>i u \ frees t = frees u" by auto lemma irelated_no_abs: assumes "t \\<^sub>i u" shows "no_abs t \ no_abs u" using assms apply (induction arbitrary: t) apply (auto elim!: nterm_to_pterm.elims) apply (fold const_nterm_def const_pterm_def free_nterm_def free_pterm_def app_pterm_def app_nterm_def) by auto lemma irelated_subst: assumes "t \\<^sub>i u" "irelated.P_env nenv penv" shows "subst t nenv \\<^sub>i subst u penv" using assms proof (induction arbitrary: nenv penv u rule: nterm_to_pterm.induct) case (1 s) then show ?case by (auto elim!: fmrel_cases[where x = s]) next case 4 from 4(2)[symmetric] show ?case apply simp apply (rule 4) apply simp using 4(3) by (simp add: fmrel_drop) qed auto lemma related_irewrite_step: assumes "name, pats, nterm_to_pterm rhs \\<^sub>i u \ u'" "t \\<^sub>i u" obtains t' where "unsplit_rule (name, pats, rhs) \ t \ t'" "t' \\<^sub>i u'" proof - let ?rhs' = "nterm_to_pterm rhs" let ?x = "name $$ pats" from assms obtain env where "match ?x u = Some env" "u' = subst ?rhs' env" unfolding irewrite_step_def by blast then obtain nenv where "match ?x t = Some nenv" "irelated.P_env nenv env" using assms by (metis Option.is_none_def not_None_eq option.rel_distinct(1) option.sel rel_option_unfold irelated.match_rel) show thesis proof show "unsplit_rule (name, pats, rhs) \ t \ subst rhs nenv" using \match ?x t = _\ by auto next show "subst rhs nenv \\<^sub>i u'" unfolding \u' = _\ using \irelated.P_env nenv env\ by (auto intro: irelated_subst) qed qed theorem (in nrules) compile_correct: assumes "compile (consts_of rs) \\<^sub>i u \ u'" "t \\<^sub>i u" "closed t" obtains t' where "rs \\<^sub>n t \ t'" "t' \\<^sub>i u'" using assms(1-3) proof (induction arbitrary: t thesis rule: irewrite.induct) case (step name irs pats rhs u u') then obtain crs where "irs = translate_crules crs" "(name, crs) |\| consts_of rs" unfolding compile_def by force moreover with step obtain rhs' where "rhs = nterm_to_pterm rhs'" "(pats, rhs') |\| crs" unfolding translate_crules_def by force ultimately obtain rule where "split_rule rule = (name, (pats, rhs'))" "rule |\| rs" unfolding consts_of_def by blast hence "nrule rule" using all_rules by blast obtain t' where "unsplit_rule (name, pats, rhs') \ t \ t'" "t' \\<^sub>i u'" using \name, pats, rhs \\<^sub>i u \ u'\ \t \\<^sub>i u\ unfolding \rhs = nterm_to_pterm rhs'\ by (elim related_irewrite_step) hence "rule \ t \ t'" using \nrule rule\ \split_rule rule = (name, (pats, rhs'))\ by (metis unsplit_split) show ?case proof (rule step.prems) show "rs \\<^sub>n t \ t'" apply (rule nrewrite.step) apply fact apply fact done next show "t' \\<^sub>i u'" by fact qed next case (beta c cs u u') then obtain pat rhs where "c = (pat, rhs)" "(pat, rhs) |\| cs" by (cases c) auto obtain v w where "t = v $\<^sub>n w" "v \\<^sub>i Pabs cs" "w \\<^sub>i u" using \t \\<^sub>i Pabs cs $\<^sub>p u\ by (auto elim: nterm_to_pterm.elims) obtain x nrhs irhs where "v = (\\<^sub>n x. nrhs)" "cs = {| (Free x, irhs) |}" "nrhs \\<^sub>i irhs" using \v \\<^sub>i Pabs cs\ by (auto elim: nterm_to_pterm.elims) hence "t = (\\<^sub>n x. nrhs) $\<^sub>n w" "\\<^sub>n x. nrhs \\<^sub>i \\<^sub>p x. irhs" unfolding \t = v $\<^sub>n w\ using \v \\<^sub>i Pabs cs\ by auto have "pat = Free x" "rhs = irhs" using \cs = {| (Free x, irhs) |}\ \(pat, rhs) |\| cs\ by auto hence "(Free x, irhs) \ u \ u'" using beta \c = _\ by simp hence "u' = subst irhs (fmap_of_list [(x, u)])" by simp show ?case proof (rule beta.prems) show "rs \\<^sub>n t \ subst nrhs (fmap_of_list [(x, w)])" unfolding \t = (\\<^sub>n x. nrhs) $\<^sub>n w\ by (rule nrewrite.beta) next show "subst nrhs (fmap_of_list [(x, w)]) \\<^sub>i u'" unfolding \u' = subst irhs _\ apply (rule irelated_subst) apply fact apply simp apply rule apply rule apply fact done qed next case ("fun" v v' u) obtain w x where "t = w $\<^sub>n x" "w \\<^sub>i v" "x \\<^sub>i u" using \t \\<^sub>i v $\<^sub>p u\ by (auto elim: nterm_to_pterm.elims) with "fun" obtain w' where "rs \\<^sub>n w \ w'" "w' \\<^sub>i v'" unfolding closed_except_def by auto show ?case proof (rule fun.prems) show "rs \\<^sub>n t \ w' $\<^sub>n x" unfolding \t = w $\<^sub>n x\ by (rule nrewrite.fun) fact next show "w' $\<^sub>n x \\<^sub>i v' $\<^sub>p u" by auto fact+ qed next case (arg u u' v) obtain w x where "t = w $\<^sub>n x" "w \\<^sub>i v" "x \\<^sub>i u" using \ t \\<^sub>i v $\<^sub>p u\ by (auto elim: nterm_to_pterm.elims) with arg obtain x' where "rs \\<^sub>n x \ x'" "x' \\<^sub>i u'" unfolding closed_except_def by auto show ?case proof (rule arg.prems) show "rs \\<^sub>n t \ w $\<^sub>n x'" unfolding \t = w $\<^sub>n x\ by (rule nrewrite.arg) fact next show "w $\<^sub>n x' \\<^sub>i v $\<^sub>p u'" by auto fact+ qed qed corollary (in nrules) compile_correct_rt: assumes "compile (consts_of rs) \\<^sub>i u \* u'" "t \\<^sub>i u" "closed t" obtains t' where "rs \\<^sub>n t \* t'" "t' \\<^sub>i u'" using assms proof (induction arbitrary: thesis t) (* FIXME clone of transform_correct_rt, maybe locale? *) case (step u' u'') obtain t' where "rs \\<^sub>n t \* t'" "t' \\<^sub>i u'" using step by blast obtain t'' where "rs \\<^sub>n t' \* t''" "t'' \\<^sub>i u''" proof (rule compile_correct) show "compile (consts_of rs) \\<^sub>i u' \ u''" "t' \\<^sub>i u'" by fact+ next show "closed t'" using \rs \\<^sub>n t \* t'\ \closed t\ by (rule nrewrite_rt_closed) qed blast show ?case proof (rule step.prems) show "rs \\<^sub>n t \* t''" using \rs \\<^sub>n t \* t'\ \rs \\<^sub>n t' \* t''\ by auto qed fact qed blast subsubsection \Completeness of translation\ lemma (in nrules) compile_complete: assumes "rs \\<^sub>n t \ t'" "closed t" shows "compile (consts_of rs) \\<^sub>i nterm_to_pterm t \ nterm_to_pterm t'" using assms proof induction case (step r t t') then obtain pat rhs' where "r = (pat, rhs')" by force then have "(pat, rhs') |\| rs" "(pat, rhs') \ t \ t'" using step by blast+ then have "nrule (pat, rhs')" using all_rules by blast then obtain name pats where "(name, (pats, rhs')) = split_rule r" "pat = name $$ pats" unfolding split_rule_def \r = _\ apply atomize_elim by (auto simp: split_beta) obtain crs where "(name, crs) |\| consts_of rs" "(pats, rhs') |\| crs" using step \_ = split_rule r\ \r = _\ by (metis consts_of_def fgroup_by_complete fst_conv snd_conv) then obtain irs where "irs = translate_crules crs" by blast then have "(name, irs) |\| compile (consts_of rs)" unfolding compile_def using \(name, _) |\| _\ by (metis fimageI id_def map_prod_simp) obtain rhs where "rhs = nterm_to_pterm rhs'" "(pats, rhs) |\| irs" using \irs = _\ \_ |\| crs\ unfolding translate_crules_def by (metis fimageI id_def map_prod_simp) from step obtain env' where "match pat t = Some env'" "t' = subst rhs' env'" unfolding \r = _\ using rewrite_step.simps by force then obtain env where "match pat (nterm_to_pterm t) = Some env" "irelated.P_env env' env" by (metis irelated.match_rel option_rel_Some1) then have "subst rhs env = nterm_to_pterm t'" unfolding \t' = _\ apply - apply (rule sym) apply (rule irelated_subst) unfolding \rhs = _\ by auto have "name, pats, rhs \\<^sub>i nterm_to_pterm t \ nterm_to_pterm t'" apply (rule irewrite_stepI) using \match _ _ = Some env\ unfolding \pat = _\ apply assumption by fact show ?case by rule fact+ next case (beta x t t') obtain c where "c = (Free x, nterm_to_pterm t)" by blast from beta have "closed (nterm_to_pterm t')" using closed_nterm_to_pterm[where t = t'] unfolding closed_except_def by auto show ?case apply simp apply rule using \c = _\ by (fastforce intro: irelated_subst[THEN sym])+ next case ("fun" t t' u) show ?case apply simp apply rule apply (rule "fun") using "fun" unfolding closed_except_def apply simp done next case (arg u u' t) show ?case apply simp apply rule apply (rule arg) using arg unfolding closed_except_def by simp qed subsubsection \Correctness of transformation\ abbreviation irules_deferred_matches :: "pterm list \ irules \ (term \ pterm) fset" where "irules_deferred_matches args \ fselect (\(pats, rhs). map_option (\env. (last pats, subst rhs env)) (matchs (butlast pats) args))" context irules begin inductive prelated :: "pterm \ pterm \ bool" ("_ \\<^sub>p _" [0,50] 50) where const: "Pconst x \\<^sub>p Pconst x" | var: "Pvar x \\<^sub>p Pvar x" | app: "t\<^sub>1 \\<^sub>p u\<^sub>1 \ t\<^sub>2 \\<^sub>p u\<^sub>2 \ t\<^sub>1 $\<^sub>p t\<^sub>2 \\<^sub>p u\<^sub>1 $\<^sub>p u\<^sub>2" | pat: "rel_fset (rel_prod (=) prelated) cs\<^sub>1 cs\<^sub>2 \ Pabs cs\<^sub>1 \\<^sub>p Pabs cs\<^sub>2" | "defer": "(name, rsi) |\| rs \ 0 < arity rsi \ rel_fset (rel_prod (=) prelated) (irules_deferred_matches args rsi) cs \ list_all closed args \ name $$ args \\<^sub>p Pabs cs" inductive_cases prelated_absE[consumes 1, case_names pat "defer"]: "t \\<^sub>p Pabs cs" lemma prelated_refl[intro!]: "t \\<^sub>p t" proof (induction t) case Pabs thus ?case by (auto simp: snds.simps intro!: prelated.pat rel_fset_refl_strong rel_prod.intros) qed (auto intro: prelated.intros) sublocale prelated: term_struct_rel prelated by standard (auto simp: const_pterm_def app_pterm_def intro: prelated.intros elim: prelated.cases) lemma prelated_pvars: assumes "t \\<^sub>p u" shows "frees t = frees u" using assms proof (induction rule: prelated.induct) case (pat cs\<^sub>1 cs\<^sub>2) show ?case apply simp apply (rule arg_cong[where f = ffUnion]) apply (rule rel_fset_image_eq) apply fact apply auto done next case ("defer" name rsi args cs) { fix pat t assume "(pat, t) |\| cs" with "defer" obtain t' where "(pat, t') |\| irules_deferred_matches args rsi" "frees t = frees t'" by (auto elim: rel_fsetE2) then obtain pats rhs env where "pat = last pats" "(pats, rhs) |\| rsi" and "matchs (butlast pats) args = Some env" "t' = subst rhs env" by auto have "closed_except rhs (freess pats)" "linears pats" - using \(pats, rhs) |\| rsi\ \(name, rsi) |\| rs\ inner by blast+ + using \(pats, rhs) |\| rsi\ \(name, rsi) |\| rs\ inner by auto have "arity_compatibles rsi" - using "defer" inner by (blast dest: fpairwiseD) + using "defer" inner + by (metis (no_types, lifting) case_prodD) have "length pats > 0" by (subst arity_compatible_length) fact+ hence "pats = butlast pats @ [last pats]" by simp note \frees t = frees t'\ also have "frees t' = frees rhs - fmdom env" unfolding \t' = _\ apply (rule subst_frees) apply (rule closed.matchs) apply fact+ done also have "\ = frees rhs - freess (butlast pats)" using \matchs _ _ = _\ by (metis matchs_dom) also have "\ |\| freess pats - freess (butlast pats)" using \closed_except _ _\ by (auto simp: closed_except_def) also have "\ = frees (last pats) |-| freess (butlast pats)" by (subst \pats = _\) (simp add: funion_fminus) also have "\ = frees (last pats)" proof (rule fminus_triv) have "fdisjnt (freess (butlast pats)) (freess [last pats])" using \linears pats\ \pats = _\ by (metis linears_appendD) thus "frees (last pats) |\| freess (butlast pats) = {||}" by (fastforce simp: fdisjnt_alt_def) qed also have "\ = frees pat" unfolding \pat = _\ .. finally have "frees t |\| frees pat" . } hence "closed (Pabs cs)" unfolding closed_except_simps by (auto simp: closed_except_def) moreover have "closed (name $$ args)" unfolding closed_list_comb by fact ultimately show ?case unfolding closed_except_def by simp qed auto corollary prelated_closed: "t \\<^sub>p u \ closed t \ closed u" unfolding closed_except_def by (auto simp: prelated_pvars) lemma prelated_no_abs_right: assumes "t \\<^sub>p u" "no_abs u" shows "t = u" using assms apply (induction rule: prelated.induct) apply auto apply (fold app_pterm_def) apply auto done corollary env_prelated_refl[intro!]: "prelated.P_env env env" by (auto intro: fmap.rel_refl) text \ The following, more general statement does not hold: @{prop "t \\<^sub>p u \ rel_option prelated.P_env (match x t) (match x u)"} If @{text t} and @{text u} are related because of the @{thm [source=true] prelated.defer} rule, they have completely different shapes. Establishing @{prop "is_abs t \ is_abs u"} as a precondition would rule out this case, but at the same time be too restrictive. Instead, we use @{thm prelated.related_match}. \ lemma prelated_subst: assumes "t\<^sub>1 \\<^sub>p t\<^sub>2" "prelated.P_env env\<^sub>1 env\<^sub>2" shows "subst t\<^sub>1 env\<^sub>1 \\<^sub>p subst t\<^sub>2 env\<^sub>2" using assms proof (induction arbitrary: env\<^sub>1 env\<^sub>2 rule: prelated.induct) case (var x) thus ?case proof (cases rule: fmrel_cases[where x = x]) case none thus ?thesis by (auto intro: prelated.var) next case (some t u) thus ?thesis by simp qed next case (pat cs\<^sub>1 cs\<^sub>2) let ?drop = "\env. \(pat::term). fmdrop_fset (frees pat) env" from pat have "prelated.P_env (?drop env\<^sub>1 pat) (?drop env\<^sub>2 pat)" for pat by blast with pat show ?case by (auto intro!: prelated.pat rel_fset_image) next case ("defer" name rsi args cs) have "name $$ args \\<^sub>p Pabs cs" apply (rule prelated.defer) apply fact+ apply (rule fset.rel_mono_strong) apply fact apply force apply fact done moreover have "closed (name $$ args)" unfolding closed_list_comb by fact ultimately have "closed (Pabs cs)" by (metis prelated_closed) let ?drop = "\env. \pat. fmdrop_fset (frees pat) env" let ?f = "\env. (\(pat, rhs). (pat, subst rhs (?drop env pat)))" have "name $$ args \\<^sub>p Pabs (?f env\<^sub>2 |`| cs)" proof (rule prelated.defer) show "(name, rsi) |\| rs" "0 < arity rsi" "list_all closed args" using "defer" by auto next { fix pat\<^sub>1 rhs\<^sub>1 fix pat\<^sub>2 rhs\<^sub>2 assume "(pat\<^sub>2, rhs\<^sub>2) |\| cs" assume "pat\<^sub>1 = pat\<^sub>2" "rhs\<^sub>1 \\<^sub>p rhs\<^sub>2" have "rhs\<^sub>1 \\<^sub>p subst rhs\<^sub>2 (fmdrop_fset (frees pat\<^sub>2) env\<^sub>2)" by (subst subst_closed_pabs) fact+ } hence "rel_fset (rel_prod (=) prelated) (id |`| irules_deferred_matches args rsi) (?f env\<^sub>2 |`| cs)" by (force intro!: rel_fset_image[OF \rel_fset _ _ _\]) thus "rel_fset (rel_prod (=) prelated) (irules_deferred_matches args rsi) (?f env\<^sub>2 |`| cs)" by simp qed moreover have "map (\t. subst t env\<^sub>1) args = args" apply (rule map_idI) apply (rule subst_closed_id) using "defer" by (simp add: list_all_iff) ultimately show ?case by (simp add: subst_list_comb) qed (auto intro: prelated.intros) lemma prelated_step: assumes "name, pats, rhs \\<^sub>i u \ u'" "t \\<^sub>p u" obtains t' where "name, pats, rhs \\<^sub>i t \ t'" "t' \\<^sub>p u'" proof - let ?lhs = "name $$ pats" from assms obtain env where "match ?lhs u = Some env" "u' = subst rhs env" unfolding irewrite_step_def by blast then obtain env' where "match ?lhs t = Some env'" "prelated.P_env env' env" using assms by (auto elim: prelated.related_match) hence "subst rhs env' \\<^sub>p subst rhs env" using assms by (auto intro: prelated_subst) show thesis proof show "name, pats, rhs \\<^sub>i t \ subst rhs env'" unfolding irewrite_step_def using \match ?lhs t = Some env'\ by simp next show "subst rhs env' \\<^sub>p u'" unfolding \u' = subst rhs env\ by fact qed qed (* FIXME write using relators *) lemma prelated_beta: \ \same problem as @{thm [source=true] prelated.related_match}\ assumes "(pat, rhs\<^sub>2) \ t\<^sub>2 \ u\<^sub>2" "rhs\<^sub>1 \\<^sub>p rhs\<^sub>2" "t\<^sub>1 \\<^sub>p t\<^sub>2" obtains u\<^sub>1 where "(pat, rhs\<^sub>1) \ t\<^sub>1 \ u\<^sub>1" "u\<^sub>1 \\<^sub>p u\<^sub>2" proof - from assms obtain env\<^sub>2 where "match pat t\<^sub>2 = Some env\<^sub>2" "u\<^sub>2 = subst rhs\<^sub>2 env\<^sub>2" by auto with assms obtain env\<^sub>1 where "match pat t\<^sub>1 = Some env\<^sub>1" "prelated.P_env env\<^sub>1 env\<^sub>2" by (auto elim: prelated.related_match) with assms have "subst rhs\<^sub>1 env\<^sub>1 \\<^sub>p subst rhs\<^sub>2 env\<^sub>2" by (auto intro: prelated_subst) show thesis proof show "(pat, rhs\<^sub>1) \ t\<^sub>1 \ subst rhs\<^sub>1 env\<^sub>1" using \match pat t\<^sub>1 = _\ by simp next show "subst rhs\<^sub>1 env\<^sub>1 \\<^sub>p u\<^sub>2" unfolding \u\<^sub>2 = _\ by fact qed qed theorem transform_correct: assumes "transform_irule_set rs \\<^sub>i u \ u'" "t \\<^sub>p u" "closed t" obtains t' where "rs \\<^sub>i t \* t'" \ \zero or one step\ and "t' \\<^sub>p u'" using assms(1-3) proof (induction arbitrary: t thesis rule: irewrite.induct) case (beta c cs\<^sub>2 u\<^sub>2 x\<^sub>2') obtain v u\<^sub>1 where "t = v $\<^sub>p u\<^sub>1" "v \\<^sub>p Pabs cs\<^sub>2" "u\<^sub>1 \\<^sub>p u\<^sub>2" using \t \\<^sub>p Pabs cs\<^sub>2 $\<^sub>p u\<^sub>2\ by cases with beta have "closed u\<^sub>1" by (simp add: closed_except_def) obtain pat rhs\<^sub>2 where "c = (pat, rhs\<^sub>2)" by (cases c) auto from \v \\<^sub>p Pabs cs\<^sub>2\ show ?case proof (cases rule: prelated_absE) case (pat cs\<^sub>1) with beta \c = _\ obtain rhs\<^sub>1 where "(pat, rhs\<^sub>1) |\| cs\<^sub>1" "rhs\<^sub>1 \\<^sub>p rhs\<^sub>2" by (auto elim: rel_fsetE2) with beta obtain x\<^sub>1' where "(pat, rhs\<^sub>1) \ u\<^sub>1 \ x\<^sub>1'" "x\<^sub>1' \\<^sub>p x\<^sub>2'" using \u\<^sub>1 \\<^sub>p u\<^sub>2\ assms \c = _\ by (auto elim: prelated_beta simp del: rewrite_step.simps) show ?thesis proof (rule beta.prems) show "rs \\<^sub>i t \* x\<^sub>1'" unfolding \t = _\ \v = _\ by (intro r_into_rtranclp irewrite.beta) fact+ next show "x\<^sub>1' \\<^sub>p x\<^sub>2'" by fact qed next case ("defer" name rsi args) with beta \c = _\ obtain rhs\<^sub>1' where "(pat, rhs\<^sub>1') |\| irules_deferred_matches args rsi" "rhs\<^sub>1' \\<^sub>p rhs\<^sub>2" by (auto elim: rel_fsetE2) then obtain env\<^sub>a rhs\<^sub>1 pats where "matchs (butlast pats) args = Some env\<^sub>a" "pat = last pats" "rhs\<^sub>1' = subst rhs\<^sub>1 env\<^sub>a" and "(pats, rhs\<^sub>1) |\| rsi" by auto hence "linears pats" - using \(name, rsi) |\| rs\ inner unfolding irules_def by blast + using \(name, rsi) |\| rs\ inner unfolding irules_def by auto have "arity_compatibles rsi" - using "defer" inner by (blast dest: fpairwiseD) + using "defer" inner + by (metis (no_types, lifting) case_prodD) have "length pats > 0" by (subst arity_compatible_length) fact+ hence "pats = butlast pats @ [pat]" unfolding \pat = _\ by simp from beta \c = _\ obtain env\<^sub>b where "match pat u\<^sub>2 = Some env\<^sub>b" "x\<^sub>2' = subst rhs\<^sub>2 env\<^sub>b" by fastforce with \u\<^sub>1 \\<^sub>p u\<^sub>2\ obtain env\<^sub>b' where "match pat u\<^sub>1 = Some env\<^sub>b'" "prelated.P_env env\<^sub>b' env\<^sub>b" by (metis prelated.related_match) have "closed_env env\<^sub>a" by (rule closed.matchs) fact+ have "closed_env env\<^sub>b'" apply (rule closed.matchs[where pats = "[pat]" and ts = "[u\<^sub>1]"]) apply simp apply fact apply simp apply fact done have "fmdom env\<^sub>a = freess (butlast pats)" by (rule matchs_dom) fact moreover have "fmdom env\<^sub>b' = frees pat" by (rule match_dom) fact moreover have "fdisjnt (freess (butlast pats)) (frees pat)" using \pats = _\ \linears pats\ by (metis freess_single linears_appendD(3)) ultimately have "fdisjnt (fmdom env\<^sub>a) (fmdom env\<^sub>b')" by simp show ?thesis proof (rule beta.prems) have "rs \\<^sub>i name $$ args $\<^sub>p u\<^sub>1 \ subst rhs\<^sub>1' env\<^sub>b'" proof (rule irewrite.step) show "(name, rsi) |\| rs" "(pats, rhs\<^sub>1) |\| rsi" by fact+ next show "name, pats, rhs\<^sub>1 \\<^sub>i name $$ args $\<^sub>p u\<^sub>1 \ subst rhs\<^sub>1' env\<^sub>b'" apply (rule irewrite_stepI) apply (fold app_pterm_def) apply (subst list_comb_snoc) apply (subst matchs_match_list_comb) apply (subst \pats = _\) apply (rule matchs_appI) apply fact apply simp apply fact unfolding \rhs\<^sub>1' = _\ apply (rule subst_indep') apply fact+ done qed thus "rs \\<^sub>i t \* subst rhs\<^sub>1' env\<^sub>b'" unfolding \t = _\ \v = _\ by (rule r_into_rtranclp) next show "subst rhs\<^sub>1' env\<^sub>b' \\<^sub>p x\<^sub>2'" unfolding \x\<^sub>2' = _\ by (rule prelated_subst) fact+ qed qed next case (step name rs\<^sub>2 pats rhs u u') then obtain rs\<^sub>1 where "rs\<^sub>2 = transform_irules rs\<^sub>1" "(name, rs\<^sub>1) |\| rs" unfolding transform_irule_set_def by force hence "arity_compatibles rs\<^sub>1" - using inner by (blast dest: fpairwiseD) + using inner + by (metis (no_types, lifting) case_prodD) show ?case proof (cases "arity rs\<^sub>1 = 0") case True hence "rs\<^sub>2 = rs\<^sub>1" unfolding \rs\<^sub>2 = _\ transform_irules_def by simp with step have "(pats, rhs) |\| rs\<^sub>1" by simp from step obtain t' where "name, pats, rhs \\<^sub>i t \ t'" "t' \\<^sub>p u'" using assms by (auto elim: prelated_step) show ?thesis proof (rule step.prems) show "rs \\<^sub>i t \* t'" by (intro conjI exI r_into_rtranclp irewrite.step) fact+ qed fact next let ?f = "\(pats, rhs). (butlast pats, last pats, rhs)" let ?grp = "fgroup_by ?f rs\<^sub>1" case False hence "rs\<^sub>2 = map_prod id Pabs |`| ?grp" unfolding \rs\<^sub>2 = _\ transform_irules_def by simp with step obtain cs where "rhs = Pabs cs" "(pats, cs) |\| ?grp" by force from step obtain env\<^sub>2 where "match (name $$ pats) u = Some env\<^sub>2" "u' = subst rhs env\<^sub>2" unfolding irewrite_step_def by auto then obtain args\<^sub>2 where "u = name $$ args\<^sub>2" "matchs pats args\<^sub>2 = Some env\<^sub>2" by (auto elim: match_list_combE) with step obtain args\<^sub>1 where "t = name $$ args\<^sub>1" "list_all2 prelated args\<^sub>1 args\<^sub>2" by (auto elim: prelated.list_combE) then obtain env\<^sub>1 where "matchs pats args\<^sub>1 = Some env\<^sub>1" "prelated.P_env env\<^sub>1 env\<^sub>2" using \matchs pats args\<^sub>2 = _\ by (metis prelated.related_matchs) hence "fmdom env\<^sub>1 = freess pats" by (auto simp: matchs_dom) obtain cs' where "u' = Pabs cs'" unfolding \u' = _\ \rhs = _\ by auto hence "cs' = (\(pat, rhs). (pat, subst rhs (fmdrop_fset (frees pat) env\<^sub>2 ))) |`| cs" using \u' = subst rhs env\<^sub>2\ unfolding \rhs = _\ by simp show ?thesis proof (rule step.prems) show "rs \\<^sub>i t \* t" by (rule rtranclp.rtrancl_refl) next show "t \\<^sub>p u'" unfolding \u' = Pabs cs'\ \t = _\ proof (intro prelated.defer rel_fsetI; safe?) show "(name, rs\<^sub>1) |\| rs" by fact next show "0 < arity rs\<^sub>1" using False by simp next show "list_all closed args\<^sub>1" using \closed t\ unfolding \t = _\ closed_list_comb . next fix pat rhs' assume "(pat, rhs') |\| irules_deferred_matches args\<^sub>1 rs\<^sub>1" then obtain pats' rhs env where "(pats', rhs) |\| rs\<^sub>1" and "matchs (butlast pats') args\<^sub>1 = Some env" "pat = last pats'" "rhs' = subst rhs env" by auto with False have "pats' \ []" using \arity_compatibles rs\<^sub>1\ by (metis list.size(3) arity_compatible_length) hence "butlast pats' @ [last pats'] = pats'" by simp from \(pats, cs) |\| ?grp\ obtain pats\<^sub>e rhs\<^sub>e where "(pats\<^sub>e, rhs\<^sub>e) |\| rs\<^sub>1" "pats = butlast pats\<^sub>e" by (auto elim: fgroup_byE2) have "patterns_compatible (butlast pats') pats" unfolding \pats = _\ apply (rule rev_accum_rel_butlast) using \(pats', rhs) |\| rs\<^sub>1\ \(pats\<^sub>e, rhs\<^sub>e) |\| rs\<^sub>1\ \(name, rs\<^sub>1) |\| rs\ inner by (blast dest: fpairwiseD) interpret irules': irules C_info "transform_irule_set rs" by (rule rules_transform) have "butlast pats' = pats" "env = env\<^sub>1" apply (rule matchs_compatible_eq) subgoal by fact subgoal apply (rule linears_butlastI) - using \(pats', rhs) |\| rs\<^sub>1\ \(name, rs\<^sub>1) |\| rs\ inner by blast + using \(pats', rhs) |\| rs\<^sub>1\ \(name, rs\<^sub>1) |\| rs\ inner by auto subgoal using \(pats, _) |\| rs\<^sub>2\ \(name, rs\<^sub>2) |\| transform_irule_set rs\ - using irules'.inner by blast + using irules'.inner by auto apply fact+ subgoal apply (rule matchs_compatible_eq) apply fact apply (rule linears_butlastI) using \(pats', rhs) |\| rs\<^sub>1\ \(name, rs\<^sub>1) |\| rs\ inner - apply blast + apply auto [] using \(pats, _) |\| rs\<^sub>2\ \(name, rs\<^sub>2) |\| transform_irule_set rs\ - using irules'.inner apply blast + using irules'.inner apply auto[] by fact+ done let ?rhs_subst = "\env. subst rhs (fmdrop_fset (frees pat) env)" have "fmdom env\<^sub>2 = freess pats" using \match (_ $$ _) _ = Some env\<^sub>2\ by (simp add: match_dom) show "fBex cs' (rel_prod (=) prelated (pat, rhs'))" unfolding \rhs' = _\ proof (rule fBexI, rule rel_prod.intros) have "fdisjnt (freess (butlast pats')) (frees (last pats'))" apply (subst freess_single[symmetric]) apply (rule linears_appendD) apply (subst \butlast pats' @ [last pats'] = pats'\) using \(pats', rhs) |\| rs\<^sub>1\ \(name, rs\<^sub>1) |\| rs\ inner - by blast + by auto show "subst rhs env \\<^sub>p ?rhs_subst env\<^sub>2" apply (rule prelated_subst) apply (rule prelated_refl) unfolding fmfilter_alt_defs apply (subst fmfilter_true) subgoal premises prems for x y using fmdomI[OF prems] unfolding \pat = _\ \fmdom env\<^sub>2 = _\ apply (subst (asm) \butlast pats' = pats\[symmetric]) using \fdisjnt (freess (butlast pats')) (frees (last pats'))\ by (auto simp: fdisjnt_alt_def) subgoal unfolding \env = _\ by fact done next have "(pat, rhs) |\| cs" unfolding \pat = _\ apply (rule fgroup_byD[where a = "(x, y)" for x y]) apply fact apply simp apply (intro conjI) apply fact apply (rule refl)+ apply fact done thus "(pat, ?rhs_subst env\<^sub>2) |\| cs'" unfolding \cs' = _\ by force qed simp next fix pat rhs' assume "(pat, rhs') |\| cs'" then obtain rhs where "(pat, rhs) |\| cs" and "rhs' = subst rhs (fmdrop_fset (frees pat) env\<^sub>2 )" unfolding \cs' = _\ by auto with \(pats, cs) |\| ?grp\ obtain pats' where "(pats', rhs) |\| rs\<^sub>1" "pats = butlast pats'" "pat = last pats'" by auto with False have "length pats' \ 0" using \arity_compatibles _\ by (metis arity_compatible_length) hence "pats' = pats @ [pat]" unfolding \pats = _\ \pat = _\ by auto moreover have "linears pats'" - using \(pats', rhs) |\| rs\<^sub>1\ \(name, rs\<^sub>1) |\| _\ inner by blast + using \(pats', rhs) |\| rs\<^sub>1\ \(name, rs\<^sub>1) |\| _\ inner by auto ultimately have "fdisjnt (fmdom env\<^sub>1) (frees pat)" unfolding \fmdom env\<^sub>1 = _\ by (auto dest: linears_appendD) let ?rhs_subst = "\env. subst rhs (fmdrop_fset (frees pat) env)" show "fBex (irules_deferred_matches args\<^sub>1 rs\<^sub>1) (\e. rel_prod (=) prelated e (pat, rhs'))" unfolding \rhs' = _\ proof (rule fBexI, rule rel_prod.intros) show "?rhs_subst env\<^sub>1 \\<^sub>p ?rhs_subst env\<^sub>2" using \prelated.P_env env\<^sub>1 env\<^sub>2\ inner by (auto intro: prelated_subst) next have "matchs (butlast pats') args\<^sub>1 = Some env\<^sub>1" using \matchs pats args\<^sub>1 = _\ \pats = _\ by simp moreover have "subst rhs env\<^sub>1 = ?rhs_subst env\<^sub>1" apply (rule arg_cong[where f = "subst rhs"]) unfolding fmfilter_alt_defs apply (rule fmfilter_true[symmetric]) using \fdisjnt (fmdom env\<^sub>1) _\ by (auto simp: fdisjnt_alt_def intro: fmdomI) ultimately show "(pat, ?rhs_subst env\<^sub>1) |\| irules_deferred_matches args\<^sub>1 rs\<^sub>1" using \(pats', rhs) |\| rs\<^sub>1\ \pat = last pats'\ by auto qed simp qed qed qed next case ("fun" v v' u) obtain w x where "t = w $\<^sub>p x" "w \\<^sub>p v" "x \\<^sub>p u" "closed w" using \t \\<^sub>p v $\<^sub>p u\ \closed t\ by cases (auto simp: closed_except_def) with "fun" obtain w' where "rs \\<^sub>i w \* w'" "w' \\<^sub>p v'" by blast show ?case proof (rule fun.prems) show "rs \\<^sub>i t \* w' $\<^sub>p x" unfolding \t = _\ by (intro irewrite.rt_comb[unfolded app_pterm_def] rtranclp.rtrancl_refl) fact next show "w' $\<^sub>p x \\<^sub>p v' $\<^sub>p u" by (rule prelated.app) fact+ qed next case (arg u u' v) obtain w x where "t = w $\<^sub>p x" "w \\<^sub>p v" "x \\<^sub>p u" "closed x" using \t \\<^sub>p v $\<^sub>p u\ \closed t\ by cases (auto simp: closed_except_def) with arg obtain x' where "rs \\<^sub>i x \* x'" "x' \\<^sub>p u'" by blast show ?case proof (rule arg.prems) show "rs \\<^sub>i t \* w $\<^sub>p x'" unfolding \t = w $\<^sub>p x\ by (intro irewrite.rt_comb[unfolded app_pterm_def] rtranclp.rtrancl_refl) fact next show "w $\<^sub>p x' \\<^sub>p v $\<^sub>p u'" by (rule prelated.app) fact+ qed qed end subsubsection \Completeness of transformation\ lemma (in irules) transform_completeness: assumes "rs \\<^sub>i t \ t'" "closed t" shows "transform_irule_set rs \\<^sub>i t \* t'" using assms proof induction case (step name irs' pats' rhs' t t') then obtain irs where "irs = transform_irules irs'" "(name, irs) |\| transform_irule_set rs" unfolding transform_irule_set_def by (metis fimageI id_apply map_prod_simp) show ?case proof (cases "arity irs' = 0") case True hence "irs = irs'" unfolding \irs = _\ unfolding transform_irules_def by simp with step have "(pats', rhs') |\| irs" "name, pats', rhs' \\<^sub>i t \ t'" by blast+ have "transform_irule_set rs \\<^sub>i t \* t'" apply (rule r_into_rtranclp) apply rule by fact+ show ?thesis by fact next let ?f = "\(pats, rhs). (butlast pats, last pats, rhs)" let ?grp = "fgroup_by ?f irs'" note closed_except_def [simp add] case False then have "irs = map_prod id Pabs |`| ?grp" unfolding \irs = _\ unfolding transform_irules_def by simp with False have "irs = transform_irules irs'" unfolding transform_irules_def by simp obtain pat pats where "pat = last pats'" "pats = butlast pats'" by blast from step False have "length pats' \ 0" using arity_compatible_length inner - by (smt fBallE prod.simps(2)) + by (smt (verit, ccfv_threshold) case_prodD) then have "pats' = pats @ [pat]" unfolding \pat = _\ \pats = _\ by simp from step have "linears pats'" - using inner fBallE - by (metis (mono_tags, lifting) old.prod.case) + using inner by auto then have "fdisjnt (freess pats) (frees pat)" unfolding \pats' = _\ using linears_appendD(3) freess_single by force from step obtain cs where "(pats, cs) |\| ?grp" unfolding \pats = _\ by (metis (no_types, lifting) fgroup_by_complete fst_conv prod.simps(2)) with step have "(pat, rhs') |\| cs" unfolding \pat = _\ \pats = _\ by (meson fgroup_byD old.prod.case) have "(pats, Pabs cs) |\| irs" using \irs = map_prod id Pabs |`| ?grp\ \(pats, cs) |\| _\ by (metis (no_types, lifting) eq_snd_iff fst_conv fst_map_prod id_def rev_fimage_eqI snd_map_prod) from step obtain env' where "match (name $$ pats') t = Some env'" "subst rhs' env' = t'" using irewrite_step_def by auto have "name $$ pats' = (name $$ pats) $ pat" unfolding \pats' = _\ by (simp add: app_term_def) then obtain t\<^sub>0 t\<^sub>1 env\<^sub>0 env\<^sub>1 where "t = t\<^sub>0 $\<^sub>p t\<^sub>1" "match (name $$ pats) t\<^sub>0 = Some env\<^sub>0" "match pat t\<^sub>1 = Some env\<^sub>1" "env' = env\<^sub>0 ++\<^sub>f env\<^sub>1" using match_appE_split[OF \match (name $$ pats') _ = _\[unfolded \name $$ pats' = _\], unfolded app_pterm_def] by blast with step have "closed t\<^sub>0" "closed t\<^sub>1" by auto then have "closed_env env\<^sub>0" "closed_env env\<^sub>1" using match_vars[OF \match _ t\<^sub>0 = _\] match_vars[OF \match _ t\<^sub>1 = _\] unfolding closed_except_def by auto obtain t\<^sub>0' where "subst (Pabs cs) env\<^sub>0 = t\<^sub>0'" by blast then obtain cs' where "t\<^sub>0' = Pabs cs'" "cs' = ((\(pat, rhs). (pat, subst rhs (fmdrop_fset (frees pat) env\<^sub>0))) |`| cs)" using subst_pterm.simps(3) by blast obtain rhs where "subst rhs' (fmdrop_fset (frees pat) env\<^sub>0) = rhs" by blast then have "(pat, rhs) |\| cs'" unfolding \cs' = _\ using \_ |\| cs\ by (metis (mono_tags, lifting) old.prod.case rev_fimage_eqI) have "env\<^sub>0 ++\<^sub>f env\<^sub>1 = (fmdrop_fset (frees pat) env\<^sub>0) ++\<^sub>f env\<^sub>1" apply (subst fmadd_drop_left_dom[symmetric]) using \match pat _ = _\ match_dom by metis have "fdisjnt (fmdom env\<^sub>0) (fmdom env\<^sub>1)" using match_dom using \match pat _ = _\ \match (name $$ pats) _ = _\ using \fdisjnt _ _\ unfolding fdisjnt_alt_def by (metis matchs_dom match_list_combE) have "subst rhs env\<^sub>1 = t'" unfolding \_ = rhs\[symmetric] unfolding \_ = t'\[symmetric] unfolding \env' = _\ unfolding \env\<^sub>0 ++\<^sub>f _ = _\ apply (subst subst_indep') using \closed_env env\<^sub>0\ apply blast using \fdisjnt (fmdom _) _\ unfolding fdisjnt_alt_def by auto show ?thesis unfolding \t = _\ apply rule apply (rule r_into_rtranclp) apply (rule irewrite.intros(3)) apply rule apply fact+ apply (rule irewrite_stepI) apply fact+ unfolding \t\<^sub>0' = _\ apply rule apply fact using \match pat t\<^sub>1 = _\ \subst rhs _ = _\ by force qed qed (auto intro: irewrite.rt_comb[unfolded app_pterm_def] intro!: irewrite.intros simp: closed_except_def) subsubsection \Computability\ export_code compile transform_irules checking Scala SML end \ No newline at end of file diff --git a/thys/Coinductive_Languages/Context_Free_Grammar.thy b/thys/Coinductive_Languages/Context_Free_Grammar.thy --- a/thys/Coinductive_Languages/Context_Free_Grammar.thy +++ b/thys/Coinductive_Languages/Context_Free_Grammar.thy @@ -1,177 +1,177 @@ section \Word Problem for Context-Free Grammars\ (*<*) theory Context_Free_Grammar imports Coinductive_Language "HOL-Library.FSet" begin (*>*) section \Context Free Languages\ text \ A context-free grammar consists of a list of productions for every nonterminal and an initial nonterminal. The productions are required to be in weak Greibach normal form, i.e. each right hand side of a production must either be empty or start with a terminal. \ abbreviation "wgreibach \ \ (case \ of (Inr N # _) \ False | _ \ True)" record ('t, 'n) cfg = init :: "'n :: finite" prod :: "'n \ ('t + 'n) list fset" context fixes G :: "('t, 'n :: finite) cfg" begin inductive in_cfl where "in_cfl [] []" | "in_cfl \ w \ in_cfl (Inl a # \) (a # w)" | "fBex (prod G N) (\\. in_cfl (\ @ \) w) \ in_cfl (Inr N # \) w" abbreviation lang_trad where "lang_trad \ {w. in_cfl [Inr (init G)] w}" fun \\<^sub>P where "\\<^sub>P [] = True" | "\\<^sub>P (Inl _ # _) = False" | "\\<^sub>P (Inr N # \) = ([] |\| prod G N \ \\<^sub>P \)" fun \
\<^sub>P where "\
\<^sub>P [] a = {||}" | "\
\<^sub>P (Inl b # \) a = (if a = b then {|\|} else {||})" | "\
\<^sub>P (Inr N # \) a = (\\. tl \ @ \) |`| ffilter (\\. \ \ [] \ hd \ = Inl a) (prod G N) |\| (if [] |\| prod G N then \
\<^sub>P \ a else {||})" primcorec subst :: "('t + 'n) list fset \ 't language" where "subst P = Lang (fBex P \\<^sub>P) (\a. subst (ffUnion ((\r. \
\<^sub>P r a) |`| P)))" inductive in_cfls where "fBex P \\<^sub>P \ in_cfls P []" | "in_cfls (ffUnion ((\\. \
\<^sub>P \ a) |`| P)) w \ in_cfls P (a # w)" inductive_cases [elim!]: "in_cfls P []" inductive_cases [elim!]: "in_cfls P (a # w)" declare inj_eq[OF bij_is_inj[OF to_language_bij], simp] lemma subst_in_cfls: "subst P = to_language {w. in_cfls P w}" by (coinduction arbitrary: P) (auto intro: in_cfls.intros) lemma \\<^sub>P_in_cfl: "\\<^sub>P \ \ in_cfl \ []" by (induct \ rule: \\<^sub>P.induct) (auto intro!: in_cfl.intros elim: fBexI[rotated]) lemma \
\<^sub>P_in_cfl: "\ |\| \
\<^sub>P \ a \ in_cfl \ w \ in_cfl \ (a # w)" proof (induct \ a arbitrary: \ w rule: \
\<^sub>P.induct) case (3 N \ a) then show ?case by (auto simp: rev_fBexI neq_Nil_conv split: if_splits intro!: in_cfl.intros elim!: rev_fBexI[of "_ # _"]) qed (auto split: if_splits intro: in_cfl.intros) lemma in_cfls_in_cfl: "in_cfls P w \ fBex P (\\. in_cfl \ w)" by (induct P w rule: in_cfls.induct) - (auto simp: \\<^sub>P_in_cfl \
\<^sub>P_in_cfl ffUnion.rep_eq fBex.rep_eq fBall.rep_eq + (auto simp: \\<^sub>P_in_cfl \
\<^sub>P_in_cfl ffUnion.rep_eq intro: in_cfl.intros elim: rev_bexI) lemma in_cfls_mono: "in_cfls P w \ P |\| Q \ in_cfls Q w" proof (induct P w arbitrary: Q rule: in_cfls.induct) case (2 a P w) from 2(3) 2(2)[of "ffUnion ((\\. local.\
\<^sub>P \ a) |`| Q)"] show ?case by (auto intro!: ffunion_mono in_cfls.intros) qed (auto intro!: in_cfls.intros) end locale cfg_wgreibach = fixes G :: "('t, 'n :: finite) cfg" assumes weakGreibach: "\N \. \ |\| prod G N \ wgreibach \" begin lemma in_cfl_in_cfls: "in_cfl G \ w \ in_cfls G {|\|} w" proof (induct \ w rule: in_cfl.induct) case (3 N \ w) then obtain \ where \: "\ |\| prod G N" and in_cfl: "in_cfl G (\ @ \) w" and in_cfls: "in_cfls G {|\ @ \|} w" by blast then show ?case proof (cases \) case [simp]: Nil from \ in_cfls show ?thesis by (cases w) (auto intro!: in_cfls.intros elim: in_cfls_mono) next case [simp]: (Cons x \) from weakGreibach[OF \] obtain a where [simp]: "x = Inl a" by (cases x) auto with \ in_cfls show ?thesis apply - apply (rule in_cfl.cases[OF in_cfl]; auto) apply (force intro: in_cfls.intros(2) elim!: in_cfls_mono) done qed qed (auto intro!: in_cfls.intros) abbreviation lang where "lang \ subst G {|[Inr (init G)]|}" lemma lang_lang_trad: "lang = to_language (lang_trad G)" proof - have "in_cfls G {|[Inr (init G)]|} w \ in_cfl G [Inr (init G)] w" for w by (auto dest: in_cfls_in_cfl in_cfl_in_cfls) then show ?thesis by (auto simp: subst_in_cfls) qed end text \ The function @{term in_language} decides the word problem for a given language. Since we can construct the language of a CFG using @{const cfg_wgreibach.lang} we obtain an executable (but not very efficient) decision procedure for CFGs for free. \ abbreviation "\ \ Inl True" abbreviation "\ \ Inl False" abbreviation "S \ Inr ()" interpretation palindromes: cfg_wgreibach "\init = (), prod = \_. {|[], [\], [\], [\, S, \], [\, S, \]|}\" by unfold_locales auto lemma "in_language palindromes.lang []" by normalization lemma "in_language palindromes.lang [True]" by normalization lemma "in_language palindromes.lang [False]" by normalization lemma "in_language palindromes.lang [True, True]" by normalization lemma "in_language palindromes.lang [True, False, True]" by normalization lemma "\ in_language palindromes.lang [True, False]" by normalization lemma "\ in_language palindromes.lang [True, False, True, False]" by normalization lemma "in_language palindromes.lang [True, False, True, True, False, True]" by normalization lemma "\ in_language palindromes.lang [True, False, True, False, False, True]" by normalization interpretation Dyck: cfg_wgreibach "\init = (), prod = \_. {|[], [\, S, \, S]|}\" by unfold_locales auto lemma "in_language Dyck.lang []" by normalization lemma "\ in_language Dyck.lang [True]" by normalization lemma "\ in_language Dyck.lang [False]" by normalization lemma "in_language Dyck.lang [True, False, True, False]" by normalization lemma "in_language Dyck.lang [True, True, False, False]" by normalization lemma "in_language Dyck.lang [True, False, True, False]" by normalization lemma "in_language Dyck.lang [True, False, True, False, True, True, False, False]" by normalization lemma "\ in_language Dyck.lang [True, False, True, True, False]" by normalization lemma "\ in_language Dyck.lang [True, True, False, False, False, True]" by normalization interpretation abSSa: cfg_wgreibach "\init = (), prod = \_. {|[], [\, \, S, S, \]|}\" by unfold_locales auto lemma "in_language abSSa.lang []" by normalization lemma "\ in_language abSSa.lang [True]" by normalization lemma "\ in_language abSSa.lang [False]" by normalization lemma "in_language abSSa.lang [True, False, True]" by normalization lemma "in_language abSSa.lang [True, False, True, False, True, True, False, True, True]" by normalization lemma "in_language abSSa.lang [True, False, True, False, True, True]" by normalization lemma "\ in_language abSSa.lang [True, False, True, True, False]" by normalization lemma "\ in_language abSSa.lang [True, True, False, False, False, True]" by normalization (*<*) end (*>*) \ No newline at end of file diff --git a/thys/Extended_Finite_State_Machine_Inference/Inference.thy b/thys/Extended_Finite_State_Machine_Inference/Inference.thy --- a/thys/Extended_Finite_State_Machine_Inference/Inference.thy +++ b/thys/Extended_Finite_State_Machine_Inference/Inference.thy @@ -1,586 +1,586 @@ chapter\EFSM Inference\ text\This chapter presents the definitions necessary for EFSM inference by state-merging.\ section\Inference by State-Merging\ text\This theory sets out the key definitions for the inference of EFSMs from system traces.\ theory Inference imports Subsumption "Extended_Finite_State_Machines.Transition_Lexorder" "HOL-Library.Product_Lexorder" begin declare One_nat_def [simp del] subsection\Transition Identifiers\ text\We first need to define the \texttt{iEFSM} data type which assigns each transition a unique identity. This is necessary because transitions may not occur uniquely in an EFSM. Assigning transitions a unique identifier enables us to look up the origin and destination states of transitions without having to pass them around in the inference functions.\ type_synonym tid = nat type_synonym tids = "tid list" type_synonym iEFSM = "(tids \ (cfstate \ cfstate) \ transition) fset" definition origin :: "tids \ iEFSM \ nat" where "origin uid t = fst (fst (snd (fthe_elem (ffilter (\x. set uid \ set (fst x)) t))))" definition dest :: "tids \ iEFSM \ nat" where "dest uid t = snd (fst (snd (fthe_elem (ffilter (\x. set uid \ set (fst x)) t))))" definition get_by_id :: "iEFSM \ tid \ transition" where "get_by_id e uid = (snd \ snd) (fthe_elem (ffilter (\(tids, _). uid \ set tids) e))" definition get_by_ids :: "iEFSM \ tids \ transition" where "get_by_ids e uid = (snd \ snd) (fthe_elem (ffilter (\(tids, _). set uid \ set tids) e))" definition uids :: "iEFSM \ nat fset" where "uids e = ffUnion (fimage (fset_of_list \ fst) e)" definition max_uid :: "iEFSM \ nat option" where "max_uid e = (let uids = uids e in if uids = {||} then None else Some (fMax uids))" definition tm :: "iEFSM \ transition_matrix" where "tm e = fimage snd e" definition all_regs :: "iEFSM \ nat set" where "all_regs e = EFSM.all_regs (tm e)" definition max_reg :: "iEFSM \ nat option" where "max_reg e = EFSM.max_reg (tm e)" definition "max_reg_total e = (case max_reg e of None \ 0 | Some r \ r)" definition max_output :: "iEFSM \ nat" where "max_output e = EFSM.max_output (tm e)" definition max_int :: "iEFSM \ int" where "max_int e = EFSM.max_int (tm e)" definition S :: "iEFSM \ nat fset" where "S m = (fimage (\(uid, (s, s'), t). s) m) |\| fimage (\(uid, (s, s'), t). s') m" lemma S_alt: "S t = EFSM.S (tm t)" apply (simp add: S_def EFSM.S_def tm_def) by force lemma to_in_S: "(\to from uid. (uid, (from, to), t) |\| xb \ to |\| S xb)" apply (simp add: S_def) by blast lemma from_in_S: "(\to from uid. (uid, (from, to), t) |\| xb \ from |\| S xb)" apply (simp add: S_def) by blast subsection\Building the PTA\ text\The first step in EFSM inference is to construct a PTA from the observed traces in the same way as for classical FSM inference. Beginning with the empty EFSM, we iteratively attempt to walk each observed trace in the model. When we reach a point where there is no available transition, one is added. For classical FSMs, this is simply an atomic label. EFSMs deal with data, so we need to add guards which test for the observed input values and outputs which produce the observed values.\ primrec make_guard :: "value list \ nat \ vname gexp list" where "make_guard [] _ = []" | "make_guard (h#t) n = (gexp.Eq (V (vname.I n)) (L h))#(make_guard t (n+1))" primrec make_outputs :: "value list \ output_function list" where "make_outputs [] = []" | "make_outputs (h#t) = (L h)#(make_outputs t)" definition max_uid_total :: "iEFSM \ nat" where "max_uid_total e = (case max_uid e of None \ 0 | Some u \ u)" definition add_transition :: "iEFSM \ cfstate \ label \ value list \ value list \ iEFSM" where "add_transition e s label inputs outputs = finsert ([max_uid_total e + 1], (s, (maxS (tm e))+1), \Label=label, Arity=length inputs, Guards=(make_guard inputs 0), Outputs=(make_outputs outputs), Updates=[]\) e" fun make_branch :: "iEFSM \ cfstate \ registers \ trace \ iEFSM" where "make_branch e _ _ [] = e" | "make_branch e s r ((label, inputs, outputs)#t) = (case (step (tm e) s r label inputs) of Some (transition, s', outputs', updated) \ if outputs' = (map Some outputs) then make_branch e s' updated t else make_branch (add_transition e s label inputs outputs) ((maxS (tm e))+1) r t | None \ make_branch (add_transition e s label inputs outputs) ((maxS (tm e))+1) r t )" primrec make_pta_aux :: "log \ iEFSM \ iEFSM" where "make_pta_aux [] e = e" | "make_pta_aux (h#t) e = make_pta_aux t (make_branch e 0 <> h)" definition "make_pta log = make_pta_aux log {||}" lemma make_pta_aux_fold [code]: "make_pta_aux l e = fold (\h e. make_branch e 0 <> h) l e" by(induct l arbitrary: e, auto) subsection\Integrating Heuristics\ text\A key contribution of the inference technique presented in \<^cite>\"foster2019"\ is the ability to introduce \emph{internal variables} to the model to generalise behaviours and allow transitions to be merged. This is done by providing the inference technique with a set of \emph{heuristics}. The aim here is not to create a ``one size fits all'' magic oracle, rather to recognise particular \emph{data usage patterns} which can be abstracted.\ type_synonym update_modifier = "tids \ tids \ cfstate \ iEFSM \ iEFSM \ iEFSM \ (transition_matrix \ bool) \ iEFSM option" definition null_modifier :: update_modifier where "null_modifier f _ _ _ _ _ _ = None" definition replace_transition :: "iEFSM \ tids \ transition \ iEFSM" where "replace_transition e uid new = (fimage (\(uids, (from, to), t). if set uid \ set uids then (uids, (from, to), new) else (uids, (from, to), t)) e)" definition replace_all :: "iEFSM \ tids list \ transition \ iEFSM" where "replace_all e ids new = fold (\id acc. replace_transition acc id new) ids e" definition replace_transitions :: "iEFSM \ (tids \ transition) list \ iEFSM" where "replace_transitions e ts = fold (\(uid, new) acc. replace_transition acc uid new) ts e" primrec try_heuristics_check :: "(transition_matrix \ bool) \ update_modifier list \ update_modifier" where "try_heuristics_check _ [] = null_modifier" | "try_heuristics_check check (h#t) = (\a b c d e f ch. case h a b c d e f ch of Some e' \ Some e' | None \ (try_heuristics_check check t) a b c d e f ch )" subsection\Scoring State Merges\ text\To tackle the state merging challenge, we need some means of determining which states are compatible for merging. Because states are merged pairwise, we additionally require a way of ordering the state merges. The potential merges are then sorted highest to lowest according to this score such that we can merge states in order of their merge score. We want to sort first by score (highest to lowest) and then by state pairs (lowest to highest) so we endup merging the states with the highest scores first and then break ties by those state pairs which are closest to the origin.\ record score = Score :: nat S1 :: cfstate S2 :: cfstate instantiation score_ext :: (linorder) linorder begin definition less_score_ext :: "'a::linorder score_ext \ 'a score_ext \ bool" where "less_score_ext t1 t2 = ((Score t2, S1 t1, S2 t1, more t1) < (Score t1, S1 t2, S2 t2, more t2) )" definition less_eq_score_ext :: "'a::linorder score_ext \ 'a::linorder score_ext \ bool" where "less_eq_score_ext s1 s2 = (s1 < s2 \ s1 = s2)" instance apply standard prefer 5 unfolding less_score_ext_def less_eq_score_ext_def using score.equality apply fastforce by auto end type_synonym scoreboard = "score fset" type_synonym strategy = "tids \ tids \ iEFSM \ nat" definition outgoing_transitions :: "cfstate \ iEFSM \ (cfstate \ transition \ tids) fset" where "outgoing_transitions s e = fimage (\(uid, (from, to), t'). (to, t', uid)) ((ffilter (\(uid, (origin, dest), t). origin = s)) e)" primrec paths_of_length :: "nat \ iEFSM \ cfstate \ tids list fset" where "paths_of_length 0 _ _ = {|[]|}" | "paths_of_length (Suc m) e s = ( let outgoing = outgoing_transitions s e; paths = ffUnion (fimage (\(d, t, id). fimage (\p. id#p) (paths_of_length m e d)) outgoing) in ffilter (\l. length l = Suc m) paths )" lemma paths_of_length_1: "paths_of_length 1 e s = fimage (\(d, t, id). [id]) (outgoing_transitions s e)" apply (simp add: One_nat_def) apply (simp add: outgoing_transitions_def comp_def One_nat_def[symmetric]) apply (rule fBall_ffilter2) defer - apply (simp add: ffilter_def ffUnion_def fBall_def Abs_fset_inverse) + apply (simp add: ffilter_def ffUnion_def Abs_fset_inverse) apply auto[1] - apply (simp add: ffilter_def ffUnion_def fBall_def Abs_fset_inverse fset_both_sides) + apply (simp add: ffilter_def ffUnion_def Abs_fset_inverse fset_both_sides) by force fun step_score :: "(tids \ tids) list \ iEFSM \ strategy \ nat" where "step_score [] _ _ = 0" | "step_score ((id1, id2)#t) e s = ( let score = s id1 id2 e in if score = 0 then 0 else score + (step_score t e s) )" lemma step_score_foldr [code]: "step_score xs e s = foldr (\(id1, id2) acc. let score = s id1 id2 e in if score = 0 then 0 else score + acc) xs 0" proof(induct xs) case Nil then show ?case by simp next case (Cons a xs) then show ?case apply (cases a, clarify) by (simp add: Let_def) qed definition score_from_list :: "tids list fset \ tids list fset \ iEFSM \ strategy \ nat" where "score_from_list P1 P2 e s = ( let pairs = fimage (\(l1, l2). zip l1 l2) (P1 |\| P2); scored_pairs = fimage (\l. step_score l e s) pairs in fSum scored_pairs )" definition k_score :: "nat \ iEFSM \ strategy \ scoreboard" where "k_score k e strat = ( let states = S e; pairs_to_score = (ffilter (\(x, y). x < y) (states |\| states)); paths = fimage (\(s1, s2). (s1, s2, paths_of_length k e s1, paths_of_length k e s2)) pairs_to_score; scores = fimage (\(s1, s2, p1, p2). \Score = score_from_list p1 p2 e strat, S1 = s1, S2 = s2\) paths in ffilter (\x. Score x > 0) scores )" definition score_state_pair :: "strategy \ iEFSM \ cfstate \ cfstate \ nat" where "score_state_pair strat e s1 s2 = ( let T1 = outgoing_transitions s1 e; T2 = outgoing_transitions s2 e in fSum (fimage (\((_, _, t1), (_, _, t2)). strat t1 t2 e) (T1 |\| T2)) )" definition score_1 :: "iEFSM \ strategy \ scoreboard" where "score_1 e strat = ( let states = S e; pairs_to_score = (ffilter (\(x, y). x < y) (states |\| states)); scores = fimage (\(s1, s2). \Score = score_state_pair strat e s1 s2, S1 = s1, S2 = s2\) pairs_to_score in ffilter (\x. Score x > 0) scores )" lemma score_1: "score_1 e s = k_score 1 e s" proof- have fprod_fimage: "\a b. ((\(_, _, id). [id]) |`| a |\| (\(_, _, id). [id]) |`| b) = fimage (\((_, _, id1), (_, _, id2)). ([id1], [id2])) (a |\| b)" apply (simp add: fimage_def fprod_def Abs_fset_inverse fset_both_sides) by force show ?thesis apply (simp add: score_1_def k_score_def Let_def comp_def) apply (rule arg_cong[of _ _ "ffilter (\x. 0 < Score x)"]) apply (rule fun_cong[of _ _ "(Inference.S e |\| Inference.S e)"]) apply (rule ext) subgoal for x apply (rule fun_cong[of _ _ "ffilter (\a. case a of (a, b) \ a < b) x"]) apply (rule arg_cong[of _ _ fimage]) apply (rule ext) subgoal for x apply (case_tac x) apply simp apply (simp add: paths_of_length_1) apply (simp add: score_state_pair_def Let_def score_from_list_def comp_def) subgoal for a b apply (rule arg_cong[of _ _ fSum]) apply (simp add: fprod_fimage) apply (rule fun_cong[of _ _ "(outgoing_transitions a e |\| outgoing_transitions b e)"]) apply (rule arg_cong[of _ _ fimage]) apply (rule ext) apply clarify by (simp add: Let_def) done done done qed fun bool2nat :: "bool \ nat" where "bool2nat True = 1" | "bool2nat False = 0" definition score_transitions :: "transition \ transition \ nat" where "score_transitions t1 t2 = ( if Label t1 = Label t2 \ Arity t1 = Arity t2 \ length (Outputs t1) = length (Outputs t2) then 1 + bool2nat (t1 = t2) + card ((set (Guards t2)) \ (set (Guards t2))) + card ((set (Updates t2)) \ (set (Updates t2))) + card ((set (Outputs t2)) \ (set (Outputs t2))) else 0 )" subsection\Merging States\ definition merge_states_aux :: "nat \ nat \ iEFSM \ iEFSM" where "merge_states_aux s1 s2 e = fimage (\(uid, (origin, dest), t). (uid, (if origin = s1 then s2 else origin , if dest = s1 then s2 else dest), t)) e" definition merge_states :: "nat \ nat \ iEFSM \ iEFSM" where "merge_states x y t = (if x > y then merge_states_aux x y t else merge_states_aux y x t)" lemma merge_states_symmetry: "merge_states x y t = merge_states y x t" by (simp add: merge_states_def) lemma merge_state_self: "merge_states s s t = t" apply (simp add: merge_states_def merge_states_aux_def) by force lemma merge_states_self_simp [code]: "merge_states x y t = (if x = y then t else if x > y then merge_states_aux x y t else merge_states_aux y x t)" apply (simp add: merge_states_def merge_states_aux_def) by force subsection\Resolving Nondeterminism\ text\Because EFSM transitions are not simply atomic actions, duplicated behaviours cannot be resolved into a single transition by simply merging destination states, as it can in classical FSM inference. It is now possible for attempts to resolve the nondeterminism introduced by merging states to fail, meaning that two states which initially seemed compatible cannot actually be merged. This is not the case in classical FSM inference.\ type_synonym nondeterministic_pair = "(cfstate \ (cfstate \ cfstate) \ ((transition \ tids) \ (transition \ tids)))" definition state_nondeterminism :: "nat \ (cfstate \ transition \ tids) fset \ nondeterministic_pair fset" where "state_nondeterminism og nt = (if size nt < 2 then {||} else ffUnion (fimage (\x. let (dest, t) = x in fimage (\y. let (dest', t') = y in (og, (dest, dest'), (t, t'))) (nt - {|x|})) nt))" lemma state_nondeterminism_empty [simp]: "state_nondeterminism a {||} = {||}" by (simp add: state_nondeterminism_def ffilter_def Set.filter_def) lemma state_nondeterminism_singledestn [simp]: "state_nondeterminism a {|x|} = {||}" by (simp add: state_nondeterminism_def ffilter_def Set.filter_def) (* For each state, get its outgoing transitions and see if there's any nondeterminism there *) definition nondeterministic_pairs :: "iEFSM \ nondeterministic_pair fset" where "nondeterministic_pairs t = ffilter (\(_, _, (t, _), (t', _)). Label t = Label t' \ Arity t = Arity t' \ choice t t') (ffUnion (fimage (\s. state_nondeterminism s (outgoing_transitions s t)) (S t)))" definition nondeterministic_pairs_labar_dest :: "iEFSM \ nondeterministic_pair fset" where "nondeterministic_pairs_labar_dest t = ffilter (\(_, (d, d'), (t, _), (t', _)). Label t = Label t' \ Arity t = Arity t' \ (choice t t' \ (Outputs t = Outputs t' \ d = d'))) (ffUnion (fimage (\s. state_nondeterminism s (outgoing_transitions s t)) (S t)))" definition nondeterministic_pairs_labar :: "iEFSM \ nondeterministic_pair fset" where "nondeterministic_pairs_labar t = ffilter (\(_, (d, d'), (t, _), (t', _)). Label t = Label t' \ Arity t = Arity t' \ (choice t t' \ Outputs t = Outputs t')) (ffUnion (fimage (\s. state_nondeterminism s (outgoing_transitions s t)) (S t)))" definition deterministic :: "iEFSM \ (iEFSM \ nondeterministic_pair fset) \ bool" where "deterministic t np = (np t = {||})" definition nondeterministic :: "iEFSM \ (iEFSM \ nondeterministic_pair fset) \ bool" where "nondeterministic t np = (\ deterministic t np)" definition insert_transition :: "tids \ cfstate \ cfstate \ transition \ iEFSM \ iEFSM" where "insert_transition uid from to t e = ( if \(uid, (from', to'), t') |\| e. from = from' \ to = to' \ t = t' then finsert (uid, (from, to), t) e else fimage (\(uid', (from', to'), t'). if from = from' \ to = to' \ t = t' then (List.union uid' uid, (from', to'), t') else (uid', (from', to'), t') ) e )" definition make_distinct :: "iEFSM \ iEFSM" where "make_distinct e = ffold_ord (\(uid, (from, to), t) acc. insert_transition uid from to t acc) e {||}" \ \When we replace one transition with another, we need to merge their uids to keep track of which\ \ \transition accounts for which action in the original traces \ definition merge_transitions_aux :: "iEFSM \ tids \ tids \ iEFSM" where "merge_transitions_aux e oldID newID = (let (uids1, (origin, dest), old) = fthe_elem (ffilter (\(uids, _). oldID = uids) e); (uids2, (origin, dest), new) = fthe_elem (ffilter (\(uids, _). newID = uids) e) in make_distinct (finsert (List.union uids1 uids2, (origin, dest), new) (e - {|(uids1, (origin, dest), old), (uids2, (origin, dest), new)|})) )" (* merge_transitions - Try dest merge transitions t1 and t2 dest help resolve nondeterminism in newEFSM. If either subsumes the other directly then the subsumed transition can simply be replaced with the subsuming one, else we try dest apply the modifier function dest resolve nondeterminism that way. *) (* @param oldEFSM - the EFSM before merging the states which caused the nondeterminism *) (* @param preDestMerge - the EFSM after merging the states which caused the nondeterminism *) (* @param newEFSM - the current EFSM with nondeterminism *) (* @param t1 - a transition dest be merged with t2 *) (* @param u1 - the unique identifier of t1 *) (* @param t2 - a transition dest be merged with t1 *) (* @param u2 - the unique identifier of t2 *) (* @param modifier - an update modifier function which tries dest generalise transitions *) definition merge_transitions :: "(cfstate \ cfstate) set \ iEFSM \ iEFSM \ iEFSM \ transition \ tids \ transition \ tids \ update_modifier \ (transition_matrix \ bool) \ iEFSM option" where "merge_transitions failedMerges oldEFSM preDestMerge destMerge t1 u1 t2 u2 modifier check = ( if \id \ set u1. directly_subsumes (tm oldEFSM) (tm destMerge) (origin [id] oldEFSM) (origin u1 destMerge) t2 t1 then \ \Replace t1 with t2\ Some (merge_transitions_aux destMerge u1 u2) else if \id \ set u2. directly_subsumes (tm oldEFSM) (tm destMerge) (origin [id] oldEFSM) (origin u2 destMerge) t1 t2 then \ \Replace t2 with t1\ Some (merge_transitions_aux destMerge u2 u1) else case modifier u1 u2 (origin u1 destMerge) destMerge preDestMerge oldEFSM check of None \ None | Some e \ Some (make_distinct e) )" definition outgoing_transitions_from :: "iEFSM \ cfstate \ transition fset" where "outgoing_transitions_from e s = fimage (\(_, _, t). t) (ffilter (\(_, (orig, _), _). orig = s) e)" definition order_nondeterministic_pairs :: "nondeterministic_pair fset \ nondeterministic_pair list" where "order_nondeterministic_pairs s = map snd (sorted_list_of_fset (fimage (\s. let (_, _, (t1, _), (t2, _)) = s in (score_transitions t1 t2, s)) s))" (* resolve_nondeterminism - tries dest resolve nondeterminism in a given iEFSM *) (* @param ((from, (dest1, dest2), ((t1, u1), (t2, u2)))#ss) - a list of nondeterministic pairs where from - nat - the state from which t1 and t2 eminate dest1 - nat - the destination state of t1 dest2 - nat - the destination state of t2 t1 - transition - a transition dest be merged with t2 t2 - transition - a transition dest be merged with t1 u1 - nat - the unique identifier of t1 u2 - nat - the unique identifier of t2 ss - list - the rest of the list *) (* @param oldEFSM - the EFSM before merging the states which caused the nondeterminism *) (* @param newEFSM - the current EFSM with nondeterminism *) (* @param m - an update modifier function which tries dest generalise transitions *) (* @param check - a function which takes an EFSM and returns a bool dest ensure that certain properties hold in the new iEFSM *) function resolve_nondeterminism :: "(cfstate \ cfstate) set \ nondeterministic_pair list \ iEFSM \ iEFSM \ update_modifier \ (transition_matrix \ bool) \ (iEFSM \ nondeterministic_pair fset) \ (iEFSM option \ (cfstate \ cfstate) set)" where "resolve_nondeterminism failedMerges [] _ newEFSM _ check np = ( if deterministic newEFSM np \ check (tm newEFSM) then Some newEFSM else None, failedMerges )" | "resolve_nondeterminism failedMerges ((from, (dest1, dest2), ((t1, u1), (t2, u2)))#ss) oldEFSM newEFSM m check np = ( if (dest1, dest2) \ failedMerges \ (dest2, dest1) \ failedMerges then (None, failedMerges) else let destMerge = merge_states dest1 dest2 newEFSM in case merge_transitions failedMerges oldEFSM newEFSM destMerge t1 u1 t2 u2 m check of None \ resolve_nondeterminism (insert (dest1, dest2) failedMerges) ss oldEFSM newEFSM m check np | Some new \ ( let newScores = order_nondeterministic_pairs (np new) in if (size new, size (S new), size (newScores)) < (size newEFSM, size (S newEFSM), size ss) then case resolve_nondeterminism failedMerges newScores oldEFSM new m check np of (Some new', failedMerges) \ (Some new', failedMerges) | (None, failedMerges) \ resolve_nondeterminism (insert (dest1, dest2) failedMerges) ss oldEFSM newEFSM m check np else (None, failedMerges) ) )" apply (clarify, metis neq_Nil_conv prod_cases3 surj_pair) by auto termination by (relation "measures [\(_, _, _, newEFSM, _). size newEFSM, \(_, _, _, newEFSM, _). size (S newEFSM), \(_, ss, _, _, _). size ss]", auto) subsection\EFSM Inference\ (* Merge - tries dest merge two states in a given iEFSM and resolve the resulting nondeterminism *) (* @param e - an iEFSM *) (* @param s1 - a state dest be merged with s2 *) (* @param s2 - a state dest be merged with s1 *) (* @param m - an update modifier function which tries dest generalise transitions *) (* @param check - a function which takes an EFSM and returns a bool dest ensure that certain properties hold in the new iEFSM *) definition merge :: "(cfstate \ cfstate) set \ iEFSM \ nat \ nat \ update_modifier \ (transition_matrix \ bool) \ (iEFSM \ nondeterministic_pair fset) \ (iEFSM option \ (cfstate \ cfstate) set)" where "merge failedMerges e s1 s2 m check np = ( if s1 = s2 \ (s1, s2) \ failedMerges \ (s2, s1) \ failedMerges then (None, failedMerges) else let e' = make_distinct (merge_states s1 s2 e) in resolve_nondeterminism failedMerges (order_nondeterministic_pairs (np e')) e e' m check np )" (* inference_step - attempt dest carry out a single step of the inference process by merging the *) (* @param e - an iEFSM dest be generalised *) (* @param ((s, s1, s2)#t) - a list of triples of the form (score, state, state) dest be merged *) (* @param m - an update modifier function which tries dest generalise transitions *) (* @param check - a function which takes an EFSM and returns a bool dest ensure that certain properties hold in the new iEFSM *) function inference_step :: "(cfstate \ cfstate) set \ iEFSM \ score fset \ update_modifier \ (transition_matrix \ bool) \ (iEFSM \ nondeterministic_pair fset) \ (iEFSM option \ (cfstate \ cfstate) set)" where "inference_step failedMerges e s m check np = ( if s = {||} then (None, failedMerges) else let h = fMin s; t = s - {|h|} in case merge failedMerges e (S1 h) (S2 h) m check np of (Some new, failedMerges) \ (Some new, failedMerges) | (None, failedMerges) \ inference_step (insert ((S1 h), (S2 h)) failedMerges) e t m check np )" by auto termination by (relation "measures [\(_, _, s, _, _, _). size s]") (auto dest!: card_minus_fMin) (* Takes an iEFSM and iterates inference_step until no further states can be successfully merged *) (* @param e - an iEFSM dest be generalised *) (* @param r - a strategy dest identify and prioritise pairs of states dest merge *) (* @param m - an update modifier function which tries dest generalise transitions *) (* @param check - a function which takes an EFSM and returns a bool dest ensure that certain properties hold in the new iEFSM *) function infer :: "(cfstate \ cfstate) set \ nat \ iEFSM \ strategy \ update_modifier \ (transition_matrix \ bool) \ (iEFSM \ nondeterministic_pair fset) \ iEFSM" where "infer failedMerges k e r m check np = ( let scores = if k = 1 then score_1 e r else (k_score k e r) in case inference_step failedMerges e (ffilter (\s. (S1 s, S2 s) \ failedMerges \ (S2 s, S1 s) \ failedMerges) scores) m check np of (None, _) \ e | (Some new, failedMerges) \ if (S new) |\| (S e) then infer failedMerges k new r m check np else e )" by auto termination apply (relation "measures [\(_, _, e, _). size (S e)]") apply simp by (metis (no_types, lifting) case_prod_conv measures_less size_fsubset) fun get_ints :: "trace \ int list" where "get_ints [] = []" | "get_ints ((_, inputs, outputs)#t) = (map (\x. case x of Num n \ n) (filter is_Num (inputs@outputs)))" definition learn :: "nat \ iEFSM \ log \ strategy \ update_modifier \ (iEFSM \ nondeterministic_pair fset) \ iEFSM" where "learn n pta l r m np = ( let check = accepts_log (set l) in (infer {} n pta r m check np) )" subsection\Evaluating Inferred Models\ text\We need a function to test the EFSMs we infer. The \texttt{test\_trace} function executes a trace in the model and outputs a more comprehensive trace such that the expected outputs and actual outputs can be compared. If a point is reached where the model does not recognise an action, the remainder of the trace forms the second element of the output pair such that we know the exact point at which the model stopped processing.\ definition i_possible_steps :: "iEFSM \ cfstate \ registers \ label \ inputs \ (tids \ cfstate \ transition) fset" where "i_possible_steps e s r l i = fimage (\(uid, (origin, dest), t). (uid, dest, t)) (ffilter (\(uid, (origin, dest::nat), t::transition). origin = s \ (Label t) = l \ (length i) = (Arity t) \ apply_guards (Guards t) (join_ir i r) ) e)" fun test_trace :: "trace \ iEFSM \ cfstate \ registers \ ((label \ inputs \ cfstate \ cfstate \ registers \ tids \ value list \ outputs) list \ trace)" where "test_trace [] _ _ _ = ([], [])" | "test_trace ((l, i, expected)#es) e s r = ( let ps = i_possible_steps e s r l i in if fis_singleton ps then let (id, s', t) = fthe_elem ps; r' = evaluate_updates t i r; actual = evaluate_outputs t i r; (est, fail) = (test_trace es e s' r') in ((l, i, s, s', r, id, expected, actual)#est, fail) else ([], (l, i, expected)#es) )" text\The \texttt{test\_log} function executes the \texttt{test\_trace} function on a collection of traces known as the \emph{test set.}\ definition test_log :: "log \ iEFSM \ ((label \ inputs \ cfstate \ cfstate \ registers \ tids \ value list \ outputs) list \ trace) list" where "test_log l e = map (\t. test_trace t e 0 <>) l" end diff --git a/thys/Extended_Finite_State_Machine_Inference/code-targets/Code_Target_FSet.thy b/thys/Extended_Finite_State_Machine_Inference/code-targets/Code_Target_FSet.thy --- a/thys/Extended_Finite_State_Machine_Inference/code-targets/Code_Target_FSet.thy +++ b/thys/Extended_Finite_State_Machine_Inference/code-targets/Code_Target_FSet.thy @@ -1,164 +1,155 @@ section\Finite Sets\ text\Here we define the operations on the \texttt{fset} datatype in terms of lists rather than sets. This allows the Scala implementation to skip a case match each time, which makes for cleaner and slightly faster code.\ theory Code_Target_FSet imports "Extended_Finite_State_Machines.FSet_Utils" begin code_datatype fset_of_list declare FSet.fset_of_list.rep_eq [code] lemma fprod_code [code]: "fprod (fset_of_list xs) (fset_of_list ys) = fset_of_list (remdups [(x, y). x \ xs, y \ ys])" apply (simp add: fprod_def fset_of_list_def fset_both_sides Abs_fset_inverse) by auto lemma fminus_fset_filter [code]: "fset_of_list A - xs = fset_of_list (remdups (filter (\x. x |\| xs) A))" by auto lemma sup_fset_fold [code]: "(fset_of_list f1) |\| (fset_of_list f2) = fset_of_list (remdups (f1@f2))" by simp lemma bot_fset [code]: "{||} = fset_of_list []" by simp lemma finsert [code]: "finsert a (fset_of_list as) = fset_of_list (List.insert a as)" by (simp add: List.insert_def finsert_absorb fset_of_list_elem) lemma ffilter_filter [code]: "ffilter f (fset_of_list as) = fset_of_list (List.filter f (remdups as))" by simp lemma fimage_map [code]: "fimage f (fset_of_list as) = fset_of_list (List.map f (remdups as))" by simp lemma ffUnion_fold [code]: "ffUnion (fset_of_list as) = fold (|\|) as {||}" by (simp add: fold_union_ffUnion) lemma fmember [code]: "a |\| (fset_of_list as) = List.member as a" by (simp add: fset_of_list_elem member_def) lemma fthe_elem [code]: "fthe_elem (fset_of_list [x]) = x" by simp lemma size [code]: "size (fset_of_list as) = length (remdups as)" proof(induct as) case Nil then show ?case by simp next case (Cons a as) then show ?case by (simp add: fset_of_list.rep_eq insert_absorb) qed lemma fMax_fold [code]: "fMax (fset_of_list (a#as)) = fold max as a" by (metis Max.set_eq_fold fMax.F.rep_eq fset_of_list.rep_eq) lemma fMin_fold [code]: "fMin (fset_of_list (h#t)) = fold min t h" apply (simp add: fset_of_list_def) by (metis Min.set_eq_fold fMin_Min fset_of_list.abs_eq list.simps(15)) lemma fremove_code [code]: "fremove a (fset_of_list A) = fset_of_list (filter (\x. x \ a) A)" apply (simp add: fremove_def minus_fset_def ffilter_def fset_both_sides Abs_fset_inverse fset_of_list.rep_eq) by auto lemma fsubseteq [code]: "(fset_of_list l) |\| A = List.list_all (\x. x |\| A) l" by (induct l, auto) lemma fsum_fold [code]: "fSum (fset_of_list l) = fold (+) (remdups l) 0" proof(induct l) case Nil then show ?case by (simp add: fsum.F.rep_eq fSum_def) next case (Cons a l) then show ?case apply simp apply standard apply (simp add: finsert_absorb fset_of_list_elem) by (simp add: add.commute fold_plus_sum_list_rev fset_of_list.rep_eq fsum.F.rep_eq fSum_def) qed lemma code_fset_eq [code]: "HOL.equal X (fset_of_list Y) \ size X = length (remdups Y) \ (\x |\| X. List.member Y x)" apply (simp only: HOL.equal_class.equal_eq fset_eq_alt) apply (simp only: size) using fmember by fastforce lemma code_fsubset [code]: "s |\| s' = (s |\| s' \ size s < size s')" apply standard apply (simp only: size_fsubset) by auto -lemma code_fset [code]: "fset (fset_of_list l) = fold insert l {}" - using fset_of_list.rep_eq union_set_fold by fastforce - -lemma code_fBall [code]: "fBall (fset_of_list l) f = list_all f l" - by (simp add: Ball_set fBall.rep_eq fset_of_list.rep_eq) - -lemma code_fBex [code]: "fBex (fset_of_list l) f = list_ex f l" - by (meson Bex_set fBexE fset_of_list_elem rev_fBexI) - definition "nativeSort = sort" code_printing constant nativeSort \ (Scala) "_.sortWith((Orderings.less))" lemma [code]: "sorted_list_of_fset (fset_of_list l) = nativeSort (remdups l)" by (simp add: nativeSort_def sorted_list_of_fset_sort) lemma [code]: "sorted_list_of_set (set l) = nativeSort (remdups l)" by (simp add: nativeSort_def sorted_list_of_set_sort_remdups) lemma [code]: "fMin (fset_of_list (h#t)) = hd (nativeSort (h#t))" by (metis fMin_Min hd_sort_Min list.distinct(1) nativeSort_def) lemma sorted_Max_Cons: "l \ [] \ sorted (a#l) \ Max (set (a#l)) = Max (set l)" using eq_iff by fastforce lemma sorted_Max: "l \ [] \ sorted l \ Max (set l) = hd (rev l)" proof(induct l) case Nil then show ?case by simp next case (Cons a l) then show ?case by (metis sorted_Max_Cons Max_singleton hd_rev last.simps list.set(1) list.simps(15) sorted_simps(2)) qed lemma [code]: "fMax (fset_of_list (h#t)) = last (nativeSort (h#t))" by (metis Max.set_eq_fold fMax_fold hd_rev list.simps(3) nativeSort_def set_empty2 set_sort sorted_Max sorted_sort) definition "list_max l = fold max l" code_printing constant list_max \ (Scala) "_.par.fold((_))(Orderings.max)" lemma [code]: "fMax (fset_of_list (h#t)) = list_max t h" by (metis fMax_fold list_max_def) definition "list_min l = fold min l" code_printing constant list_min \ (Scala) "_.par.fold((_))(Orderings.min)" lemma [code]: "fMin (fset_of_list (h#t)) = list_min t h" by (metis fMin_fold list_min_def) lemma fis_singleton_code [code]: "fis_singleton s = (size s = 1)" apply (simp add: fis_singleton_def is_singleton_def) by (simp add: card_Suc_eq) end diff --git a/thys/Extended_Finite_State_Machines/EFSM.thy b/thys/Extended_Finite_State_Machines/EFSM.thy --- a/thys/Extended_Finite_State_Machines/EFSM.thy +++ b/thys/Extended_Finite_State_Machines/EFSM.thy @@ -1,1492 +1,1492 @@ section \Extended Finite State Machines\ text\This theory defines extended finite state machines as presented in \<^cite>\"foster2018"\. States are indexed by natural numbers, however, since transition matrices are implemented by finite sets, the number of reachable states in $S$ is necessarily finite. For ease of implementation, we implicitly make the initial state zero for all EFSMs. This allows EFSMs to be represented purely by their transition matrix which, in this implementation, is a finite set of tuples of the form $((s_1, s_2), t)$ in which $s_1$ is the origin state, $s_2$ is the destination state, and $t$ is a transition.\ theory EFSM imports "HOL-Library.FSet" Transition FSet_Utils begin declare One_nat_def [simp del] type_synonym cfstate = nat type_synonym inputs = "value list" type_synonym outputs = "value option list" type_synonym action = "(label \ inputs)" type_synonym execution = "action list" type_synonym observation = "outputs list" type_synonym transition_matrix = "((cfstate \ cfstate) \ transition) fset" no_notation relcomp (infixr "O" 75) and comp (infixl "o" 55) type_synonym event = "(label \ inputs \ value list)" type_synonym trace = "event list" type_synonym log = "trace list" definition Str :: "string \ value" where "Str s \ value.Str (String.implode s)" lemma str_not_num: "Str s \ Num x1" by (simp add: Str_def) definition S :: "transition_matrix \ nat fset" where "S m = (fimage (\((s, s'), t). s) m) |\| fimage (\((s, s'), t). s') m" lemma S_ffUnion: "S e = ffUnion (fimage (\((s, s'), _). {|s, s'|}) e)" unfolding S_def by(induct e, auto) subsection\Possible Steps\ text\From a given state, the possible steps for a given action are those transitions with labels which correspond to the action label, arities which correspond to the number of inputs, and guards which are satisfied by those inputs.\ definition possible_steps :: "transition_matrix \ cfstate \ registers \ label \ inputs \ (cfstate \ transition) fset" where "possible_steps e s r l i = fimage (\((origin, dest), t). (dest, t)) (ffilter (\((origin, dest), t). origin = s \ (Label t) = l \ (length i) = (Arity t) \ apply_guards (Guards t) (join_ir i r)) e)" lemma possible_steps_finsert: "possible_steps (finsert ((s, s'), t) e) ss r l i = ( if s = ss \ (Label t) = l \ (length i) = (Arity t) \ apply_guards (Guards t) (join_ir i r) then finsert (s', t) (possible_steps e s r l i) else possible_steps e ss r l i )" by (simp add: possible_steps_def ffilter_finsert) lemma split_origin: "ffilter (\((origin, dest), t). origin = s \ Label t = l \ can_take_transition t i r) e = ffilter (\((origin, dest), t). Label t = l \ can_take_transition t i r) (ffilter (\((origin, dest), t). origin = s) e)" by auto lemma split_label: "ffilter (\((origin, dest), t). origin = s \ Label t = l \ can_take_transition t i r) e = ffilter (\((origin, dest), t). origin = s \ can_take_transition t i r) (ffilter (\((origin, dest), t). Label t = l) e)" by auto lemma possible_steps_empty_guards_false: "\((s1, s2), t) |\| ffilter (\((origin, dest), t). Label t = l) e. \can_take_transition t i r \ possible_steps e s r l i = {||}" apply (simp add: possible_steps_def can_take[symmetric] split_label) - by (simp add: Abs_ffilter fBall_def Ball_def) + by (simp add: Abs_ffilter Ball_def) lemma fmember_possible_steps: "(s', t) |\| possible_steps e s r l i = (((s, s'), t) \ {((origin, dest), t) \ fset e. origin = s \ Label t = l \ length i = Arity t \ apply_guards (Guards t) (join_ir i r)})" apply (simp add: possible_steps_def ffilter_def fimage_def Abs_fset_inverse) by force lemma possible_steps_alt_aux: "possible_steps e s r l i = {|(d, t)|} \ ffilter (\((origin, dest), t). origin = s \ Label t = l \ length i = Arity t \ apply_guards (Guards t) (join_ir i r)) e = {|((s, d), t)|}" proof(induct e) case empty then show ?case by (simp add: fempty_not_finsert possible_steps_def) next case (insert x e) then show ?case apply (case_tac x) subgoal for a b apply (case_tac a) subgoal for aa _ apply (simp add: possible_steps_def) apply (simp add: ffilter_finsert) apply (case_tac "aa = s \ Label b = l \ length i = Arity b \ apply_guards (Guards b) (join_ir i r)") by auto done done qed lemma possible_steps_alt: "(possible_steps e s r l i = {|(d, t)|}) = (ffilter (\((origin, dest), t). origin = s \ Label t = l \ length i = Arity t \ apply_guards (Guards t) (join_ir i r)) e = {|((s, d), t)|})" apply standard apply (simp add: possible_steps_alt_aux) by (simp add: possible_steps_def) lemma possible_steps_alt3: "(possible_steps e s r l i = {|(d, t)|}) = (ffilter (\((origin, dest), t). origin = s \ Label t = l \ can_take_transition t i r) e = {|((s, d), t)|})" apply standard apply (simp add: possible_steps_alt_aux can_take) by (simp add: possible_steps_def can_take) lemma possible_steps_alt_atom: "(possible_steps e s r l i = {|dt|}) = (ffilter (\((origin, dest), t). origin = s \ Label t = l \ can_take_transition t i r) e = {|((s, fst dt), snd dt)|})" apply (cases dt) by (simp add: possible_steps_alt can_take_transition_def can_take_def) lemma possible_steps_alt2: "(possible_steps e s r l i = {|(d, t)|}) = ( (ffilter (\((origin, dest), t). Label t = l \ length i = Arity t \ apply_guards (Guards t) (join_ir i r)) (ffilter (\((origin, dest), t). origin = s) e) = {|((s, d), t)|}))" apply (simp add: possible_steps_alt) apply (simp only: filter_filter) apply (rule arg_cong [of "(\((origin, dest), t). origin = s \ Label t = l \ length i = Arity t \ apply_guards (Guards t) (join_ir i r))"]) by (rule ext, auto) lemma possible_steps_single_out: "ffilter (\((origin, dest), t). origin = s) e = {|((s, d), t)|} \ Label t = l \ length i = Arity t \ apply_guards (Guards t) (join_ir i r) \ possible_steps e s r l i = {|(d, t)|}" apply (simp add: possible_steps_alt2 Abs_ffilter) by blast lemma possible_steps_singleton: "(possible_steps e s r l i = {|(d, t)|}) = ({((origin, dest), t) \ fset e. origin = s \ Label t = l \ length i = Arity t \ apply_guards (Guards t) (join_ir i r)} = {((s, d), t)})" apply (simp add: possible_steps_alt Abs_ffilter Set.filter_def) by fast lemma possible_steps_apply_guards: "possible_steps e s r l i = {|(s', t)|} \ apply_guards (Guards t) (join_ir i r)" apply (simp add: possible_steps_singleton) by auto lemma possible_steps_empty: "(possible_steps e s r l i = {||}) = (\((origin, dest), t) \ fset e. origin \ s \ Label t \ l \ \ can_take_transition t i r)" apply (simp add: can_take_transition_def can_take_def) apply (simp add: possible_steps_def Abs_ffilter Set.filter_def) by auto lemma singleton_dest: assumes "fis_singleton (possible_steps e s r aa b)" and "fthe_elem (possible_steps e s r aa b) = (baa, aba)" shows "((s, baa), aba) |\| e" using assms apply (simp add: fis_singleton_fthe_elem) using possible_steps_alt_aux by force lemma no_outgoing_transitions: "ffilter (\((s', _), _). s = s') e = {||} \ possible_steps e s r l i = {||}" apply (simp add: possible_steps_def) by (smt (verit, best) case_prod_beta eq_ffilter ffilter_empty ffmember_filter) lemma ffilter_split: "ffilter (\((origin, dest), t). origin = s \ Label t = l \ length i = Arity t \ apply_guards (Guards t) (join_ir i r)) e = ffilter (\((origin, dest), t). Label t = l \ length i = Arity t \ apply_guards (Guards t) (join_ir i r)) (ffilter (\((origin, dest), t). origin = s) e)" by auto lemma one_outgoing_transition: defines "outgoing s \ (\((origin, dest), t). origin = s)" assumes prem: "size (ffilter (outgoing s) e) = 1" shows "size (possible_steps e s r l i) \ 1" proof- have less_eq_1: "\x::nat. (x \ 1) = (x = 1 \ x = 0)" by auto have size_empty: "\f. (size f = 0) = (f = {||})" subgoal for f by (induct f, auto) done show ?thesis using prem apply (simp only: possible_steps_def) apply (rule fimage_size_le) apply (simp only: ffilter_split outgoing_def[symmetric]) by (metis (no_types, lifting) size_ffilter) qed subsection\Choice\ text\Here we define the \texttt{choice} operator which determines whether or not two transitions are nondeterministic.\ definition choice :: "transition \ transition \ bool" where "choice t t' = (\ i r. apply_guards (Guards t) (join_ir i r) \ apply_guards (Guards t') (join_ir i r))" definition choice_alt :: "transition \ transition \ bool" where "choice_alt t t' = (\ i r. apply_guards (Guards t@Guards t') (join_ir i r))" lemma choice_alt: "choice t t' = choice_alt t t'" by (simp add: choice_def choice_alt_def apply_guards_append) lemma choice_symmetry: "choice x y = choice y x" using choice_def by auto definition deterministic :: "transition_matrix \ bool" where "deterministic e = (\s r l i. size (possible_steps e s r l i) \ 1)" lemma deterministic_alt_aux: "size (possible_steps e s r l i) \ 1 =( possible_steps e s r l i = {||} \ (\s' t. ffilter (\((origin, dest), t). origin = s \ Label t = l \ length i = Arity t \ apply_guards (Guards t) (join_ir i r)) e = {|((s, s'), t)|}))" apply (case_tac "size (possible_steps e s r l i) = 0") apply (simp add: fset_equiv) apply (case_tac "possible_steps e s r l i = {||}") apply simp apply (simp only: possible_steps_alt[symmetric]) by (metis le_neq_implies_less le_numeral_extra(4) less_one prod.collapse size_fsingleton) lemma deterministic_alt: "deterministic e = ( \s r l i. possible_steps e s r l i = {||} \ (\s' t. ffilter (\((origin, dest), t). origin = s \ (Label t) = l \ (length i) = (Arity t) \ apply_guards (Guards t) (join_ir i r)) e = {|((s, s'), t)|}) )" using deterministic_alt_aux by (simp add: deterministic_def) lemma size_le_1: "size f \ 1 = (f = {||} \ (\e. f = {|e|}))" apply standard apply (metis bot.not_eq_extremum gr_implies_not0 le_neq_implies_less less_one size_fsingleton size_fsubset) by auto lemma ffilter_empty_if: "\x |\| xs. \ P x \ ffilter P xs = {||}" by auto lemma empty_ffilter: "ffilter P xs = {||} = (\x |\| xs. \ P x)" by auto lemma all_states_deterministic: "(\s l i r. ffilter (\((origin, dest), t). origin = s \ (Label t) = l \ can_take_transition t i r) e = {||} \ (\x. ffilter (\((origin, dest), t). origin = s \ (Label t) = l \ can_take_transition t i r) e = {|x|}) ) \ deterministic e" unfolding deterministic_def apply clarify subgoal for s r l i apply (erule_tac x=s in allE) apply (erule_tac x=l in allE) apply (erule_tac x=i in allE) apply (erule_tac x=r in allE) apply (simp only: size_le_1) apply (erule disjE) apply (rule_tac disjI1) apply (simp add: possible_steps_def can_take_transition_def can_take_def) apply (erule exE) subgoal for x apply (case_tac x) subgoal for a b apply (case_tac a) apply simp apply (induct e) apply auto[1] subgoal for _ _ _ ba apply (rule disjI2) apply (rule_tac x=ba in exI) apply (rule_tac x=b in exI) by (simp add: possible_steps_def can_take_transition_def[symmetric] can_take_def[symmetric]) done done done done lemma deterministic_finsert: "\i r l. \((a, b), t) |\| ffilter (\((origin, dest), t). origin = s) (finsert ((s, s'), t') e). Label t = l \ can_take_transition t i r \ \ can_take_transition t' i r \ deterministic e \ deterministic (finsert ((s, s'), t') e)" apply (simp add: deterministic_def possible_steps_finsert can_take del: size_fset_overloaded_simps) apply clarify subgoal for r i apply (erule_tac x=s in allE) apply (erule_tac x=r in allE) apply (erule_tac x="Label t'" in allE) apply (erule_tac x=i in allE) apply (erule_tac x=r in allE) apply (erule_tac x=i in allE) apply (erule_tac x="Label t'" in allE) by auto done lemma ffilter_fBall: "(\x |\| xs. P x) = (ffilter P xs = xs)" by auto lemma fsubset_if: "\x. x |\| f1 \ x |\| f2 \ f1 |\| f2" by auto lemma in_possible_steps: "(((s, s'), t)|\|e \ Label t = l \ can_take_transition t i r) = ((s', t) |\| possible_steps e s r l i)" apply (simp add: fmember_possible_steps) by (simp add: can_take_def can_take_transition_def) lemma possible_steps_can_take_transition: "(s2, t1) |\| possible_steps e1 s1 r l i \ can_take_transition t1 i r" using in_possible_steps by blast lemma not_deterministic: "\s l i r. \d1 d2 t1 t2. d1 \ d2 \ t1 \ t2 \ ((s, d1), t1) |\| e \ ((s, d2), t2) |\| e \ Label t1 = Label t2 \ can_take_transition t1 i r \ can_take_transition t2 i r \ \deterministic e" apply (simp add: deterministic_def not_le del: size_fset_overloaded_simps) apply clarify subgoal for s i r d1 d2 t1 t2 apply (rule_tac x=s in exI) apply (rule_tac x=r in exI) apply (rule_tac x="Label t1" in exI) apply (rule_tac x=i in exI) apply (case_tac "(d1, t1) |\| possible_steps e s r (Label t1) i") defer using in_possible_steps apply blast apply (case_tac "(d2, t2) |\| possible_steps e s r (Label t1) i") apply (metis fempty_iff fsingleton_iff not_le_imp_less prod.inject size_le_1) using in_possible_steps by force done lemma not_deterministic_conv: "\deterministic e \ \s l i r. \d1 d2 t1 t2. (d1 \ d2 \ t1 \ t2) \ ((s, d1), t1) |\| e \ ((s, d2), t2) |\| e \ Label t1 = Label t2 \ can_take_transition t1 i r \ can_take_transition t2 i r" apply (simp add: deterministic_def not_le del: size_fset_overloaded_simps) apply clarify subgoal for s r l i apply (case_tac "\e1 e2 f'. e1 \ e2 \ possible_steps e s r l i = finsert e1 (finsert e2 f')") defer using size_gt_1 apply blast apply (erule exE)+ subgoal for e1 e2 f' apply (case_tac e1, case_tac e2) subgoal for a b aa ba apply (simp del: size_fset_overloaded_simps) apply (rule_tac x=s in exI) apply (rule_tac x=i in exI) apply (rule_tac x=r in exI) apply (rule_tac x=a in exI) apply (rule_tac x=aa in exI) apply (rule_tac x=b in exI) apply (rule_tac x=ba in exI) by (metis finsertCI in_possible_steps) done done done lemma deterministic_if: "\s l i r. \d1 d2 t1 t2. (d1 \ d2 \ t1 \ t2) \ ((s, d1), t1) |\| e \ ((s, d2), t2) |\| e \ Label t1 = Label t2 \ can_take_transition t1 i r \ can_take_transition t2 i r \ deterministic e" using not_deterministic_conv by blast lemma "\l i r. (\((s, s'), t) |\| e. Label t = l \ can_take_transition t i r \ (\t' s''. ((s, s''), t') |\| e \ (s' \ s'' \ t' \ t) \ Label t' = l \ can_take_transition t' i r)) \ deterministic e" apply (simp add: deterministic_def del: size_fset_overloaded_simps) apply (rule allI)+ apply (simp only: size_le_1 possible_steps_empty) apply (case_tac "\t s'. ((s, s'), t)|\|e \ Label t = l \ can_take_transition t i r") defer apply fastforce apply (rule disjI2) apply clarify apply (rule_tac x="(s', t)" in exI) apply standard defer apply (meson fempty_fsubsetI finsert_fsubset in_possible_steps) apply standard apply (case_tac x) apply (simp add: in_possible_steps[symmetric]) apply (erule_tac x="Label t" in allE) apply (erule_tac x=i in allE) apply (erule_tac x=r in allE) apply (erule_tac x="((s, s'), t)" in fBallE) defer apply simp apply simp apply (erule_tac x=b in allE) apply simp apply (erule_tac x=a in allE) by simp definition "outgoing_transitions e s = ffilter (\((o, _), _). o = s) e" lemma in_outgoing: "((s1, s2), t) |\| outgoing_transitions e s = (((s1, s2), t) |\| e \ s1 = s)" by (auto simp add: outgoing_transitions_def) lemma outgoing_transitions_deterministic: "\s. \((s1, s2), t) |\| outgoing_transitions e s. \((s1', s2'), t') |\| outgoing_transitions e s. s2 \ s2' \ t \ t' \ Label t = Label t' \ \ choice t t' \ deterministic e" apply (rule deterministic_if) apply simp apply (rule allI) subgoal for s apply (erule_tac x=s in allE) - apply (simp add: fBall_def Ball_def) + apply (simp add: Ball_def) apply (rule allI)+ subgoal for i r d1 d2 t1 apply (erule_tac x=s in allE) apply (erule_tac x=d1 in allE) apply (erule_tac x=t1 in allE) apply (rule impI, rule allI) subgoal for t2 apply (case_tac "((s, d1), t1) \ fset (outgoing_transitions e s)") apply simp apply (erule_tac x=s in allE) apply (erule_tac x=d2 in allE) apply (erule_tac x=t2 in allE) apply (simp add: outgoing_transitions_def choice_def can_take) apply meson by (simp add: outgoing_transitions_def) done done done lemma outgoing_transitions_deterministic2: "(\s a b ba aa bb bc. ((a, b), ba) |\| outgoing_transitions e s \ ((aa, bb), bc) |\| (outgoing_transitions e s) - {|((a, b), ba)|} \ b \ bb \ ba \ bc \ \choice ba bc) \ deterministic e" apply (rule outgoing_transitions_deterministic) by blast lemma outgoing_transitions_fprod_deterministic: "(\s b ba bb bc. (((s, b), ba), ((s, bb), bc)) \ fset (outgoing_transitions e s) \ fset (outgoing_transitions e s) \ b \ bb \ ba \ bc \ Label ba = Label bc \ \choice ba bc) \ deterministic e" apply (rule outgoing_transitions_deterministic) apply clarify by (metis SigmaI in_outgoing) text\The \texttt{random\_member} function returns a random member from a finite set, or \texttt{None}, if the set is empty.\ definition random_member :: "'a fset \ 'a option" where "random_member f = (if f = {||} then None else Some (Eps (\x. x |\| f)))" lemma random_member_nonempty: "s \ {||} = (random_member s \ None)" by (simp add: random_member_def) lemma random_member_singleton [simp]: "random_member {|a|} = Some a" by (simp add: random_member_def) lemma random_member_is_member: "random_member ss = Some s \ s |\| ss" apply (simp add: random_member_def) by (metis equalsffemptyI option.distinct(1) option.inject verit_sko_ex_indirect) lemma random_member_None[simp]: "random_member ss = None = (ss = {||})" by (simp add: random_member_def) lemma random_member_empty[simp]: "random_member {||} = None" by simp definition step :: "transition_matrix \ cfstate \ registers \ label \ inputs \ (transition \ cfstate \ outputs \ registers) option" where "step e s r l i = (case random_member (possible_steps e s r l i) of None \ None | Some (s', t) \ Some (t, s', evaluate_outputs t i r, evaluate_updates t i r) )" lemma possible_steps_not_empty_iff: "step e s r a b \ None \ \aa ba. (aa, ba) |\| possible_steps e s r a b" apply (simp add: step_def) apply (case_tac "possible_steps e s r a b") apply (simp add: random_member_def) by auto lemma step_member: "step e s r l i = Some (t, s', p, r') \ (s', t) |\| possible_steps e s r l i" apply (simp add: step_def) apply (case_tac "random_member (possible_steps e s r l i)") apply simp subgoal for a by (case_tac a, simp add: random_member_is_member) done lemma step_outputs: "step e s r l i = Some (t, s', p, r') \ evaluate_outputs t i r = p" apply (simp add: step_def) apply (case_tac "random_member (possible_steps e s r l i)") by auto lemma step: "possibilities = (possible_steps e s r l i) \ random_member possibilities = Some (s', t) \ evaluate_outputs t i r = p \ evaluate_updates t i r = r' \ step e s r l i = Some (t, s', p, r')" by (simp add: step_def) lemma step_None: "step e s r l i = None = (possible_steps e s r l i = {||})" by (simp add: step_def prod.case_eq_if random_member_def) lemma step_Some: "step e s r l i = Some (t, s', p, r') = ( random_member (possible_steps e s r l i) = Some (s', t) \ evaluate_outputs t i r = p \ evaluate_updates t i r = r' )" apply (simp add: step_def) apply (case_tac "random_member (possible_steps e s r l i)") apply simp subgoal for a by (case_tac a, auto) done lemma no_possible_steps_1: "possible_steps e s r l i = {||} \ step e s r l i = None" by (simp add: step_def random_member_def) subsection\Execution Observation\ text\One of the key features of this formalisation of EFSMs is their ability to produce \emph{outputs}, which represent function return values. When action sequences are executed in an EFSM, they produce a corresponding \emph{observation}.\ text_raw\\snip{observe}{1}{2}{%\ fun observe_execution :: "transition_matrix \ cfstate \ registers \ execution \ outputs list" where "observe_execution _ _ _ [] = []" | "observe_execution e s r ((l, i)#as) = ( let viable = possible_steps e s r l i in if viable = {||} then [] else let (s', t) = Eps (\x. x |\| viable) in (evaluate_outputs t i r)#(observe_execution e s' (evaluate_updates t i r) as) )" text_raw\}%endsnip\ lemma observe_execution_step_def: "observe_execution e s r ((l, i)#as) = ( case step e s r l i of None \ []| Some (t, s', p, r') \ p#(observe_execution e s' r' as) )" apply (simp add: step_def) apply (case_tac "possible_steps e s r l i") apply simp subgoal for x S' apply (simp add: random_member_def) apply (case_tac "SOME xa. xa = x \ xa |\| S'") by simp done lemma observe_execution_first_outputs_equiv: "observe_execution e1 s1 r1 ((l, i) # ts) = observe_execution e2 s2 r2 ((l, i) # ts) \ step e1 s1 r1 l i = Some (t, s', p, r') \ \(s2', t2)|\|possible_steps e2 s2 r2 l i. evaluate_outputs t2 i r2 = p" apply (simp only: observe_execution_step_def) apply (case_tac "step e2 s2 r2 l i") apply simp subgoal for a apply simp apply (case_tac a) apply clarsimp by (meson step_member case_prodI rev_fBexI step_outputs) done lemma observe_execution_step: "step e s r (fst h) (snd h) = Some (t, s', p, r') \ observe_execution e s' r' es = obs \ observe_execution e s r (h#es) = p#obs" apply (cases h, simp add: step_def) apply (case_tac "possible_steps e s r a b = {||}") apply simp subgoal for a b apply (case_tac "SOME x. x |\| possible_steps e s r a b") by (simp add: random_member_def) done lemma observe_execution_possible_step: "possible_steps e s r (fst h) (snd h) = {|(s', t)|} \ apply_outputs (Outputs t) (join_ir (snd h) r) = p \ apply_updates (Updates t) (join_ir (snd h) r) r = r' \ observe_execution e s' r' es = obs \ observe_execution e s r (h#es) = p#obs" by (simp add: observe_execution_step step) lemma observe_execution_no_possible_step: "possible_steps e s r (fst h) (snd h) = {||} \ observe_execution e s r (h#es) = []" by (cases h, simp) lemma observe_execution_no_possible_steps: "possible_steps e1 s1 r1 (fst h) (snd h) = {||} \ possible_steps e2 s2 r2 (fst h) (snd h) = {||} \ (observe_execution e1 s1 r1 (h#t)) = (observe_execution e2 s2 r2 (h#t))" by (simp add: observe_execution_no_possible_step) lemma observe_execution_one_possible_step: "possible_steps e1 s1 r (fst h) (snd h) = {|(s1', t1)|} \ possible_steps e2 s2 r (fst h) (snd h) = {|(s2', t2)|} \ apply_outputs (Outputs t1) (join_ir (snd h) r) = apply_outputs (Outputs t2) (join_ir (snd h) r) \ apply_updates (Updates t1) (join_ir (snd h) r) r = r' \ apply_updates (Updates t2) (join_ir (snd h) r) r = r' \ (observe_execution e1 s1' r' t) = (observe_execution e2 s2' r' t) \ (observe_execution e1 s1 r (h#t)) = (observe_execution e2 s2 r (h#t))" by (simp add: observe_execution_possible_step) subsubsection\Utilities\ text\Here we define some utility functions to access the various key properties of a given EFSM.\ definition max_reg :: "transition_matrix \ nat option" where "max_reg e = (let maxes = (fimage (\(_, t). Transition.max_reg t) e) in if maxes = {||} then None else fMax maxes)" definition enumerate_ints :: "transition_matrix \ int set" where "enumerate_ints e = \ (image (\(_, t). Transition.enumerate_ints t) (fset e))" definition max_int :: "transition_matrix \ int" where "max_int e = Max (insert 0 (enumerate_ints e))" definition max_output :: "transition_matrix \ nat" where "max_output e = fMax (fimage (\(_, t). length (Outputs t)) e)" definition all_regs :: "transition_matrix \ nat set" where "all_regs e = \ (image (\(_, t). enumerate_regs t) (fset e))" text_raw\\snip{finiteRegs}{1}{2}{%\ lemma finite_all_regs: "finite (all_regs e)" text_raw\}%endsnip\ apply (simp add: all_regs_def enumerate_regs_def) apply clarify apply standard apply (rule finite_UnI)+ using GExp.finite_enumerate_regs apply blast using AExp.finite_enumerate_regs apply blast apply (simp add: AExp.finite_enumerate_regs prod.case_eq_if) by auto definition max_input :: "transition_matrix \ nat option" where "max_input e = fMax (fimage (\(_, t). Transition.max_input t) e)" fun maxS :: "transition_matrix \ nat" where "maxS t = (if t = {||} then 0 else fMax ((fimage (\((origin, dest), t). origin) t) |\| (fimage (\((origin, dest), t). dest) t)))" subsection\Execution Recognition\ text\The \texttt{recognises} function returns true if the given EFSM recognises a given execution. That is, the EFSM is able to respond to each event in sequence. There is no restriction on the outputs produced. When a recognised execution is observed, it produces an accepted trace of the EFSM.\ text_raw\\snip{recognises}{1}{2}{%\ inductive recognises_execution :: "transition_matrix \ nat \ registers \ execution \ bool" where base [simp]: "recognises_execution e s r []" | step: "\(s', T) |\| possible_steps e s r l i. recognises_execution e s' (evaluate_updates T i r) t \ recognises_execution e s r ((l, i)#t)" text_raw\}%endsnip\ abbreviation "recognises e t \ recognises_execution e 0 <> t" definition "E e = {x. recognises e x}" lemma no_possible_steps_rejects: "possible_steps e s r l i = {||} \ \ recognises_execution e s r ((l, i)#t)" apply clarify by (rule recognises_execution.cases, auto) lemma recognises_step_equiv: "recognises_execution e s r ((l, i)#t) = (\(s', T) |\| possible_steps e s r l i. recognises_execution e s' (evaluate_updates T i r) t)" apply standard apply (rule recognises_execution.cases) by (auto simp: recognises_execution.step) fun recognises_prim :: "transition_matrix \ nat \ registers \ execution \ bool" where "recognises_prim e s r [] = True" | "recognises_prim e s r ((l, i)#t) = ( let poss_steps = possible_steps e s r l i in (\(s', T) |\| poss_steps. recognises_prim e s' (evaluate_updates T i r) t) )" lemma recognises_prim [code]: "recognises_execution e s r t = recognises_prim e s r t" proof(induct t arbitrary: r s) case Nil then show ?case by simp next case (Cons h t) then show ?case apply (cases h) apply simp apply standard apply (rule recognises_execution.cases, simp) apply simp apply auto[1] using recognises_execution.step by blast qed lemma recognises_single_possible_step: assumes "possible_steps e s r l i = {|(s', t)|}" and "recognises_execution e s' (evaluate_updates t i r) trace" shows "recognises_execution e s r ((l, i)#trace)" apply (rule recognises_execution.step) using assms by auto lemma recognises_single_possible_step_atomic: assumes "possible_steps e s r (fst h) (snd h) = {|(s', t)|}" and "recognises_execution e s' (apply_updates (Updates t) (join_ir (snd h) r) r) trace" shows "recognises_execution e s r (h#trace)" by (metis assms prod.collapse recognises_single_possible_step) lemma recognises_must_be_possible_step: "recognises_execution e s r (h # t) \ \aa ba. (aa, ba) |\| possible_steps e s r (fst h) (snd h)" using recognises_step_equiv by fastforce lemma recognises_possible_steps_not_empty: "recognises_execution e s r (h#t) \ possible_steps e s r (fst h) (snd h) \ {||}" apply (rule recognises_execution.cases) by auto lemma recognises_must_be_step: "recognises_execution e s r (h#ts) \ \t s' p d'. step e s r (fst h) (snd h) = Some (t, s', p, d')" apply (cases h) subgoal for a b apply (simp add: recognises_step_equiv step_def) apply clarify apply (case_tac "(possible_steps e s r a b)") apply (simp add: random_member_def) apply (simp add: random_member_def) subgoal for _ _ x S' apply (case_tac "SOME xa. xa = x \ xa |\| S'") by simp done done lemma recognises_cons_step: "recognises_execution e s r (h # t) \ step e s r (fst h) (snd h) \ None" by (simp add: recognises_must_be_step) lemma no_step_none: "step e s r aa ba = None \ \ recognises_execution e s r ((aa, ba) # p)" using recognises_cons_step by fastforce lemma step_none_rejects: "step e s r (fst h) (snd h) = None \ \ recognises_execution e s r (h#t)" using no_step_none surjective_pairing by fastforce lemma trace_reject: "(\ recognises_execution e s r ((l, i)#t)) = (possible_steps e s r l i = {||} \ (\(s', T) |\| possible_steps e s r l i. \ recognises_execution e s' (evaluate_updates T i r) t))" using recognises_prim by fastforce lemma trace_reject_no_possible_steps_atomic: "possible_steps e s r (fst a) (snd a) = {||} \ \ recognises_execution e s r (a#t)" using recognises_possible_steps_not_empty by auto lemma trace_reject_later: "\(s', T) |\| possible_steps e s r l i. \ recognises_execution e s' (evaluate_updates T i r) t \ \ recognises_execution e s r ((l, i)#t)" using trace_reject by auto lemma recognition_prefix_closure: "recognises_execution e s r (t@t') \ recognises_execution e s r t" proof(induct t arbitrary: s r) case (Cons a t) then show ?case apply (cases a, clarsimp) apply (rule recognises_execution.cases) apply simp apply simp by (rule recognises_execution.step, auto) qed auto lemma rejects_prefix: "\ recognises_execution e s r t \ \ recognises_execution e s r (t @ t')" using recognition_prefix_closure by blast lemma recognises_head: "recognises_execution e s r (h#t) \ recognises_execution e s r [h]" by (simp add: recognition_prefix_closure) subsubsection\Trace Acceptance\ text\The \texttt{accepts} function returns true if the given EFSM accepts a given trace. That is, the EFSM is able to respond to each event in sequence \emph{and} is able to produce the expected output. Accepted traces represent valid runs of an EFSM.\ text_raw\\snip{accepts}{1}{2}{%\ inductive accepts_trace :: "transition_matrix \ cfstate \ registers \ trace \ bool" where base [simp]: "accepts_trace e s r []" | step: "\(s', T) |\| possible_steps e s r l i. evaluate_outputs T i r = map Some p \ accepts_trace e s' (evaluate_updates T i r) t \ accepts_trace e s r ((l, i, p)#t)" text_raw\}%endsnip\ text_raw\\snip{T}{1}{2}{%\ definition T :: "transition_matrix \ trace set" where "T e = {t. accepts_trace e 0 <> t}" text_raw\}%endsnip\ abbreviation "rejects_trace e s r t \ \ accepts_trace e s r t" lemma accepts_trace_step: "accepts_trace e s r ((l, i, p)#t) = (\(s', T) |\| possible_steps e s r l i. evaluate_outputs T i r = map Some p \ accepts_trace e s' (evaluate_updates T i r) t)" apply standard by (rule accepts_trace.cases, auto simp: accepts_trace.step) lemma accepts_trace_exists_possible_step: "accepts_trace e1 s1 r1 ((aa, b, c) # t) \ \(s1', t1)|\|possible_steps e1 s1 r1 aa b. evaluate_outputs t1 b r1 = map Some c" using accepts_trace_step by auto lemma rejects_trace_step: "rejects_trace e s r ((l, i, p)#t) = ( (\(s', T) |\| possible_steps e s r l i. evaluate_outputs T i r \ map Some p \ rejects_trace e s' (evaluate_updates T i r) t) )" apply (simp add: accepts_trace_step) by auto definition accepts_log :: "trace set \ transition_matrix \ bool" where "accepts_log l e = (\t \ l. accepts_trace e 0 <> t)" text_raw\\snip{prefixClosure}{1}{2}{%\ lemma prefix_closure: "accepts_trace e s r (t@t') \ accepts_trace e s r t" text_raw\}%endsnip\ proof(induct t arbitrary: s r) next case (Cons a t) then show ?case apply (cases a, clarsimp) apply (simp add: accepts_trace_step) by auto qed auto text\For code generation, it is much more efficient to re-implement the \texttt{accepts\_trace} function primitively than to use the code generator's default setup for inductive definitions.\ fun accepts_trace_prim :: "transition_matrix \ cfstate \ registers \ trace \ bool" where "accepts_trace_prim _ _ _ [] = True" | "accepts_trace_prim e s r ((l, i, p)#t) = ( let poss_steps = possible_steps e s r l i in if fis_singleton poss_steps then let (s', T) = fthe_elem poss_steps in if evaluate_outputs T i r = map Some p then accepts_trace_prim e s' (evaluate_updates T i r) t else False else (\(s', T) |\| poss_steps. evaluate_outputs T i r = (map Some p) \ accepts_trace_prim e s' (evaluate_updates T i r) t))" lemma accepts_trace_prim [code]: "accepts_trace e s r l = accepts_trace_prim e s r l" proof(induct l arbitrary: s r) case (Cons a l) then show ?case apply (cases a) apply (simp add: accepts_trace_step Let_def fis_singleton_alt) by auto qed auto subsection\EFSM Comparison\ text\Here, we define some different metrics of EFSM equality.\ subsubsection\State Isomporphism\ text\Two EFSMs are isomorphic with respect to states if there exists a bijective function between the state names of the two EFSMs, i.e. the only difference between the two models is the way the states are indexed.\ definition isomorphic :: "transition_matrix \ transition_matrix \ bool" where "isomorphic e1 e2 = (\f. bij f \ (\((s1, s2), t) |\| e1. ((f s1, f s2), t) |\| e2))" subsubsection\Register Isomporphism\ text\Two EFSMs are isomorphic with respect to registers if there exists a bijective function between the indices of the registers in the two EFSMs, i.e. the only difference between the two models is the way the registers are indexed.\ definition rename_regs :: "(nat \ nat) \ transition_matrix \ transition_matrix" where "rename_regs f e = fimage (\(tf, t). (tf, Transition.rename_regs f t)) e" definition eq_upto_rename_strong :: "transition_matrix \ transition_matrix \ bool" where "eq_upto_rename_strong e1 e2 = (\f. bij f \ rename_regs f e1 = e2)" subsubsection\Trace Simulation\ text\An EFSM, $e_1$ simulates another EFSM $e_2$ if there is a function between the states of the states of $e_1$ and $e_1$ such that in each state, if $e_1$ can respond to the event and produce the correct output, so can $e_2$.\ text_raw\\snip{traceSim}{1}{2}{%\ inductive trace_simulation :: "(cfstate \ cfstate) \ transition_matrix \ cfstate \ registers \ transition_matrix \ cfstate \ registers \ trace \ bool" where base: "s2 = f s1 \ trace_simulation f e1 s1 r1 e2 s2 r2 []" | step: "s2 = f s1 \ \(s1', t1) |\| ffilter (\(s1', t1). evaluate_outputs t1 i r1 = map Some o) (possible_steps e1 s1 r1 l i). \(s2', t2) |\| possible_steps e2 s2 r2 l i. evaluate_outputs t2 i r2 = map Some o \ trace_simulation f e1 s1' (evaluate_updates t1 i r1) e2 s2' (evaluate_updates t2 i r2) es \ trace_simulation f e1 s1 r1 e2 s2 r2 ((l, i, o)#es)" text_raw\}%endsnip\ lemma trace_simulation_step: "trace_simulation f e1 s1 r1 e2 s2 r2 ((l, i, o)#es) = ( (s2 = f s1) \ (\(s1', t1) |\| ffilter (\(s1', t1). evaluate_outputs t1 i r1 = map Some o) (possible_steps e1 s1 r1 l i). (\(s2', t2) |\| possible_steps e2 s2 r2 l i. evaluate_outputs t2 i r2 = map Some o \ trace_simulation f e1 s1' (evaluate_updates t1 i r1) e2 s2' (evaluate_updates t2 i r2) es)) )" apply standard apply (rule trace_simulation.cases, simp+) apply (rule trace_simulation.step) apply simp - by blast + by fastforce lemma trace_simulation_step_none: "s2 = f s1 \ \(s1', t1) |\| possible_steps e1 s1 r1 l i. evaluate_outputs t1 i r1 = map Some o \ trace_simulation f e1 s1 r1 e2 s2 r2 ((l, i, o)#es)" apply (rule trace_simulation.step) apply simp apply (case_tac "ffilter (\(s1', t1). evaluate_outputs t1 i r1 = map Some o) (possible_steps e1 s1 r1 l i)") apply simp by fastforce definition "trace_simulates e1 e2 = (\f. \t. trace_simulation f e1 0 <> e2 0 <> t)" lemma rejects_trace_simulation: "rejects_trace e2 s2 r2 t \ accepts_trace e1 s1 r1 t \ \trace_simulation f e1 s1 r1 e2 s2 r2 t" proof(induct t arbitrary: s1 r1 s2 r2) case Nil then show ?case using accepts_trace.base by blast next case (Cons a t) then show ?case apply (cases a) apply (simp add: rejects_trace_step) apply (simp add: accepts_trace_step) apply clarify apply (rule trace_simulation.cases) apply simp apply simp apply clarsimp subgoal for l o _ _ i - by fastforce + by blast done qed lemma accepts_trace_simulation: "accepts_trace e1 s1 r1 t \ trace_simulation f e1 s1 r1 e2 s2 r2 t \ accepts_trace e2 s2 r2 t" using rejects_trace_simulation by blast lemma simulates_trace_subset: "trace_simulates e1 e2 \ T e1 \ T e2" using T_def accepts_trace_simulation trace_simulates_def by fastforce subsubsection\Trace Equivalence\ text\Two EFSMs are trace equivalent if they accept the same traces. This is the intuitive definition of ``observable equivalence'' between the behaviours of the two models. If two EFSMs are trace equivalent, there is no trace which can distinguish the two.\ text_raw\\snip{traceEquiv}{1}{2}{%\ definition "trace_equivalent e1 e2 = (T e1 = T e2)" text_raw\}%endsnip\ text_raw\\snip{simEquiv}{1}{2}{%\ lemma simulation_implies_trace_equivalent: "trace_simulates e1 e2 \ trace_simulates e2 e1 \ trace_equivalent e1 e2" text_raw\}%endsnip\ using simulates_trace_subset trace_equivalent_def by auto lemma trace_equivalent_reflexive: "trace_equivalent e1 e1" by (simp add: trace_equivalent_def) lemma trace_equivalent_symmetric: "trace_equivalent e1 e2 = trace_equivalent e2 e1" using trace_equivalent_def by auto lemma trace_equivalent_transitive: "trace_equivalent e1 e2 \ trace_equivalent e2 e3 \ trace_equivalent e1 e3" by (simp add: trace_equivalent_def) text\Two EFSMs are trace equivalent if they accept the same traces.\ lemma trace_equivalent: "\t. accepts_trace e1 0 <> t = accepts_trace e2 0 <> t \ trace_equivalent e1 e2" by (simp add: T_def trace_equivalent_def) lemma accepts_trace_step_2: "(s2', t2) |\| possible_steps e2 s2 r2 l i \ accepts_trace e2 s2' (evaluate_updates t2 i r2) t \ evaluate_outputs t2 i r2 = map Some p \ accepts_trace e2 s2 r2 ((l, i, p)#t)" by (rule accepts_trace.step, auto) subsubsection\Execution Simulation\ text\Execution simulation is similar to trace simulation but for executions rather than traces. Execution simulation has no notion of ``expected'' output. It simply requires that the simulating EFSM must be able to produce equivalent output for each action.\ text_raw\\snip{execSim}{1}{2}{%\ inductive execution_simulation :: "(cfstate \ cfstate) \ transition_matrix \ cfstate \ registers \ transition_matrix \ cfstate \ registers \ execution \ bool" where base: "s2 = f s1 \ execution_simulation f e1 s1 r1 e2 s2 r2 []" | step: "s2 = f s1 \ \(s1', t1) |\| (possible_steps e1 s1 r1 l i). \(s2', t2) |\| possible_steps e2 s2 r2 l i. evaluate_outputs t1 i r1 = evaluate_outputs t2 i r2 \ execution_simulation f e1 s1' (evaluate_updates t1 i r1) e2 s2' (evaluate_updates t2 i r2) es \ execution_simulation f e1 s1 r1 e2 s2 r2 ((l, i)#es)" text_raw\}%endsnip\ definition "execution_simulates e1 e2 = (\f. \t. execution_simulation f e1 0 <> e2 0 <> t)" lemma execution_simulation_step: "execution_simulation f e1 s1 r1 e2 s2 r2 ((l, i)#es) = (s2 = f s1 \ (\(s1', t1) |\| (possible_steps e1 s1 r1 l i). (\(s2', t2) |\| possible_steps e2 s2 r2 l i. evaluate_outputs t1 i r1 = evaluate_outputs t2 i r2 \ execution_simulation f e1 s1' (evaluate_updates t1 i r1) e2 s2' (evaluate_updates t2 i r2) es)) )" apply standard apply (rule execution_simulation.cases) apply simp apply simp apply simp apply (rule execution_simulation.step) apply simp by blast text_raw\\snip{execTraceSim}{1}{2}{%\ lemma execution_simulation_trace_simulation: "execution_simulation f e1 s1 r1 e2 s2 r2 (map (\(l, i, o). (l, i)) t) \ trace_simulation f e1 s1 r1 e2 s2 r2 t" text_raw\}%endsnip\ proof(induct t arbitrary: s1 s2 r1 r2) case Nil then show ?case apply (rule execution_simulation.cases) apply (simp add: trace_simulation.base) by simp next case (Cons a t) then show ?case apply (cases a, clarsimp) apply (rule execution_simulation.cases) apply simp apply simp apply (rule trace_simulation.step) apply simp apply clarsimp subgoal for _ _ _ aa ba apply (erule_tac x="(aa, ba)" in fBallE) apply clarsimp apply blast by simp done qed lemma execution_simulates_trace_simulates: "execution_simulates e1 e2 \ trace_simulates e1 e2" apply (simp add: execution_simulates_def trace_simulates_def) using execution_simulation_trace_simulation by blast subsubsection\Executional Equivalence\ text\Two EFSMs are executionally equivalent if there is no execution which can distinguish between the two. That is, for every execution, they must produce equivalent outputs.\ text_raw\\snip{execEquiv}{1}{2}{%\ inductive executionally_equivalent :: "transition_matrix \ cfstate \ registers \ transition_matrix \ cfstate \ registers \ execution \ bool" where base [simp]: "executionally_equivalent e1 s1 r1 e2 s2 r2 []" | step: "\(s1', t1) |\| possible_steps e1 s1 r1 l i. \(s2', t2) |\| possible_steps e2 s2 r2 l i. evaluate_outputs t1 i r1 = evaluate_outputs t2 i r2 \ executionally_equivalent e1 s1' (evaluate_updates t1 i r1) e2 s2' (evaluate_updates t2 i r2) es \ \(s2', t2) |\| possible_steps e2 s2 r2 l i. \(s1', t1) |\| possible_steps e1 s1 r1 l i. evaluate_outputs t1 i r1 = evaluate_outputs t2 i r2 \ executionally_equivalent e1 s1' (evaluate_updates t1 i r1) e2 s2' (evaluate_updates t2 i r2) es \ executionally_equivalent e1 s1 r1 e2 s2 r2 ((l, i)#es)" text_raw\}%endsnip\ lemma executionally_equivalent_step: "executionally_equivalent e1 s1 r1 e2 s2 r2 ((l, i)#es) = ( (\(s1', t1) |\| (possible_steps e1 s1 r1 l i). (\(s2', t2) |\| possible_steps e2 s2 r2 l i. evaluate_outputs t1 i r1 = evaluate_outputs t2 i r2 \ executionally_equivalent e1 s1' (evaluate_updates t1 i r1) e2 s2' (evaluate_updates t2 i r2) es)) \ (\(s2', t2) |\| (possible_steps e2 s2 r2 l i). (\(s1', t1) |\| possible_steps e1 s1 r1 l i. evaluate_outputs t1 i r1 = evaluate_outputs t2 i r2 \ executionally_equivalent e1 s1' (evaluate_updates t1 i r1) e2 s2' (evaluate_updates t2 i r2) es)))" apply standard apply (rule executionally_equivalent.cases) apply simp apply simp apply simp by (rule executionally_equivalent.step, auto) lemma execution_end: "possible_steps e1 s1 r1 l i = {||} \ possible_steps e2 s2 r2 l i = {||} \ executionally_equivalent e1 s1 r1 e2 s2 r2 ((l, i)#es)" by (simp add: executionally_equivalent_step) lemma possible_steps_disparity: "possible_steps e1 s1 r1 l i \ {||} \ possible_steps e2 s2 r2 l i = {||} \ \executionally_equivalent e1 s1 r1 e2 s2 r2 ((l, i)#es)" by (simp add: executionally_equivalent_step, auto) lemma executionally_equivalent_acceptance_map: "executionally_equivalent e1 s1 r1 e2 s2 r2 (map (\(l, i, o). (l, i)) t) \ accepts_trace e2 s2 r2 t = accepts_trace e1 s1 r1 t" proof(induct t arbitrary: s1 s2 r1 r2) case (Cons a t) then show ?case apply (cases a, simp) apply (rule executionally_equivalent.cases) apply simp apply simp apply clarsimp apply standard subgoal for i p l apply (rule accepts_trace.cases) apply simp apply simp apply clarsimp subgoal for aa b apply (rule accepts_trace.step) apply (erule_tac x="(aa, b)" in fBallE[of "possible_steps e2 s2 r2 l i"]) defer apply simp apply simp by blast done apply (rule accepts_trace.cases) apply simp apply simp apply clarsimp subgoal for _ _ _ aa b apply (rule accepts_trace.step) apply (erule_tac x="(aa, b)" in fBallE) defer apply simp apply simp by fastforce done qed auto lemma executionally_equivalent_acceptance: "\x. executionally_equivalent e1 s1 r1 e2 s2 r2 x \ accepts_trace e1 s1 r1 t \ accepts_trace e2 s2 r2 t" using executionally_equivalent_acceptance_map by blast lemma executionally_equivalent_trace_equivalent: "\x. executionally_equivalent e1 0 <> e2 0 <> x \ trace_equivalent e1 e2" apply (rule trace_equivalent) apply clarify subgoal for t apply (erule_tac x="map (\(l, i, o). (l, i)) t" in allE) by (simp add: executionally_equivalent_acceptance_map) done lemma executionally_equivalent_symmetry: "executionally_equivalent e1 s1 r1 e2 s2 r2 x \ executionally_equivalent e2 s2 r2 e1 s1 r1 x" proof(induct x arbitrary: s1 s2 r1 r2) case (Cons a x) then show ?case apply (cases a, clarsimp) apply (simp add: executionally_equivalent_step) apply standard apply (rule fBallI) apply clarsimp subgoal for aa b aaa ba apply (erule_tac x="(aaa, ba)" in fBallE[of "possible_steps e2 s2 r2 aa b"]) by (force, simp) apply (rule fBallI) apply clarsimp subgoal for aa b aaa ba apply (erule_tac x="(aaa, ba)" in fBallE) by (force, simp) done qed auto lemma executionally_equivalent_transitivity: "executionally_equivalent e1 s1 r1 e2 s2 r2 x \ executionally_equivalent e2 s2 r2 e3 s3 r3 x \ executionally_equivalent e1 s1 r1 e3 s3 r3 x" proof(induct x arbitrary: s1 s2 s3 r1 r2 r3) case (Cons a x) then show ?case apply (cases a, clarsimp) apply (simp add: executionally_equivalent_step) apply clarsimp apply standard apply (rule fBallI) apply clarsimp subgoal for aa b ab ba apply (erule_tac x="(ab, ba)" in fBallE[of "possible_steps e1 s1 r1 aa b"]) prefer 2 apply simp apply simp apply (erule fBexE) subgoal for x apply (case_tac x) apply simp by blast done apply (rule fBallI) apply clarsimp subgoal for aa b ab ba apply (erule_tac x="(ab, ba)" in fBallE[of "possible_steps e3 s3 r3 aa b"]) prefer 2 apply simp apply simp apply (erule fBexE) subgoal for x apply (case_tac x) apply clarsimp subgoal for aaa baa apply (erule_tac x="(aaa, baa)" in fBallE[of "possible_steps e2 s2 r2 aa b"]) prefer 2 apply simp apply simp by blast done done done qed auto subsection\Reachability\ text\Here, we define the function \texttt{visits} which returns true if the given execution leaves the given EFSM in the given state.\ text_raw\\snip{reachable}{1}{2}{%\ inductive visits :: "cfstate \ transition_matrix \ cfstate \ registers \ execution \ bool" where base [simp]: "visits s e s r []" | step: "\(s', T) |\| possible_steps e s r l i. visits target e s' (evaluate_updates T i r) t \ visits target e s r ((l, i)#t)" definition "reachable s e = (\t. visits s e 0 <> t)" text_raw\}%endsnip\ lemma no_further_steps: "s \ s' \ \ visits s e s' r []" apply safe apply (rule visits.cases) by auto lemma visits_base: "visits target e s r [] = (s = target)" by (metis visits.base no_further_steps) lemma visits_step: "visits target e s r (h#t) = (\(s', T) |\| possible_steps e s r (fst h) (snd h). visits target e s' (evaluate_updates T (snd h) r) t)" apply standard apply (rule visits.cases) apply simp+ apply (cases h) using visits.step by auto lemma reachable_initial: "reachable 0 e" apply (simp add: reachable_def) apply (rule_tac x="[]" in exI) by simp lemma visits_finsert: "visits s e s' r t \ visits s (finsert ((aa, ba), b) e) s' r t" proof(induct t arbitrary: s' r) case Nil then show ?case by (simp add: visits_base) next case (Cons a t) then show ?case apply (simp add: visits_step) apply (erule fBexE) apply (rule_tac x=x in fBexI) apply auto[1] by (simp add: possible_steps_finsert) qed lemma reachable_finsert: "reachable s e \ reachable s (finsert ((aa, ba), b) e)" apply (simp add: reachable_def) by (meson visits_finsert) lemma reachable_finsert_contra: "\ reachable s (finsert ((aa, ba), b) e) \ \reachable s e" using reachable_finsert by blast lemma visits_empty: "visits s e s' r [] = (s = s')" apply standard by (rule visits.cases, auto) definition "remove_state s e = ffilter (\((from, to), t). from \ s \ to \ s) e" text_raw\\snip{obtainable}{1}{2}{%\ inductive "obtains" :: "cfstate \ registers \ transition_matrix \ cfstate \ registers \ execution \ bool" where base [simp]: "obtains s r e s r []" | step: "\(s'', T) |\| possible_steps e s' r' l i. obtains s r e s'' (evaluate_updates T i r') t \ obtains s r e s' r' ((l, i)#t)" definition "obtainable s r e = (\t. obtains s r e 0 <> t)" text_raw\}%endsnip\ lemma obtains_obtainable: "obtains s r e 0 <> t \ obtainable s r e" apply (simp add: obtainable_def) by auto lemma obtains_base: "obtains s r e s' r' [] = (s = s' \ r = r')" apply standard by (rule obtains.cases, auto) lemma obtains_step: "obtains s r e s' r' ((l, i)#t) = (\(s'', T) |\| possible_steps e s' r' l i. obtains s r e s'' (evaluate_updates T i r') t)" apply standard by (rule obtains.cases, auto simp add: obtains.step) lemma obtains_recognises: "obtains s c e s' r t \ recognises_execution e s' r t" proof(induct t arbitrary: s' r) case Nil then show ?case by (simp add: obtains_base) next case (Cons a t) then show ?case apply (cases a) apply simp apply (rule obtains.cases) apply simp apply simp apply clarsimp using recognises_execution.step by fastforce qed lemma ex_comm4: "(\c1 s a b. (a, b) \ fset (possible_steps e s' r l i) \ obtains s c1 e a (evaluate_updates b i r) t) = (\a b s c1. (a, b) \ fset (possible_steps e s' r l i) \ obtains s c1 e a (evaluate_updates b i r) t)" by auto lemma recognises_execution_obtains: "recognises_execution e s' r t \ \c1 s. obtains s c1 e s' r t" proof(induct t arbitrary: s' r) case Nil then show ?case by (simp add: obtains_base) next case (Cons a t) then show ?case apply (cases a) apply (simp add: obtains_step) apply (rule recognises_execution.cases) apply simp apply simp apply clarsimp - apply (simp add: fBex_def Bex_def ex_comm4) + apply (simp add: Bex_def ex_comm4) subgoal for _ _ aa ba apply (rule_tac x=aa in exI) apply (rule_tac x=ba in exI) apply simp by blast done qed lemma obtainable_empty_efsm: "obtainable s c {||} = (s=0 \ c = <>)" apply (simp add: obtainable_def) apply standard apply (metis ffilter_empty no_outgoing_transitions no_step_none obtains.cases obtains_recognises step_None) using obtains_base by blast lemma obtains_visits: "obtains s r e s' r' t \ visits s e s' r' t" proof(induct t arbitrary: s' r') case Nil then show ?case by (simp add: obtains_base) next case (Cons a t) then show ?case apply (cases a) apply (rule obtains.cases) apply simp apply simp apply clarsimp apply (rule visits.step) by auto qed lemma unobtainable_if: "\ visits s e s' r' t \ \ obtains s r e s' r' t" using obtains_visits by blast lemma obtainable_if_unreachable: "\reachable s e \ \obtainable s r e" by (simp add: reachable_def obtainable_def unobtainable_if) lemma obtains_step_append: "obtains s r e s' r' t \ (s'', ta) |\| possible_steps e s r l i \ obtains s'' (evaluate_updates ta i r) e s' r' (t @ [(l, i)])" proof(induct t arbitrary: s' r') case Nil then show ?case apply (simp add: obtains_base) apply (rule obtains.step) apply (rule_tac x="(s'', ta)" in fBexI) by auto next case (Cons a t) then show ?case apply simp apply (rule obtains.cases) apply simp apply simp apply clarsimp apply (rule obtains.step) by auto qed lemma reachable_if_obtainable_step: "obtainable s r e \ \l i t. (s', t) |\| possible_steps e s r l i \ reachable s' e" apply (simp add: reachable_def obtainable_def) apply clarify subgoal for t l i apply (rule_tac x="t@[(l, i)]" in exI) using obtains_step_append unobtainable_if by blast done lemma possible_steps_remove_unreachable: "obtainable s r e \ \ reachable s' e \ possible_steps (remove_state s' e) s r l i = possible_steps e s r l i" apply standard apply (simp add: fsubset_eq) apply (rule fBallI) apply clarsimp apply (metis ffmember_filter in_possible_steps remove_state_def) apply (simp add: fsubset_eq) apply (rule fBallI) apply clarsimp subgoal for a b apply (case_tac "a = s'") using reachable_if_obtainable_step apply blast apply (simp add: remove_state_def) by (metis (mono_tags, lifting) ffmember_filter in_possible_steps obtainable_if_unreachable old.prod.case) done text_raw\\snip{removeUnreachableArb}{1}{2}{%\ lemma executionally_equivalent_remove_unreachable_state_arbitrary: "obtainable s r e \ \ reachable s' e \ executionally_equivalent e s r (remove_state s' e) s r x" text_raw\}%endsnip\ proof(induct x arbitrary: s r) case (Cons a x) then show ?case apply (cases a, simp) apply (rule executionally_equivalent.step) apply (simp add: possible_steps_remove_unreachable) apply standard apply clarsimp subgoal for aa b ab ba apply (rule_tac x="(ab, ba)" in fBexI) apply (metis (mono_tags, lifting) obtainable_def obtains_step_append case_prodI) apply simp done apply (rule fBallI) apply clarsimp apply (rule_tac x="(ab, ba)" in fBexI) apply simp apply (metis obtainable_def obtains_step_append possible_steps_remove_unreachable) by (simp add: possible_steps_remove_unreachable) qed auto text_raw\\snip{removeUnreachable}{1}{2}{%\ lemma executionally_equivalent_remove_unreachable_state: "\ reachable s' e \ executionally_equivalent e 0 <> (remove_state s' e) 0 <> x" text_raw\}%endsnip\ by (meson executionally_equivalent_remove_unreachable_state_arbitrary obtains.simps obtains_obtainable) subsection\Transition Replacement\ text\Here, we define the function \texttt{replace} to replace one transition with another, and prove some of its properties.\ definition "replace e1 old new = fimage (\x. if x = old then new else x) e1" lemma replace_finsert: "replace (finsert ((aaa, baa), b) e1) old new = (if ((aaa, baa), b) = old then (finsert new (replace e1 old new)) else (finsert ((aaa, baa), b) (replace e1 old new)))" by (simp add: replace_def) lemma possible_steps_replace_unchanged: "((s, aa), ba) \ ((s1, s2), t1) \ (aa, ba) |\| possible_steps e1 s r l i \ (aa, ba) |\| possible_steps (replace e1 ((s1, s2), t1) ((s1, s2), t2)) s r l i" by (simp add: in_possible_steps[symmetric] replace_def) end diff --git a/thys/Extended_Finite_State_Machines/EFSM_LTL.thy b/thys/Extended_Finite_State_Machines/EFSM_LTL.thy --- a/thys/Extended_Finite_State_Machines/EFSM_LTL.thy +++ b/thys/Extended_Finite_State_Machines/EFSM_LTL.thy @@ -1,281 +1,281 @@ section\LTL for EFSMs\ text\This theory builds off the \texttt{Linear\_Temporal\_Logic\_on\_Streams} theory from the HOL library and defines functions to ease the expression of LTL properties over EFSMs. Since the LTL operators effectively act over traces of models we must find a way to express models as streams.\ theory EFSM_LTL imports "Extended_Finite_State_Machines.EFSM" "HOL-Library.Linear_Temporal_Logic_on_Streams" begin text_raw\\snip{statedef}{1}{2}{%\ record state = statename :: "nat option" datastate :: registers action :: action "output" :: outputs text_raw\}%endsnip\ text_raw\\snip{whitebox}{1}{2}{%\ type_synonym whitebox_trace = "state stream" text_raw\}%endsnip\ type_synonym property = "whitebox_trace \ bool" abbreviation label :: "state \ String.literal" where "label s \ fst (action s)" abbreviation inputs :: "state \ value list" where "inputs s \ snd (action s)" text_raw\\snip{ltlStep}{1}{2}{%\ fun ltl_step :: "transition_matrix \ cfstate option \ registers \ action \ (nat option \ outputs \ registers)" where "ltl_step _ None r _ = (None, [], r)" | "ltl_step e (Some s) r (l, i) = (let possibilities = possible_steps e s r l i in if possibilities = {||} then (None, [], r) else let (s', t) = Eps (\x. x |\| possibilities) in (Some s', (evaluate_outputs t i r), (evaluate_updates t i r)) )" text_raw\}%endsnip\ lemma ltl_step_singleton: "\t. possible_steps e n r (fst v) (snd v) = {|(aa, t)|} \ evaluate_outputs t (snd v) r = b \ evaluate_updates t (snd v) r = c\ ltl_step e (Some n) r v = (Some aa, b, c)" apply (cases v) by auto lemma ltl_step_none: "possible_steps e s r a b = {||} \ ltl_step e (Some s) r (a, b) = (None, [], r)" by simp lemma ltl_step_none_2: "possible_steps e s r (fst ie) (snd ie) = {||} \ ltl_step e (Some s) r ie = (None, [], r)" by (metis ltl_step_none prod.exhaust_sel) lemma ltl_step_alt: "ltl_step e (Some s) r t = ( let possibilities = possible_steps e s r (fst t) (snd t) in if possibilities = {||} then (None, [], r) else let (s', t') = Eps (\x. x |\| possibilities) in (Some s', (apply_outputs (Outputs t') (join_ir (snd t) r)), (apply_updates (Updates t') (join_ir (snd t) r) r)) )" by (case_tac t, simp add: Let_def) lemma ltl_step_some: assumes "possible_steps e s r l i = {|(s', t)|}" and "evaluate_outputs t i r = p" and "evaluate_updates t i r = r'" shows "ltl_step e (Some s) r (l, i) = (Some s', p, r')" by (simp add: assms) lemma ltl_step_cases: assumes invalid: "P (None, [], r)" and valid: "\(s', t) |\| (possible_steps e s r l i). P (Some s', (evaluate_outputs t i r), (evaluate_updates t i r))" shows "P (ltl_step e (Some s) r (l, i))" apply simp apply (case_tac "possible_steps e s r l i") apply (simp add: invalid) apply simp subgoal for x S' apply (case_tac "SOME xa. xa = x \ xa |\| S'") apply simp apply (insert assms(2)) - apply (simp add: fBall_def Ball_def) + apply (simp add: Ball_def) by (metis (mono_tags, lifting) fst_conv prod.case_eq_if snd_conv someI_ex) done text\The \texttt{make\_full\_observation} function behaves similarly to \texttt{observe\_execution} from the \texttt{EFSM} theory. The main difference in behaviour is what is recorded. While the observe execution function simply observes an execution of the EFSM to produce the corresponding output for each action, the intention here is to record every detail of execution, including the values of internal variables. Thinking of each action as a step forward in time, there are five components which characterise a given point in the execution of an EFSM. At each point, the model has a current control state and data state. Each action has a label and some input parameters, and its execution may produce some observableoutput. It is therefore sufficient to provide a stream of 5-tuples containing the current control state, data state, the label and inputs of the action, and computed output. The make full observation function can then be defined as in Figure 9.1, with an additional function watch defined on top of this which starts the make full observation off in the initial control state with the empty data state. Careful inspection of the definition reveals another way that \texttt{make\_full\_observation} differs from \texttt{observe\_execution}. Rather than taking a cfstate, it takes a cfstate option. The reason for this is that we need to make our EFSM models complete. That is, we need them to be able to respond to every action from every state like a DFA. If a model does not recognise a given action in a given state, we cannot simply stop processing because we are working with necessarily infinite traces. Since these traces are generated by observing action sequences, the make full observation function must keep processing whether there is a viable transition or not. To support this, the make full observation adds an implicit ``sink state'' to every EFSM it processes by lifting control flow state indices from \texttt{nat} to \texttt{nat option} such that state $n$ is seen as state \texttt{Some} $n$. The control flow state \texttt{None} represents a sink state. If a model is unable to recognise a particular action from its current state, it moves into the \texttt{None} state. From here, the behaviour is constant for the rest of the time --- the control flow state remains None; the data state does not change, and no output is produced.\ text_raw\\snip{makeFullObservation}{1}{2}{%\ primcorec make_full_observation :: "transition_matrix \ cfstate option \ registers \ outputs \ action stream \ whitebox_trace" where "make_full_observation e s d p i = ( let (s', o', d') = ltl_step e s d (shd i) in \statename = s, datastate = d, action=(shd i), output = p\##(make_full_observation e s' d' o' (stl i)) )" text_raw\}%endsnip\ text_raw\\snip{watch}{1}{2}{%\ abbreviation watch :: "transition_matrix \ action stream \ whitebox_trace" where "watch e i \ (make_full_observation e (Some 0) <> [] i)" text_raw\}%endsnip\ subsection\Expressing Properties\ text\In order to simplify the expression and understanding of properties, this theory defines a number of named functions which can be used to express certain properties of EFSMs.\ subsubsection\State Equality\ text\The \textsc{state\_eq} takes a cfstate option representing a control flow state index and returns true if this is the control flow state at the head of the full observation.\ abbreviation state_eq :: "cfstate option \ whitebox_trace \ bool" where "state_eq v s \ statename (shd s) = v" lemma state_eq_holds: "state_eq s = holds (\x. statename x = s)" apply (rule ext) by (simp add: holds_def) lemma state_eq_None_not_Some: "state_eq None s \ \ state_eq (Some n) s" by simp subsubsection\Label Equality\ text\The \textsc{label\_eq} function takes a string and returns true if this is equal to the label at the head of the full observation.\ abbreviation "label_eq v s \ fst (action (shd s)) = (String.implode v)" lemma watch_label: "label_eq l (watch e t) = (fst (shd t) = String.implode l)" by simp subsubsection\Input Equality\ text\The \textsc{input\_eq} function takes a value list and returns true if this is equal to the input at the head of the full observation.\ abbreviation "input_eq v s \ inputs (shd s) = v" subsubsection\Action Equality\ text\The \textsc{action\_eq} function takes a (label, value list) pair and returns true if this is equal to the action at the head of the full observation. This effectively combines \texttt{label\_eq} and \texttt{input\_eq} into one function.\ abbreviation "action_eq e \ label_eq (fst e) aand input_eq (snd e)" subsubsection\Output Equality\ text\The \textsc{output\_eq} function takes a takes a value option list and returns true if this is equal to the output at the head of the full observation.\ abbreviation "output_eq v s \ output (shd s) = v" text_raw\\snip{ltlVName}{1}{2}{%\ datatype ltl_vname = Ip nat | Op nat | Rg nat text_raw\}%endsnip\ subsubsection\Checking Arbitrary Expressions\ text\The \textsc{check\_exp} function takes a guard expression and returns true if the guard expression evaluates to true in the given state.\ type_synonym ltl_gexp = "ltl_vname gexp" definition join_iro :: "value list \ registers \ outputs \ ltl_vname datastate" where "join_iro i r p = (\x. case x of Rg n \ r $ n | Ip n \ Some (i ! n) | Op n \ p ! n )" lemma join_iro_R [simp]: "join_iro i r p (Rg n) = r $ n" by (simp add: join_iro_def) abbreviation "check_exp g s \ (gval g (join_iro (snd (action (shd s))) (datastate (shd s)) (output (shd s))) = trilean.true)" lemma alw_ev: "alw f = not (ev (\s. \f s))" by simp lemma alw_state_eq_smap: "alw (state_eq s) ss = alw (\ss. shd ss = s) (smap statename ss)" apply standard apply (simp add: alw_iff_sdrop ) by (simp add: alw_mono alw_smap ) subsection\Sink State\ text\Once the sink state is entered, it cannot be left and there are no outputs or updates henceforth.\ lemma shd_state_is_none: "(state_eq None) (make_full_observation e None r p t)" by simp lemma unfold_observe_none: "make_full_observation e None d p t = (\statename = None, datastate = d, action=(shd t), output = p\##(make_full_observation e None d [] (stl t)))" by (simp add: stream.expand) lemma once_none_always_none_aux: assumes "\ p r i. j = (make_full_observation e None r p) i" shows "alw (state_eq None) j" using assms apply coinduct apply simp by fastforce lemma once_none_always_none: "alw (state_eq None) (make_full_observation e None r p t)" using once_none_always_none_aux by blast lemma once_none_nxt_always_none: "alw (nxt (state_eq None)) (make_full_observation e None r p t)" using once_none_always_none by (simp add: alw_iff_sdrop del: sdrop.simps) lemma snth_sconst: "(\i. s !! i = h) = (s = sconst h)" by (auto simp add: sconst_alt sset_range) lemma alw_sconst: "(alw (\xs. shd xs = h) t) = (t = sconst h)" by (simp add: snth_sconst[symmetric] alw_iff_sdrop) lemma smap_statename_None: "smap statename (make_full_observation e None r p i) = sconst None" by (meson EFSM_LTL.alw_sconst alw_state_eq_smap once_none_always_none) lemma alw_not_some: "alw (\xs. statename (shd xs) \ Some s) (make_full_observation e None r p t)" by (metis (mono_tags, lifting) alw_mono once_none_always_none option.distinct(1) ) lemma state_none: "((state_eq None) impl nxt (state_eq None)) (make_full_observation e s r p t)" by simp lemma state_none_2: "(state_eq None) (make_full_observation e s r p t) \ (state_eq None) (make_full_observation e s r p (stl t))" by simp lemma no_output_none_aux: assumes "\ p r i. j = (make_full_observation e None r []) i" shows "alw (output_eq []) j" using assms apply coinduct apply simp by fastforce lemma no_output_none: "nxt (alw (output_eq [])) (make_full_observation e None r p t)" using no_output_none_aux by auto lemma nxt_alw: "nxt (alw P) s \ alw (nxt P) s" by (simp add: alw_iff_sdrop) lemma no_output_none_nxt: "alw (nxt (output_eq [])) (make_full_observation e None r p t)" using nxt_alw no_output_none by blast lemma no_output_none_if_empty: "alw (output_eq []) (make_full_observation e None r [] t)" by (metis (mono_tags, lifting) alw_nxt make_full_observation.simps(1) no_output_none state.select_convs(4)) lemma no_updates_none_aux: assumes "\ p i. j = (make_full_observation e None r p) i" shows "alw (\x. datastate (shd x) = r) j" using assms apply coinduct by fastforce lemma no_updates_none: "alw (\x. datastate (shd x) = r) (make_full_observation e None r p t)" using no_updates_none_aux by blast lemma action_components: "(label_eq l aand input_eq i) s = (action (shd s) = (String.implode l, i))" by (metis fst_conv prod.collapse snd_conv) end diff --git a/thys/FO_Theory_Rewriting/FOL_Extra.thy b/thys/FO_Theory_Rewriting/FOL_Extra.thy --- a/thys/FO_Theory_Rewriting/FOL_Extra.thy +++ b/thys/FO_Theory_Rewriting/FOL_Extra.thy @@ -1,405 +1,405 @@ theory FOL_Extra imports Type_Instances_Impl "FOL-Fitting.FOL_Fitting" "HOL-Library.FSet" begin section \Additional support for FOL-Fitting\ subsection \Iff\ definition Iff where "Iff p q = And (Impl p q) (Impl q p)" lemma eval_Iff: "eval e f g (Iff p q) \ (eval e f g p \ eval e f g q)" by (auto simp: Iff_def) subsection \Replacement of subformulas\ datatype ('a, 'b) ctxt = Hole | And1 "('a, 'b) ctxt" "('a, 'b) form" | And2 "('a, 'b) form" "('a, 'b) ctxt" | Or1 "('a, 'b) ctxt" "('a, 'b) form" | Or2 "('a, 'b) form" "('a, 'b) ctxt" | Impl1 "('a, 'b) ctxt" "('a, 'b) form" | Impl2 "('a, 'b) form" "('a, 'b) ctxt" | Neg1 "('a, 'b) ctxt" | Forall1 "('a, 'b) ctxt" | Exists1 "('a, 'b) ctxt" primrec apply_ctxt :: "('a, 'b) ctxt \ ('a, 'b) form \ ('a, 'b) form" where "apply_ctxt Hole p = p" | "apply_ctxt (And1 c v) p = And (apply_ctxt c p) v" | "apply_ctxt (And2 u c) p = And u (apply_ctxt c p)" | "apply_ctxt (Or1 c v) p = Or (apply_ctxt c p) v" | "apply_ctxt (Or2 u c) p = Or u (apply_ctxt c p)" | "apply_ctxt (Impl1 c v) p = Impl (apply_ctxt c p) v" | "apply_ctxt (Impl2 u c) p = Impl u (apply_ctxt c p)" | "apply_ctxt (Neg1 c) p = Neg (apply_ctxt c p)" | "apply_ctxt (Forall1 c) p = Forall (apply_ctxt c p)" | "apply_ctxt (Exists1 c) p = Exists (apply_ctxt c p)" lemma replace_subformula: assumes "\e. eval e f g (Iff p q)" shows "eval e f g (Iff (apply_ctxt c p) (apply_ctxt c q))" by (induct c arbitrary: e) (auto simp: assms[unfolded eval_Iff] Iff_def) subsection \Propositional identities\ lemma prop_ids: "eval e f g (Iff (And p q) (And q p))" "eval e f g (Iff (Or p q) (Or q p))" "eval e f g (Iff (Or p (Or q r)) (Or (Or p q) r))" "eval e f g (Iff (And p (And q r)) (And (And p q) r))" "eval e f g (Iff (Neg (Or p q)) (And (Neg p) (Neg q)))" "eval e f g (Iff (Neg (And p q)) (Or (Neg p) (Neg q)))" (* ... *) by (auto simp: Iff_def) subsection \de Bruijn index manipulation for formulas; cf. @{term liftt}\ primrec liftti :: "nat \ 'a term \ 'a term" where "liftti i (Var j) = (if i > j then Var j else Var (Suc j))" | "liftti i (App f ts) = App f (map (liftti i) ts)" lemma liftts_def': "liftts ts = map liftt ts" by (induct ts) auto text \@{term liftt} is a special case of @{term liftti}\ lemma lifttti_0: "liftti 0 t = liftt t" by (induct t) (auto simp: liftts_def') primrec lifti :: "nat \ ('a, 'b) form \ ('a, 'b) form" where "lifti i FF = FF" | "lifti i TT = TT" | "lifti i (Pred b ts) = Pred b (map (liftti i) ts)" | "lifti i (And p q) = And (lifti i p) (lifti i q)" | "lifti i (Or p q) = Or (lifti i p) (lifti i q)" | "lifti i (Impl p q) = Impl (lifti i p) (lifti i q)" | "lifti i (Neg p) = Neg (lifti i p)" | "lifti i (Forall p) = Forall (lifti (Suc i) p)" | "lifti i (Exists p) = Exists (lifti (Suc i) p)" abbreviation lift where "lift \ lifti 0" text \interaction of @{term lifti} and @{term eval}\ lemma evalts_def': "evalts e f ts = map (evalt e f) ts" by (induct ts) auto lemma evalt_liftti: "evalt (e\i:z\) f (liftti i t) = evalt e f t" by (induct t) (auto simp: evalts_def' cong: map_cong) lemma eval_lifti [simp]: "eval (e\i:z\) f g (lifti i p) = eval e f g p" by (induct p arbitrary: e i) (auto simp: evalt_liftti evalts_def' comp_def) subsection \Quantifier Identities\ lemma quant_ids: "eval e f g (Iff (Neg (Exists p)) (Forall (Neg p)))" "eval e f g (Iff (Neg (Forall p)) (Exists (Neg p)))" "eval e f g (Iff (And p (Forall q)) (Forall (And (lift p) q)))" "eval e f g (Iff (And p (Exists q)) (Exists (And (lift p) q)))" "eval e f g (Iff (Or p (Forall q)) (Forall (Or (lift p) q)))" "eval e f g (Iff (Or p (Exists q)) (Exists (Or (lift p) q)))" (* ... *) by (auto simp: Iff_def) (* We'd need a bit of more machinery to deal with "\x y. P(x,y) \ \y x. P(x, y)": * swapping of de Bruijn indices (perhaps arbitrary permutation?) *) subsection \Function symbols and predicates, with arities.\ primrec predas_form :: "('a, 'b) form \ ('b \ nat) set" where "predas_form FF = {}" | "predas_form TT = {}" | "predas_form (Pred b ts) = {(b, length ts)}" | "predas_form (And p q) = predas_form p \ predas_form q" | "predas_form (Or p q) = predas_form p \ predas_form q" | "predas_form (Impl p q) = predas_form p \ predas_form q" | "predas_form (Neg p) = predas_form p" | "predas_form (Forall p) = predas_form p" | "predas_form (Exists p) = predas_form p" primrec funas_term :: "'a term \ ('a \ nat) set" where "funas_term (Var x) = {}" | "funas_term (App f ts) = {(f, length ts)} \ \(set (map funas_term ts))" primrec terms_form :: "('a, 'b) form \ 'a term set" where "terms_form FF = {}" | "terms_form TT = {}" | "terms_form (Pred b ts) = set ts" | "terms_form (And p q) = terms_form p \ terms_form q" | "terms_form (Or p q) = terms_form p \ terms_form q" | "terms_form (Impl p q) = terms_form p \ terms_form q" | "terms_form (Neg p) = terms_form p" | "terms_form (Forall p) = terms_form p" | "terms_form (Exists p) = terms_form p" definition funas_form :: "('a, 'b) form \ ('a \ nat) set" where "funas_form f \ \(funas_term ` terms_form f)" subsection \Negation Normal Form\ inductive is_nnf :: "('a, 'b) form \ bool" where "is_nnf TT" | "is_nnf FF" | "is_nnf (Pred p ts)" | "is_nnf (Neg (Pred p ts))" | "is_nnf p \ is_nnf q \ is_nnf (And p q)" | "is_nnf p \ is_nnf q \ is_nnf (Or p q)" | "is_nnf p \ is_nnf (Forall p)" | "is_nnf p \ is_nnf (Exists p)" primrec nnf' :: "bool \ ('a, 'b) form \ ('a, 'b) form" where "nnf' b TT = (if b then TT else FF)" | "nnf' b FF = (if b then FF else TT)" | "nnf' b (Pred p ts) = (if b then id else Neg) (Pred p ts)" | "nnf' b (And p q) = (if b then And else Or) (nnf' b p) (nnf' b q)" | "nnf' b (Or p q) = (if b then Or else And) (nnf' b p) (nnf' b q)" | "nnf' b (Impl p q) = (if b then Or else And) (nnf' (\ b) p) (nnf' b q)" | "nnf' b (Neg p) = nnf' (\ b) p" | "nnf' b (Forall p) = (if b then Forall else Exists) (nnf' b p)" | "nnf' b (Exists p) = (if b then Exists else Forall) (nnf' b p)" lemma eval_nnf': "eval e f g (nnf' b p) \ (eval e f g p \ b)" by (induct p arbitrary: e b) auto lemma is_nnf_nnf': "is_nnf (nnf' b p)" by (induct p arbitrary: b) (auto intro: is_nnf.intros) abbreviation nnf where "nnf \ nnf' True" lemmas nnf_simpls [simp] = eval_nnf'[where b = True, unfolded eq_True] is_nnf_nnf'[where b = True] subsection \Reasoning modulo ACI01\ datatype ('a, 'b) form_aci = TT_aci | FF_aci | Pred_aci bool 'b "'a term list" | And_aci "('a, 'b) form_aci fset" | Or_aci "('a, 'b) form_aci fset" | Forall_aci "('a, 'b) form_aci" | Exists_aci "('a, 'b) form_aci" text \evaluation, see @{const eval}\ primrec eval_aci :: \(nat \ 'c) \ ('a \ 'c list \ 'c) \ ('b \ 'c list \ bool) \ ('a, 'b) form_aci \ bool\ where "eval_aci e f g FF_aci \ False" | "eval_aci e f g TT_aci \ True" | "eval_aci e f g (Pred_aci b a ts) \ (g a (evalts e f ts) \ b)" | "eval_aci e f g (And_aci ps) \ fBall (fimage (eval_aci e f g) ps) id" | "eval_aci e f g (Or_aci ps) \ fBex (fimage (eval_aci e f g) ps) id" | "eval_aci e f g (Forall_aci p) \ (\z. eval_aci (e\0:z\) f g p)" | "eval_aci e f g (Exists_aci p) \ (\z. eval_aci (e\0:z\) f g p)" text \smart constructor: conjunction\ fun and_aci where "and_aci FF_aci _ = FF_aci" | "and_aci _ FF_aci = FF_aci" | "and_aci TT_aci q = q" | "and_aci p TT_aci = p" | "and_aci (And_aci ps) (And_aci qs) = And_aci (ps |\| qs)" | "and_aci (And_aci ps) q = And_aci (ps |\| {|q|})" | "and_aci p (And_aci qs) = And_aci ({|p|} |\| qs)" | "and_aci p q = (if p = q then p else And_aci {|p,q|})" lemma eval_and_aci [simp]: "eval_aci e f g (and_aci p q) \ eval_aci e f g p \ eval_aci e f g q" - by (cases "(p, q)" rule: and_aci.cases) (simp_all add: fBall_funion, meson+) + by (cases "(p, q)" rule: and_aci.cases) (simp_all add: ball_Un, meson+) declare and_aci.simps [simp del] text \smart constructor: disjunction\ fun or_aci where "or_aci TT_aci _ = TT_aci" | "or_aci _ TT_aci = TT_aci" | "or_aci FF_aci q = q" | "or_aci p FF_aci = p" | "or_aci (Or_aci ps) (Or_aci qs) = Or_aci (ps |\| qs)" | "or_aci (Or_aci ps) q = Or_aci (ps |\| {|q|})" | "or_aci p (Or_aci qs) = Or_aci ({|p|} |\| qs)" | "or_aci p q = (if p = q then p else Or_aci {|p,q|})" lemma eval_or_aci [simp]: "eval_aci e f g (or_aci p q) \ eval_aci e f g p \ eval_aci e f g q" - by (cases "(p, q)" rule: or_aci.cases) (simp_all add: fBex_funion, meson+) + by (cases "(p, q)" rule: or_aci.cases) (simp_all add: bex_Un, meson+) declare or_aci.simps [simp del] text \convert negation normal form to ACIU01 normal form\ fun nnf_to_aci :: "('a, 'b) form \ ('a, 'b) form_aci" where "nnf_to_aci FF = FF_aci" | "nnf_to_aci TT = TT_aci" | "nnf_to_aci (Pred b ts) = Pred_aci True b ts" | "nnf_to_aci (Neg (Pred b ts)) = Pred_aci False b ts" | "nnf_to_aci (And p q) = and_aci (nnf_to_aci p) (nnf_to_aci q)" | "nnf_to_aci (Or p q) = or_aci (nnf_to_aci p) (nnf_to_aci q)" | "nnf_to_aci (Forall p) = Forall_aci (nnf_to_aci p)" | "nnf_to_aci (Exists p) = Exists_aci (nnf_to_aci p)" | "nnf_to_aci _ = undefined" (* the remaining cases are impossible for NNFs *) lemma eval_nnf_to_aci: "is_nnf p \ eval_aci e f g (nnf_to_aci p) \ eval e f g p" by (induct p arbitrary: e rule: is_nnf.induct) simp_all subsection \A (mostly) Propositional Equivalence Check\ text \We reason modulo $\forall = \neg\exists\neg$, de Morgan, double negation, and ACUI01 of $\vee$ and $\wedge$, by converting to negation normal form, and then collapsing conjunctions and disjunctions taking units, absorption, commutativity, associativity, and idempotence into account. We only need soundness for a certifier.\ lemma check_equivalence_by_nnf_aci: "nnf_to_aci (nnf p) = nnf_to_aci (nnf q) \ eval e f g p \ eval e f g q" by (metis eval_nnf_to_aci is_nnf_nnf' eval_nnf') subsection \Reasoning modulo ACI01\ datatype ('a, 'b) form_list_aci = TT_aci | FF_aci | Pred_aci bool 'b "'a term list" | And_aci "('a, 'b) form_list_aci list" | Or_aci "('a, 'b) form_list_aci list" | Forall_aci "('a, 'b) form_list_aci" | Exists_aci "('a, 'b) form_list_aci" text \evaluation, see @{const eval}\ fun eval_list_aci :: \(nat \ 'c) \ ('a \ 'c list \ 'c) \ ('b \ 'c list \ bool) \ ('a, 'b) form_list_aci \ bool\ where "eval_list_aci e f g FF_aci \ False" | "eval_list_aci e f g TT_aci \ True" | "eval_list_aci e f g (Pred_aci b a ts) \ (g a (evalts e f ts) \ b)" | "eval_list_aci e f g (And_aci ps) \ list_all (\ fm. eval_list_aci e f g fm) ps" | "eval_list_aci e f g (Or_aci ps) \ list_ex (\ fm. eval_list_aci e f g fm) ps" | "eval_list_aci e f g (Forall_aci p) \ (\z. eval_list_aci (e\0:z\) f g p)" | "eval_list_aci e f g (Exists_aci p) \ (\z. eval_list_aci (e\0:z\) f g p)" text \smart constructor: conjunction\ fun and_list_aci where "and_list_aci FF_aci _ = FF_aci" | "and_list_aci _ FF_aci = FF_aci" | "and_list_aci TT_aci q = q" | "and_list_aci p TT_aci = p" | "and_list_aci (And_aci ps) (And_aci qs) = And_aci (remdups (ps @ qs))" | "and_list_aci (And_aci ps) q = And_aci (List.insert q ps)" | "and_list_aci p (And_aci qs) = And_aci (List.insert p qs)" | "and_list_aci p q = (if p = q then p else And_aci [p,q])" lemma eval_and_list_aci [simp]: "eval_list_aci e f g (and_list_aci p q) \ eval_list_aci e f g p \ eval_list_aci e f g q" apply (cases "(p, q)" rule: and_list_aci.cases) apply (simp_all add: list.pred_set list_ex_iff) apply blast+ done declare and_list_aci.simps [simp del] text \smart constructor: disjunction\ fun or_list_aci where "or_list_aci TT_aci _ = TT_aci" | "or_list_aci _ TT_aci = TT_aci" | "or_list_aci FF_aci q = q" | "or_list_aci p FF_aci = p" | "or_list_aci (Or_aci ps) (Or_aci qs) = Or_aci (remdups (ps @ qs))" | "or_list_aci (Or_aci ps) q = Or_aci (List.insert q ps)" | "or_list_aci p (Or_aci qs) = Or_aci (List.insert p qs)" | "or_list_aci p q = (if p = q then p else Or_aci [p,q])" lemma eval_or_list_aci [simp]: "eval_list_aci e f g (or_list_aci p q) \ eval_list_aci e f g p \ eval_list_aci e f g q" by (cases "(p, q)" rule: or_list_aci.cases) (simp_all add: list.pred_set list_ex_iff, blast+) declare or_list_aci.simps [simp del] text \convert negation normal form to ACIU01 normal form\ fun nnf_to_list_aci :: "('a, 'b) form \ ('a, 'b) form_list_aci" where "nnf_to_list_aci FF = FF_aci" | "nnf_to_list_aci TT = TT_aci" | "nnf_to_list_aci (Pred b ts) = Pred_aci True b ts" | "nnf_to_list_aci (Neg (Pred b ts)) = Pred_aci False b ts" | "nnf_to_list_aci (And p q) = and_list_aci (nnf_to_list_aci p) (nnf_to_list_aci q)" | "nnf_to_list_aci (Or p q) = or_list_aci (nnf_to_list_aci p) (nnf_to_list_aci q)" | "nnf_to_list_aci (Forall p) = Forall_aci (nnf_to_list_aci p)" | "nnf_to_list_aci (Exists p) = Exists_aci (nnf_to_list_aci p)" | "nnf_to_list_aci _ = undefined" (* the remaining cases are impossible for NNFs *) lemma eval_nnf_to_list_aci: "is_nnf p \ eval_list_aci e f g (nnf_to_list_aci p) \ eval e f g p" by (induct p arbitrary: e rule: is_nnf.induct) simp_all subsection \A (mostly) Propositional Equivalence Check\ text \We reason modulo $\forall = \neg\exists\neg$, de Morgan, double negation, and ACUI01 of $\vee$ and $\wedge$, by converting to negation normal form, and then collapsing conjunctions and disjunctions taking units, absorption, commutativity, associativity, and idempotence into account. We only need soundness for a certifier.\ derive linorder "term" derive compare "term" derive linorder form_list_aci derive compare form_list_aci fun ord_form_list_aci where "ord_form_list_aci TT_aci = TT_aci" | "ord_form_list_aci FF_aci = FF_aci" | "ord_form_list_aci (Pred_aci bool b ts) = Pred_aci bool b ts" | "ord_form_list_aci (And_aci fm) = (And_aci (sort (map ord_form_list_aci fm)))" | "ord_form_list_aci (Or_aci fm) = (Or_aci (sort (map ord_form_list_aci fm)))" | "ord_form_list_aci (Forall_aci fm) = (Forall_aci (ord_form_list_aci fm))" | "ord_form_list_aci (Exists_aci fm) = Exists_aci (ord_form_list_aci fm)" lemma and_list_aci_simps: "and_list_aci TT_aci q = q" "and_list_aci q FF_aci = FF_aci" by (cases q, auto simp add: and_list_aci.simps)+ lemma ord_form_list_idemp: "ord_form_list_aci (ord_form_list_aci q) = ord_form_list_aci q" apply (induct q) apply (auto simp: list.set_map) apply (smt imageE list.set_map map_idI set_sort sorted_sort_id sorted_sort_key)+ done lemma eval_lsit_aci_ord_form_list_aci: "eval_list_aci e f g (ord_form_list_aci p) \ eval_list_aci e f g p" by (induct p arbitrary: e) (auto simp: list.pred_set list_ex_iff) lemma check_equivalence_by_nnf_sortedlist_aci: "ord_form_list_aci (nnf_to_list_aci (nnf p)) = ord_form_list_aci (nnf_to_list_aci (nnf q)) \ eval e f g p \ eval e f g q" by (metis eval_nnf_to_list_aci eval_lsit_aci_ord_form_list_aci is_nnf_nnf' eval_nnf') hide_type (open) "term" hide_const (open) Var hide_type (open) ctxt end \ No newline at end of file diff --git a/thys/Higher_Order_Terms/Term_to_Nterm.thy b/thys/Higher_Order_Terms/Term_to_Nterm.thy --- a/thys/Higher_Order_Terms/Term_to_Nterm.thy +++ b/thys/Higher_Order_Terms/Term_to_Nterm.thy @@ -1,657 +1,665 @@ chapter \Converting between \term\s and \nterm\s\ theory Term_to_Nterm imports Fresh_Class Find_First Term Nterm begin section \\\\-equivalence\ inductive alpha_equiv :: "(name, name) fmap \ nterm \ nterm \ bool" where const: "alpha_equiv env (Nconst x) (Nconst x)" | var1: "x |\| fmdom env \ x |\| fmran env \ alpha_equiv env (Nvar x) (Nvar x)" | var2: "fmlookup env x = Some y \ alpha_equiv env (Nvar x) (Nvar y)" | abs: "alpha_equiv (fmupd x y env) n1 n2 \ alpha_equiv env (\\<^sub>n x. n1) (\\<^sub>n y. n2)" | app: "alpha_equiv env n1 n2 \ alpha_equiv env m1 m2 \ alpha_equiv env (n1 $\<^sub>n m1) (n2 $\<^sub>n m2)" code_pred alpha_equiv . abbreviation alpha_eq :: "nterm \ nterm \ bool" (infixl "\\<^sub>\" 50) where "alpha_eq n1 n2 \ alpha_equiv fmempty n1 n2" lemma alpha_equiv_refl[intro?]: assumes "fmpred (=) \" shows "alpha_equiv \ t t" using assms proof (induction t arbitrary: \) case Napp show ?case apply (rule alpha_equiv.app; rule Napp) using Napp.prems unfolding fdisjnt_alt_def by auto qed (auto simp: fdisjnt_alt_def intro: alpha_equiv.intros) corollary alpha_eq_refl: "alpha_eq t t" by (auto intro: alpha_equiv_refl) section \From @{typ term} to @{typ nterm}\ fun term_to_nterm :: "name list \ term \ (name, nterm) state" where "term_to_nterm _ (Const name) = State_Monad.return (Nconst name)" | "term_to_nterm _ (Free name) = State_Monad.return (Nvar name)" | "term_to_nterm \ (Bound n) = State_Monad.return (Nvar (\ ! n))" | "term_to_nterm \ (\ t) = do { n \ fresh_create; e \ term_to_nterm (n # \) t; State_Monad.return (\\<^sub>n n. e) }" | "term_to_nterm \ (t\<^sub>1 $ t\<^sub>2) = do { e\<^sub>1 \ term_to_nterm \ t\<^sub>1; e\<^sub>2 \ term_to_nterm \ t\<^sub>2; State_Monad.return (e\<^sub>1 $\<^sub>n e\<^sub>2) }" lemmas term_to_nterm_induct = term_to_nterm.induct[case_names const free bound abs app] lemma term_to_nterm: assumes "no_abs t" shows "fst (run_state (term_to_nterm \ t) x) = convert_term t" using assms apply (induction arbitrary: x) apply auto by (auto simp: free_term_def free_nterm_def const_term_def const_nterm_def app_term_def app_nterm_def split_beta split: prod.splits) definition term_to_nterm' :: "term \ nterm" where "term_to_nterm' t = frun_fresh (term_to_nterm [] t) (frees t)" lemma term_to_nterm_mono: "mono_state (term_to_nterm \ x)" by (induction \ x rule: term_to_nterm.induct) (auto intro: bind_mono_strong) lemma term_to_nterm_vars0: assumes "wellformed' (length \) t" shows "frees (fst (run_state (term_to_nterm \ t) s)) |\| frees t |\| fset_of_list \" using assms proof (induction \ t arbitrary: s rule: term_to_nterm_induct) case (bound \ i) hence "\ ! i |\| fset_of_list \" including fset.lifting by transfer auto thus ?case by (auto simp: State_Monad.return_def) next case (abs \ t) let ?x = "next s" from abs have "frees (fst (run_state (term_to_nterm (?x # \) t) ?x)) |\| frees t |\| {|?x|} |\| fset_of_list \" by simp thus ?case by (auto simp: create_alt_def split_beta) qed (auto simp: split_beta) corollary term_to_nterm_vars: assumes "wellformed t" shows "frees (fresh_frun (term_to_nterm [] t) F) |\| frees t" proof - let ?\ = "[]" from assms have "wellformed' (length ?\) t" by simp hence "frees (fst (run_state (term_to_nterm ?\ t) (fNext F))) |\| (frees t |\| fset_of_list ?\)" by (rule term_to_nterm_vars0) thus ?thesis by (simp add: fresh_fNext_def fresh_frun_def) qed corollary term_to_nterm_closed: "closed t \ wellformed t \ closed (term_to_nterm' t)" using term_to_nterm_vars[where F = "frees t" and t = t, simplified] unfolding closed_except_def term_to_nterm'_def by (simp add: fresh_frun_def) lemma term_to_nterm_consts: "pred_state (\t'. consts t' = consts t) (term_to_nterm \ t)" apply (rule pred_stateI) apply (induction t arbitrary: \) apply (auto split: prod.splits) done section \From @{typ nterm} to @{typ term}\ fun nterm_to_term :: "name list \ nterm \ term" where "nterm_to_term \ (Nconst name) = Const name" | "nterm_to_term \ (Nvar name) = (case find_first name \ of Some n \ Bound n | None \ Free name)" | "nterm_to_term \ (t $\<^sub>n u) = nterm_to_term \ t $ nterm_to_term \ u" | "nterm_to_term \ (\\<^sub>n x. t) = \ nterm_to_term (x # \) t" lemma nterm_to_term: assumes "no_abs t" "fdisjnt (fset_of_list \) (frees t)" shows "nterm_to_term \ t = convert_term t" using assms proof (induction arbitrary: \) case (free name) then show ?case apply simp apply (auto simp: free_nterm_def free_term_def fdisjnt_alt_def split: option.splits) apply (rule find_first_none) by (metis fset_of_list_elem) next case (const name) show ?case apply simp by (simp add: const_nterm_def const_term_def) next case (app t\<^sub>1 t\<^sub>2) then have "nterm_to_term \ t\<^sub>1 = convert_term t\<^sub>1" "nterm_to_term \ t\<^sub>2 = convert_term t\<^sub>2" by (auto simp: fdisjnt_alt_def finter_funion_distrib) then show ?case apply simp by (simp add: app_nterm_def app_term_def) qed abbreviation "nterm_to_term' \ nterm_to_term []" lemma nterm_to_term': "no_abs t \ nterm_to_term' t = convert_term t" by (auto simp: fdisjnt_alt_def nterm_to_term) lemma nterm_to_term_frees[simp]: "frees (nterm_to_term \ t) = frees t - fset_of_list \" proof (induction t arbitrary: \) case (Nvar name) show ?case proof (cases "find_first name \") case None hence "name |\| fset_of_list \" including fset.lifting by transfer (metis find_first_some option.distinct(1)) with None show ?thesis by auto next case (Some u) hence "name |\| fset_of_list \" including fset.lifting by transfer (metis find_first_none option.distinct(1)) with Some show ?thesis by auto qed qed (auto split: option.splits) section \Correctness\ text \Some proofs in this section have been contributed by Yu Zhang.\ lemma term_to_nterm_nterm_to_term0: assumes "wellformed' (length \) t" "fdisjnt (fset_of_list \) (frees t)" "distinct \" assumes "fBall (frees t |\| fset_of_list \) (\x. x \ s)" shows "nterm_to_term \ (fst (run_state (term_to_nterm \ t) s)) = t" using assms proof (induction \ t arbitrary: s rule: term_to_nterm_induct) case (free \ name) hence "fdisjnt (fset_of_list \) {|name|}" by simp hence "name \ set \" including fset.lifting by transfer' (simp add: disjnt_def) hence "find_first name \ = None" by (rule find_first_none) thus ?case by (simp add: State_Monad.return_def) next case (bound \ i) thus ?case by (simp add: State_Monad.return_def find_first_some_index) next case (app \ t\<^sub>1 t\<^sub>2) have "fdisjnt (fset_of_list \) (frees t\<^sub>1)" apply (rule fdisjnt_subset_right[where N = "frees t\<^sub>1 |\| frees t\<^sub>2"]) using app by auto have "fdisjnt (fset_of_list \) (frees t\<^sub>2)" apply (rule fdisjnt_subset_right[where N = "frees t\<^sub>1 |\| frees t\<^sub>2"]) using app by auto have s: "s \ snd (run_state (term_to_nterm \ t\<^sub>1) s)" apply (rule state_io_relD[OF term_to_nterm_mono]) apply (rule surjective_pairing) done show ?case apply (auto simp: split_beta) subgoal apply (rule app) subgoal using app by simp subgoal by fact subgoal by fact using app by auto subgoal apply (rule app) subgoal using app by simp subgoal by fact subgoal by fact using app s by force+ done next case (abs \ t) have "next s |\| frees t |\| fset_of_list \" - using abs(5)[simplified] - by (rule next_ge_fall) + using abs(5) next_ge_fall by auto have "nterm_to_term (next s # \) (fst (run_state (term_to_nterm (next s # \) t) (next s))) = t" proof (subst abs) show "wellformed' (length (next s # \)) t" using abs by auto show "fdisjnt (fset_of_list (next s # \)) (frees t)" apply simp apply (rule fdisjnt_insert) using \next s |\| frees t |\| fset_of_list \\ abs by auto show "distinct (next s # \)" apply simp apply rule using \next s |\| frees t |\| fset_of_list \\ apply (simp add: fset_of_list_elem) apply fact done - show "fBall (frees t |\| fset_of_list (next s # \)) (\x. x \ next s)" - using abs(5) apply simp - apply (rule fBall_pred_weaken[where P = "\x. x \ s"]) - subgoal - apply simp - by (meson dual_order.strict_implies_order dual_order.strict_trans2 fresh_class.next_ge) - subgoal by assumption - done + have "fBall (frees t |\| fset_of_list \) (\x. x \ next s)" + proof (rule fBall_pred_weaken) + show "fBall (frees t |\| fset_of_list \) (\x. x \ s)" + using abs(5) by simp + next + show "\x. x |\| frees t |\| fset_of_list \ \ x \ s \ x \ next s" + by (metis Fresh_Class.next_ge dual_order.strict_trans less_eq_name_def) + qed + thus "fBall (frees t |\| fset_of_list (next s # \)) (\x. x \ next s)" + by simp qed auto thus ?case by (auto simp: split_beta create_alt_def) qed (auto simp: State_Monad.return_def) lemma term_to_nterm_nterm_to_term: assumes "wellformed t" "frees t |\| S" shows "nterm_to_term' (frun_fresh (term_to_nterm [] t) (S |\| Q)) = t" proof (rule term_to_nterm_nterm_to_term0) show "wellformed' (length []) t" using assms by simp next show "fdisjnt (fset_of_list []) (frees t)" unfolding fdisjnt_alt_def by simp next have "fBall (S |\| Q) (\x. x < fresh.fNext next default (S |\| Q))" by (metis fNext_geq_not_member fresh_fNext_def le_less_linear fBallI) hence "fBall (S |\| Q) (\x. x \ fresh.fNext next default (S |\| Q))" by (meson fBall_pred_weaken less_eq_name_def) thus "fBall (frees t |\| fset_of_list []) (\x. x \ fresh.fNext next default (S |\| Q))" using \frees t |\| S\ by auto qed simp corollary term_to_nterm_nterm_to_term_simple: assumes "wellformed t" shows "nterm_to_term' (term_to_nterm' t) = t" unfolding term_to_nterm'_def using assms by (metis order_refl sup.idem term_to_nterm_nterm_to_term) lemma nterm_to_term_eq: assumes "frees u |\| fset_of_list (common_prefix \ \')" shows "nterm_to_term \ u = nterm_to_term \' u" using assms proof (induction u arbitrary: \ \') case (Nvar name) hence "name \ set (common_prefix \ \')" unfolding frees_nterm.simps including fset.lifting by transfer' simp thus ?case by (auto simp: common_prefix_find) next case (Nabs x t) hence *: "frees t - {|x|} |\| fset_of_list (common_prefix \ \')" by simp show ?case apply simp apply (rule Nabs) using * Nabs by auto qed auto corollary nterm_to_term_eq_closed: "closed t \ nterm_to_term \ t = nterm_to_term \' t" by (rule nterm_to_term_eq) (auto simp: closed_except_def) lemma nterm_to_term_wellformed: "wellformed' (length \) (nterm_to_term \ t)" proof (induction t arbitrary: \) case (Nabs x t) hence "wellformed' (Suc (length \)) (nterm_to_term (x # \) t)" by (metis length_Cons) thus ?case by auto qed (auto simp: find_first_correct split: option.splits) corollary nterm_to_term_closed_wellformed: "closed t \ wellformed (nterm_to_term \ t)" by (metis Ex_list_of_length nterm_to_term_eq_closed nterm_to_term_wellformed) lemma nterm_to_term_term_to_nterm: assumes "frees t |\| fset_of_list \" "length \ = length \'" shows "alpha_equiv (fmap_of_list (zip \ \')) t (fst (run_state (term_to_nterm \' (nterm_to_term \ t)) s))" using assms proof (induction \ t arbitrary: s \' rule:nterm_to_term.induct) case (4 \ x t) show ?case apply (simp add: split_beta) apply (rule alpha_equiv.abs) using "4.IH"[where \' = "next s # \'"] "4.prems" by (fastforce simp: create_alt_def intro: alpha_equiv.intros) qed (force split: option.splits intro: find_first_some intro!: alpha_equiv.intros simp: fset_of_list_elem find_first_in_map split_beta fdisjnt_alt_def)+ corollary nterm_to_term_term_to_nterm': "closed t \ t \\<^sub>\ term_to_nterm' (nterm_to_term' t)" unfolding term_to_nterm'_def closed_except_def apply (rule nterm_to_term_term_to_nterm[where \ = "[]" and \' = "[]", simplified]) by auto context begin private lemma term_to_nterm_alpha_equiv0: "length \1 = length \2 \ distinct \1 \ distinct \2 \ wellformed' (length \1) t1 \ fresh_fin (frees t1 |\| fset_of_list \1) s1 \ fdisjnt (fset_of_list \1) (frees t1) \ fresh_fin (frees t1 |\| fset_of_list \2) s2 \ fdisjnt (fset_of_list \2) (frees t1) \ alpha_equiv (fmap_of_list (zip \1 \2)) (fst( run_state (term_to_nterm \1 t1) s1)) (fst ( run_state (term_to_nterm \2 t1) s2))" proof (induction \1 t1 arbitrary: \2 s1 s2 rule:term_to_nterm_induct) case (free \1 name) then have "name |\| fmdom (fmap_of_list (zip \1 \2))" unfolding fdisjnt_alt_def by force moreover have "name |\| fmran (fmap_of_list (zip \1 \2))" apply rule apply (subst (asm) fmran_of_list) apply (subst (asm) fset_of_list_map[symmetric]) apply (subst (asm) distinct_clearjunk_id) subgoal apply (subst map_fst_zip) apply fact apply fact done apply (subst (asm) map_snd_zip) apply fact using free unfolding fdisjnt_alt_def by fastforce ultimately show ?case by (force intro:alpha_equiv.intros) next case (abs \ t) have *: "next s1 > s1" "next s2 > s2" using next_ge by force+ - with abs have "next s1 \ set \" "next s2 \ set \2" - by (metis fBall_funion fset_of_list_elem next_ge_fall)+ - have "fresh_fin (frees t |\| fset_of_list \) (next s1)" + from abs.prems(5,7) have "next s1 \ set \" "next s2 \ set \2" + unfolding fBall_funion + by (metis fset_of_list_elem next_ge_fall)+ + moreover have "fresh_fin (frees t |\| fset_of_list \) (next s1)" "fresh_fin (frees t |\| fset_of_list \2) (next s2)" using * abs by (smt dual_order.trans fBall_pred_weaken frees_term.simps(3) less_imp_le)+ - have "fdisjnt (finsert (next s1) (fset_of_list \)) (frees t)" + moreover have "fdisjnt (finsert (next s1) (fset_of_list \)) (frees t)" "fdisjnt (finsert (next s2) (fset_of_list \2)) (frees t)" unfolding fdisjnt_alt_def using abs frees_term.simps by (metis fdisjnt_alt_def finter_finsert_right funionCI inf_commute next_ge_fall)+ - have "wellformed' (Suc (length \2)) t" + moreover have "wellformed' (Suc (length \2)) t" using wellformed'.simps abs by (metis Suc_eq_plus1) - show ?case - apply (auto simp: split_beta create_alt_def) - apply rule - apply (rule abs.IH[of _ "next s2 # \2", simplified]) - by (fact | rule)+ + ultimately show ?case + using abs.prems(1,2,3) + by (auto simp: split_beta create_alt_def + intro: alpha_equiv.abs abs.IH[of _ "next s2 # \2", simplified]) next case (app \1 t\<^sub>1 t\<^sub>2) hence "wellformed' (length \1) t\<^sub>1" "wellformed' (length \1) t\<^sub>2" and "fresh_fin (frees t\<^sub>1 |\| fset_of_list \1) s1" "fresh_fin (frees t\<^sub>1 |\| fset_of_list \2) s2" and "fdisjnt (fset_of_list \1) (frees t\<^sub>1)" "fdisjnt (fset_of_list \2) (frees t\<^sub>1)" and "fdisjnt (fset_of_list \1) (frees t\<^sub>2)" "fdisjnt (fset_of_list \2) (frees t\<^sub>2)" using app unfolding fdisjnt_alt_def by (auto simp: inf_sup_distrib1) have "s1 \ snd (run_state (term_to_nterm \1 t\<^sub>1) s1)" "s2 \ snd (run_state (term_to_nterm \2 t\<^sub>1) s2)" using term_to_nterm_mono by (simp add: state_io_rel_def)+ hence "fresh_fin (frees t\<^sub>2 |\| fset_of_list \1) (snd (run_state (term_to_nterm \1 t\<^sub>1) s1))" using \fresh_fin (frees (t\<^sub>1 $ t\<^sub>2) |\| fset_of_list \1) s1\ by force have "fresh_fin (frees t\<^sub>2 |\| fset_of_list \2) (snd (run_state (term_to_nterm \2 t\<^sub>1) s2))" apply rule using app frees_term.simps \s2 \ _\ dual_order.trans - by (metis fbspec funion_iff) + by (metis funion_iff) show ?case apply (auto simp: split_beta create_alt_def) apply (rule alpha_equiv.app) subgoal - apply (rule app) - apply (simp | fact)+ - done + using app.IH + using \fBall (frees t\<^sub>1 |\| fset_of_list \1) (\y. y \ s1)\ + \fBall (frees t\<^sub>1 |\| fset_of_list \2) (\y. y \ s2)\ + \fdisjnt (fset_of_list \1) (frees t\<^sub>1)\ + \fdisjnt (fset_of_list \2) (frees t\<^sub>1)\ \wellformed' (length \1) t\<^sub>1\ + app.prems(1) app.prems(2) app.prems(3) by blast subgoal - apply (rule app) - apply (simp | fact)+ - done + using app.IH + using \fBall (frees t\<^sub>2 |\| fset_of_list \1) (\y. y \ snd (run_state (term_to_nterm \1 t\<^sub>1) s1))\ + \fBall (frees t\<^sub>2 |\| fset_of_list \2) (\y. y \ snd (run_state (term_to_nterm \2 t\<^sub>1) s2))\ + \fdisjnt (fset_of_list \1) (frees t\<^sub>2)\ + \fdisjnt (fset_of_list \2) (frees t\<^sub>2)\ + \wellformed' (length \1) t\<^sub>2\ + app.prems(1) app.prems(2) app.prems(3) by blast done qed (force intro: alpha_equiv.intros simp: fmlookup_of_list in_set_zip)+ lemma term_to_nterm_alpha_equiv: assumes "length \1 = length \2" "distinct \1" "distinct \2" "closed t" assumes "wellformed' (length \1) t" assumes "fresh_fin (fset_of_list \1) s1" "fresh_fin (fset_of_list \2) s2" shows "alpha_equiv (fmap_of_list (zip \1 \2)) (fst (run_state (term_to_nterm \1 t) s1)) (fst (run_state (term_to_nterm \2 t) s2))" \ \An instantiated version of this lemma with @{prop \\1 = []\} and @{prop \\2 = []\} would not make sense because then it would just be a special case of @{thm [source=true] alpha_eq_refl}.\ using assms by (simp add: fdisjnt_alt_def closed_except_def term_to_nterm_alpha_equiv0) end global_interpretation nrelated: term_struct_rel_strong "(\t n. t = nterm_to_term \ n)" for \ proof (standard, goal_cases) case (5 name t) then show ?case by (cases t) (auto simp: const_term_def const_nterm_def split: option.splits) next case (6 u\<^sub>1 u\<^sub>2 t) then show ?case by (cases t) (auto simp: app_term_def app_nterm_def split: option.splits) qed (auto simp: const_term_def const_nterm_def app_term_def app_nterm_def) lemma env_nrelated_closed: assumes "nrelated.P_env \ env nenv" "closed_env nenv" shows "nrelated.P_env \' env nenv" proof fix x from assms have "rel_option (\t n. t = nterm_to_term \ n) (fmlookup env x) (fmlookup nenv x)" by auto thus "rel_option (\t n. t = nterm_to_term \' n) (fmlookup env x) (fmlookup nenv x)" using assms by (cases rule: option.rel_cases) (auto dest: fmdomI simp: nterm_to_term_eq_closed) qed lemma nrelated_subst: assumes "nrelated.P_env \ env nenv" "closed_env nenv" "fdisjnt (fset_of_list \) (fmdom nenv)" shows "subst (nterm_to_term \ t) env = nterm_to_term \ (subst t nenv)" using assms proof (induction t arbitrary: \ env nenv) case (Nvar name) thus ?case proof (cases rule: fmrel_cases[where x = name]) case (some t\<^sub>1 t\<^sub>2) with Nvar have "name |\| fset_of_list \" unfolding fdisjnt_alt_def by (auto dest: fmdomI) hence "find_first name \ = None" including fset.lifting by transfer' (simp add: find_first_none) with some show ?thesis by auto qed (auto split: option.splits) next case (Nabs x t) show ?case apply simp apply (subst subst_drop[symmetric, where x = x]) subgoal by simp apply (rule Nabs) using Nabs unfolding fdisjnt_alt_def by (auto intro: env_nrelated_closed) qed auto lemma nterm_to_term_insert_dupl: assumes "y \ set (take n \)" "n \ length \" shows "nterm_to_term \ t = incr_bounds (- 1) (Suc n) (nterm_to_term (insert_nth n y \) t)" using assms proof (induction t arbitrary: n \) case (Nvar name) show ?case proof (cases "y = name") case True with Nvar obtain i where "find_first name \ = Some i" "i < n" by (auto elim: find_first_some_strong) hence "find_first name (take n \) = Some i" by (rule find_first_prefix) show ?thesis apply simp apply (subst \find_first name \ = Some i\) apply simp apply (subst find_first_append) apply (subst \find_first name (take n \) = Some i\) apply simp using \i < n\ by simp next case False show ?thesis apply (simp del: insert_nth_take_drop) apply (subst find_first_insert_nth_neq) subgoal using False by simp by (cases "find_first name \") auto qed next case (Nabs x t) show ?case apply simp apply (subst Nabs(1)[where n = "Suc n"]) using Nabs by auto qed auto lemma nterm_to_term_bounds_dupl: assumes "i < length \" "j < length \" "i < j" assumes "\ ! i = \ ! j" shows "j |\| bounds (nterm_to_term \ t)" using assms proof (induction t arbitrary: \ i j) case (Nvar name) show ?case proof (cases "find_first name \") case (Some k) show ?thesis proof assume "j |\| bounds (nterm_to_term \ (Nvar name))" with Some have "find_first name \ = Some j" by simp moreover have "find_first name \ \ Some j" proof (rule find_first_later) show "i < length \" "j < length \" "i < j" by fact+ next show "\ ! j = name" by (rule find_first_correct) fact thus "\ ! i = name" using Nvar by simp qed ultimately show False by blast qed qed simp next case (Nabs x t) show ?case proof assume "j |\| bounds (nterm_to_term \ (\\<^sub>n x. t))" then obtain j' where "j' |\| bounds (nterm_to_term (x # \) t)" "j' > 0" "j = j' - 1" by auto hence "Suc j |\| bounds (nterm_to_term (x # \) t)" by simp moreover have "Suc j |\| bounds (nterm_to_term (x # \) t)" proof (rule Nabs) show "Suc i < length (x # \)" "Suc j < length (x # \)" "Suc i < Suc j" "(x # \) ! Suc i = (x # \) ! Suc j" using Nabs by simp+ qed ultimately show False by blast qed qed auto fun subst_single :: "nterm \ name \ nterm \ nterm" where "subst_single (Nvar s) s' t' = (if s = s' then t' else Nvar s)" | "subst_single (t\<^sub>1 $\<^sub>n t\<^sub>2) s' t' = subst_single t\<^sub>1 s' t' $\<^sub>n subst_single t\<^sub>2 s' t'" | "subst_single (\\<^sub>n x. t) s' t' = (\\<^sub>n x. (if x = s' then t else subst_single t s' t'))" | "subst_single t _ _ = t" lemma subst_single_eq: "subst_single t s t' = subst t (fmap_of_list [(s, t')])" proof (induction t) case (Nabs x t) then show ?case by (cases "x = s") (simp add: fmfilter_alt_defs)+ qed auto lemma nterm_to_term_subst_replace_bound: assumes "closed u'" "n \ length \" "x \ set (take n \)" shows "nterm_to_term \ (subst_single u x u') = replace_bound n (nterm_to_term (insert_nth n x \) u) (nterm_to_term \ u')" using assms proof (induction u arbitrary: n \) case (Nvar name) note insert_nth_take_drop[simp del] show ?case proof (cases "name = x") case True thus ?thesis using Nvar apply (simp add: find_first_insert_nth_eq) apply (subst incr_bounds_eq[where k = 0]) subgoal by simp apply (rule nterm_to_term_closed_wellformed) by auto next case False thus ?thesis apply auto apply (subst find_first_insert_nth_neq) subgoal by simp by (cases "find_first name \") auto qed next case (Nabs y t) note insert_nth_take_drop[simp del] show ?case proof (cases "x = y") case True have "nterm_to_term (y # \) t = replace_bound (Suc n) (nterm_to_term (y # insert_nth n y \) t) (nterm_to_term \ u')" proof (subst replace_bound_eq) show "Suc n |\| bounds (nterm_to_term (y # insert_nth n y \) t)" apply (rule nterm_to_term_bounds_dupl[where i = 0]) subgoal by simp subgoal using Nabs(3) by (simp add: insert_nth_take_drop) subgoal by simp apply simp apply (subst nth_insert_nth_index_eq) using Nabs by auto show "nterm_to_term (y # \) t = incr_bounds (- 1) (Suc n + 1) (nterm_to_term (y # insert_nth n y \) t)" apply (subst nterm_to_term_insert_dupl[where \ = "y # \" and y = y and n = "Suc n"]) using Nabs by auto qed with True show ?thesis by auto next case False have "nterm_to_term (y # \) (subst_single t x u') = replace_bound (Suc n) (nterm_to_term (y # insert_nth n x \) t) (nterm_to_term \ u')" apply (subst Nabs(1)[of "Suc n"]) subgoal by fact subgoal using Nabs by simp subgoal using False Nabs by simp apply (subst nterm_to_term_eq_closed[where t = u']) using Nabs by auto with False show ?thesis by auto qed qed auto corollary nterm_to_term_subst_\: assumes "closed u'" shows "nterm_to_term \ (subst u (fmap_of_list [(x, u')])) = nterm_to_term (x # \) u [nterm_to_term \ u']\<^sub>\" using assms by (rule nterm_to_term_subst_replace_bound[where n = 0, simplified, unfolded subst_single_eq]) 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,1100 @@ (* 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) 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) 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 + have "w \\<^sub>n ltln_of_dnf \ \ (\\ \ fset ` \. \\ \ \. w \\<^sub>n \)" + unfolding ltln_of_dnf_def + proof (rule Or\<^sub>n_And\<^sub>n_image_semantics) + show "finite (fset ` \)" + using assms by blast + next + show "\\. \ \ fset ` \ \ finite \" + by auto + qed then show ?thesis by (metis image_iff) 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 + have "\ \\<^sub>P ltln_of_dnf \ \ (\\ \ fset ` \. \\ \ \. \ \\<^sub>P \)" + unfolding ltln_of_dnf_def + proof (rule Or\<^sub>n_And\<^sub>n_image_prop_semantics) + show "finite (fset ` \)" + using assms by blast + next + show "\\. \ \ fset ` \ \ finite \" + by auto + qed then show ?thesis by (metis image_iff) 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 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, simp add: product_singleton_singleton) lemma fold_union: "Finite_Set.fold (\x. (\) {x}) {} (fset x) = fset x" by (induction x) (simp_all add: 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+ 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/Monomorphic_Monad/Monomorphic_Monad.thy b/thys/Monomorphic_Monad/Monomorphic_Monad.thy --- a/thys/Monomorphic_Monad/Monomorphic_Monad.thy +++ b/thys/Monomorphic_Monad/Monomorphic_Monad.thy @@ -1,4345 +1,4345 @@ (* Title: Monomorphic_Monad.thy Author: Andreas Lochbihler, ETH Zurich *) theory Monomorphic_Monad imports "HOL-Probability.Probability" "HOL-Library.Multiset" "HOL-Library.Countable_Set_Type" begin section \Preliminaries\ lemma (in comp_fun_idem) fold_set_union: "\ finite A; finite B \ \ Finite_Set.fold f x (A \ B) = Finite_Set.fold f (Finite_Set.fold f x A) B" by(induction A arbitrary: x rule: finite_induct)(simp_all add: fold_insert_idem2 del: fold_insert_idem) lemma (in comp_fun_idem) ffold_set_union: "ffold f x (A |\| B) = ffold f (ffold f x A) B" including fset.lifting by(transfer fixing: f)(rule fold_set_union) lemma relcompp_top_top [simp]: "top OO top = top" by(auto simp add: fun_eq_iff) attribute_setup locale_witness = \Scan.succeed Locale.witness_add\ named_theorems monad_unfold "Defining equations for overloaded monad operations" context includes lifting_syntax begin inductive rel_itself :: "'a itself \ 'b itself \ bool" where "rel_itself TYPE(_) TYPE(_)" lemma type_parametric [transfer_rule]: "rel_itself TYPE('a) TYPE('b)" by(simp add: rel_itself.simps) lemma plus_multiset_parametric [transfer_rule]: "(rel_mset A ===> rel_mset A ===> rel_mset A) (+) (+)" apply(rule rel_funI)+ subgoal premises prems using prems by induction(auto intro: rel_mset_Plus) done lemma Mempty_parametric [transfer_rule]: "rel_mset A {#} {#}" by(fact rel_mset_Zero) lemma fold_mset_parametric: assumes 12: "(A ===> B ===> B) f1 f2" and "comp_fun_commute f1" "comp_fun_commute f2" shows "(B ===> rel_mset A ===> B) (fold_mset f1) (fold_mset f2)" proof(rule rel_funI)+ interpret f1: comp_fun_commute f1 by fact interpret f2: comp_fun_commute f2 by fact show "B (fold_mset f1 z1 X) (fold_mset f2 z2 Y)" if "rel_mset A X Y" "B z1 z2" for z1 z2 X Y using that(1) by(induction R\A X Y)(simp_all add: that(2) 12[THEN rel_funD, THEN rel_funD]) qed lemma rel_fset_induct [consumes 1, case_names empty step, induct pred: rel_fset]: assumes XY: "rel_fset A X Y" and empty: "P {||} {||}" and step: "\X Y x y. \ rel_fset A X Y; P X Y; A x y; x |\| X \ y |\| Y \ \ P (finsert x X) (finsert y Y)" shows "P X Y" proof - from XY obtain Z where X: "X = fst |`| Z" and Y: "Y = snd |`| Z" and Z: "fBall Z (\(x, y). A x y)" - unfolding fset.in_rel by(auto simp add: fBall.rep_eq) + unfolding fset.in_rel by auto from Z show ?thesis unfolding X Y proof(induction Z) case (insert xy Z) obtain x y where [simp]: "xy = (x, y)" by(cases xy) show ?case using insert apply(cases "x |\| fst |`| Z \ y |\| snd |`| Z") apply(simp add: finsert_absorb) - apply(auto intro!: step simp add: fset.in_rel fBall.rep_eq; blast) + apply(auto intro!: step simp add: fset.in_rel; blast) done qed(simp add: assms) qed lemma ffold_parametric: assumes 12: "(A ===> B ===> B) f1 f2" and "comp_fun_idem f1" "comp_fun_idem f2" shows "(B ===> rel_fset A ===> B) (ffold f1) (ffold f2)" proof(rule rel_funI)+ interpret f1: comp_fun_idem f1 by fact interpret f2: comp_fun_idem f2 by fact show "B (ffold f1 z1 X) (ffold f2 z2 Y)" if "rel_fset A X Y" "B z1 z2" for z1 z2 X Y using that(1) by(induction)(simp_all add: that(2) 12[THEN rel_funD, THEN rel_funD]) qed end lemma rel_set_Grp: "rel_set (BNF_Def.Grp A f) = BNF_Def.Grp {X. X \ A} (image f)" by(auto simp add: fun_eq_iff Grp_def rel_set_def) context includes cset.lifting begin lemma cUNION_assoc: "cUNION (cUNION A f) g = cUNION A (\x. cUNION (f x) g)" by transfer auto lemma cUnion_cempty [simp]: "cUnion cempty = cempty" by transfer simp lemma cUNION_cempty [simp]: "cUNION cempty f = cempty" by simp lemma cUnion_cinsert: "cUnion (cinsert x A) = cUn x (cUnion A)" by transfer simp lemma cUNION_cinsert: "cUNION (cinsert x A) f = cUn (f x) (cUNION A f)" by (simp add: cUnion_cinsert) lemma cUnion_csingle [simp]: "cUnion (csingle x) = x" by (simp add: cUnion_cinsert) lemma cUNION_csingle [simp]: "cUNION (csingle x) f = f x" by simp lemma cUNION_csingle2 [simp]: "cUNION A csingle = A" by (fact cUN_csingleton) lemma cUNION_cUn: "cUNION (cUn A B) f = cUn (cUNION A f) (cUNION B f)" by simp lemma cUNION_parametric [transfer_rule]: includes lifting_syntax shows "(rel_cset A ===> (A ===> rel_cset B) ===> rel_cset B) cUNION cUNION" unfolding rel_fun_def by transfer(blast intro: rel_set_UNION) end locale three = fixes tytok :: "'a itself" assumes ex_three: "\x y z :: 'a. x \ y \ x \ z \ y \ z" begin definition threes :: "'a \ 'a \ 'a" where "threes = (SOME (x, y, z). x \ y \ x \ z \ y \ z)" definition three\<^sub>1 :: 'a ("\<^bold>1") where "\<^bold>1 = fst threes" definition three\<^sub>2 :: 'a ("\<^bold>2") where "\<^bold>2 = fst (snd threes)" definition three\<^sub>3 :: 'a ("\<^bold>3") where "\<^bold>3 = snd (snd (threes))" lemma three_neq_aux: "\<^bold>1 \ \<^bold>2" "\<^bold>1 \ \<^bold>3" "\<^bold>2 \ \<^bold>3" proof - have "\<^bold>1 \ \<^bold>2 \ \<^bold>1 \ \<^bold>3 \ \<^bold>2 \ \<^bold>3" unfolding three\<^sub>1_def three\<^sub>2_def three\<^sub>3_def threes_def split_def by(rule someI_ex)(use ex_three in auto) then show "\<^bold>1 \ \<^bold>2" "\<^bold>1 \ \<^bold>3" "\<^bold>2 \ \<^bold>3" by simp_all qed lemmas three_neq [simp] = three_neq_aux three_neq_aux[symmetric] inductive rel_12_23 :: "'a \ 'a \ bool" where "rel_12_23 \<^bold>1 \<^bold>2" | "rel_12_23 \<^bold>2 \<^bold>3" lemma bi_unique_rel_12_23 [simp, transfer_rule]: "bi_unique rel_12_23" by(auto simp add: bi_unique_def rel_12_23.simps) inductive rel_12_21 :: "'a \ 'a \ bool" where "rel_12_21 \<^bold>1 \<^bold>2" | "rel_12_21 \<^bold>2 \<^bold>1" lemma bi_unique_rel_12_21 [simp, transfer_rule]: "bi_unique rel_12_21" by(auto simp add: bi_unique_def rel_12_21.simps) end lemma bernoulli_pmf_0: "bernoulli_pmf 0 = return_pmf False" by(rule pmf_eqI)(simp split: split_indicator) lemma bernoulli_pmf_1: "bernoulli_pmf 1 = return_pmf True" by(rule pmf_eqI)(simp split: split_indicator) lemma bernoulli_Not: "map_pmf Not (bernoulli_pmf r) = bernoulli_pmf (1 - r)" apply(rule pmf_eqI) apply(rewrite in "pmf _ \ = _" not_not[symmetric]) apply(subst pmf_map_inj') apply(simp_all add: inj_on_def bernoulli_pmf.rep_eq min_def max_def) done lemma pmf_eqI_avoid: "p = q" if "\i. i \ x \ pmf p i = pmf q i" proof(rule pmf_eqI) show "pmf p i = pmf q i" for i proof(cases "i = x") case [simp]: True have "pmf p i = measure_pmf.prob p {i}" by(simp add: measure_pmf_single) also have "\ = 1 - measure_pmf.prob p (UNIV - {i})" by(subst measure_pmf.prob_compl[unfolded space_measure_pmf]) simp_all also have "measure_pmf.prob p (UNIV - {i}) = measure_pmf.prob q (UNIV - {i})" unfolding integral_pmf[symmetric] by(rule Bochner_Integration.integral_cong)(auto intro: that) also have "1 - \ = measure_pmf.prob q {i}" by(subst measure_pmf.prob_compl[unfolded space_measure_pmf]) simp_all also have "\ = pmf q i" by(simp add: measure_pmf_single) finally show ?thesis . next case False then show ?thesis by(rule that) qed qed section \Locales for monomorphic monads\ subsection \Plain monad\ type_synonym ('a, 'm) bind = "'m \ ('a \ 'm) \ 'm" type_synonym ('a, 'm) return = "'a \ 'm" locale monad_base = fixes return :: "('a, 'm) return" and bind :: "('a, 'm) bind" begin primrec sequence :: "'m list \ ('a list \ 'm) \ 'm" where "sequence [] f = f []" | "sequence (x # xs) f = bind x (\a. sequence xs (f \ (#) a))" definition lift :: "('a \ 'a) \ 'm \ 'm" where "lift f x = bind x (\x. return (f x))" end declare monad_base.sequence.simps [code] monad_base.lift_def [code] context includes lifting_syntax begin lemma sequence_parametric [transfer_rule]: "((M ===> (A ===> M) ===> M) ===> list_all2 M ===> (list_all2 A ===> M) ===> M) monad_base.sequence monad_base.sequence" unfolding monad_base.sequence_def[abs_def] by transfer_prover lemma lift_parametric [transfer_rule]: "((A ===> M) ===> (M ===> (A ===> M) ===> M) ===> (A ===> A) ===> M ===> M) monad_base.lift monad_base.lift" unfolding monad_base.lift_def by transfer_prover end locale monad = monad_base return bind for return :: "('a, 'm) return" and bind :: "('a, 'm) bind" + assumes bind_assoc: "\(x :: 'm) f g. bind (bind x f) g = bind x (\y. bind (f y) g)" and return_bind: "\x f. bind (return x) f = f x" and bind_return: "\x. bind x return = x" begin lemma bind_lift [simp]: "bind (lift f x) g = bind x (g \ f)" by(simp add: lift_def bind_assoc return_bind o_def) lemma lift_bind [simp]: "lift f (bind m g) = bind m (\x. lift f (g x))" by(simp add: lift_def bind_assoc) end subsection \State\ type_synonym ('s, 'm) get = "('s \ 'm) \ 'm" type_synonym ('s, 'm) put = "'s \ 'm \ 'm" locale monad_state_base = monad_base return bind for return :: "('a, 'm) return" and bind :: "('a, 'm) bind" + fixes get :: "('s, 'm) get" and put :: "('s, 'm) put" begin definition update :: "('s \ 's) \ 'm \ 'm" where "update f m = get (\s. put (f s) m)" end declare monad_state_base.update_def [code] lemma update_parametric [transfer_rule]: includes lifting_syntax shows "(((S ===> M) ===> M) ===> (S ===> M ===> M) ===> (S ===> S) ===> M ===> M) monad_state_base.update monad_state_base.update" unfolding monad_state_base.update_def by transfer_prover locale monad_state = monad_state_base return bind get put + monad return bind for return :: "('a, 'm) return" and bind :: "('a, 'm) bind" and get :: "('s, 'm) get" and put :: "('s, 'm) put" + assumes put_get: "\f. put s (get f) = put s (f s)" and get_get: "\f. get (\s. get (f s)) = get (\s. f s s)" and put_put: "put s (put s' m) = put s' m" and get_put: "get (\s. put s m) = m" and get_const: "\m. get (\_. m) = m" and bind_get: "\f g. bind (get f) g = get (\s. bind (f s) g)" and bind_put: "\f. bind (put s m) f = put s (bind m f)" begin lemma put_update: "put s (update f m) = put (f s) m" by(simp add: update_def put_get put_put) lemma update_put: "update f (put s m) = put s m" by(simp add: update_def put_put get_const) lemma bind_update: "bind (update f m) g = update f (bind m g)" by(simp add: update_def bind_get bind_put) lemma update_get: "update f (get g) = get (update f \ g \ f)" by(simp add: update_def put_get get_get o_def) lemma update_const: "update (\_. s) m = put s m" by(simp add: update_def get_const) lemma update_update: "update f (update g m) = update (g \ f) m" by(simp add: update_def put_get put_put) lemma update_id: "update id m = m" by(simp add: update_def get_put) end subsection \Failure\ type_synonym 'm fail = "'m" locale monad_fail_base = monad_base return bind for return :: "('a, 'm) return" and bind :: "('a, 'm) bind" + fixes fail :: "'m fail" begin definition assert :: "('a \ bool) \ 'm \ 'm" where "assert P m = bind m (\x. if P x then return x else fail)" end declare monad_fail_base.assert_def [code] lemma assert_parametric [transfer_rule]: includes lifting_syntax shows "((A ===> M) ===> (M ===> (A ===> M) ===> M) ===> M ===> (A ===> (=)) ===> M ===> M) monad_fail_base.assert monad_fail_base.assert" unfolding monad_fail_base.assert_def by transfer_prover locale monad_fail = monad_fail_base return bind fail + monad return bind for return :: "('a, 'm) return" and bind :: "('a, 'm) bind" and fail :: "'m fail" + assumes fail_bind: "\f. bind fail f = fail" begin lemma assert_fail: "assert P fail = fail" by(simp add: assert_def fail_bind) end subsection \Exception\ type_synonym 'm catch = "'m \ 'm \ 'm" locale monad_catch_base = monad_fail_base return bind fail for return :: "('a, 'm) return" and bind :: "('a, 'm) bind" and fail :: "'m fail" + fixes catch :: "'m catch" locale monad_catch = monad_catch_base return bind fail catch + monad_fail return bind fail for return :: "('a, 'm) return" and bind :: "('a, 'm) bind" and fail :: "'m fail" and catch :: "'m catch" + assumes catch_return: "catch (return x) m = return x" and catch_fail: "catch fail m = m" and catch_fail2: "catch m fail = m" and catch_assoc: "catch (catch m m') m'' = catch m (catch m' m'')" locale monad_catch_state = monad_catch return bind fail catch + monad_state return bind get put for return :: "('a, 'm) return" and bind :: "('a, 'm) bind" and fail :: "'m fail" and catch :: "'m catch" and get :: "('s, 'm) get" and put :: "('s, 'm) put" + assumes catch_get: "catch (get f) m = get (\s. catch (f s) m)" and catch_put: "catch (put s m) m' = put s (catch m m')" begin lemma catch_update: "catch (update f m) m' = update f (catch m m')" by(simp add: update_def catch_get catch_put) end subsection \Reader\ text \As ask takes a continuation, we have to restate the monad laws for ask\ type_synonym ('r, 'm) ask = "('r \ 'm) \ 'm" locale monad_reader_base = monad_base return bind for return :: "('a, 'm) return" and bind :: "('a, 'm) bind" + fixes ask :: "('r, 'm) ask" locale monad_reader = monad_reader_base return bind ask + monad return bind for return :: "('a, 'm) return" and bind :: "('a, 'm) bind" and ask :: "('r, 'm) ask" + assumes ask_ask: "\f. ask (\r. ask (f r)) = ask (\r. f r r)" and ask_const: "ask (\_. m) = m" and bind_ask: "\f g. bind (ask f) g = ask (\r. bind (f r) g)" and bind_ask2: "\f. bind m (\x. ask (f x)) = ask (\r. bind m (\x. f x r))" begin lemma ask_bind: "ask (\r. bind (f r) (g r)) = bind (ask f) (\x. ask (\r. g r x))" by(simp add: bind_ask bind_ask2 ask_ask) end locale monad_reader_state = monad_reader return bind ask + monad_state return bind get put for return :: "('a, 'm) return" and bind :: "('a, 'm) bind" and ask :: "('r, 'm) ask" and get :: "('s, 'm) get" and put :: "('s, 'm) put" + assumes ask_get: "\f. ask (\r. get (f r)) = get (\s. ask (\r. f r s))" and put_ask: "\f. put s (ask f) = ask (\r. put s (f r))" subsection \Probability\ type_synonym ('p, 'm) sample = "'p pmf \ ('p \ 'm) \ 'm" locale monad_prob_base = monad_base return bind for return :: "('a, 'm) return" and bind :: "('a, 'm) bind" + fixes sample :: "('p, 'm) sample" locale monad_prob = monad return bind + monad_prob_base return bind sample for return :: "('a, 'm) return" and bind :: "('a, 'm) bind" and sample :: "('p, 'm) sample" + assumes sample_const: "\p m. sample p (\_. m) = m" and sample_return_pmf: "\x f. sample (return_pmf x) f = f x" and sample_bind_pmf: "\p f g. sample (bind_pmf p f) g = sample p (\x. sample (f x) g)" and sample_commute: "\p q f. sample p (\x. sample q (f x)) = sample q (\y. sample p (\x. f x y))" \ \We'd like to state that we can combine independent samples rather than just commute them, but that's not possible with a monomorphic sampling operation\ and bind_sample1: "\p f g. bind (sample p f) g = sample p (\x. bind (f x) g)" and bind_sample2: "\m f p. bind m (\y. sample p (f y)) = sample p (\x. bind m (\y. f y x))" and sample_parametric: "\R. bi_unique R \ rel_fun (rel_pmf R) (rel_fun (rel_fun R (=)) (=)) sample sample" begin lemma sample_cong: "(\x. x \ set_pmf p \ f x = g x) \ sample p f = sample q g" if "p = q" by(rule sample_parametric[where R="eq_onp (\x. x \ set_pmf p)", THEN rel_funD, THEN rel_funD]) (simp_all add: bi_unique_def eq_onp_def rel_fun_def pmf.rel_refl_strong that) end text \We can implement binary probabilistic choice using @{term sample} provided that the sample space contains at least three elements.\ locale monad_prob3 = monad_prob return bind sample + three "TYPE('p)" for return :: "('a, 'm) return" and bind :: "('a, 'm) bind" and sample :: "('p, 'm) sample" begin definition pchoose :: "real \ 'm \ 'm \ 'm" where "pchoose r m m' = sample (map_pmf (\b. if b then \<^bold>1 else \<^bold>2) (bernoulli_pmf r)) (\x. if x = \<^bold>1 then m else m')" abbreviation pchoose_syntax :: "'m \ real \ 'm \ 'm" ("_ \ _ \ _" [100, 0, 100] 99) where "m \ r \ m' \ pchoose r m m'" lemma pchoose_0: "m \ 0 \ m' = m'" by(simp add: pchoose_def bernoulli_pmf_0 sample_return_pmf) lemma pchoose_1: "m \ 1 \ m' = m" by(simp add: pchoose_def bernoulli_pmf_1 sample_return_pmf) lemma pchoose_idemp: "m \ r \ m = m" by(simp add: pchoose_def sample_const) lemma pchoose_bind1: "bind (m \ r \ m') f = bind m f \ r \ bind m' f" by(simp add: pchoose_def bind_sample1 if_distrib[where f="\m. bind m _"]) lemma pchoose_bind2: "bind m (\x. f x \ p \ g x) = bind m f \ p \ bind m g" by(auto simp add: pchoose_def bind_sample2 intro!: arg_cong2[where f=sample]) lemma pchoose_commute: "m \ 1 - r \ m' = m' \ r \ m" apply(simp add: pchoose_def bernoulli_Not[symmetric] pmf.map_comp o_def) apply(rule sample_parametric[where R=rel_12_21, THEN rel_funD, THEN rel_funD]) subgoal by(simp) subgoal by(rule pmf.map_transfer[where Rb="(=)", THEN rel_funD, THEN rel_funD]) (simp_all add: rel_fun_def rel_12_21.simps pmf.rel_eq) subgoal by(simp add: rel_fun_def rel_12_21.simps) done lemma pchoose_assoc: "m \ p \ (m' \ q \ m'') = (m \ r \ m') \ s \ m''" (is "?lhs = ?rhs") if "min 1 (max 0 p) = min 1 (max 0 r) * min 1 (max 0 s)" and "1 - min 1 (max 0 s) = (1 - min 1 (max 0 p)) * (1 - min 1 (max 0 q))" proof - let ?f = "(\x. if x = \<^bold>1 then m else if x = \<^bold>2 then m' else m'')" let ?p = "bind_pmf (map_pmf (\b. if b then \<^bold>1 else \<^bold>2) (bernoulli_pmf p)) (\x. if x = \<^bold>1 then return_pmf \<^bold>1 else map_pmf (\b. if b then \<^bold>2 else \<^bold>3) (bernoulli_pmf q))" let ?q = "bind_pmf (map_pmf (\b. if b then \<^bold>1 else \<^bold>2) (bernoulli_pmf s)) (\x. if x = \<^bold>1 then map_pmf (\b. if b then \<^bold>1 else \<^bold>2) (bernoulli_pmf r) else return_pmf \<^bold>3)" have [simp]: "{x. \ x} = {False}" "{x. x} = {True}" by auto have "?lhs = sample ?p ?f" by(auto simp add: pchoose_def sample_bind_pmf if_distrib[where f="\x. sample x _"] sample_return_pmf rel_fun_def rel_12_23.simps pmf.rel_eq cong: if_cong intro!: sample_cong[OF refl] sample_parametric[where R="rel_12_23", THEN rel_funD, THEN rel_funD] pmf.map_transfer[where Rb="(=)", THEN rel_funD, THEN rel_funD]) also have "?p = ?q" proof(rule pmf_eqI_avoid) fix i :: "'p" assume "i \ \<^bold>2" then consider (one) "i = \<^bold>1" | (three) "i = \<^bold>3" | (other) "i \ \<^bold>1" "i \ \<^bold>2" "i \ \<^bold>3" by metis then show "pmf ?p i = pmf ?q i" proof cases case [simp]: one have "pmf ?p i = measure_pmf.expectation (map_pmf (\b. if b then \<^bold>1 else \<^bold>2) (bernoulli_pmf p)) (indicator {\<^bold>1})" unfolding pmf_bind by(rule arg_cong2[where f=measure_pmf.expectation, OF refl])(auto simp add: fun_eq_iff pmf_eq_0_set_pmf) also have "\ = min 1 (max 0 p)" by(simp add: vimage_def)(simp add: measure_pmf_single bernoulli_pmf.rep_eq) also have "\ = min 1 (max 0 s) * min 1 (max 0 r)" using that(1) by simp also have "\ = measure_pmf.expectation (bernoulli_pmf s) (\x. indicator {True} x * pmf (map_pmf (\b. if b then \<^bold>1 else \<^bold>2) (bernoulli_pmf r)) \<^bold>1)" by(simp add: pmf_map vimage_def measure_pmf_single)(simp add: bernoulli_pmf.rep_eq) also have "\ = pmf ?q i" unfolding pmf_bind integral_map_pmf by(rule arg_cong2[where f=measure_pmf.expectation, OF refl])(auto simp add: fun_eq_iff pmf_eq_0_set_pmf) finally show ?thesis . next case [simp]: three have "pmf ?p i = measure_pmf.expectation (bernoulli_pmf p) (\x. indicator {False} x * pmf (map_pmf (\b. if b then \<^bold>2 else \<^bold>3) (bernoulli_pmf q)) \<^bold>3)" unfolding pmf_bind integral_map_pmf by(rule arg_cong2[where f=measure_pmf.expectation, OF refl])(auto simp add: fun_eq_iff pmf_eq_0_set_pmf) also have "\ = (1 - min 1 (max 0 p)) * (1 - min 1 (max 0 q))" by(simp add: pmf_map vimage_def measure_pmf_single)(simp add: bernoulli_pmf.rep_eq) also have "\ = 1 - min 1 (max 0 s)" using that(2) by simp also have "\ = measure_pmf.expectation (map_pmf (\b. if b then \<^bold>1 else \<^bold>2) (bernoulli_pmf s)) (indicator {\<^bold>2})" by(simp add: vimage_def)(simp add: measure_pmf_single bernoulli_pmf.rep_eq) also have "\ = pmf ?q i" unfolding pmf_bind by(rule Bochner_Integration.integral_cong_AE)(auto simp add: fun_eq_iff pmf_eq_0_set_pmf AE_measure_pmf_iff) finally show ?thesis . next case other then have "pmf ?p i = 0" "pmf ?q i = 0" by(auto simp add: pmf_eq_0_set_pmf) then show ?thesis by simp qed qed also have "sample ?q ?f = ?rhs" by(auto simp add: pchoose_def sample_bind_pmf if_distrib[where f="\x. sample x _"] sample_return_pmf cong: if_cong intro!: sample_cong[OF refl]) finally show ?thesis . qed lemma pchoose_assoc': "m \ p \ (m' \ q \ m'') = (m \ r \ m') \ s \ m''" if "p = r * s" and "1 - s = (1 - p) * (1 - q)" and "0 \ p" "p \ 1" "0 \ q" "q \ 1" "0 \ r" "r \ 1" "0 \ s" "s \ 1" by(rule pchoose_assoc; use that in \simp add: min_def max_def\) end locale monad_state_prob = monad_state return bind get put + monad_prob return bind sample for return :: "('a, 'm) return" and bind :: "('a, 'm) bind" and get :: "('s, 'm) get" and put :: "('s, 'm) put" and sample :: "('p, 'm) sample" + assumes sample_get: "sample p (\x. get (f x)) = get (\s. sample p (\x. f x s))" begin lemma sample_put: "sample p (\x. put s (m x)) = put s (sample p m)" proof - fix UU have "sample p (\x. put s (m x)) = sample p (\x. bind (put s (return UU)) (\_. m x))" by(simp add: bind_put return_bind) also have "\ = bind (put s (return UU)) (\_. sample p m)" by(simp add: bind_sample2) also have "\ = put s (sample p m)" by(simp add: bind_put return_bind) finally show ?thesis . qed lemma sample_update: "sample p (\x. update f (m x)) = update f (sample p m)" by(simp add: update_def sample_get sample_put) end subsection \Nondeterministic choice\ subsubsection \Binary choice\ type_synonym 'm alt = "'m \ 'm \ 'm" locale monad_alt_base = monad_base return bind for return :: "('a, 'm) return" and bind :: "('a, 'm) bind" + fixes alt :: "'m alt" locale monad_alt = monad return bind + monad_alt_base return bind alt for return :: "('a, 'm) return" and bind :: "('a, 'm) bind" and alt :: "'m alt" + \ \Laws taken from Gibbons, Hinze: Just do it\ assumes alt_assoc: "alt (alt m1 m2) m3 = alt m1 (alt m2 m3)" and bind_alt1: "bind (alt m m') f = alt (bind m f) (bind m' f)" locale monad_fail_alt = monad_fail return bind fail + monad_alt return bind alt for return :: "('a, 'm) return" and bind :: "('a, 'm) bind" and fail :: "'m fail" and alt :: "'m alt" + assumes alt_fail1: "alt fail m = m" and alt_fail2: "alt m fail = m" begin lemma assert_alt: "assert P (alt m m') = alt (assert P m) (assert P m')" by(simp add: assert_def bind_alt1) end locale monad_state_alt = monad_state return bind get put + monad_alt return bind alt for return :: "('a, 'm) return" and bind :: "('a, 'm) bind" and get :: "('s, 'm) get" and put :: "('s, 'm) put" and alt :: "'m alt" + assumes alt_get: "alt (get f) (get g) = get (\x. alt (f x) (g x))" and alt_put: "alt (put s m) (put s m') = put s (alt m m')" \ \Unlike for @{term sample}, we must require both @{text alt_get} and @{text alt_put} because we do not require that @{term bind} right-distributes over @{term alt}.\ begin lemma alt_update: "alt (update f m) (update f m') = update f (alt m m')" by(simp add: update_def alt_get alt_put) end subsubsection \Countable choice\ type_synonym ('c, 'm) altc = "'c cset \ ('c \ 'm) \ 'm" locale monad_altc_base = monad_base return bind for return :: "('a, 'm) return" and bind :: "('a, 'm) bind" + fixes altc :: "('c, 'm) altc" begin definition fail :: "'m fail" where "fail = altc cempty (\_. undefined)" end declare monad_altc_base.fail_def [code] locale monad_altc = monad return bind + monad_altc_base return bind altc for return :: "('a, 'm) return" and bind :: "('a, 'm) bind" and altc :: "('c, 'm) altc" + assumes bind_altc1: "\C g f. bind (altc C g) f = altc C (\c. bind (g c) f)" and altc_single: "\x f. altc (csingle x) f = f x" and altc_cUNION: "\C f g. altc (cUNION C f) g = altc C (\x. altc (f x) g)" \ \We do not assume @{text altc_const} like for @{text sample} because the choice set might be empty\ and altc_parametric: "\R. bi_unique R \ rel_fun (rel_cset R) (rel_fun (rel_fun R (=)) (=)) altc altc" begin lemma altc_cong: "cBall C (\x. f x = g x) \ altc C f = altc C g" apply(rule altc_parametric[where R="eq_onp (\x. cin x C)", THEN rel_funD, THEN rel_funD]) subgoal by(simp add: bi_unique_def eq_onp_def) subgoal by(simp add: cset.rel_eq_onp eq_onp_same_args pred_cset_def cin_def) subgoal by(simp add: rel_fun_def eq_onp_def cBall_def cin_def) done lemma monad_fail [locale_witness]: "monad_fail return bind fail" proof show "bind fail f = fail" for f by(simp add: fail_def bind_altc1 cong: altc_cong) qed end text \We can implement \alt\ via \altc\ only if we know that there are sufficiently many elements in the choice type @{typ 'c}. For the associativity law, we need at least three elements.\ locale monad_altc3 = monad_altc return bind altc + three "TYPE('c)" for return :: "('a, 'm) return" and bind :: "('a, 'm) bind" and altc :: "('c, 'm) altc" begin definition alt :: "'m alt" where "alt m1 m2 = altc (cinsert \<^bold>1 (csingle \<^bold>2)) (\c. if c = \<^bold>1 then m1 else m2)" lemma monad_alt: "monad_alt return bind alt" proof show "bind (alt m m') f = alt (bind m f) (bind m' f)" for m m' f by(simp add: alt_def bind_altc1 if_distrib[where f="\m. bind m _"]) fix m1 m2 m3 :: 'm let ?C = "cUNION (cinsert \<^bold>1 (csingle \<^bold>2)) (\c. if c = \<^bold>1 then cinsert \<^bold>1 (csingle \<^bold>2) else csingle \<^bold>3)" let ?D = "cUNION (cinsert \<^bold>1 (csingle \<^bold>2)) (\c. if c = \<^bold>1 then csingle \<^bold>1 else cinsert \<^bold>2 (csingle \<^bold>3))" let ?f = "\c. if c = \<^bold>1 then m1 else if c = \<^bold>2 then m2 else m3" have "alt (alt m1 m2) m3 = altc ?C ?f" by (simp only: altc_cUNION) (auto simp add: alt_def altc_single intro!: altc_cong) also have "?C = ?D" including cset.lifting by transfer(auto simp add: insert_commute) also have "altc ?D ?f = alt m1 (alt m2 m3)" apply (simp only: altc_cUNION) apply (clarsimp simp add: alt_def altc_single intro!: altc_cong) apply (rule altc_parametric [where R="conversep rel_12_23", THEN rel_funD, THEN rel_funD]) subgoal by simp subgoal including cset.lifting by transfer (simp add: rel_set_def rel_12_23.simps) subgoal by (simp add: rel_fun_def rel_12_23.simps) done finally show "alt (alt m1 m2) m3 = alt m1 (alt m2 m3)" . qed end locale monad_state_altc = monad_state return bind get put + monad_altc return bind altc for return :: "('a, 'm) return" and bind :: "('a, 'm) bind" and get :: "('s, 'm) get" and put :: "('s, 'm) put" and altc :: "('c, 'm) altc" + assumes altc_get: "\C f. altc C (\c. get (f c)) = get (\s. altc C (\c. f c s))" and altc_put: "\C f. altc C (\c. put s (f c)) = put s (altc C f)" subsection \Writer monad\ type_synonym ('w, 'm) tell = "'w \ 'm \ 'm" locale monad_writer_base = monad_base return bind for return :: "('a, 'm) return" and bind :: "('a, 'm) bind" + fixes tell :: "('w, 'm) tell" locale monad_writer = monad_writer_base return bind tell + monad return bind for return :: "('a, 'm) return" and bind :: "('a, 'm) bind" and tell :: "('w, 'm) tell" + assumes bind_tell: "\w m f. bind (tell w m) f = tell w (bind m f)" subsection \Resumption monad\ type_synonym ('o, 'i, 'm) pause = "'o \ ('i \ 'm) \ 'm" locale monad_resumption_base = monad_base return bind for return :: "('a, 'm) return" and bind :: "('a, 'm) bind" + fixes pause :: "('o, 'i, 'm) pause" locale monad_resumption = monad_resumption_base return bind pause + monad return bind for return :: "('a, 'm) return" and bind :: "('a, 'm) bind" and pause :: "('o, 'i, 'm) pause" + assumes bind_pause: "bind (pause out c) f = pause out (\i. bind (c i) f)" subsection \Commutative monad\ locale monad_commute = monad return bind for return :: "('a, 'm) return" and bind :: "('a, 'm) bind" + assumes bind_commute: "bind m (\x. bind m' (f x)) = bind m' (\y. bind m (\x. f x y))" subsection \Discardable monad\ locale monad_discard = monad return bind for return :: "('a, 'm) return" and bind :: "('a, 'm) bind" + assumes bind_const: "bind m (\_. m') = m'" subsection \Duplicable monad\ locale monad_duplicate = monad return bind for return :: "('a, 'm) return" and bind :: "('a, 'm) bind" + assumes bind_duplicate: "bind m (\x. bind m (f x)) = bind m (\x. f x x)" section \Monad implementations\ subsection \Identity monad\ text \We need a type constructor such that we can overload the monad operations\ datatype 'a id = return_id ("extract": 'a) lemmas return_id_parametric = id.ctr_transfer lemma rel_id_unfold: "rel_id A (return_id x) m' \ (\x'. m' = return_id x' \ A x x')" "rel_id A m (return_id x') \ (\x. m = return_id x \ A x x')" subgoal by(cases m'; simp) subgoal by(cases m; simp) done lemma rel_id_expand: "M (extract m) (extract m') \ rel_id M m m'" by(cases m; cases m'; simp) subsubsection \Plain monad\ primrec bind_id :: "('a, 'a id) bind" where "bind_id (return_id x) f = f x" lemma extract_bind [simp]: "extract (bind_id x f) = extract (f (extract x))" by(cases x) simp lemma bind_id_parametric [transfer_rule]: includes lifting_syntax shows "(rel_id A ===> (A ===> rel_id A) ===> rel_id A) bind_id bind_id" unfolding bind_id_def by transfer_prover lemma monad_id [locale_witness]: "monad return_id bind_id" proof show "bind_id (bind_id x f) g = bind_id x (\x. bind_id (f x) g)" for x :: "'a id" and f :: "'a \ 'a id" and g :: "'a \ 'a id" by(rule id.expand) simp show "bind_id (return_id x) f = f x" for f :: "'a \ 'a id" and x by(rule id.expand) simp show "bind_id x return_id = x" for x :: "'a id" by(rule id.expand) simp qed lemma monad_commute_id [locale_witness]: "monad_commute return_id bind_id" proof show "bind_id m (\x. bind_id m' (f x)) = bind_id m' (\y. bind_id m (\x. f x y))" for m m' :: "'a id" and f by(rule id.expand) simp qed lemma monad_discard_id [locale_witness]: "monad_discard return_id bind_id" proof show "bind_id m (\_. m') = m'" for m m' :: "'a id" by(rule id.expand) simp qed lemma monad_duplicate_id [locale_witness]: "monad_duplicate return_id bind_id" proof show "bind_id m (\x. bind_id m (f x)) = bind_id m (\x. f x x)" for m :: "'a id" and f by(rule id.expand) simp qed subsection \Probability monad\ text \We don't know of a sensible probability monad transformer, so we define the plain probability monad.\ type_synonym 'a prob = "'a pmf" lemma monad_prob [locale_witness]: "monad return_pmf bind_pmf" by unfold_locales(simp_all add: bind_assoc_pmf bind_return_pmf bind_return_pmf') lemma monad_prob_prob [locale_witness]: "monad_prob return_pmf bind_pmf bind_pmf" including lifting_syntax proof show "bind_pmf p (\_. m) = m" for p :: "'b pmf" and m :: "'a prob" by(rule bind_pmf_const) show "bind_pmf (return_pmf x) f = f x" for f :: "'b \ 'a prob" and x by(rule bind_return_pmf) show "bind_pmf (bind_pmf p f) g = bind_pmf p (\x. bind_pmf (f x) g)" for p :: "'b pmf" and f :: "'b \ 'b pmf" and g :: "'b \ 'a prob" by(rule bind_assoc_pmf) show "bind_pmf p (\x. bind_pmf q (f x)) = bind_pmf q (\y. bind_pmf p (\x. f x y))" for p q :: "'b pmf" and f :: "'b \ 'b \ 'a prob" by(rule bind_commute_pmf) show "bind_pmf (bind_pmf p f) g = bind_pmf p (\x. bind_pmf (f x) g)" for p :: "'b pmf" and f :: "'b \ 'a prob" and g :: "'a \ 'a prob" by(simp add: bind_assoc_pmf) show "bind_pmf m (\y. bind_pmf p (f y)) = bind_pmf p (\x. bind_pmf m (\y. f y x))" for m :: "'a prob" and p :: "'b pmf" and f :: "'a \ 'b \ 'a prob" by(rule bind_commute_pmf) show "(rel_pmf R ===> (R ===> (=)) ===> (=)) bind_pmf bind_pmf" for R :: "'b \ 'b \ bool" by transfer_prover qed lemma monad_commute_prob [locale_witness]: "monad_commute return_pmf bind_pmf" proof show "bind_pmf m (\x. bind_pmf m' (f x)) = bind_pmf m' (\y. bind_pmf m (\x. f x y))" for m m' :: "'a prob" and f :: "'a \ 'a \ 'a prob" by(rule bind_commute_pmf) qed lemma monad_discard_prob [locale_witness]: "monad_discard return_pmf bind_pmf" proof show "bind_pmf m (\_. m') = m'" for m m' :: "'a pmf" by(simp) qed subsection \Resumption\ text \ We cannot define a resumption monad transformer because the codatatype recursion would have to go through a type variable. If we plug in something like unbounded non-determinism, then the HOL type does not exist. \ codatatype ('o, 'i, 'a) resumption = is_Done: Done (result: 'a) | Pause ("output": 'o) (resume: "'i \ ('o, 'i, 'a) resumption") subsubsection \Plain monad\ definition return_resumption :: "'a \ ('o, 'i, 'a) resumption" where "return_resumption = Done" primcorec bind_resumption :: "('o, 'i, 'a) resumption \ ('a \ ('o, 'i, 'a) resumption) \ ('o, 'i, 'a) resumption" where "bind_resumption m f = (if is_Done m then f (result m) else Pause (output m) (\i. bind_resumption (resume m i) f))" definition pause_resumption :: "'o \ ('i \ ('o, 'i, 'a) resumption) \ ('o, 'i, 'a) resumption" where "pause_resumption = Pause" lemma is_Done_return_resumption [simp]: "is_Done (return_resumption x)" by(simp add: return_resumption_def) lemma result_return_resumption [simp]: "result (return_resumption x) = x" by(simp add: return_resumption_def) lemma monad_resumption [locale_witness]: "monad return_resumption bind_resumption" proof show "bind_resumption (bind_resumption x f) g = bind_resumption x (\y. bind_resumption (f y) g)" for x :: "('o, 'i, 'a) resumption" and f g by(coinduction arbitrary: x f g rule: resumption.coinduct_strong) auto show "bind_resumption (return_resumption x) f = f x" for x and f :: "'a \ ('o, 'i, 'a) resumption" by(rule resumption.expand)(simp_all add: return_resumption_def) show "bind_resumption x return_resumption = x" for x :: "('o, 'i, 'a) resumption" by(coinduction arbitrary: x rule: resumption.coinduct_strong) auto qed lemma monad_resumption_resumption [locale_witness]: "monad_resumption return_resumption bind_resumption pause_resumption" proof show "bind_resumption (pause_resumption out c) f = pause_resumption out (\i. bind_resumption (c i) f)" for out c and f :: "'a \ ('o, 'i, 'a) resumption" by(rule resumption.expand)(simp_all add: pause_resumption_def) qed subsection \Failure and exception monad transformer\ text \ The phantom type variable @{typ 'a} is needed to avoid hidden polymorphism when overloading the monad operations for the failure monad transformer. \ datatype (plugins del: transfer) (phantom_optionT: 'a, set_optionT: 'm) optionT = OptionT (run_option: 'm) for rel: rel_optionT' map: map_optionT' text \ We define our own relator and mapper such that the phantom variable does not need any relation. \ lemma phantom_optionT [simp]: "phantom_optionT x = {}" by(cases x) simp context includes lifting_syntax begin lemma rel_optionT'_phantom: "rel_optionT' A = rel_optionT' top" by(auto 4 4 intro: optionT.rel_mono antisym optionT.rel_mono_strong) lemma map_optionT'_phantom: "map_optionT' f = map_optionT' undefined" by(auto 4 4 intro: optionT.map_cong) definition map_optionT :: "('m \ 'm') \ ('a, 'm) optionT \ ('b, 'm') optionT" where "map_optionT = map_optionT' undefined" definition rel_optionT :: "('m \ 'm' \ bool) \ ('a, 'm) optionT \ ('b, 'm') optionT \ bool" where "rel_optionT = rel_optionT' top" lemma rel_optionTE: assumes "rel_optionT M m m'" obtains x y where "m = OptionT x" "m' = OptionT y" "M x y" using assms by(cases m; cases m'; simp add: rel_optionT_def) lemma rel_optionT_simps [simp]: "rel_optionT M (OptionT m) (OptionT m') \ M m m'" by(simp add: rel_optionT_def) lemma rel_optionT_eq [relator_eq]: "rel_optionT (=) = (=)" by(auto simp add: fun_eq_iff rel_optionT_def intro: optionT.rel_refl_strong elim: optionT.rel_cases) lemma rel_optionT_mono [relator_mono]: "rel_optionT A \ rel_optionT B" if "A \ B" by(simp add: rel_optionT_def optionT.rel_mono that) lemma rel_optionT_distr [relator_distr]: "rel_optionT A OO rel_optionT B = rel_optionT (A OO B)" by(simp add: rel_optionT_def optionT.rel_compp[symmetric]) lemma rel_optionT_Grp: "rel_optionT (BNF_Def.Grp A f) = BNF_Def.Grp {x. set_optionT x \ A} (map_optionT f)" by(simp add: rel_optionT_def rel_optionT'_phantom[of "BNF_Def.Grp UNIV undefined", symmetric] optionT.rel_Grp map_optionT_def) lemma OptionT_parametric [transfer_rule]: "(M ===> rel_optionT M) OptionT OptionT" by(simp add: rel_fun_def rel_optionT_def) lemma run_option_parametric [transfer_rule]: "(rel_optionT M ===> M) run_option run_option" by(auto simp add: rel_fun_def rel_optionT_def elim: optionT.rel_cases) lemma case_optionT_parametric [transfer_rule]: "((M ===> X) ===> rel_optionT M ===> X) case_optionT case_optionT" by(auto simp add: rel_fun_def rel_optionT_def split: optionT.split) lemma rec_optionT_parametric [transfer_rule]: "((M ===> X) ===> rel_optionT M ===> X) rec_optionT rec_optionT" by(auto simp add: rel_fun_def elim: rel_optionTE) end subsubsection \Plain monad, failure, and exceptions\ context fixes return :: "('a option, 'm) return" and bind :: "('a option, 'm) bind" begin definition return_option :: "('a, ('a, 'm) optionT) return" where "return_option x = OptionT (return (Some x))" primrec bind_option :: "('a, ('a, 'm) optionT) bind" where [code_unfold, monad_unfold]: "bind_option (OptionT x) f = OptionT (bind x (\x. case x of None \ return (None :: 'a option) | Some y \ run_option (f y)))" definition fail_option :: "('a, 'm) optionT fail" where [code_unfold, monad_unfold]: "fail_option = OptionT (return None)" definition catch_option :: "('a, 'm) optionT catch" where "catch_option m h = OptionT (bind (run_option m) (\x. if x = None then run_option h else return x))" lemma run_bind_option: "run_option (bind_option x f) = bind (run_option x) (\x. case x of None \ return None | Some y \ run_option (f y))" by(cases x) simp lemma run_return_option [simp]: "run_option (return_option x) = return (Some x)" by(simp add: return_option_def) lemma run_fail_option [simp]: "run_option fail_option = return None" by(simp add: fail_option_def) lemma run_catch_option [simp]: "run_option (catch_option m1 m2) = bind (run_option m1) (\x. if x = None then run_option m2 else return x)" by(simp add: catch_option_def) context assumes monad: "monad return bind" begin interpretation monad return bind by(rule monad) lemma monad_optionT [locale_witness]: "monad return_option bind_option" (is "monad ?return ?bind") proof show "?bind (?bind x f) g = ?bind x (\x. ?bind (f x) g)" for x f g by(rule optionT.expand)(auto simp add: bind_assoc run_bind_option return_bind intro!: arg_cong2[where f=bind] split: option.split) show "?bind (?return x) f = f x" for f x by(rule optionT.expand)(simp add: run_bind_option return_bind return_option_def) show "?bind x ?return = x" for x by(rule optionT.expand)(simp add: run_bind_option option.case_distrib[symmetric] case_option_id bind_return cong del: option.case_cong) qed lemma monad_fail_optionT [locale_witness]: "monad_fail return_option bind_option fail_option" proof show "bind_option fail_option f = fail_option" for f by(rule optionT.expand)(simp add: run_bind_option return_bind) qed lemma monad_catch_optionT [locale_witness]: "monad_catch return_option bind_option fail_option catch_option" proof show "catch_option (return_option x) m = return_option x" for x m by(rule optionT.expand)(simp add: return_bind) show "catch_option fail_option m = m" for m by(rule optionT.expand)(simp add: return_bind) show "catch_option m fail_option = m" for m by(rule optionT.expand)(simp add: bind_return if_distrib[where f="return", symmetric] cong del: if_weak_cong) show "catch_option (catch_option m m') m'' = catch_option m (catch_option m' m'')" for m m' m'' by(rule optionT.expand)(auto simp add: bind_assoc fun_eq_iff return_bind intro!: arg_cong2[where f=bind]) qed end subsubsection \Reader\ context fixes ask :: "('r, 'm) ask" begin definition ask_option :: "('r, ('a, 'm) optionT) ask" where [code_unfold, monad_unfold]: "ask_option f = OptionT (ask (\r. run_option (f r)))" lemma run_ask_option [simp]: "run_option (ask_option f) = ask (\r. run_option (f r))" by(simp add: ask_option_def) lemma monad_reader_optionT [locale_witness]: assumes "monad_reader return bind ask" shows "monad_reader return_option bind_option ask_option" proof - interpret monad_reader return bind ask by(fact assms) show ?thesis proof show "ask_option (\r. ask_option (f r)) = ask_option (\r. f r r)" for f by(rule optionT.expand)(simp add: ask_ask) show "ask_option (\_. x) = x" for x by(rule optionT.expand)(simp add: ask_const) show "bind_option (ask_option f) g = ask_option (\r. bind_option (f r) g)" for f g by(rule optionT.expand)(simp add: bind_ask run_bind_option) show "bind_option m (\x. ask_option (f x)) = ask_option (\r. bind_option m (\x. f x r))" for m f by(rule optionT.expand)(auto simp add: bind_ask2[symmetric] run_bind_option ask_const del: ext intro!: arg_cong2[where f=bind] ext split: option.split) qed qed end subsubsection \State\ context fixes get :: "('s, 'm) get" and put :: "('s, 'm) put" begin definition get_option :: "('s, ('a, 'm) optionT) get" where "get_option f = OptionT (get (\s. run_option (f s)))" primrec put_option :: "('s, ('a, 'm) optionT) put" where "put_option s (OptionT m) = OptionT (put s m)" lemma run_get_option [simp]: "run_option (get_option f) = get (\s. run_option (f s))" by(simp add: get_option_def) lemma run_put_option [simp]: "run_option (put_option s m) = put s (run_option m)" by(cases m)(simp) context assumes state: "monad_state return bind get put" begin interpretation monad_state return bind get put by(fact state) lemma monad_state_optionT [locale_witness]: "monad_state return_option bind_option get_option put_option" proof show "put_option s (get_option f) = put_option s (f s)" for s f by(rule optionT.expand)(simp add: put_get) show "get_option (\s. get_option (f s)) = get_option (\s. f s s)" for f by(rule optionT.expand)(simp add: get_get) show "put_option s (put_option s' m) = put_option s' m" for s s' m by(rule optionT.expand)(simp add: put_put) show "get_option (\s. put_option s m) = m" for m by(rule optionT.expand)(simp add: get_put) show "get_option (\_. m) = m" for m by(rule optionT.expand)(simp add: get_const) show "bind_option (get_option f) g = get_option (\s. bind_option (f s) g)" for f g by(rule optionT.expand)(simp add: bind_get run_bind_option) show "bind_option (put_option s m) f = put_option s (bind_option m f)" for s m f by(rule optionT.expand)(simp add: bind_put run_bind_option) qed lemma monad_catch_state_optionT [locale_witness]: "monad_catch_state return_option bind_option fail_option catch_option get_option put_option" proof show "catch_option (get_option f) m = get_option (\s. catch_option (f s) m)" for f m by(rule optionT.expand)(simp add: bind_get) show "catch_option (put_option s m) m' = put_option s (catch_option m m')" for s m m' by(rule optionT.expand)(simp add: bind_put) qed end subsubsection \Probability\ definition altc_sample_option :: "('x \ ('b \ 'm) \ 'm) \ 'x \ ('b \ ('a, 'm) optionT) \ ('a, 'm) optionT" where "altc_sample_option altc_sample p f = OptionT (altc_sample p (\x. run_option (f x)))" lemma run_altc_sample_option [simp]: "run_option (altc_sample_option altc_sample p f) = altc_sample p (\x. run_option (f x))" by(simp add: altc_sample_option_def) context fixes sample :: "('p, 'm) sample" begin abbreviation sample_option :: "('p, ('a, 'm) optionT) sample" where "sample_option \ altc_sample_option sample" lemma monad_prob_optionT [locale_witness]: assumes "monad_prob return bind sample" shows "monad_prob return_option bind_option sample_option" proof - interpret monad_prob return bind sample by(fact assms) note sample_parametric[transfer_rule] show ?thesis including lifting_syntax proof show "sample_option p (\_. x) = x" for p x by(rule optionT.expand)(simp add: sample_const) show "sample_option (return_pmf x) f = f x" for f x by(rule optionT.expand)(simp add: sample_return_pmf) show "sample_option (bind_pmf p f) g = sample_option p (\x. sample_option (f x) g)" for p f g by(rule optionT.expand)(simp add: sample_bind_pmf) show "sample_option p (\x. sample_option q (f x)) = sample_option q (\y. sample_option p (\x. f x y))" for p q f by(rule optionT.expand)(auto intro!: sample_commute) show "bind_option (sample_option p f) g = sample_option p (\x. bind_option (f x) g)" for p f g by(rule optionT.expand)(auto simp add: bind_sample1 run_bind_option) show "bind_option m (\y. sample_option p (f y)) = sample_option p (\x. bind_option m (\y. f y x))" for m p f by(rule optionT.expand)(auto simp add: bind_sample2[symmetric] run_bind_option sample_const del: ext intro!: arg_cong2[where f=bind] ext split: option.split) show "(rel_pmf R ===> (R ===> (=)) ===> (=)) sample_option sample_option" if [transfer_rule]: "bi_unique R" for R unfolding altc_sample_option_def by transfer_prover qed qed lemma monad_state_prob_optionT [locale_witness]: assumes "monad_state_prob return bind get put sample" shows "monad_state_prob return_option bind_option get_option put_option sample_option" proof - interpret monad_state_prob return bind get put sample by fact show ?thesis proof show "sample_option p (\x. get_option (f x)) = get_option (\s. sample_option p (\x. f x s))" for p f by(rule optionT.expand)(simp add: sample_get) qed qed end subsubsection \Writer\ context fixes tell :: "('w, 'm) tell" begin definition tell_option :: "('w, ('a, 'm) optionT) tell" where "tell_option w m = OptionT (tell w (run_option m))" lemma run_tell_option [simp]: "run_option (tell_option w m) = tell w (run_option m)" by(simp add: tell_option_def) lemma monad_writer_optionT [locale_witness]: assumes "monad_writer return bind tell" shows "monad_writer return_option bind_option tell_option" proof - interpret monad_writer return bind tell by fact show ?thesis proof show "bind_option (tell_option w m) f = tell_option w (bind_option m f)" for w m f by(rule optionT.expand)(simp add: run_bind_option bind_tell) qed qed end subsubsection \Binary Non-determinism\ context fixes alt :: "'m alt" begin definition alt_option :: "('a, 'm) optionT alt" where "alt_option m1 m2 = OptionT (alt (run_option m1) (run_option m2))" lemma run_alt_option [simp]: "run_option (alt_option m1 m2) = alt (run_option m1) (run_option m2)" by(simp add: alt_option_def) lemma monad_alt_optionT [locale_witness]: assumes "monad_alt return bind alt" shows "monad_alt return_option bind_option alt_option" proof - interpret monad_alt return bind alt by fact show ?thesis proof show "alt_option (alt_option m1 m2) m3 = alt_option m1 (alt_option m2 m3)" for m1 m2 m3 by(rule optionT.expand)(simp add: alt_assoc) show "bind_option (alt_option m m') f = alt_option (bind_option m f) (bind_option m' f)" for m m' f by(rule optionT.expand)(simp add: bind_alt1 run_bind_option) qed qed text \ The @{term fail} of @{typ "(_, _) optionT"} does not combine with @{term "alt"} of the inner monad because @{typ "(_, _) optionT"} injects failures with @{term "return None"} into the inner monad. \ lemma monad_state_alt_optionT [locale_witness]: assumes "monad_state_alt return bind get put alt" shows "monad_state_alt return_option bind_option get_option put_option alt_option" proof - interpret monad_state_alt return bind get put alt by fact show ?thesis proof show "alt_option (get_option f) (get_option g) = get_option (\x. alt_option (f x) (g x))" for f g by(rule optionT.expand)(simp add: alt_get) show "alt_option (put_option s m) (put_option s m') = put_option s (alt_option m m')" for s m m' by(rule optionT.expand)(simp add: alt_put) qed qed end subsubsection \Countable Non-determinism\ context fixes altc :: "('c, 'm) altc" begin abbreviation altc_option :: "('c, ('a, 'm) optionT) altc" where "altc_option \ altc_sample_option altc" lemma monad_altc_optionT [locale_witness]: assumes "monad_altc return bind altc" shows "monad_altc return_option bind_option altc_option" proof - interpret monad_altc return bind altc by fact note altc_parametric[transfer_rule] show ?thesis including lifting_syntax proof show "bind_option (altc_option C g) f = altc_option C (\c. bind_option (g c) f)" for C g f by(rule optionT.expand)(simp add: run_bind_option bind_altc1 o_def) show "altc_option (csingle x) f = f x" for x f by(rule optionT.expand)(simp add: bind_altc1 altc_single) show "altc_option (cUNION C f) g = altc_option C (\x. altc_option (f x) g)" for C f g by(rule optionT.expand)(simp add: bind_altc1 altc_cUNION o_def) show "(rel_cset R ===> (R ===> (=)) ===> (=)) altc_option altc_option" if [transfer_rule]: "bi_unique R" for R unfolding altc_sample_option_def by transfer_prover qed qed lemma monad_altc3_optionT [locale_witness]: assumes "monad_altc3 return bind altc" shows "monad_altc3 return_option bind_option altc_option" proof - interpret monad_altc3 return bind altc by fact show ?thesis .. qed lemma monad_state_altc_optionT [locale_witness]: assumes "monad_state_altc return bind get put altc" shows "monad_state_altc return_option bind_option get_option put_option altc_option" proof - interpret monad_state_altc return bind get put altc by fact show ?thesis proof show "altc_option C (\c. get_option (f c)) = get_option (\s. altc_option C (\c. f c s))" for C f by(rule optionT.expand)(simp add: o_def altc_get) show "altc_option C (\c. put_option s (f c)) = put_option s (altc_option C f)" for C s f by(rule optionT.expand)(simp add: o_def altc_put) qed qed end end subsubsection \Resumption\ context fixes pause :: "('o, 'i, 'm) pause" begin definition pause_option :: "('o, 'i, ('a, 'm) optionT) pause" where "pause_option out c = OptionT (pause out (\i. run_option (c i)))" lemma run_pause_option [simp]: "run_option (pause_option out c) = pause out (\i. run_option (c i))" by(simp add: pause_option_def) lemma monad_resumption_optionT [locale_witness]: assumes "monad_resumption return bind pause" shows "monad_resumption return_option bind_option pause_option" proof - interpret monad_resumption return bind pause by fact show ?thesis proof show "bind_option (pause_option out c) f = pause_option out (\i. bind_option (c i) f)" for out c f by(rule optionT.expand)(simp add: bind_pause run_bind_option) qed qed end subsubsection \Commutativity\ lemma monad_commute_optionT [locale_witness]: assumes "monad_commute return bind" and "monad_discard return bind" shows "monad_commute return_option bind_option" proof - interpret monad_commute return bind by fact interpret monad_discard return bind by fact show ?thesis proof fix m m' f have "run_option (bind_option m (\x. bind_option m' (f x))) = bind (run_option m) (\x. bind (run_option m') (\y. case (x, y) of (Some x', Some y') \ run_option (f x' y') | _ \ return None))" by(auto simp add: run_bind_option bind_const cong del: option.case_cong del: ext intro!: arg_cong2[where f=bind] ext split: option.split) also have "\ = bind (run_option m') (\y. bind (run_option m) (\x. case (x, y) of (Some x', Some y') \ run_option (f x' y') | _ \ return None))" by(rule bind_commute) also have "\ = run_option (bind_option m' (\y. bind_option m (\x. f x y)))" by(auto simp add: run_bind_option bind_const case_option_collapse cong del: option.case_cong del: ext intro!: arg_cong2[where f=bind] ext split: option.split) finally show "bind_option m (\x. bind_option m' (f x)) = bind_option m' (\y. bind_option m (\x. f x y))" by(rule optionT.expand) qed qed subsubsection \Duplicability\ lemma monad_duplicate_optionT [locale_witness]: assumes "monad_duplicate return bind" and "monad_discard return bind" shows "monad_duplicate return_option bind_option" proof - interpret monad_duplicate return bind by fact interpret monad_discard return bind by fact show ?thesis proof fix m f have "run_option (bind_option m (\x. bind_option m (f x))) = bind (run_option m) (\x. bind (run_option m) (\y. case x of None \ return None | Some x' \ (case y of None \ return None | Some y' \ run_option (f x' y'))))" by(auto intro!: arg_cong2[where f=bind] simp add: fun_eq_iff bind_const run_bind_option split: option.split) also have "\ = run_option (bind_option m (\x. f x x))" by(simp add: bind_duplicate run_bind_option cong: option.case_cong) finally show "bind_option m (\x. bind_option m (f x)) = bind_option m (\x. f x x)" by(rule optionT.expand) qed qed end subsubsection \Parametricity\ context includes lifting_syntax begin lemma return_option_parametric [transfer_rule]: "((rel_option A ===> M) ===> A ===> rel_optionT M) return_option return_option" unfolding return_option_def by transfer_prover lemma bind_option_parametric [transfer_rule]: "((rel_option A ===> M) ===> (M ===> (rel_option A ===> M) ===> M) ===> rel_optionT M ===> (A ===> rel_optionT M) ===> rel_optionT M) bind_option bind_option" unfolding bind_option_def by transfer_prover lemma fail_option_parametric [transfer_rule]: "((rel_option A ===> M) ===> rel_optionT M) fail_option fail_option" unfolding fail_option_def by transfer_prover lemma catch_option_parametric [transfer_rule]: "((rel_option A ===> M) ===> (M ===> (rel_option A ===> M) ===> M) ===> rel_optionT M ===> rel_optionT M ===> rel_optionT M) catch_option catch_option" unfolding catch_option_def Option.is_none_def[symmetric] by transfer_prover lemma ask_option_parametric [transfer_rule]: "(((R ===> M) ===> M) ===> (R ===> rel_optionT M) ===> rel_optionT M) ask_option ask_option" unfolding ask_option_def by transfer_prover lemma get_option_parametric [transfer_rule]: "(((S ===> M) ===> M) ===> (S ===> rel_optionT M) ===> rel_optionT M) get_option get_option" unfolding get_option_def by transfer_prover lemma put_option_parametric [transfer_rule]: "((S ===> M ===> M) ===> S ===> rel_optionT M ===> rel_optionT M) put_option put_option" unfolding put_option_def by transfer_prover lemma altc_sample_option_parametric [transfer_rule]: "((A ===> (P ===> M) ===> M) ===> A ===> (P ===> rel_optionT M) ===> rel_optionT M) altc_sample_option altc_sample_option" unfolding altc_sample_option_def by transfer_prover lemma alt_option_parametric [transfer_rule]: "((M ===> M ===> M) ===> rel_optionT M ===> rel_optionT M ===> rel_optionT M) alt_option alt_option" unfolding alt_option_def by transfer_prover lemma tell_option_parametric [transfer_rule]: "((W ===> M ===> M) ===> W ===> rel_optionT M ===> rel_optionT M) tell_option tell_option" unfolding tell_option_def by transfer_prover lemma pause_option_parametric [transfer_rule]: "((Out ===> (In ===> M) ===> M) ===> Out ===> (In ===> rel_optionT M) ===> rel_optionT M) pause_option pause_option" unfolding pause_option_def by transfer_prover end subsection \Reader monad transformer\ datatype ('r, 'm) envT = EnvT (run_env: "'r \ 'm") context includes lifting_syntax begin definition rel_envT :: "('r \ 'r' \ bool) \ ('m \ 'm' \ bool) \ ('r, 'm) envT \ ('r', 'm') envT \ bool" where "rel_envT R M = BNF_Def.vimage2p run_env run_env (R ===> M)" lemma rel_envTI [intro!]: "(R ===> M) f g \ rel_envT R M (EnvT f) (EnvT g)" by(simp add: rel_envT_def BNF_Def.vimage2p_def) lemma rel_envT_simps: "rel_envT R M (EnvT f) (EnvT g) \ (R ===> M) f g" by(simp add: rel_envT_def BNF_Def.vimage2p_def) lemma rel_envTE [cases pred]: assumes "rel_envT R M m m'" obtains f g where "m = EnvT f" "m' = EnvT g" "(R ===> M) f g" using assms by(cases m; cases m'; auto simp add: rel_envT_simps) lemma rel_envT_eq [relator_eq]: "rel_envT (=) (=) = (=)" by(auto simp add: rel_envT_def rel_fun_eq BNF_Def.vimage2p_def fun_eq_iff intro: envT.expand) lemma rel_envT_mono [relator_mono]: "\ R \ R'; M \ M' \ \ rel_envT R' M \ rel_envT R M'" by(simp add: rel_envT_def predicate2I vimage2p_mono fun_mono) lemma EnvT_parametric [transfer_rule]: "((R ===> M) ===> rel_envT R M) EnvT EnvT" by(simp add: rel_funI rel_envT_simps) lemma run_env_parametric [transfer_rule]: "(rel_envT R M ===> R ===> M) run_env run_env" by(auto elim!: rel_envTE) lemma rec_envT_parametric [transfer_rule]: "(((R ===> M) ===> X) ===> rel_envT R M ===> X) rec_envT rec_envT" by(auto 4 4 elim!: rel_envTE dest: rel_funD) lemma case_envT_parametric [transfer_rule]: "(((R ===> M) ===> X) ===> rel_envT R M ===> X) case_envT case_envT" by(auto 4 4 elim!: rel_envTE dest: rel_funD) end subsubsection \Plain monad and ask\ context fixes return :: "('a, 'm) return" and bind :: "('a, 'm) bind" begin definition return_env :: "('a, ('r, 'm) envT) return" where "return_env x = EnvT (\_. return x)" primrec bind_env :: "('a, ('r, 'm) envT) bind" where "bind_env (EnvT x) f = EnvT (\r. bind (x r) (\y. run_env (f y) r))" definition ask_env :: "('r, ('r, 'm) envT) ask" where "ask_env f = EnvT (\r. run_env (f r) r)" lemma run_bind_env [simp]: "run_env (bind_env x f) r = bind (run_env x r) (\y. run_env (f y) r)" by(cases x) simp lemma run_return_env [simp]: "run_env (return_env x) r = return x" by(simp add: return_env_def) lemma run_ask_env [simp]: "run_env (ask_env f) r = run_env (f r) r" by(simp add: ask_env_def) context assumes monad: "monad return bind" begin interpretation monad return "bind :: ('a, 'm) bind" by(fact monad) lemma monad_envT [locale_witness]: "monad return_env bind_env" proof show "bind_env (bind_env x f) g = bind_env x (\x. bind_env (f x) g)" for x :: "('r, 'm) envT" and f :: "'a \ ('r, 'm) envT" and g :: "'a \ ('r, 'm) envT" by(rule envT.expand)(auto simp add: bind_assoc return_bind) show "bind_env (return_env x) f = f x" for f :: "'a \ ('r, 'm) envT" and x by(rule envT.expand)(simp add: return_bind return_env_def) show "bind_env x (return_env :: ('a, ('r, 'm) envT) return) = x" for x :: "('r, 'm) envT" by(rule envT.expand)(simp add: bind_return fun_eq_iff) qed lemma monad_reader_envT [locale_witness]: "monad_reader return_env bind_env ask_env" proof show "ask_env (\r. ask_env (f r)) = ask_env (\r. f r r)" for f :: "'r \ 'r \ ('r, 'm) envT" by(rule envT.expand)(auto simp add: fun_eq_iff) show "ask_env (\_. x) = x" for x :: "('r, 'm) envT" by(rule envT.expand)(auto simp add: fun_eq_iff) show "bind_env (ask_env f) g = ask_env (\r. bind_env (f r) g)" for f :: "'r \ ('r, 'm) envT" and g by(rule envT.expand)(auto simp add: fun_eq_iff) show "bind_env m (\x. ask_env (f x)) = ask_env (\r. bind_env m (\x. f x r))" for m :: "('r, 'm) envT" and f by(rule envT.expand)(auto simp add: fun_eq_iff) qed end subsubsection \Failure\ context fixes fail :: "'m fail" begin definition fail_env :: "('r, 'm) envT fail" where "fail_env = EnvT (\r. fail)" lemma run_fail_env [simp]: "run_env fail_env r = fail" by(simp add: fail_env_def) lemma monad_fail_envT [locale_witness]: assumes "monad_fail return bind fail" shows "monad_fail return_env bind_env fail_env" proof - interpret monad_fail return bind fail by(fact assms) have "bind_env fail_env f = fail_env" for f :: "'a \ ('r, 'm) envT" by(rule envT.expand)(simp add: fun_eq_iff fail_bind) then show ?thesis by unfold_locales qed context fixes catch :: "'m catch" begin definition catch_env :: "('r, 'm) envT catch" where "catch_env m1 m2 = EnvT (\r. catch (run_env m1 r) (run_env m2 r))" lemma run_catch_env [simp]: "run_env (catch_env m1 m2) r = catch (run_env m1 r) (run_env m2 r)" by(simp add: catch_env_def) lemma monad_catch_envT [locale_witness]: assumes "monad_catch return bind fail catch" shows "monad_catch return_env bind_env fail_env catch_env" proof - interpret monad_catch return bind fail catch by fact show ?thesis proof show "catch_env (return_env x) m = return_env x" for x and m :: "('r, 'm) envT" by(rule envT.expand)(simp add: fun_eq_iff catch_return) show "catch_env fail_env m = m" for m :: "('r, 'm) envT" by(rule envT.expand)(simp add: fun_eq_iff catch_fail) show "catch_env m fail_env = m" for m :: "('r, 'm) envT" by(rule envT.expand)(simp add: fun_eq_iff catch_fail2) show "catch_env (catch_env m m') m'' = catch_env m (catch_env m' m'')" for m m' m'' :: "('r, 'm) envT" by(rule envT.expand)(simp add: fun_eq_iff catch_assoc) qed qed end end subsubsection \State\ context fixes get :: "('s, 'm) get" and put :: "('s, 'm) put" begin definition get_env :: "('s, ('r, 'm) envT) get" where "get_env f = EnvT (\r. get (\s. run_env (f s) r))" definition put_env :: "('s, ('r, 'm) envT) put" where "put_env s m = EnvT (\r. put s (run_env m r))" lemma run_get_env [simp]: "run_env (get_env f) r = get (\s. run_env (f s) r)" by(simp add: get_env_def) lemma run_put_env [simp]: "run_env (put_env s m) r = put s (run_env m r)" by(simp add: put_env_def) lemma monad_state_envT [locale_witness]: assumes "monad_state return bind get put" shows "monad_state return_env bind_env get_env put_env" proof - interpret monad_state return bind get put by(fact assms) show ?thesis proof show "put_env s (get_env f) = put_env s (f s)" for s :: 's and f :: "'s \ ('r, 'm) envT" by(rule envT.expand)(simp add: fun_eq_iff put_get) show "get_env (\s. get_env (f s)) = get_env (\s. f s s)" for f :: "'s \ 's \ ('r, 'm) envT" by(rule envT.expand)(simp add: fun_eq_iff get_get) show "put_env s (put_env s' m) = put_env s' m" for s s' :: 's and m :: "('r, 'm) envT" by(rule envT.expand)(simp add: fun_eq_iff put_put) show "get_env (\s. put_env s m) = m" for m :: "('r, 'm) envT" by(rule envT.expand)(simp add: fun_eq_iff get_put) show "get_env (\_. m) = m" for m :: "('r, 'm) envT" by(rule envT.expand)(simp add: fun_eq_iff get_const) show "bind_env (get_env f) g = get_env (\s. bind_env (f s) g)" for f :: "'s \ ('r, 'm) envT" and g by(rule envT.expand)(simp add: fun_eq_iff bind_get) show "bind_env (put_env s m) f = put_env s (bind_env m f)" for s and m :: "('r, 'm) envT" and f by(rule envT.expand)(simp add: fun_eq_iff bind_put) qed qed subsubsection \Probability\ context fixes sample :: "('p, 'm) sample" begin definition sample_env :: "('p, ('r, 'm) envT) sample" where "sample_env p f = EnvT (\r. sample p (\x. run_env (f x) r))" lemma run_sample_env [simp]: "run_env (sample_env p f) r = sample p (\x. run_env (f x) r)" by(simp add: sample_env_def) lemma monad_prob_envT [locale_witness]: assumes "monad_prob return bind sample" shows "monad_prob return_env bind_env sample_env" proof - interpret monad_prob return bind sample by(fact assms) note sample_parametric[transfer_rule] show ?thesis including lifting_syntax proof show "sample_env p (\_. x) = x" for p :: "'p pmf" and x :: "('r, 'm) envT" by(rule envT.expand)(simp add: fun_eq_iff sample_const) show "sample_env (return_pmf x) f = f x" for f :: "'p \ ('r, 'm) envT" and x by(rule envT.expand)(simp add: fun_eq_iff sample_return_pmf) show "sample_env (bind_pmf p f) g = sample_env p (\x. sample_env (f x) g)" for f and g :: "'p \ ('r, 'm) envT" and p by(rule envT.expand)(simp add: fun_eq_iff sample_bind_pmf) show "sample_env p (\x. sample_env q (f x)) = sample_env q (\y. sample_env p (\x. f x y))" for p q :: "'p pmf" and f :: "'p \ 'p \ ('r, 'm) envT" by(rule envT.expand)(auto simp add: fun_eq_iff intro: sample_commute) show "bind_env (sample_env p f) g = sample_env p (\x. bind_env (f x) g)" for p and f :: "'p \ ('r, 'm) envT" and g by(rule envT.expand)(simp add: fun_eq_iff bind_sample1) show "bind_env m (\y. sample_env p (f y)) = sample_env p (\x. bind_env m (\y. f y x))" for m p and f :: "'a \ 'p \ ('r, 'm) envT" by(rule envT.expand)(simp add: fun_eq_iff bind_sample2) show "(rel_pmf R ===> (R ===> (=)) ===> (=)) sample_env sample_env" if [transfer_rule]: "bi_unique R" for R unfolding sample_env_def by transfer_prover qed qed lemma monad_state_prob_envT [locale_witness]: assumes "monad_state_prob return bind get put sample" shows "monad_state_prob return_env bind_env get_env put_env sample_env" proof - interpret monad_state_prob return bind get put sample by fact show ?thesis proof show "sample_env p (\x. get_env (f x)) = get_env (\s. sample_env p (\x. f x s))" for p and f :: "'p \ 's \ ('r, 'm) envT" by(rule envT.expand)(simp add: fun_eq_iff sample_get) qed qed end subsubsection \Binary Non-determinism\ context fixes alt :: "'m alt" begin definition alt_env :: "('r, 'm) envT alt" where "alt_env m1 m2 = EnvT (\r. alt (run_env m1 r) (run_env m2 r))" lemma run_alt_env [simp]: "run_env (alt_env m1 m2) r = alt (run_env m1 r) (run_env m2 r)" by(simp add: alt_env_def) lemma monad_alt_envT [locale_witness]: assumes "monad_alt return bind alt" shows "monad_alt return_env bind_env alt_env" proof - interpret monad_alt return bind alt by fact show ?thesis proof show "alt_env (alt_env m1 m2) m3 = alt_env m1 (alt_env m2 m3)" for m1 m2 m3 :: "('r, 'm) envT" by(rule envT.expand)(simp add: fun_eq_iff alt_assoc) show "bind_env (alt_env m m') f = alt_env (bind_env m f) (bind_env m' f)" for m m' :: "('r, 'm) envT" and f by(rule envT.expand)(simp add: fun_eq_iff bind_alt1) qed qed lemma monad_fail_alt_envT [locale_witness]: fixes fail assumes "monad_fail_alt return bind fail alt" shows "monad_fail_alt return_env bind_env (fail_env fail) alt_env" proof - interpret monad_fail_alt return bind fail alt by fact show ?thesis proof show "alt_env (fail_env fail) m = m" for m :: "('r, 'm) envT" by(rule envT.expand)(simp add: alt_fail1 fun_eq_iff) show "alt_env m (fail_env fail) = m" for m :: "('r, 'm) envT" by(rule envT.expand)(simp add: alt_fail2 fun_eq_iff) qed qed lemma monad_state_alt_envT [locale_witness]: assumes "monad_state_alt return bind get put alt" shows "monad_state_alt return_env bind_env get_env put_env alt_env" proof - interpret monad_state_alt return bind get put alt by fact show ?thesis proof show "alt_env (get_env f) (get_env g) = get_env (\x. alt_env (f x) (g x))" for f g :: "'s \ ('b, 'm) envT" by(rule envT.expand)(simp add: fun_eq_iff alt_get) show "alt_env (put_env s m) (put_env s m') = put_env s (alt_env m m')" for s and m m' :: "('b, 'm) envT" by(rule envT.expand)(simp add: fun_eq_iff alt_put) qed qed end subsubsection \Countable Non-determinism\ context fixes altc :: "('c, 'm) altc" begin definition altc_env :: "('c, ('r, 'm) envT) altc" where "altc_env C f = EnvT (\r. altc C (\c. run_env (f c) r))" lemma run_altc_env [simp]: "run_env (altc_env C f) r = altc C (\c. run_env (f c) r)" by(simp add: altc_env_def) lemma monad_altc_envT [locale_witness]: assumes "monad_altc return bind altc" shows "monad_altc return_env bind_env altc_env" proof - interpret monad_altc return bind altc by fact note altc_parametric[transfer_rule] show ?thesis including lifting_syntax proof show "bind_env (altc_env C g) f = altc_env C (\c. bind_env (g c) f)" for C g and f :: "'a \ ('b, 'm) envT" by(rule envT.expand)(simp add: fun_eq_iff bind_altc1) show "altc_env (csingle x) f = f x" for x and f :: "'c \ ('b, 'm) envT" by(rule envT.expand)(simp add: fun_eq_iff altc_single) show "altc_env (cUNION C f) g = altc_env C (\x. altc_env (f x) g)" for C f and g :: "'c \ ('b, 'm) envT" by(rule envT.expand)(simp add: fun_eq_iff altc_cUNION) show "(rel_cset R ===> (R ===> (=)) ===> (=)) altc_env altc_env" if [transfer_rule]: "bi_unique R" for R unfolding altc_env_def by transfer_prover qed qed lemma monad_altc3_envT [locale_witness]: assumes "monad_altc3 return bind altc" shows "monad_altc3 return_env bind_env altc_env" proof - interpret monad_altc3 return bind altc by fact show ?thesis .. qed lemma monad_state_altc_envT [locale_witness]: assumes "monad_state_altc return bind get put altc" shows "monad_state_altc return_env bind_env get_env put_env altc_env" proof - interpret monad_state_altc return bind get put altc by fact show ?thesis proof show "altc_env C (\c. get_env (f c)) = get_env (\s. altc_env C (\c. f c s))" for C and f :: "'c \ 's \ ('b, 'm) envT" by(rule envT.expand)(simp add: fun_eq_iff altc_get) show "altc_env C (\c. put_env s (f c)) = put_env s (altc_env C f)" for C s and f :: "'c \ ('b, 'm) envT" by(rule envT.expand)(simp add: fun_eq_iff altc_put) qed qed end end subsubsection \Resumption\ context fixes pause :: "('o, 'i, 'm) pause" begin definition pause_env :: "('o, 'i, ('r, 'm) envT) pause" where "pause_env out c = EnvT (\r. pause out (\i. run_env (c i) r))" lemma run_pause_env [simp]: "run_env (pause_env out c) r = pause out (\i. run_env (c i) r)" by(simp add: pause_env_def) lemma monad_resumption_envT [locale_witness]: assumes "monad_resumption return bind pause" shows "monad_resumption return_env bind_env pause_env" proof - interpret monad_resumption return bind pause by fact show ?thesis proof show "bind_env (pause_env out c) f = pause_env out (\i. bind_env (c i) f)" for out f and c :: "'i \ ('r, 'm) envT" by(rule envT.expand)(simp add: fun_eq_iff bind_pause) qed qed end subsubsection \Writer\ context fixes tell :: "('w, 'm) tell" begin definition tell_env :: "('w, ('r, 'm) envT) tell" where "tell_env w m = EnvT (\r. tell w (run_env m r))" lemma run_tell_env [simp]: "run_env (tell_env w m) r = tell w (run_env m r)" by(simp add: tell_env_def) lemma monad_writer_envT [locale_witness]: assumes "monad_writer return bind tell" shows "monad_writer return_env bind_env tell_env" proof - interpret monad_writer return bind tell by fact show ?thesis proof show "bind_env (tell_env w m) f = tell_env w (bind_env m f)" for w and m :: "('r, 'm) envT" and f by(rule envT.expand)(simp add: bind_tell fun_eq_iff) qed qed end subsubsection \Commutativity\ lemma monad_commute_envT [locale_witness]: assumes "monad_commute return bind" shows "monad_commute return_env bind_env" proof - interpret monad_commute return bind by fact show ?thesis proof show "bind_env m (\x. bind_env m' (f x)) = bind_env m' (\y. bind_env m (\x. f x y))" for f and m m' :: "('r, 'm) envT" by(rule envT.expand)(auto simp add: fun_eq_iff intro: bind_commute) qed qed subsubsection \Discardability\ lemma monad_discard_envT [locale_witness]: assumes "monad_discard return bind" shows "monad_discard return_env bind_env" proof - interpret monad_discard return bind by fact show ?thesis proof show "bind_env m (\_. m') = m'" for m m' :: "('r, 'm) envT" by(rule envT.expand)(simp add: fun_eq_iff bind_const) qed qed subsubsection \Duplicability\ lemma monad_duplicate_envT [locale_witness]: assumes "monad_duplicate return bind" shows "monad_duplicate return_env bind_env" proof - interpret monad_duplicate return bind by fact show ?thesis proof show "bind_env m (\x. bind_env m (f x)) = bind_env m (\x. f x x)" for m :: "('b, 'm) envT" and f by(rule envT.expand)(simp add: fun_eq_iff bind_duplicate) qed qed end subsubsection \Parametricity\ context includes lifting_syntax begin lemma return_env_parametric [transfer_rule]: "((A ===> M) ===> A ===> rel_envT R M) return_env return_env" unfolding return_env_def by transfer_prover lemma bind_env_parametric [transfer_rule]: "((M ===> (A ===> M) ===> M) ===> rel_envT R M ===> (A ===> rel_envT R M) ===> rel_envT R M) bind_env bind_env" unfolding bind_env_def by transfer_prover lemma ask_env_parametric [transfer_rule]: "((R ===> rel_envT R M) ===> rel_envT R M) ask_env ask_env" unfolding ask_env_def by transfer_prover lemma fail_env_parametric [transfer_rule]: "(M ===> rel_envT R M) fail_env fail_env" unfolding fail_env_def by transfer_prover lemma catch_env_parametric [transfer_rule]: "((M ===> M ===> M) ===> rel_envT R M ===> rel_envT R M ===> rel_envT R M) catch_env catch_env" unfolding catch_env_def by transfer_prover lemma get_env_parametric [transfer_rule]: "(((S ===> M) ===> M) ===> (S ===> rel_envT R M) ===> rel_envT R M) get_env get_env" unfolding get_env_def by transfer_prover lemma put_env_parametric [transfer_rule]: "((S ===> M ===> M) ===> S ===> rel_envT R M ===> rel_envT R M) put_env put_env" unfolding put_env_def by transfer_prover lemma sample_env_parametric [transfer_rule]: "((rel_pmf P ===> (P ===> M) ===> M) ===> rel_pmf P ===> (P ===> rel_envT R M) ===> rel_envT R M) sample_env sample_env" unfolding sample_env_def by transfer_prover lemma alt_env_parametric [transfer_rule]: "((M ===> M ===> M) ===> rel_envT R M ===> rel_envT R M ===> rel_envT R M) alt_env alt_env" unfolding alt_env_def by transfer_prover lemma altc_env_parametric [transfer_rule]: "((rel_cset C ===> (C ===> M) ===> M) ===> rel_cset C ===> (C ===> rel_envT R M) ===> rel_envT R M) altc_env altc_env" unfolding altc_env_def by transfer_prover lemma pause_env_parametric [transfer_rule]: "((Out ===> (In ===> M) ===> M) ===> Out ===> (In ===> rel_envT R M) ===> rel_envT R M) pause_env pause_env" unfolding pause_env_def by transfer_prover lemma tell_env_parametric [transfer_rule]: "((W ===> M ===> M) ===> W ===> rel_envT R M ===> rel_envT R M) tell_env tell_env" unfolding tell_env_def by transfer_prover end subsection \Unbounded non-determinism\ abbreviation (input) return_set :: "('a, 'a set) return" where "return_set x \ {x}" abbreviation (input) bind_set :: "('a, 'a set) bind" where "bind_set \ \A f. \ (f ` A)" abbreviation (input) fail_set :: "'a set fail" where "fail_set \ {}" abbreviation (input) alt_set :: "'a set alt" where "alt_set \ (\)" abbreviation (input) altc_set :: "('c, 'a set) altc" where "altc_set C \ \f. \ (f ` rcset C)" lemma monad_set [locale_witness]: "monad return_set bind_set" by unfold_locales auto lemma monad_fail_set [locale_witness]: "monad_fail return_set bind_set fail_set" by unfold_locales auto lemma monad_lift_set [simp]: "monad_base.lift return_set bind_set = image" by(auto simp add: monad_base.lift_def o_def fun_eq_iff) lemma monad_alt_set [locale_witness]: "monad_alt return_set bind_set alt_set" by unfold_locales auto lemma monad_altc_set [locale_witness]: "monad_altc return_set bind_set altc_set" including cset.lifting lifting_syntax proof show "(rel_cset R ===> (R ===> (=)) ===> (=)) (\C f. \ (f ` rcset C)) (\C f. \ (f ` rcset C))" for R by transfer_prover qed(transfer; auto; fail)+ lemma monad_altc3_set [locale_witness]: "monad_altc3 return_set bind_set (altc_set :: ('c, 'a set) altc)" if [locale_witness]: "three TYPE('c)" .. subsection \Non-determinism transformer\ datatype (plugins del: transfer) (phantom_nondetT: 'a, set_nondetT: 'm) nondetT = NondetT (run_nondet: 'm) for map: map_nondetT' rel: rel_nondetT' text \ We define our own relator and mapper such that the phantom variable does not need any relation. \ lemma phantom_nondetT [simp]: "phantom_nondetT x = {}" by(cases x) simp context includes lifting_syntax begin lemma rel_nondetT'_phantom: "rel_nondetT' A = rel_nondetT' top" by(auto 4 4 intro: nondetT.rel_mono antisym nondetT.rel_mono_strong) lemma map_nondetT'_phantom: "map_nondetT' f = map_nondetT' undefined" by(auto 4 4 intro: nondetT.map_cong) definition map_nondetT :: "('m \ 'm') \ ('a, 'm) nondetT \ ('b, 'm') nondetT" where "map_nondetT = map_nondetT' undefined" definition rel_nondetT :: "('m \ 'm' \ bool) \ ('a, 'm) nondetT \ ('b, 'm') nondetT \ bool" where "rel_nondetT = rel_nondetT' top" lemma rel_nondetTE: assumes "rel_nondetT M m m'" obtains x y where "m = NondetT x" "m' = NondetT y" "M x y" using assms by(cases m; cases m'; simp add: rel_nondetT_def) lemma rel_nondetT_simps [simp]: "rel_nondetT M (NondetT m) (NondetT m') \ M m m'" by(simp add: rel_nondetT_def) lemma rel_nondetT_unfold: "\m m'. rel_nondetT M (NondetT m) m' \ (\m''. m' = NondetT m'' \ M m m'')" "\m m'. rel_nondetT M m (NondetT m') \ (\m''. m = NondetT m'' \ M m'' m')" subgoal for m m' by(cases m'; simp) subgoal for m m' by(cases m; simp) done lemma rel_nondetT_expand: "M (run_nondet m) (run_nondet m') \ rel_nondetT M m m'" by(cases m; cases m'; simp) lemma rel_nondetT_eq [relator_eq]: "rel_nondetT (=) = (=)" by(auto simp add: fun_eq_iff rel_nondetT_def intro: nondetT.rel_refl_strong elim: nondetT.rel_cases) lemma rel_nondetT_mono [relator_mono]: "rel_nondetT A \ rel_nondetT B" if "A \ B" by(simp add: rel_nondetT_def nondetT.rel_mono that) lemma rel_nondetT_distr [relator_distr]: "rel_nondetT A OO rel_nondetT B = rel_nondetT (A OO B)" by(simp add: rel_nondetT_def nondetT.rel_compp[symmetric]) lemma rel_nondetT_Grp: "rel_nondetT (BNF_Def.Grp A f) = BNF_Def.Grp {x. set_nondetT x \ A} (map_nondetT f)" by(simp add: rel_nondetT_def rel_nondetT'_phantom[of "BNF_Def.Grp UNIV undefined", symmetric] nondetT.rel_Grp map_nondetT_def) lemma NondetT_parametric [transfer_rule]: "(M ===> rel_nondetT M) NondetT NondetT" by(simp add: rel_fun_def rel_nondetT_def) lemma run_nondet_parametric [transfer_rule]: "(rel_nondetT M ===> M) run_nondet run_nondet" by(auto simp add: rel_fun_def rel_nondetT_def elim: nondetT.rel_cases) lemma case_nondetT_parametric [transfer_rule]: "((M ===> X) ===> rel_nondetT M ===> X) case_nondetT case_nondetT" by(auto simp add: rel_fun_def rel_nondetT_def split: nondetT.split) lemma rec_nondetT_parametric [transfer_rule]: "((M ===> X) ===> rel_nondetT M ===> X) rec_nondetT rec_nondetT" by(auto simp add: rel_fun_def elim: rel_nondetTE) end subsubsection \Generic implementation\ type_synonym ('a, 'm, 's) merge = "'s \ ('a \ 'm) \ 'm" locale nondetM_base = monad_base return bind for return :: "('s, 'm) return" and bind :: "('s, 'm) bind" and merge :: "('a, 'm, 's) merge" and empty :: "'s" and single :: "'a \ 's" and union :: "'s \ 's \ 's" (infixl "\<^bold>\" 65) begin definition return_nondet :: "('a, ('a, 'm) nondetT) return" where "return_nondet x = NondetT (return (single x))" definition bind_nondet :: "('a, ('a, 'm) nondetT) bind" where "bind_nondet m f = NondetT (bind (run_nondet m) (\A. merge A (run_nondet \ f)))" definition fail_nondet :: "('a, 'm) nondetT fail" where "fail_nondet = NondetT (return empty)" definition alt_nondet :: "('a, 'm) nondetT alt" where "alt_nondet m1 m2 = NondetT (bind (run_nondet m1) (\A. bind (run_nondet m2) (\B. return (A \<^bold>\ B))))" definition get_nondet :: "('state, 'm) get \ ('state, ('a, 'm) nondetT) get" where "get_nondet get f = NondetT (get (\s. run_nondet (f s)))" for get definition put_nondet :: "('state, 'm) put \ ('state, ('a, 'm) nondetT) put" where "put_nondet put s m = NondetT (put s (run_nondet m))" for put definition ask_nondet :: "('r, 'm) ask \ ('r, ('a, 'm) nondetT) ask" where "ask_nondet ask f = NondetT (ask (\r. run_nondet (f r)))" text \ The canonical lift of sampling into @{typ "(_, _) nondetT"} does not satisfy @{const monad_prob}, because sampling does not distribute over bind backwards. Intuitively, if we sample first, then the same sample is used in all non-deterministic choices. But if we sample later, each non-deterministic choice may sample a different value. \ lemma run_return_nondet [simp]: "run_nondet (return_nondet x) = return (single x)" by(simp add: return_nondet_def) lemma run_bind_nondet [simp]: "run_nondet (bind_nondet m f) = bind (run_nondet m) (\A. merge A (run_nondet \ f))" by(simp add: bind_nondet_def) lemma run_fail_nondet [simp]: "run_nondet fail_nondet = return empty" by(simp add: fail_nondet_def) lemma run_alt_nondet [simp]: "run_nondet (alt_nondet m1 m2) = bind (run_nondet m1) (\A. bind (run_nondet m2) (\B. return (A \<^bold>\ B)))" by(simp add: alt_nondet_def) lemma run_get_nondet [simp]: "run_nondet (get_nondet get f) = get (\s. run_nondet (f s))" for get by(simp add: get_nondet_def) lemma run_put_nondet [simp]: "run_nondet (put_nondet put s m) = put s (run_nondet m)" for put by(simp add: put_nondet_def) lemma run_ask_nondet [simp]: "run_nondet (ask_nondet ask f) = ask (\r. run_nondet (f r))" for ask by(simp add: ask_nondet_def) end lemma bind_nondet_cong [cong]: "nondetM_base.bind_nondet bind merge = nondetM_base.bind_nondet bind merge" for bind merge .. lemmas [code] = nondetM_base.return_nondet_def nondetM_base.bind_nondet_def nondetM_base.fail_nondet_def nondetM_base.alt_nondet_def nondetM_base.get_nondet_def nondetM_base.put_nondet_def nondetM_base.ask_nondet_def locale nondetM = nondetM_base return bind merge empty single union + monad_commute return bind for return :: "('s, 'm) return" and bind :: "('s, 'm) bind" and merge :: "('a, 'm, 's) merge" and empty :: "'s" and single :: "'a \ 's" and union :: "'s \ 's \ 's" (infixl "\<^bold>\" 65) + assumes bind_merge_merge: "\y f g. bind (merge y f) (\A. merge A g) = merge y (\x. bind (f x) (\A. merge A g))" and merge_empty: "\f. merge empty f = return empty" and merge_single: "\x f. merge (single x) f = f x" and merge_single2: "\A. merge A (\x. return (single x)) = return A" and merge_union: "\A B f. merge (A \<^bold>\ B) f = bind (merge A f) (\A'. bind (merge B f) (\B'. return (A' \<^bold>\ B')))" and union_assoc: "\A B C. (A \<^bold>\ B) \<^bold>\ C = A \<^bold>\ (B \<^bold>\ C)" and empty_union: "\A. empty \<^bold>\ A = A" and union_empty: "\A. A \<^bold>\ empty = A" begin lemma monad_nondetT [locale_witness]: "monad return_nondet bind_nondet" proof show "bind_nondet (bind_nondet x f) g = bind_nondet x (\y. bind_nondet (f y) g)" for x f g by(rule nondetT.expand)(simp add: bind_assoc bind_merge_merge o_def) show "bind_nondet (return_nondet x) f = f x" for x f by(rule nondetT.expand)(simp add: return_bind merge_single) show "bind_nondet x return_nondet = x" for x by(rule nondetT.expand)(simp add: bind_return o_def merge_single2) qed lemma monad_fail_nondetT [locale_witness]: "monad_fail return_nondet bind_nondet fail_nondet" proof show "bind_nondet fail_nondet f = fail_nondet" for f by(rule nondetT.expand)(simp add: return_bind merge_empty) qed lemma monad_alt_nondetT [locale_witness]: "monad_alt return_nondet bind_nondet alt_nondet" proof show "alt_nondet (alt_nondet m1 m2) m3 = alt_nondet m1 (alt_nondet m2 m3)" for m1 m2 m3 by(rule nondetT.expand)(simp add: bind_assoc return_bind union_assoc) show "bind_nondet (alt_nondet m m') f = alt_nondet (bind_nondet m f) (bind_nondet m' f)" for m m' f apply(rule nondetT.expand) apply(simp add: bind_assoc return_bind) apply(subst (2) bind_commute) apply(simp add: merge_union) done qed lemma monad_fail_alt_nondetT [locale_witness]: "monad_fail_alt return_nondet bind_nondet fail_nondet alt_nondet" proof show "alt_nondet fail_nondet m = m" for m by(rule nondetT.expand)(simp add: return_bind bind_return empty_union) show "alt_nondet m fail_nondet = m" for m by(rule nondetT.expand)(simp add: return_bind bind_return union_empty) qed lemma monad_state_nondetT [locale_witness]: \ \It's not really sensible to assume a commutative state monad, but let's prove it anyway ...\ fixes get put assumes "monad_state return bind get put" shows "monad_state return_nondet bind_nondet (get_nondet get) (put_nondet put)" proof - interpret monad_state return bind get put by fact show ?thesis proof show "put_nondet put s (get_nondet get f) = put_nondet put s (f s)" for s f by(rule nondetT.expand)(simp add: put_get) show "get_nondet get (\s. get_nondet get (f s)) = get_nondet get (\s. f s s)" for f by(rule nondetT.expand)(simp add: get_get) show "put_nondet put s (put_nondet put s' m) = put_nondet put s' m" for s s' m by(rule nondetT.expand)(simp add: put_put) show "get_nondet get (\s. put_nondet put s m) = m" for m by(rule nondetT.expand)(simp add: get_put) show "get_nondet get (\_. m) = m" for m by(rule nondetT.expand)(simp add: get_const) show "bind_nondet (get_nondet get f) g = get_nondet get (\s. bind_nondet (f s) g)" for f g by(rule nondetT.expand)(simp add: bind_get) show "bind_nondet (put_nondet put s m) f = put_nondet put s (bind_nondet m f)" for s m f by(rule nondetT.expand)(simp add: bind_put) qed qed lemma monad_state_alt_nondetT [locale_witness]: fixes get put assumes "monad_state return bind get put" shows "monad_state_alt return_nondet bind_nondet (get_nondet get) (put_nondet put) alt_nondet" proof - interpret monad_state return bind get put by fact show ?thesis proof show "alt_nondet (get_nondet get f) (get_nondet get g) = get_nondet get (\x. alt_nondet (f x) (g x))" for f g apply(rule nondetT.expand; simp) apply(subst bind_get) apply(subst (1 2) bind_commute) apply(simp add: bind_get get_get) done show "alt_nondet (put_nondet put s m) (put_nondet put s m') = put_nondet put s (alt_nondet m m')" for s m m' apply(rule nondetT.expand; simp) apply(subst bind_put) apply(subst (1 2) bind_commute) apply(simp add: bind_put put_put) done qed qed end lemmas nondetM_lemmas = nondetM.monad_nondetT nondetM.monad_fail_nondetT nondetM.monad_alt_nondetT nondetM.monad_fail_alt_nondetT nondetM.monad_state_nondetT locale nondetM_ask = nondetM return bind merge empty single union for return :: "('s, 'm) return" and bind :: "('s, 'm) bind" and ask :: "('r, 'm) ask" and merge :: "('a, 'm, 's) merge" and empty :: "'s" and single :: "'a \ 's" and union :: "'s \ 's \ 's" (infixl "\<^bold>\" 65) + assumes monad_reader: "monad_reader return bind ask" assumes merge_ask: "\A (f :: 'a \ 'r \ ('a, 'm) nondetT). merge A (\x. ask (\r. run_nondet (f x r))) = ask (\r. merge A (\x. run_nondet (f x r)))" begin interpretation monad_reader return bind ask by(fact monad_reader) lemma monad_reader_nondetT: "monad_reader return_nondet bind_nondet (ask_nondet ask)" proof show "ask_nondet ask (\r. ask_nondet ask (f r)) = ask_nondet ask (\r. f r r)" for f by(rule nondetT.expand)(simp add: ask_ask) show "ask_nondet ask (\_. m) = m" for m by(rule nondetT.expand)(simp add: ask_const) show "bind_nondet (ask_nondet ask f) g = ask_nondet ask (\r. bind_nondet (f r) g)" for f g by(rule nondetT.expand)(simp add: bind_ask) show "bind_nondet m (\x. ask_nondet ask (f x)) = ask_nondet ask (\r. bind_nondet m (\x. f x r))" for f m by(rule nondetT.expand)(simp add: bind_ask2[symmetric] o_def merge_ask) qed end lemmas nondetM_ask_lemmas = nondetM_ask.monad_reader_nondetT subsubsection \Parametricity\ context includes lifting_syntax begin lemma return_nondet_parametric [transfer_rule]: "((S ===> M) ===> (A ===> S) ===> A ===> rel_nondetT M) nondetM_base.return_nondet nondetM_base.return_nondet" unfolding nondetM_base.return_nondet_def by transfer_prover lemma bind_nondet_parametric [transfer_rule]: "((M ===> (S ===> M) ===> M) ===> (S ===> (A ===> M) ===> M) ===> rel_nondetT M ===> (A ===> rel_nondetT M) ===> rel_nondetT M) nondetM_base.bind_nondet nondetM_base.bind_nondet" unfolding nondetM_base.bind_nondet_def by transfer_prover lemma fail_nondet_parametric [transfer_rule]: "((S ===> M) ===> S ===> rel_nondetT M) nondetM_base.fail_nondet nondetM_base.fail_nondet" unfolding nondetM_base.fail_nondet_def by transfer_prover lemma alt_nondet_parametric [transfer_rule]: "((S ===> M) ===> (M ===> (S ===> M) ===> M) ===> (S ===> S ===> S) ===> rel_nondetT M ===> rel_nondetT M ===> rel_nondetT M) nondetM_base.alt_nondet nondetM_base.alt_nondet" unfolding nondetM_base.alt_nondet_def by transfer_prover lemma get_nondet_parametric [transfer_rule]: "(((S ===> M) ===> M) ===> (S ===> rel_nondetT M) ===> rel_nondetT M) nondetM_base.get_nondet nondetM_base.get_nondet" unfolding nondetM_base.get_nondet_def by transfer_prover lemma put_nondet_parametric [transfer_rule]: "((S ===> M ===> M) ===> S ===> rel_nondetT M ===> rel_nondetT M) nondetM_base.put_nondet nondetM_base.put_nondet" unfolding nondetM_base.put_nondet_def by transfer_prover lemma ask_nondet_parametric [transfer_rule]: "(((R ===> M) ===> M) ===> (R ===> rel_nondetT M) ===> rel_nondetT M) nondetM_base.ask_nondet nondetM_base.ask_nondet" unfolding nondetM_base.ask_nondet_def by transfer_prover end subsubsection \Implementation using lists\ context fixes return :: "('a list, 'm) return" and bind :: "('a list, 'm) bind" and lunionM lUnionM defines "lunionM m1 m2 \ bind m1 (\A. bind m2 (\B. return (A @ B)))" and "lUnionM ms \ foldr lunionM ms (return [])" begin definition lmerge :: "'a list \ ('a \ 'm) \ 'm" where "lmerge A f = lUnionM (map f A)" context assumes "monad_commute return bind" begin interpretation monad_commute return bind by fact interpretation nondetM_base return bind lmerge "[]" "\x. [x]" "(@)" . lemma lUnionM_empty [simp]: "lUnionM [] = return []" by(simp add: lUnionM_def) lemma lUnionM_Cons [simp]: "lUnionM (x # M) = lunionM x (lUnionM M)" for x M by(simp add: lUnionM_def) lemma lunionM_return_empty1 [simp]: "lunionM (return []) x = x" for x by(simp add: lunionM_def return_bind bind_return) lemma lunionM_return_empty2 [simp]: "lunionM x (return []) = x" for x by(simp add: lunionM_def return_bind bind_return) lemma lunionM_return_return [simp]: "lunionM (return A) (return B) = return (A @ B)" for A B by(simp add: lunionM_def return_bind) lemma lunionM_assoc: "lunionM (lunionM x y) z = lunionM x (lunionM y z)" for x y z by(simp add: lunionM_def bind_assoc return_bind) lemma lunionM_lUnionM1: "lunionM (lUnionM A) x = foldr lunionM A x" for A x by(induction A arbitrary: x)(simp_all add: lunionM_assoc) lemma lUnionM_append [simp]: "lUnionM (A @ B) = lunionM (lUnionM A) (lUnionM B)" for A B by(subst lunionM_lUnionM1)(simp add: lUnionM_def) lemma lUnionM_return [simp]: "lUnionM (map (\x. return [x]) A) = return A" for A by(induction A) simp_all lemma bind_lunionM: "bind (lunionM m m') f = lunionM (bind m f) (bind m' f)" if "\A B. f (A @ B) = bind (f A) (\x. bind (f B) (\y. return (x @ y)))" for m m' f apply(simp add: bind_assoc return_bind lunionM_def that) apply(subst (2) bind_commute) apply simp done lemma list_nondetM: "nondetM return bind lmerge [] (\x. [x]) (@)" proof show "bind (lmerge y f) (\A. lmerge A g) = lmerge y (\x. bind (f x) (\A. lmerge A g))" for y f g apply(induction y) apply(simp_all add: lmerge_def return_bind) apply(subst bind_lunionM; simp add: lunionM_def o_def) done show "lmerge [] f = return []" for f by(simp add: lmerge_def) show "lmerge [x] f = f x" for x f by(simp add: lmerge_def) show "lmerge A (\x. return [x]) = return A" for A by(simp add: lmerge_def) show "lmerge (A @ B) f = bind (lmerge A f) (\A'. bind (lmerge B f) (\B'. return (A' @ B')))" for f A B by(simp add: lmerge_def lunionM_def) qed simp_all lemma list_nondetM_ask: notes list_nondetM[locale_witness] assumes [locale_witness]: "monad_reader return bind ask" shows "nondetM_ask return bind ask lmerge [] (\x. [x]) (@)" proof interpret monad_reader return bind ask by fact show "lmerge A (\x. ask (\r. run_nondet (f x r))) = ask (\r. lmerge A (\x. run_nondet (f x r)))" for A and f :: "'a \ 'b \ ('a, 'm) nondetT" unfolding lmerge_def by(induction A)(simp_all add: ask_const lunionM_def bind_ask bind_ask2 ask_ask) qed lemmas list_nondetMs [locale_witness] = nondetM_lemmas[OF list_nondetM] nondetM_ask_lemmas[OF list_nondetM_ask] end end lemma lmerge_parametric [transfer_rule]: includes lifting_syntax shows "((list_all2 A ===> M) ===> (M ===> (list_all2 A ===> M) ===> M) ===> list_all2 A ===> (A ===> M) ===> M) lmerge lmerge" unfolding lmerge_def by transfer_prover subsubsection \Implementation using multisets\ context fixes return :: "('a multiset, 'm) return" and bind :: "('a multiset, 'm) bind" and munionM mUnionM defines "munionM m1 m2 \ bind m1 (\A. bind m2 (\B. return (A + B)))" and "mUnionM \ fold_mset munionM (return {#})" begin definition mmerge :: "'a multiset \ ('a \ 'm) \ 'm" where "mmerge A f = mUnionM (image_mset f A)" context assumes "monad_commute return bind" begin interpretation monad_commute return bind by fact interpretation nondetM_base return bind mmerge "{#}" "\x. {#x#}" "(+)" . lemma munionM_comp_fun_commute: "comp_fun_commute munionM" apply(unfold_locales) apply(simp add: fun_eq_iff bind_assoc return_bind munionM_def) apply(subst bind_commute) apply(simp add: union_ac) done interpretation comp_fun_commute munionM by(rule munionM_comp_fun_commute) lemma mUnionM_empty [simp]: "mUnionM {#} = return {#}" by(simp add: mUnionM_def) lemma mUnionM_add_mset [simp]: "mUnionM (add_mset x M) = munionM x (mUnionM M)" for x M by(simp add: mUnionM_def) lemma munionM_return_empty1 [simp]: "munionM (return {#}) x = x" for x by(simp add: munionM_def return_bind bind_return) lemma munionM_return_empty2 [simp]: "munionM x (return {#}) = x" for x by(simp add: munionM_def return_bind bind_return) lemma munionM_return_return [simp]: "munionM (return A) (return B) = return (A + B)" for A B by(simp add: munionM_def return_bind) lemma munionM_assoc: "munionM (munionM x y) z = munionM x (munionM y z)" for x y z by(simp add: munionM_def bind_assoc return_bind add.assoc) lemma munionM_commute: "munionM x y = munionM y x" for x y unfolding munionM_def by(subst bind_commute)(simp add: add.commute) lemma munionM_mUnionM1: "munionM (mUnionM A) x = fold_mset munionM x A" for A x by(induction A arbitrary: x)(simp_all add: munionM_assoc) lemma munionM_mUnionM2: "munionM x (mUnionM A) = fold_mset munionM x A" for x A by(subst munionM_commute)(rule munionM_mUnionM1) lemma mUnionM_add [simp]: "mUnionM (A + B) = munionM (mUnionM A) (mUnionM B)" for A B by(subst munionM_mUnionM2)(simp add: mUnionM_def) lemma mUnionM_return [simp]: "mUnionM (image_mset (\x. return {#x#}) A) = return A" for A by(induction A) simp_all lemma bind_munionM: "bind (munionM m m') f = munionM (bind m f) (bind m' f)" if "\A B. f (A + B) = bind (f A) (\x. bind (f B) (\y. return (x + y)))" for m m' f apply(simp add: bind_assoc return_bind munionM_def that) apply(subst (2) bind_commute) apply simp done lemma mset_nondetM: "nondetM return bind mmerge {#} (\x. {#x#}) (+)" proof show "bind (mmerge y f) (\A. mmerge A g) = mmerge y (\x. bind (f x) (\A. mmerge A g))" for y f g apply(induction y) apply(simp_all add: return_bind mmerge_def) apply(subst bind_munionM; simp add: munionM_def o_def) done show "mmerge {#} f = return {#}" for f by(simp add: mmerge_def) show "mmerge {#x#} f = f x" for x f by(simp add: mmerge_def) show "mmerge A (\x. return {#x#}) = return A" for A by(simp add: mmerge_def) show "mmerge (A + B) f = bind (mmerge A f) (\A'. bind (mmerge B f) (\B'. return (A' + B')))" for f A B by(simp add: mmerge_def munionM_def) qed simp_all lemma mset_nondetM_ask: notes mset_nondetM[locale_witness] assumes [locale_witness]: "monad_reader return bind ask" shows "nondetM_ask return bind ask mmerge {#} (\x. {#x#}) (+)" proof interpret monad_reader return bind ask by fact show "mmerge A (\x. ask (\r. run_nondet (f x r))) = ask (\r. mmerge A (\x. run_nondet (f x r)))" for A and f :: "'a \ 'b \ ('a, 'm) nondetT" unfolding mmerge_def by(induction A)(simp_all add: ask_const munionM_def bind_ask bind_ask2 ask_ask) qed lemmas mset_nondetMs [locale_witness] = nondetM_lemmas[OF mset_nondetM] nondetM_ask_lemmas[OF mset_nondetM_ask] end end lemma mmerge_parametric: includes lifting_syntax assumes return [transfer_rule]: "(rel_mset A ===> M) return1 return2" and bind [transfer_rule]: "(M ===> (rel_mset A ===> M) ===> M) bind1 bind2" and comm1: "monad_commute return1 bind1" and comm2: "monad_commute return2 bind2" shows "(rel_mset A ===> (A ===> M) ===> M) (mmerge return1 bind1) (mmerge return2 bind2)" unfolding mmerge_def apply(rule rel_funI)+ apply(drule (1) multiset.map_transfer[THEN rel_funD, THEN rel_funD]) apply(rule fold_mset_parametric[OF _ munionM_comp_fun_commute[OF comm1] munionM_comp_fun_commute[OF comm2], THEN rel_funD, THEN rel_funD, rotated -1], assumption) subgoal premises [transfer_rule] by transfer_prover subgoal premises by transfer_prover done subsubsection \Implementation using finite sets\ context fixes return :: "('a fset, 'm) return" and bind :: "('a fset, 'm) bind" and funionM fUnionM defines "funionM m1 m2 \ bind m1 (\A. bind m2 (\B. return (A |\| B)))" and "fUnionM \ ffold funionM (return {||})" begin definition fmerge :: "'a fset \ ('a \ 'm) \ 'm" where "fmerge A f = fUnionM (fimage f A)" context assumes "monad_commute return bind" and "monad_duplicate return bind" begin interpretation monad_commute return bind by fact interpretation monad_duplicate return bind by fact interpretation nondetM_base return bind fmerge "{||}" "\x. {|x|}" "(|\|)" . lemma funionM_comp_fun_commute: "comp_fun_commute funionM" apply(unfold_locales) apply(simp add: fun_eq_iff bind_assoc return_bind funionM_def) apply(subst bind_commute) apply(simp add: funion_ac) done interpretation comp_fun_commute funionM by(rule funionM_comp_fun_commute) lemma funionM_comp_fun_idem: "comp_fun_idem funionM" by(unfold_locales)(simp add: fun_eq_iff funionM_def bind_assoc bind_duplicate return_bind) interpretation comp_fun_idem funionM by(rule funionM_comp_fun_idem) lemma fUnionM_empty [simp]: "fUnionM {||} = return {||}" by(simp add: fUnionM_def) lemma fUnionM_finset [simp]: "fUnionM (finsert x M) = funionM x (fUnionM M)" for x M by(simp add: fUnionM_def) lemma funionM_return_empty1 [simp]: "funionM (return {||}) x = x" for x by(simp add: funionM_def return_bind bind_return) lemma funionM_return_empty2 [simp]: "funionM x (return {||}) = x" for x by(simp add: funionM_def return_bind bind_return) lemma funionM_return_return [simp]: "funionM (return A) (return B) = return (A |\| B)" for A B by(simp add: funionM_def return_bind) lemma funionM_assoc: "funionM (funionM x y) z = funionM x (funionM y z)" for x y z by(simp add: funionM_def bind_assoc return_bind funion_assoc) lemma funionM_commute: "funionM x y = funionM y x" for x y unfolding funionM_def by(subst bind_commute)(simp add: funion_commute) lemma funionM_fUnionM1: "funionM (fUnionM A) x = ffold funionM x A" for A x by(induction A arbitrary: x)(simp_all add: funionM_assoc) lemma funionM_fUnionM2: "funionM x (fUnionM A) = ffold funionM x A" for x A by(subst funionM_commute)(rule funionM_fUnionM1) lemma fUnionM_funion [simp]: "fUnionM (A |\| B) = funionM (fUnionM A) (fUnionM B)" for A B by(subst funionM_fUnionM2)(simp add: fUnionM_def ffold_set_union) lemma fUnionM_return [simp]: "fUnionM (fimage (\x. return {|x|}) A) = return A" for A by(induction A) simp_all lemma bind_funionM: "bind (funionM m m') f = funionM (bind m f) (bind m' f)" if "\A B. f (A |\| B) = bind (f A) (\x. bind (f B) (\y. return (x |\| y)))" for m m' f apply(simp add: bind_assoc return_bind funionM_def that) apply(subst (2) bind_commute) apply simp done lemma fUnionM_return_fempty [simp]: "fUnionM (fimage (\x. return {||}) A) = return {||}" for A by(induction A) simp_all lemma funionM_bind: "funionM (bind m f) (bind m g) = bind m (\x. funionM (f x) (g x))" for m f g unfolding funionM_def bind_assoc by(subst bind_commute)(simp add: bind_duplicate) lemma fUnionM_funionM: "fUnionM ((\y. funionM (f y) (g y)) |`| A) = funionM (fUnionM (f |`| A)) (fUnionM (g |`| A))" for f g A by(induction A)(simp_all add: funionM_assoc funionM_commute fun_left_comm) lemma fset_nondetM: "nondetM return bind fmerge {||} (\x. {|x|}) (|\|)" proof show "bind (fmerge y f) (\A. fmerge A g) = fmerge y (\x. bind (f x) (\A. fmerge A g))" for y f g apply(induction y) apply(simp_all add: return_bind fmerge_def) apply(subst bind_funionM; simp add: funionM_def o_def fimage_funion) done show "fmerge {||} f = return {||}" for f by(simp add: fmerge_def) show "fmerge {|x|} f = f x" for x f by(simp add: fmerge_def) show "fmerge A (\x. return {|x|}) = return A" for A by(simp add: fmerge_def) show "fmerge (A |\| B) f = bind (fmerge A f) (\A'. bind (fmerge B f) (\B'. return (A' |\| B')))" for f A B by(simp add: fmerge_def funionM_def fimage_funion) qed auto lemma fset_nondetM_ask: notes fset_nondetM[locale_witness] assumes [locale_witness]: "monad_reader return bind ask" shows "nondetM_ask return bind ask fmerge {||} (\x. {|x|}) (|\|)" proof interpret monad_reader return bind ask by fact show "fmerge A (\x. ask (\r. run_nondet (f x r))) = ask (\r. fmerge A (\x. run_nondet (f x r)))" for A and f :: "'a \ 'b \ ('a, 'm) nondetT" unfolding fmerge_def by(induction A)(simp_all add: ask_const funionM_def bind_ask bind_ask2 ask_ask) qed lemmas fset_nondetMs [locale_witness] = nondetM_lemmas[OF fset_nondetM] nondetM_ask_lemmas[OF fset_nondetM_ask] context assumes "monad_discard return bind" begin interpretation monad_discard return bind by fact lemma fmerge_bind: "fmerge A (\x. bind m' (\A'. fmerge A' (f x))) = bind m' (\A'. fmerge A (\x. fmerge A' (f x)))" by(induction A)(simp_all add: fmerge_def bind_const funionM_bind) lemma fmerge_commute: "fmerge A (\x. fmerge B (f x)) = fmerge B (\y. fmerge A (\x. f x y))" by(induction A)(simp_all add: fmerge_def fUnionM_funionM) lemma monad_commute_nondetT_fset [locale_witness]: "monad_commute return_nondet bind_nondet" proof show "bind_nondet m (\x. bind_nondet m' (f x)) = bind_nondet m' (\y. bind_nondet m (\x. f x y))" for m m' f apply(rule nondetT.expand) apply(simp add: o_def) apply(subst fmerge_bind) apply(subst bind_commute) apply(subst fmerge_commute) apply(subst fmerge_bind[symmetric]) apply(rule refl) done qed end end end lemma fmerge_parametric: includes lifting_syntax assumes return [transfer_rule]: "(rel_fset A ===> M) return1 return2" and bind [transfer_rule]: "(M ===> (rel_fset A ===> M) ===> M) bind1 bind2" and comm1: "monad_commute return1 bind1" "monad_duplicate return1 bind1" and comm2: "monad_commute return2 bind2" "monad_duplicate return2 bind2" shows "(rel_fset A ===> (A ===> M) ===> M) (fmerge return1 bind1) (fmerge return2 bind2)" unfolding fmerge_def apply(rule rel_funI)+ apply(drule (1) fset.map_transfer[THEN rel_funD, THEN rel_funD]) apply(rule ffold_parametric[OF _ funionM_comp_fun_idem[OF comm1] funionM_comp_fun_idem[OF comm2], THEN rel_funD, THEN rel_funD, rotated -1], assumption) subgoal premises [transfer_rule] by transfer_prover subgoal premises by transfer_prover done subsubsection \Implementation using countable sets\ text \For non-finite choices, we cannot generically construct the merge operation. So we formalize in a locale what can be proven generically and then prove instances of the locale for concrete locale implementations. We need two separate merge parameters because we must merge effects over choices (type @{typ 'c}) and effects over the non-deterministic results (type @{typ 'a}) of computations. \ locale cset_nondetM_base = nondetM_base return bind merge cempty csingle cUn for return :: "('a cset, 'm) return" and bind :: "('a cset, 'm) bind" and merge :: "('a, 'm, 'a cset) merge" and mergec :: "('c, 'm, 'c cset) merge" begin definition altc_nondet :: "('c, ('a, 'm) nondetT) altc" where "altc_nondet A f = NondetT (mergec A (run_nondet \ f))" lemma run_altc_nondet [simp]: "run_nondet (altc_nondet A f) = mergec A (run_nondet \ f)" by(simp add: altc_nondet_def) end locale cset_nondetM = cset_nondetM_base return bind merge mergec + monad_commute return bind + monad_duplicate return bind for return :: "('a cset, 'm) return" and bind :: "('a cset, 'm) bind" and merge :: "('a, 'm, 'a cset) merge" and mergec :: "('c, 'm, 'c cset) merge" + assumes bind_merge_merge: "\y f g. bind (merge y f) (\A. merge A g) = merge y (\x. bind (f x) (\A. merge A g))" and merge_empty: "\f. merge cempty f = return cempty" and merge_single: "\x f. merge (csingle x) f = f x" and merge_single2: "\A. merge A (\x. return (csingle x)) = return A" and merge_union: "\A B f. merge (cUn A B) f = bind (merge A f) (\A'. bind (merge B f) (\B'. return (cUn A' B')))" and bind_mergec_merge: "\y f g. bind (mergec y f) (\A. merge A g) = mergec y (\x. bind (f x) (\A. merge A g))" and mergec_single: "\x f. mergec (csingle x) f = f x" and mergec_UNION: "\C f g. mergec (cUNION C f) g = mergec C (\x. mergec (f x) g)" and mergec_parametric [transfer_rule]: "\R. bi_unique R \ rel_fun (rel_cset R) (rel_fun (rel_fun R (=)) (=)) mergec mergec" begin interpretation nondetM return bind merge cempty csingle cUn by(unfold_locales; (rule bind_merge_merge merge_empty merge_single merge_single2 merge_union | simp add: cUn_assoc)?) sublocale nondet: monad_altc return_nondet bind_nondet altc_nondet including lifting_syntax proof show "bind_nondet (altc_nondet C g) f = altc_nondet C (\c. bind_nondet (g c) f)" for C g f by(rule nondetT.expand)(simp add: bind_mergec_merge o_def) show "altc_nondet (csingle x) f = f x" for x f by(rule nondetT.expand)(simp add: mergec_single) show "altc_nondet (cUNION C f) g = altc_nondet C (\x. altc_nondet (f x) g)" for C f g by(rule nondetT.expand)(simp add: o_def mergec_UNION) show "(rel_cset R ===> (R ===> (=)) ===> (=)) altc_nondet altc_nondet" if [transfer_rule]: "bi_unique R" for R unfolding altc_nondet_def by(transfer_prover) qed end locale cset_nondetM3 = cset_nondetM return bind merge mergec + three "TYPE('c)" for return :: "('a cset, 'm) return" and bind :: "('a cset, 'm) bind" and merge :: "('a, 'm, 'a cset) merge" and mergec :: "('c, 'm, 'c cset) merge" begin interpretation nondet: monad_altc3 return_nondet bind_nondet altc_nondet .. end paragraph \Identity monad\ definition merge_id :: "('c, 'a cset id, 'c cset) merge" where "merge_id A f = return_id (cUNION A (extract \ f))" lemma extract_merge_id [simp]: "extract (merge_id A f) = cUNION A (extract \ f)" by(simp add: merge_id_def) lemma merge_id_parametric [transfer_rule]: includes lifting_syntax shows "(rel_cset A ===> (A ===> rel_id (rel_cset A)) ===> rel_id (rel_cset A)) merge_id merge_id" unfolding merge_id_def by transfer_prover lemma cset_nondetM_id [locale_witness]: "cset_nondetM return_id bind_id merge_id merge_id" including lifting_syntax proof(unfold_locales) show "bind_id (merge_id y f) (\A. merge_id A g) = merge_id y (\x. bind_id (f x) (\A. merge_id A g))" for y and f :: "'c \ 'd cset id" and g by(rule id.expand)(simp add: o_def cUNION_assoc) then show "bind_id (merge_id y f) (\A. merge_id A g) = merge_id y (\x. bind_id (f x) (\A. merge_id A g))" for y and f :: "'c \ 'd cset id" and g by this show "merge_id cempty f = return_id cempty" for f :: "'a \ 'a cset id" by(rule id.expand) simp show "merge_id (csingle x) f = f x" for x and f :: "'c \ 'a cset id" by(rule id.expand) simp then show "merge_id (csingle x) f = f x" for x and f :: "'c \ 'a cset id" by this show "merge_id A (\x. return_id (csingle x)) = return_id A" for A :: "'a cset" by(rule id.expand)(simp add: o_def) show "merge_id (cUn A B) f = bind_id (merge_id A f) (\A'. bind_id (merge_id B f) (\B'. return_id (cUn A' B')))" for A B and f :: "'a \ 'a cset id" by(rule id.expand)(simp add: cUNION_cUn) show "merge_id (cUNION C f) g = merge_id C (\x. merge_id (f x) g)" for C and f :: "'b \ 'b cset" and g :: "'b \ 'a cset id" by(rule id.expand)(simp add: o_def cUNION_assoc) show "(rel_cset R ===> (R ===> (=)) ===> (=)) merge_id merge_id" if "bi_unique R" for R :: "'b \ 'b \ bool" unfolding merge_id_def by transfer_prover qed paragraph \Reader monad transformer\ definition merge_env :: "('c, 'm, 'c cset) merge \ ('c, ('r, 'm) envT, 'c cset) merge" where "merge_env merge A f = EnvT (\r. merge A (\a. run_env (f a) r))" for merge lemma run_merge_env [simp]: "run_env (merge_env merge A f) r = merge A (\a. run_env (f a) r)" for merge by(simp add: merge_env_def) lemma merge_env_parametric [transfer_rule]: includes lifting_syntax shows "((rel_cset C ===> (C ===> M) ===> M) ===> rel_cset C ===> (C ===> rel_envT R M) ===> rel_envT R M) merge_env merge_env" unfolding merge_env_def by transfer_prover lemma cset_nondetM_envT [locale_witness]: fixes return :: "('a cset, 'm) return" and bind :: "('a cset, 'm) bind" and merge :: "('a, 'm, 'a cset) merge" and mergec :: "('c, 'm, 'c cset) merge" assumes "cset_nondetM return bind merge mergec" shows "cset_nondetM (return_env return) (bind_env bind) (merge_env merge) (merge_env mergec)" proof - interpret cset_nondetM return bind merge by fact show ?thesis including lifting_syntax proof show "bind_env bind (merge_env merge y f) (\A. merge_env merge A g) = merge_env merge y (\x. bind_env bind (f x) (\A. merge_env merge A g))" for y and f :: "'a \ ('b, 'm) envT" and g by(rule envT.expand)(simp add: fun_eq_iff cUNION_assoc bind_merge_merge) show "merge_env merge cempty f = return_env return cempty" for f :: "'a \ ('b, 'm) envT" by(rule envT.expand)(simp add: fun_eq_iff merge_empty) show "merge_env merge (csingle x) f = f x" for f :: "'a \ ('b, 'm) envT" and x by(rule envT.expand)(simp add: fun_eq_iff merge_single) show "merge_env merge A (\x. return_env return (csingle x)) = return_env return A" for A by(rule envT.expand)(simp add: fun_eq_iff merge_single2) show "merge_env merge (cUn A B) f = bind_env bind (merge_env merge A f) (\A'. bind_env bind (merge_env merge B f) (\B'. return_env return (cUn A' B')))" for A B and f :: "'a \ ('b, 'm) envT" by(rule envT.expand)(simp add: fun_eq_iff merge_union) show "bind_env bind (merge_env mergec y f) (\A. merge_env merge A g) = merge_env mergec y (\x. bind_env bind (f x) (\A. merge_env merge A g))" for y and f :: "'c \ ('b, 'm) envT" and g by(rule envT.expand)(simp add: fun_eq_iff cUNION_assoc bind_mergec_merge) show "merge_env mergec (csingle x) f = f x" for f :: "'c \ ('b, 'm) envT" and x by(rule envT.expand)(simp add: fun_eq_iff mergec_single) show "merge_env mergec (cUNION C f) g = merge_env mergec C (\x. merge_env mergec (f x) g)" for C f and g :: "'c \ ('b, 'm) envT" by(rule envT.expand)(simp add: fun_eq_iff mergec_UNION) show "(rel_cset R ===> (R ===> (=)) ===> (=)) (merge_env mergec) (merge_env mergec)" if [transfer_rule]: "bi_unique R" for R unfolding merge_env_def by transfer_prover qed qed (* paragraph \Exception monad transformer\ text \Failure in the non-determinism transformer is absorbed by choice. So there is no way to handle failures. In contrast, in this combination, we choose that failures abort all choices.\ context fixes return :: "('a cset option, 'm) return" and bind :: "('a cset option, 'm) bind" and merge :: "('a option, 'm, 'a option cset) merge" begin definition merge_option :: "('a, ('a cset, 'm) optionT, 'a cset) merge" where "merge_option A f = OptionT (merge (cimage Some A) (\x. case x of None \ return None | Some x' \ run_option (f x')))" lemma run_merge_optionT [simp]: "run_option (merge_option A f) = merge (cimage Some A) (\x. case x of None \ return None | Some x' \ run_option (f x'))" by(simp add: merge_option_def) definition return_optionT_cset :: "('a option cset, 'm) return" where "return_optionT_cset A = (if cin None A then return None else return (Some (cimage the A)))" definition bind_optionT_cset :: "('a option cset, 'm) bind" where "bind_optionT_cset m f = bind m (\x. case x of None \ return None | Some A \ f (cimage Some A))" (* lemma assumes "monad return_optionT_cset bind_optionT_cset" shows "monad return bind" proof - interpret monad return_optionT_cset bind_optionT_cset by fact show ?thesis proof(unfold_locales) show "bind (return x) f = f x" for x f using return_bind[of "case x of None \ csingle None | Some A \ cimage Some A" ff] apply(simp add: bind_optionT_cset_def return_optionT_cset_def split: option.split_asm) subgoal for m f g using bind_assoc[of m "\A. if cin None A then return None else f (Some (cimage the A))" "\A. if cin None A then return None else g (Some (cimage the A))"] apply(simp add: bind_optionT_cset_def) apply(subst (asm) if_distrib[where f="\m. bind m _"]) *) lemma cset_nondetM_optionT [locale_witness]: assumes "monad_commute return bind" and "monad_discard return bind" and "monad_duplicate return bind" and "cset_nondetM return_optionT_cset bind_optionT_cset merge" shows "cset_nondetM (return_option return) (bind_option return bind) merge_option" proof - interpret monad_commute return bind by fact interpret monad_discard return bind by fact interpret monad_duplicate return bind by fact interpret *: cset_nondetM return_optionT_cset bind_optionT_cset merge by fact show ?thesis proof show "bind_option return bind (merge_option y f) (\A. merge_option A g) = merge_option y (\x. bind_option return bind (f x) (\A. merge_option A g))" for y f g apply(rule optionT.expand) apply(simp add: run_bind_option cong del: option.case_cong) *) subsection \State transformer\ datatype ('s, 'm) stateT = StateT (run_state: "'s \ 'm") for rel: rel_stateT' text \ We define a more general relator for @{typ "(_, _) stateT"} than the one generated by the datatype package such that we can also show parametricity in the state. \ context includes lifting_syntax begin definition rel_stateT :: "('s \ 's' \ bool) \ ('m \ 'm' \ bool) \ ('s, 'm) stateT \ ('s', 'm') stateT \ bool" where "rel_stateT S M m m' \ (S ===> M) (run_state m) (run_state m')" lemma rel_stateT_eq [relator_eq]: "rel_stateT (=) (=) = (=)" by(auto simp add: rel_stateT_def fun_eq_iff rel_fun_eq intro: stateT.expand) lemma rel_stateT_mono [relator_mono]: "\ S' \ S; M \ M' \ \ rel_stateT S M \ rel_stateT S' M'" by(rule predicate2I)(simp add: rel_stateT_def fun_mono[THEN predicate2D]) lemma StateT_parametric [transfer_rule]: "((S ===> M) ===> rel_stateT S M) StateT StateT" by(auto simp add: rel_stateT_def) lemma run_state_parametric [transfer_rule]: "(rel_stateT S M ===> S ===> M) run_state run_state" by(auto simp add: rel_stateT_def) lemma case_stateT_parametric [transfer_rule]: "(((S ===> M) ===> A) ===> rel_stateT S M ===> A) case_stateT case_stateT" by(auto 4 3 split: stateT.split simp add: rel_stateT_def del: rel_funI intro!: rel_funI dest: rel_funD) lemma rec_stateT_parametric [transfer_rule]: "(((S ===> M) ===> A) ===> rel_stateT S M ===> A) rec_stateT rec_stateT" apply(rule rel_funI)+ subgoal for \ m m' by(cases m; cases m')(auto 4 3 simp add: rel_stateT_def del: rel_funI intro!: rel_funI dest: rel_funD) done lemma rel_stateT_Grp: "rel_stateT (=) (BNF_Def.Grp UNIV f) = BNF_Def.Grp UNIV (map_stateT f)" by(auto simp add: fun_eq_iff Grp_def rel_stateT_def rel_fun_def stateT.map_sel intro: stateT.expand) end subsubsection \Plain monad, get, and put\ context fixes return :: "('a \ 's, 'm) return" and bind :: "('a \ 's, 'm) bind" begin primrec bind_state :: "('a, ('s, 'm) stateT) bind" where "bind_state (StateT x) f = StateT (\s. bind (x s) (\(a, s'). run_state (f a) s'))" definition return_state :: "('a, ('s, 'm) stateT) return" where "return_state x = StateT (\s. return (x, s))" definition get_state :: "('s, ('s, 'm) stateT) get" where "get_state f = StateT (\s. run_state (f s) s)" primrec put_state :: "('s, ('s, 'm) stateT) put" where "put_state s (StateT f) = StateT (\_. f s)" lemma run_put_state [simp]: "run_state (put_state s m) s' = run_state m s" by(cases m) simp lemma run_get_state [simp]: "run_state (get_state f) s = run_state (f s) s" by(simp add: get_state_def) lemma run_bind_state [simp]: "run_state (bind_state x f) s = bind (run_state x s) (\(a, s'). run_state (f a) s')" by(cases x)(simp) lemma run_return_state [simp]: "run_state (return_state x) s = return (x, s)" by(simp add: return_state_def) context assumes monad: "monad return bind" begin interpretation monad return bind by(fact monad) lemma monad_stateT [locale_witness]: "monad return_state bind_state" (is "monad ?return ?bind") proof show "?bind (?bind x f) g = ?bind x (\x. ?bind (f x) g)" for x and f g :: "'a \ ('s, 'm) stateT" by(rule stateT.expand ext)+(simp add: bind_assoc split_def) show "?bind (?return x) f = f x" for f :: "'a \ ('s, 'm) stateT" and x by(rule stateT.expand ext)+(simp add: return_bind) show "?bind x ?return = x" for x by(rule stateT.expand ext)+(simp add: bind_return) qed lemma monad_state_stateT [locale_witness]: "monad_state return_state bind_state get_state put_state" proof show "put_state s (get_state f) = put_state s (f s)" for f :: "'s \ ('s, 'm) stateT" and s :: "'s" by(rule stateT.expand)(simp add: get_state_def fun_eq_iff) show "get_state (\s. get_state (f s)) = get_state (\s. f s s)" for f :: "'s \ 's \ ('s, 'm) stateT" by(rule stateT.expand)(simp add: fun_eq_iff) show "put_state s (put_state s' m) = put_state s' m" for s s' :: 's and m :: "('s, 'm) stateT" by(rule stateT.expand)(simp add: fun_eq_iff) show "get_state (\s :: 's. put_state s m) = m" for m :: "('s, 'm) stateT" by(rule stateT.expand)(simp add: fun_eq_iff) show "get_state (\_. m) = m" for m :: "('s, 'm) stateT" by(rule stateT.expand)(simp add: fun_eq_iff) show "bind_state (get_state f) g = get_state (\s. bind_state (f s) g)" for f g by(rule stateT.expand)(simp add: fun_eq_iff) show "bind_state (put_state s m) f = put_state s (bind_state m f)" for s :: 's and m f by(rule stateT.expand)(simp add: fun_eq_iff) qed end text \ We cannot define a generic lifting operation for state like in Haskell. If we separate the monad type variable from the element type variable, then \lift\ should have type \'a 'm => (('a \ 's) 'm) stateT\, but this means that the type of results must change, which does not work for monomorphic monads. Instead, we must lift all operations individually. \lift_definition\ does not work because the monad transformer type is typically larger than the base type, but \lift_definition\ only works if the lifted type is no bigger. \ subsubsection \Failure\ context fixes fail :: "'m fail" begin definition fail_state :: "('s, 'm) stateT fail" where "fail_state = StateT (\s. fail)" lemma run_fail_state [simp]: "run_state fail_state s = fail" by(simp add: fail_state_def) lemma monad_fail_stateT [locale_witness]: assumes "monad_fail return bind fail" shows "monad_fail return_state bind_state fail_state" (is "monad_fail ?return ?bind ?fail") proof - interpret monad_fail return bind fail by(fact assms) have "?bind ?fail f = ?fail" for f by(rule stateT.expand)(simp add: fun_eq_iff fail_bind) then show ?thesis by unfold_locales qed notepad begin text \ \catch\ cannot be lifted through the state monad according to @{term monad_catch_state} because there is now way to communicate the state updates to the handler. \ fix catch :: "'m catch" assume "monad_catch return bind fail catch" then interpret monad_catch return bind fail catch . define catch_state :: "('s, 'm) stateT catch" where "catch_state m1 m2 = StateT (\s. catch (run_state m1 s) (run_state m2 s))" for m1 m2 have "monad_catch return_state bind_state fail_state catch_state" by(unfold_locales; rule stateT.expand; simp add: fun_eq_iff catch_state_def catch_return catch_fail catch_fail2 catch_assoc) end end subsubsection \Reader\ context fixes ask :: "('r, 'm) ask" begin definition ask_state :: "('r, ('s, 'm) stateT) ask" where "ask_state f = StateT (\s. ask (\r. run_state (f r) s))" lemma run_ask_state [simp]: "run_state (ask_state f) s = ask (\r. run_state (f r) s)" by(simp add: ask_state_def) lemma monad_reader_stateT [locale_witness]: assumes "monad_reader return bind ask" shows "monad_reader return_state bind_state ask_state" proof - interpret monad_reader return bind ask by(fact assms) show ?thesis proof show "ask_state (\r. ask_state (f r)) = ask_state (\r. f r r)" for f :: "'r \ 'r \ ('s, 'm) stateT" by(rule stateT.expand)(simp add: fun_eq_iff ask_ask) show "ask_state (\_. x) = x" for x by(rule stateT.expand)(simp add: fun_eq_iff ask_const) show "bind_state (ask_state f) g = ask_state (\r. bind_state (f r) g)" for f g by(rule stateT.expand)(simp add: fun_eq_iff bind_ask) show "bind_state m (\x. ask_state (f x)) = ask_state (\r. bind_state m (\x. f x r))" for m f by(rule stateT.expand)(simp add: fun_eq_iff bind_ask2 split_def) qed qed lemma monad_reader_state_stateT [locale_witness]: assumes "monad_reader return bind ask" shows "monad_reader_state return_state bind_state ask_state get_state put_state" proof - interpret monad_reader return bind ask by(fact assms) show ?thesis proof show "ask_state (\r. get_state (f r)) = get_state (\s. ask_state (\r. f r s))" for f by(rule stateT.expand)(simp add: fun_eq_iff) show "put_state m (ask_state f) = ask_state (\r. put_state m (f r))" for m f by(rule stateT.expand)(simp add: fun_eq_iff) qed qed end subsubsection \Probability\ definition altc_sample_state :: "('x \ ('b \ 'm) \ 'm) \ 'x \ ('b \ ('s, 'm) stateT) \ ('s, 'm) stateT" where "altc_sample_state altc_sample p f = StateT (\s. altc_sample p (\x. run_state (f x) s))" lemma run_altc_sample_state [simp]: "run_state (altc_sample_state altc_sample p f) s = altc_sample p (\x. run_state (f x) s)" by(simp add: altc_sample_state_def) context fixes sample :: "('p, 'm) sample" begin abbreviation sample_state :: "('p, ('s, 'm) stateT) sample" where "sample_state \ altc_sample_state sample" context assumes "monad_prob return bind sample" begin interpretation monad_prob return bind sample by(fact) lemma monad_prob_stateT [locale_witness]: "monad_prob return_state bind_state sample_state" including lifting_syntax proof note sample_parametric[transfer_rule] show "sample_state p (\_. x) = x" for p x by(rule stateT.expand)(simp add: fun_eq_iff sample_const) show "sample_state (return_pmf x) f = f x" for f x by(rule stateT.expand)(simp add: fun_eq_iff sample_return_pmf) show "sample_state (bind_pmf p f) g = sample_state p (\x. sample_state (f x) g)" for p f g by(rule stateT.expand)(simp add: fun_eq_iff sample_bind_pmf) show "sample_state p (\x. sample_state q (f x)) = sample_state q (\y. sample_state p (\x. f x y))" for p q f by(rule stateT.expand)(auto simp add: fun_eq_iff intro: sample_commute) show "bind_state (sample_state p f) g = sample_state p (\x. bind_state (f x) g)" for p f g by(rule stateT.expand)(simp add: fun_eq_iff bind_sample1) show "bind_state m (\y. sample_state p (f y)) = sample_state p (\x. bind_state m (\y. f y x))" for m p f by(rule stateT.expand)(simp add: fun_eq_iff bind_sample2 split_def) show "(rel_pmf R ===> (R ===> (=)) ===> (=)) sample_state sample_state" if [transfer_rule]: "bi_unique R" for R unfolding altc_sample_state_def by transfer_prover qed lemma monad_state_prob_stateT [locale_witness]: "monad_state_prob return_state bind_state get_state put_state sample_state" proof show "sample_state p (\x. get_state (f x)) = get_state (\s. sample_state p (\x. f x s))" for p f by(rule stateT.expand)(simp add: fun_eq_iff) qed end end subsubsection \Writer\ context fixes tell :: "('w, 'm) tell" begin definition tell_state :: "('w, ('s, 'm) stateT) tell" where "tell_state w m = StateT (\s. tell w (run_state m s))" lemma run_tell_state [simp]: "run_state (tell_state w m) s = tell w (run_state m s)" by(simp add: tell_state_def) lemma monad_writer_stateT [locale_witness]: assumes "monad_writer return bind tell" shows "monad_writer return_state bind_state tell_state" proof - interpret monad_writer return bind tell by(fact assms) show ?thesis proof show "bind_state (tell_state w m) f = tell_state w (bind_state m f)" for w m f by(rule stateT.expand)(simp add: bind_tell fun_eq_iff) qed qed end subsubsection \Binary Non-determinism\ context fixes alt :: "'m alt" begin definition alt_state :: "('s, 'm) stateT alt" where "alt_state m1 m2 = StateT (\s. alt (run_state m1 s) (run_state m2 s))" lemma run_alt_state [simp]: "run_state (alt_state m1 m2) s = alt (run_state m1 s) (run_state m2 s)" by(simp add: alt_state_def) context assumes "monad_alt return bind alt" begin interpretation monad_alt return bind alt by fact lemma monad_alt_stateT [locale_witness]: "monad_alt return_state bind_state alt_state" proof show "alt_state (alt_state m1 m2) m3 = alt_state m1 (alt_state m2 m3)" for m1 m2 m3 by(rule stateT.expand)(simp add: alt_assoc fun_eq_iff) show "bind_state (alt_state m m') f = alt_state (bind_state m f) (bind_state m' f)" for m m' f by(rule stateT.expand)(simp add: bind_alt1 fun_eq_iff) qed lemma monad_state_alt_stateT [locale_witness]: "monad_state_alt return_state bind_state get_state put_state alt_state" proof show "alt_state (get_state f) (get_state g) = get_state (\x. alt_state (f x) (g x))" for f g by(rule stateT.expand)(simp add: fun_eq_iff) show "alt_state (put_state s m) (put_state s m') = put_state s (alt_state m m')" for s m m' by(rule stateT.expand)(simp add: fun_eq_iff) qed end lemma monad_fail_alt_stateT [locale_witness]: fixes fail assumes "monad_fail_alt return bind fail alt" shows "monad_fail_alt return_state bind_state (fail_state fail) alt_state" proof - interpret monad_fail_alt return bind fail alt by fact show ?thesis proof show "alt_state (fail_state fail) m = m" for m by(rule stateT.expand)(simp add: fun_eq_iff alt_fail1) show "alt_state m (fail_state fail) = m" for m by(rule stateT.expand)(simp add: fun_eq_iff alt_fail2) qed qed end subsubsection \Countable Non-determinism\ context fixes altc :: "('c, 'm) altc" begin abbreviation altc_state :: "('c, ('s, 'm) stateT) altc" where "altc_state \ altc_sample_state altc" context includes lifting_syntax assumes "monad_altc return bind altc" begin interpretation monad_altc return bind altc by fact lemma monad_altc_stateT [locale_witness]: "monad_altc return_state bind_state altc_state" proof note altc_parametric[transfer_rule] show "bind_state (altc_state C g) f = altc_state C (\c. bind_state (g c) f)" for C g f by(rule stateT.expand)(simp add: fun_eq_iff bind_altc1) show "altc_state (csingle x) f = f x" for x f by(rule stateT.expand)(simp add: fun_eq_iff altc_single) show "altc_state (cUNION C f) g = altc_state C (\x. altc_state (f x) g)" for C f g by(rule stateT.expand)(simp add: fun_eq_iff altc_cUNION) show "(rel_cset R ===> (R ===> (=)) ===> (=)) altc_state altc_state" if [transfer_rule]: "bi_unique R" for R unfolding altc_sample_state_def by transfer_prover qed lemma monad_state_altc_stateT [locale_witness]: "monad_state_altc return_state bind_state get_state put_state altc_state" proof show "altc_state C (\c. get_state (f c)) = get_state (\s. altc_state C (\c. f c s))" for C and f :: "'c \ 's \ ('s, 'm) stateT" by(rule stateT.expand)(simp add: fun_eq_iff) show "altc_state C (\c. put_state s (f c)) = put_state s (altc_state C f)" for C s and f :: "'c \ ('s, 'm) stateT" by(rule stateT.expand)(simp add: fun_eq_iff) qed end lemma monad_altc3_stateT [locale_witness]: assumes "monad_altc3 return bind altc" shows "monad_altc3 return_state bind_state altc_state" proof - interpret monad_altc3 return bind altc by fact show ?thesis .. qed end subsubsection \Resumption\ context fixes pause :: "('o, 'i, 'm) pause" begin definition pause_state :: "('o, 'i, ('s, 'm) stateT) pause" where "pause_state out c = StateT (\s. pause out (\i. run_state (c i) s))" lemma run_pause_state [simp]: "run_state (pause_state out c) s = pause out (\i. run_state (c i) s)" by(simp add: pause_state_def) lemma monad_resumption_stateT [locale_witness]: assumes "monad_resumption return bind pause" shows "monad_resumption return_state bind_state pause_state" proof - interpret monad_resumption return bind pause by fact show ?thesis proof show "bind_state (pause_state out c) f = pause_state out (\i. bind_state (c i) f)" for out c f by(rule stateT.expand)(simp add: fun_eq_iff bind_pause) qed qed end end subsubsection \Parametricity\ context includes lifting_syntax begin lemma return_state_parametric [transfer_rule]: "((rel_prod A S ===> M) ===> A ===> rel_stateT S M) return_state return_state" unfolding return_state_def by transfer_prover lemma bind_state_parametric [transfer_rule]: "((M ===> (rel_prod A S ===> M) ===> M) ===> rel_stateT S M ===> (A ===> rel_stateT S M) ===> rel_stateT S M) bind_state bind_state" unfolding bind_state_def by transfer_prover lemma get_state_parametric [transfer_rule]: "((S ===> rel_stateT S M) ===> rel_stateT S M) get_state get_state" unfolding get_state_def by transfer_prover lemma put_state_parametric [transfer_rule]: "(S ===> rel_stateT S M ===> rel_stateT S M) put_state put_state" unfolding put_state_def by transfer_prover lemma fail_state_parametric [transfer_rule]: "(M ===> rel_stateT S M) fail_state fail_state" unfolding fail_state_def by transfer_prover lemma ask_state_parametric [transfer_rule]: "(((R ===> M) ===> M) ===> (R ===> rel_stateT S M) ===> rel_stateT S M) ask_state ask_state" unfolding ask_state_def by transfer_prover lemma altc_sample_state_parametric [transfer_rule]: "((X ===> (P ===> M) ===> M) ===> X ===> (P ===> rel_stateT S M) ===> rel_stateT S M) altc_sample_state altc_sample_state" unfolding altc_sample_state_def by transfer_prover lemma tell_state_parametric [transfer_rule]: "((W ===> M ===> M) ===> W ===> rel_stateT S M ===> rel_stateT S M) tell_state tell_state" unfolding tell_state_def by transfer_prover lemma alt_state_parametric [transfer_rule]: "((M ===> M ===> M) ===> rel_stateT S M ===> rel_stateT S M ===> rel_stateT S M) alt_state alt_state" unfolding alt_state_def by transfer_prover lemma pause_state_parametric [transfer_rule]: "((Out ===> (In ===> M) ===> M) ===> Out ===> (In ===> rel_stateT S M) ===> rel_stateT S M) pause_state pause_state" unfolding pause_state_def by transfer_prover end subsection \Writer monad transformer\ text \ We implement a simple writer monad which collects all the output in a list. It would also have been possible to use a monoid instead. The phantom type variables @{typ 'a} and @{typ 'w} are needed to avoid hidden polymorphism when overloading the monad operations for the writer monad transformer. \ datatype ('w, 'a, 'm) writerT = WriterT (run_writer: 'm) context fixes return :: "('a \ 'w list, 'm) return" and bind :: "('a \ 'w list, 'm) bind" begin definition return_writer :: "('a, ('w, 'a, 'm) writerT) return" where "return_writer x = WriterT (return (x, []))" definition bind_writer :: "('a, ('w, 'a, 'm) writerT) bind" where "bind_writer m f = WriterT (bind (run_writer m) (\(a, ws). bind (run_writer (f a)) (\(b, ws'). return (b, ws @ ws'))))" definition tell_writer :: "('w, ('w, 'a, 'm) writerT) tell" where "tell_writer w m = WriterT (bind (run_writer m) (\(a, ws). return (a, w # ws)))" lemma run_return_writer [simp]: "run_writer (return_writer x) = return (x, [])" by(simp add: return_writer_def) lemma run_bind_writer [simp]: "run_writer (bind_writer m f) = bind (run_writer m) (\(a, ws). bind (run_writer (f a)) (\(b, ws'). return (b, ws @ ws')))" by(simp add: bind_writer_def) lemma run_tell_writer [simp]: "run_writer (tell_writer w m) = bind (run_writer m) (\(a, ws). return (a, w # ws))" by(simp add: tell_writer_def) context assumes "monad return bind" begin interpretation monad return bind by fact lemma monad_writerT [locale_witness]: "monad return_writer bind_writer" proof show "bind_writer (bind_writer x f) g = bind_writer x (\y. bind_writer (f y) g)" for x f g by(rule writerT.expand)(simp add: bind_assoc split_def return_bind) show "bind_writer (return_writer x) f = f x" for x f by(rule writerT.expand)(simp add: bind_return return_bind) show "bind_writer x return_writer = x" for x by(rule writerT.expand)(simp add: bind_return return_bind) qed lemma monad_writer_writerT [locale_witness]: "monad_writer return_writer bind_writer tell_writer" proof show "bind_writer (tell_writer w m) f = tell_writer w (bind_writer m f)" for w m f by(rule writerT.expand)(simp add: bind_assoc split_def return_bind) qed end subsubsection \Failure\ context fixes fail :: "'m fail" begin definition fail_writer :: "('w, 'a, 'm) writerT fail" where "fail_writer = WriterT fail" lemma run_fail_writer [simp]: "run_writer fail_writer = fail" by(simp add: fail_writer_def) lemma monad_fail_writerT [locale_witness]: assumes "monad_fail return bind fail" shows "monad_fail return_writer bind_writer fail_writer" proof - interpret monad_fail return bind fail by fact show ?thesis proof show "bind_writer fail_writer f = fail_writer" for f by(rule writerT.expand)(simp add: fail_bind) qed qed text \ Just like for the state monad, we cannot lift @{term catch} because the output before the failure would be lost. \ subsubsection \State\ context fixes get :: "('s, 'm) get" and put :: "('s, 'm) put" begin definition get_writer :: "('s, ('w, 'a, 'm) writerT) get" where "get_writer f = WriterT (get (\s. run_writer (f s)))" definition put_writer :: "('s, ('w, 'a, 'm) writerT) put" where "put_writer s m = WriterT (put s (run_writer m))" lemma run_get_writer [simp]: "run_writer (get_writer f) = get (\s. run_writer (f s))" by(simp add: get_writer_def) lemma run_put_writer [simp]: "run_writer (put_writer s m) = put s (run_writer m)" by(simp add: put_writer_def) lemma monad_state_writerT [locale_witness]: assumes "monad_state return bind get put" shows "monad_state return_writer bind_writer get_writer put_writer" proof - interpret monad_state return bind get put by fact show ?thesis proof show "put_writer s (get_writer f) = put_writer s (f s)" for s f by(rule writerT.expand)(simp add: put_get) show "get_writer (\s. get_writer (f s)) = get_writer (\s. f s s)" for f by(rule writerT.expand)(simp add: get_get) show "put_writer s (put_writer s' m) = put_writer s' m" for s s' m by(rule writerT.expand)(simp add: put_put) show "get_writer (\s. put_writer s m) = m" for m by(rule writerT.expand)(simp add: get_put) show "get_writer (\_. m) = m" for m by(rule writerT.expand)(simp add: get_const) show "bind_writer (get_writer f) g = get_writer (\s. bind_writer (f s) g)" for f g by(rule writerT.expand)(simp add: bind_get) show "bind_writer (put_writer s m) f = put_writer s (bind_writer m f)" for s m f by(rule writerT.expand)(simp add: bind_put) qed qed subsubsection \Probability\ definition altc_sample_writer :: "('x \ ('b \ 'm) \ 'm) \ 'x \ ('b \ ('w, 'a, 'm) writerT) \ ('w, 'a, 'm) writerT" where "altc_sample_writer altc_sample p f = WriterT (altc_sample p (\p. run_writer (f p)))" lemma run_altc_sample_writer [simp]: "run_writer (altc_sample_writer altc_sample p f) = altc_sample p (\p. run_writer (f p))" by(simp add: altc_sample_writer_def) context fixes sample :: "('p, 'm) sample" begin abbreviation sample_writer :: "('p, ('w, 'a, 'm) writerT) sample" where "sample_writer \ altc_sample_writer sample" lemma monad_prob_writerT [locale_witness]: assumes "monad_prob return bind sample" shows "monad_prob return_writer bind_writer sample_writer" proof - interpret monad_prob return bind sample by fact note sample_parametric[transfer_rule] show ?thesis including lifting_syntax proof show "sample_writer p (\_. m) = m" for p m by(rule writerT.expand)(simp add: sample_const) show "sample_writer (return_pmf x) f = f x" for x f by(rule writerT.expand)(simp add: sample_return_pmf) show "sample_writer (p \ f) g = sample_writer p (\x. sample_writer (f x) g)" for p f g by(rule writerT.expand)(simp add: sample_bind_pmf) show "sample_writer p (\x. sample_writer q (f x)) = sample_writer q (\y. sample_writer p (\x. f x y))" for p q f by(rule writerT.expand)(auto intro: sample_commute) show "bind_writer (sample_writer p f) g = sample_writer p (\x. bind_writer (f x) g)" for p f g by(rule writerT.expand)(simp add: bind_sample1) show "bind_writer m (\y. sample_writer p (f y)) = sample_writer p (\x. bind_writer m (\y. f y x))" for m p f by(rule writerT.expand)(simp add: bind_sample2[symmetric] bind_sample1 split_def) show "(rel_pmf R ===> (R ===> (=)) ===> (=)) sample_writer sample_writer" if [transfer_rule]: "bi_unique R" for R unfolding altc_sample_writer_def by transfer_prover qed qed lemma monad_state_prob_writerT [locale_witness]: assumes "monad_state_prob return bind get put sample" shows "monad_state_prob return_writer bind_writer get_writer put_writer sample_writer" proof - interpret monad_state_prob return bind get put sample by fact show ?thesis proof show "sample_writer p (\x. get_writer (f x)) = get_writer (\s. sample_writer p (\x. f x s))" for p f by(rule writerT.expand)(simp add: sample_get) qed qed end subsubsection \Reader\ context fixes ask :: "('r, 'm) ask" begin definition ask_writer :: "('r, ('w, 'a, 'm) writerT) ask" where "ask_writer f = WriterT (ask (\r. run_writer (f r)))" lemma run_ask_writer [simp]: "run_writer (ask_writer f) = ask (\r. run_writer (f r))" by(simp add: ask_writer_def) lemma monad_reader_writerT [locale_witness]: assumes "monad_reader return bind ask" shows "monad_reader return_writer bind_writer ask_writer" proof - interpret monad_reader return bind ask by fact show ?thesis proof show "ask_writer (\r. ask_writer (f r)) = ask_writer (\r. f r r)" for f by(rule writerT.expand)(simp add: ask_ask) show "ask_writer (\_. m) = m" for m by(rule writerT.expand)(simp add: ask_const) show "bind_writer (ask_writer f) g = ask_writer (\r. bind_writer (f r) g)" for f g by(rule writerT.expand)(simp add: bind_ask) show "bind_writer m (\x. ask_writer (f x)) = ask_writer (\r. bind_writer m (\x. f x r))" for m f by(rule writerT.expand)(simp add: split_def bind_ask2[symmetric] bind_ask) qed qed lemma monad_reader_state_writerT [locale_witness]: assumes "monad_reader_state return bind ask get put" shows "monad_reader_state return_writer bind_writer ask_writer get_writer put_writer" proof - interpret monad_reader_state return bind ask get put by fact show ?thesis proof show "ask_writer (\r. get_writer (f r)) = get_writer (\s. ask_writer (\r. f r s))" for f by(rule writerT.expand)(simp add: ask_get) show "put_writer s (ask_writer f) = ask_writer (\r. put_writer s (f r))" for s f by(rule writerT.expand)(simp add: put_ask) qed qed end subsubsection \Resumption\ context fixes pause :: "('o, 'i, 'm) pause" begin definition pause_writer :: "('o, 'i, ('w, 'a, 'm) writerT) pause" where "pause_writer out c = WriterT (pause out (\input. run_writer (c input)))" lemma run_pause_writer [simp]: "run_writer (pause_writer out c) = pause out (\input. run_writer (c input))" by(simp add: pause_writer_def) lemma monad_resumption_writerT [locale_witness]: assumes "monad_resumption return bind pause" shows "monad_resumption return_writer bind_writer pause_writer" proof - interpret monad_resumption return bind pause by fact show ?thesis proof show "bind_writer (pause_writer out c) f = pause_writer out (\i. bind_writer (c i) f)" for out c f by(rule writerT.expand)(simp add: bind_pause) qed qed end subsubsection \Binary Non-determinism\ context fixes alt :: "'m alt" begin definition alt_writer :: "('w, 'a, 'm) writerT alt" where "alt_writer m m' = WriterT (alt (run_writer m) (run_writer m'))" lemma run_alt_writer [simp]: "run_writer (alt_writer m m') = alt (run_writer m) (run_writer m')" by(simp add: alt_writer_def) lemma monad_alt_writerT [locale_witness]: assumes "monad_alt return bind alt" shows "monad_alt return_writer bind_writer alt_writer" proof - interpret monad_alt return bind alt by fact show ?thesis proof show "alt_writer (alt_writer m1 m2) m3 = alt_writer m1 (alt_writer m2 m3)" for m1 m2 m3 by(rule writerT.expand)(simp add: alt_assoc) show "bind_writer (alt_writer m m') f = alt_writer (bind_writer m f) (bind_writer m' f)" for m m' f by(rule writerT.expand)(simp add: bind_alt1) qed qed lemma monad_fail_alt_writerT [locale_witness]: assumes "monad_fail_alt return bind fail alt" shows "monad_fail_alt return_writer bind_writer fail_writer alt_writer" proof - interpret monad_fail_alt return bind fail alt by fact show ?thesis proof show "alt_writer fail_writer m = m" for m by(rule writerT.expand)(simp add: alt_fail1) show "alt_writer m fail_writer = m" for m by(rule writerT.expand)(simp add: alt_fail2) qed qed lemma monad_state_alt_writerT [locale_witness]: assumes "monad_state_alt return bind get put alt" shows "monad_state_alt return_writer bind_writer get_writer put_writer alt_writer" proof - interpret monad_state_alt return bind get put alt by fact show ?thesis proof show "alt_writer (get_writer f) (get_writer g) = get_writer (\x. alt_writer (f x) (g x))" for f g by(rule writerT.expand)(simp add: alt_get) show "alt_writer (put_writer s m) (put_writer s m') = put_writer s (alt_writer m m')" for s m m' by(rule writerT.expand)(simp add: alt_put) qed qed end subsubsection \Countable Non-determinism\ context fixes altc :: "('c, 'm) altc" begin abbreviation altc_writer :: "('c, ('w, 'a, 'm) writerT) altc" where "altc_writer \ altc_sample_writer altc" lemma monad_altc_writerT [locale_witness]: assumes "monad_altc return bind altc" shows "monad_altc return_writer bind_writer altc_writer" proof - interpret monad_altc return bind altc by fact note altc_parametric[transfer_rule] show ?thesis including lifting_syntax proof show "bind_writer (altc_writer C g) f = altc_writer C (\c. bind_writer (g c) f)" for C g f by(rule writerT.expand)(simp add: bind_altc1 o_def) show "altc_writer (csingle x) f = f x" for x f by(rule writerT.expand)(simp add: altc_single) show "altc_writer (cUNION C f) g = altc_writer C (\x. altc_writer (f x) g)" for C f g by(rule writerT.expand)(simp add: altc_cUNION o_def) show "(rel_cset R ===> (R ===> (=)) ===> (=)) altc_writer altc_writer" if [transfer_rule]: "bi_unique R" for R unfolding altc_sample_writer_def by transfer_prover qed qed lemma monad_altc3_writerT [locale_witness]: assumes "monad_altc3 return bind altc" shows "monad_altc3 return_writer bind_writer altc_writer" proof - interpret monad_altc3 return bind altc by fact show ?thesis .. qed lemma monad_state_altc_writerT [locale_witness]: assumes "monad_state_altc return bind get put altc" shows "monad_state_altc return_writer bind_writer get_writer put_writer altc_writer" proof - interpret monad_state_altc return bind get put altc by fact show ?thesis proof show "altc_writer C (\c. get_writer (f c)) = get_writer (\s. altc_writer C (\c. f c s))" for C and f :: "'c \ 's \ ('w, 'a, 'm) writerT" by(rule writerT.expand)(simp add: o_def altc_get) show "altc_writer C (\c. put_writer s (f c)) = put_writer s (altc_writer C f)" for C s and f :: "'c \ ('w, 'a, 'm) writerT" by(rule writerT.expand)(simp add: o_def altc_put) qed qed end end end end subsubsection \Parametricity\ context includes lifting_syntax begin lemma return_writer_parametric [transfer_rule]: "((rel_prod A (list_all2 W) ===> M) ===> A ===> rel_writerT W A M) return_writer return_writer" unfolding return_writer_def by transfer_prover lemma bind_writer_parametric [transfer_rule]: "((rel_prod A (list_all2 W) ===> M) ===> (M ===> (rel_prod A (list_all2 W) ===> M) ===> M) ===> rel_writerT W A M ===> (A ===> rel_writerT W A M) ===> rel_writerT W A M) bind_writer bind_writer" unfolding bind_writer_def by transfer_prover lemma tell_writer_parametric [transfer_rule]: "((rel_prod A (list_all2 W) ===> M) ===> (M ===> (rel_prod A (list_all2 W) ===> M) ===> M) ===> W ===> rel_writerT W A M ===> rel_writerT W A M) tell_writer tell_writer" unfolding tell_writer_def by transfer_prover lemma ask_writer_parametric [transfer_rule]: "(((R ===> M) ===> M) ===> (R ===> rel_writerT W A M) ===> rel_writerT W A M) ask_writer ask_writer" unfolding ask_writer_def by transfer_prover lemma fail_writer_parametric [transfer_rule]: "(M ===> rel_writerT W A M) fail_writer fail_writer" unfolding fail_writer_def by transfer_prover lemma get_writer_parametric [transfer_rule]: "(((S ===> M) ===> M) ===> (S ===> rel_writerT W A M) ===> rel_writerT W A M) get_writer get_writer" unfolding get_writer_def by transfer_prover lemma put_writer_parametric [transfer_rule]: "((S ===> M ===> M) ===> S ===> rel_writerT W A M ===> rel_writerT W A M) put_writer put_writer" unfolding put_writer_def by transfer_prover lemma altc_sample_writer_parametric [transfer_rule]: "((X ===> (P ===> M) ===> M) ===> X ===> (P ===> rel_writerT W A M) ===> rel_writerT W A M) altc_sample_writer altc_sample_writer" unfolding altc_sample_writer_def by transfer_prover lemma alt_writer_parametric [transfer_rule]: "((M ===> M ===> M) ===> rel_writerT W A M ===> rel_writerT W A M ===> rel_writerT W A M) alt_writer alt_writer" unfolding alt_writer_def by transfer_prover lemma pause_writer_parametric [transfer_rule]: "((Out ===> (In ===> M) ===> M) ===> Out ===> (In ===> rel_writerT W A M) ===> rel_writerT W A M) pause_writer pause_writer" unfolding pause_writer_def by transfer_prover end subsection \Continuation monad transformer\ datatype ('a, 'm) contT = ContT (run_cont: "('a \ 'm) \ 'm") subsubsection \CallCC\ type_synonym ('a, 'm) callcc = "(('a \ 'm) \ 'm) \ 'm" definition callcc_cont :: "('a, ('a, 'm) contT) callcc" where "callcc_cont f = ContT (\k. run_cont (f (\x. ContT (\_. k x))) k)" lemma run_callcc_cont [simp]: "run_cont (callcc_cont f) k = run_cont (f (\x. ContT (\_. k x))) k" by(simp add: callcc_cont_def) subsubsection \Plain monad\ definition return_cont :: "('a, ('a, 'm) contT) return" where "return_cont x = ContT (\k. k x)" definition bind_cont :: "('a, ('a, 'm) contT) bind" where "bind_cont m f = ContT (\k. run_cont m (\x. run_cont (f x) k))" lemma run_return_cont [simp]: "run_cont (return_cont x) k = k x" by(simp add: return_cont_def) lemma run_bind_cont [simp]: "run_cont (bind_cont m f) k = run_cont m (\x. run_cont (f x) k)" by(simp add: bind_cont_def) lemma monad_contT [locale_witness]: "monad return_cont bind_cont" (is "monad ?return ?bind") proof show "?bind (?bind x f) g = ?bind x (\x. ?bind (f x) g)" for x f g by(rule contT.expand)(simp add: fun_eq_iff) show "?bind (?return x) f = f x" for f x by(rule contT.expand)(simp add: fun_eq_iff) show "?bind x ?return = x" for x by(rule contT.expand)(simp add: fun_eq_iff) qed subsubsection \Failure\ context fixes fail :: "'m fail" begin definition fail_cont :: "('a, 'm) contT fail" where "fail_cont = ContT (\_. fail)" lemma run_fail_cont [simp]: "run_cont fail_cont k = fail" by(simp add: fail_cont_def) lemma monad_fail_contT [locale_witness]: "monad_fail return_cont bind_cont fail_cont" proof show "bind_cont fail_cont f = fail_cont" for f :: "'a \ ('a, 'm) contT" by(rule contT.expand)(simp add: fun_eq_iff) qed end subsubsection \State\ context fixes get :: "('s, 'm) get" and put :: "('s, 'm) put" begin definition get_cont :: "('s, ('a, 'm) contT) get" where "get_cont f = ContT (\k. get (\s. run_cont (f s) k))" definition put_cont :: "('s, ('a, 'm) contT) put" where "put_cont s m = ContT (\k. put s (run_cont m k))" lemma run_get_cont [simp]: "run_cont (get_cont f) k = get (\s. run_cont (f s) k)" by(simp add: get_cont_def) lemma run_put_cont [simp]: "run_cont (put_cont s m) k = put s (run_cont m k)" by(simp add: put_cont_def) lemma monad_state_contT [locale_witness]: assumes "monad_state return' bind' get put" \ \We don't need the plain monad operations for lifting.\ shows "monad_state return_cont bind_cont get_cont (put_cont :: ('s, ('a, 'm) contT) put)" (is "monad_state ?return ?bind ?get ?put") proof - interpret monad_state return' bind' get put by(fact assms) show ?thesis proof show "put_cont s (get_cont f) = put_cont s (f s)" for s :: 's and f :: "'s \ ('a, 'm) contT" by(rule contT.expand)(simp add: put_get fun_eq_iff) show "get_cont (\s. get_cont (f s)) = get_cont (\s. f s s)" for f :: "'s \ 's \ ('a, 'm) contT" by(rule contT.expand)(simp add: get_get fun_eq_iff) show "put_cont s (put_cont s' m) = put_cont s' m" for s s' and m :: "('a, 'm) contT" by(rule contT.expand)(simp add: put_put fun_eq_iff) show "get_cont (\s. put_cont s m) = m" for m :: "('a, 'm) contT" by(rule contT.expand)(simp add: get_put fun_eq_iff) show "get_cont (\_. m) = m" for m :: "('a, 'm) contT" by(rule contT.expand)(simp add: get_const fun_eq_iff) show "bind_cont (get_cont f) g = get_cont (\s. bind_cont (f s) g)" for f :: "'s \ ('a, 'm) contT" and g by(rule contT.expand)(simp add: fun_eq_iff) show "bind_cont (put_cont s m) f = put_cont s (bind_cont m f)" for s and m :: "('a, 'm) contT" and f by(rule contT.expand)(simp add: fun_eq_iff) qed qed end section \Locales for monad homomorphisms\ locale monad_hom = m1: monad return1 bind1 + m2: monad return2 bind2 for return1 :: "('a, 'm1) return" and bind1 :: "('a, 'm1) bind" and return2 :: "('a, 'm2) return" and bind2 :: "('a, 'm2) bind" and h :: "'m1 \ 'm2" + assumes hom_return: "\x. h (return1 x) = return2 x" and hom_bind: "\x f. h (bind1 x f) = bind2 (h x) (h \ f)" begin lemma hom_lift [simp]: "h (m1.lift f m) = m2.lift f (h m)" by(simp add: m1.lift_def m2.lift_def hom_bind hom_return o_def) end locale monad_state_hom = m1: monad_state return1 bind1 get1 put1 + m2: monad_state return2 bind2 get2 put2 + monad_hom return1 bind1 return2 bind2 h for return1 :: "('a, 'm1) return" and bind1 :: "('a, 'm1) bind" and get1 :: "('s, 'm1) get" and put1 :: "('s, 'm1) put" and return2 :: "('a, 'm2) return" and bind2 :: "('a, 'm2) bind" and get2 :: "('s, 'm2) get" and put2 :: "('s, 'm2) put" and h :: "'m1 \ 'm2" + assumes hom_get [simp]: "h (get1 f) = get2 (h \ f)" and hom_put [simp]: "h (put1 s m) = put2 s (h m)" locale monad_fail_hom = m1: monad_fail return1 bind1 fail1 + m2: monad_fail return2 bind2 fail2 + monad_hom return1 bind1 return2 bind2 h for return1 :: "('a, 'm1) return" and bind1 :: "('a, 'm1) bind" and fail1 :: "'m1 fail" and return2 :: "('a, 'm2) return" and bind2 :: "('a, 'm2) bind" and fail2 :: "'m2 fail" and h :: "'m1 \ 'm2" + assumes hom_fail [simp]: "h fail1 = fail2" locale monad_catch_hom = m1: monad_catch return1 bind1 fail1 catch1 + m2: monad_catch return2 bind2 fail2 catch2 + monad_fail_hom return1 bind1 fail1 return2 bind2 fail2 h for return1 :: "('a, 'm1) return" and bind1 :: "('a, 'm1) bind" and fail1 :: "'m1 fail" and catch1 :: "'m1 catch" and return2 :: "('a, 'm2) return" and bind2 :: "('a, 'm2) bind" and fail2 :: "'m2 fail" and catch2 :: "'m2 catch" and h :: "'m1 \ 'm2" + assumes hom_catch [simp]: "h (catch1 m1 m2) = catch2 (h m1) (h m2)" locale monad_reader_hom = m1: monad_reader return1 bind1 ask1 + m2: monad_reader return2 bind2 ask2 + monad_hom return1 bind1 return2 bind2 h for return1 :: "('a, 'm1) return" and bind1 :: "('a, 'm1) bind" and ask1 :: "('r, 'm1) ask" and return2 :: "('a, 'm2) return" and bind2 :: "('a, 'm2) bind" and ask2 :: "('r, 'm2) ask" and h :: "'m1 \ 'm2" + assumes hom_ask [simp]: "h (ask1 f) = ask2 (h \ f)" locale monad_prob_hom = m1: monad_prob return1 bind1 sample1 + m2: monad_prob return2 bind2 sample2 + monad_hom return1 bind1 return2 bind2 h for return1 :: "('a, 'm1) return" and bind1 :: "('a, 'm1) bind" and sample1 :: "('p, 'm1) sample" and return2 :: "('a, 'm2) return" and bind2 :: "('a, 'm2) bind" and sample2 :: "('p, 'm2) sample" and h :: "'m1 \ 'm2" + assumes hom_sample [simp]: "h (sample1 p f) = sample2 p (h \ f)" locale monad_alt_hom = m1: monad_alt return1 bind1 alt1 + m2: monad_alt return2 bind2 alt2 + monad_hom return1 bind1 return2 bind2 h for return1 :: "('a, 'm1) return" and bind1 :: "('a, 'm1) bind" and alt1 :: "'m1 alt" and return2 :: "('a, 'm2) return" and bind2 :: "('a, 'm2) bind" and alt2 :: "'m2 alt" and h :: "'m1 \ 'm2" + assumes hom_alt [simp]: "h (alt1 m m') = alt2 (h m) (h m')" locale monad_altc_hom = m1: monad_altc return1 bind1 altc1 + m2: monad_altc return2 bind2 altc2 + monad_hom return1 bind1 return2 bind2 h for return1 :: "('a, 'm1) return" and bind1 :: "('a, 'm1) bind" and altc1 :: "('c, 'm1) altc" and return2 :: "('a, 'm2) return" and bind2 :: "('a, 'm2) bind" and altc2 :: "('c, 'm2) altc" and h :: "'m1 \ 'm2" + assumes hom_altc [simp]: "h (altc1 C f) = altc2 C (h \ f)" locale monad_writer_hom = m1: monad_writer return1 bind1 tell1 + m2: monad_writer return2 bind2 tell2 + monad_hom return1 bind1 return2 bind2 h for return1 :: "('a, 'm1) return" and bind1 :: "('a, 'm1) bind" and tell1 :: "('w, 'm1) tell" and return2 :: "('a, 'm2) return" and bind2 :: "('a, 'm2) bind" and tell2 :: "('w, 'm2) tell" and h :: "'m1 \ 'm2" + assumes hom_tell [simp]: "h (tell1 w m) = tell2 w (h m)" locale monad_resumption_hom = m1: monad_resumption return1 bind1 pause1 + m2: monad_resumption return2 bind2 pause2 + monad_hom return1 bind1 return2 bind2 h for return1 :: "('a, 'm1) return" and bind1 :: "('a, 'm1) bind" and pause1 :: "('o, 'i, 'm1) pause" and return2 :: "('a, 'm2) return" and bind2 :: "('a, 'm2) bind" and pause2 :: "('o, 'i, 'm2) pause" and h :: "'m1 \ 'm2" + assumes hom_pause [simp]: "h (pause1 out c) = pause2 out (h \ c)" section \Switching between monads\ text \Homomorphisms are functional relations between monads. In general, it is more convenient to use arbitrary relations as embeddings because arbitrary relations allow us to change the type of values in a monad. As different monad transformers change the value type in different ways, the embeddings must also support such changes in values. \ context includes lifting_syntax begin subsection \Embedding Identity into Probability\ named_theorems cr_id_prob_transfer definition prob_of_id :: "'a id \ 'a prob" where "prob_of_id m = return_pmf (extract m)" lemma monad_id_prob_hom [locale_witness]: "monad_hom return_id bind_id return_pmf bind_pmf prob_of_id" proof show "prob_of_id (return_id x) = return_pmf x" for x :: 'a by(simp add: prob_of_id_def) show "prob_of_id (bind_id x f) = prob_of_id x \ prob_of_id \ f" for x :: "'a id" and f by(simp add: prob_of_id_def bind_return_pmf) qed inductive cr_id_prob :: "('a \ 'b \ bool) \ 'a id \ 'b prob \ bool" for A where "A x y \ cr_id_prob A (return_id x) (return_pmf y)" inductive_simps cr_id_prob_simps [simp]: "cr_id_prob A (return_id x) (return_pmf y)" lemma cr_id_prob_return [cr_id_prob_transfer]: "(A ===> cr_id_prob A) return_id return_pmf" by(simp add: rel_fun_def) lemma cr_id_prob_bind [cr_id_prob_transfer]: "(cr_id_prob A ===> (A ===> cr_id_prob B) ===> cr_id_prob B) bind_id bind_pmf" by(auto simp add: rel_fun_def bind_return_pmf elim!: cr_id_prob.cases) lemma cr_id_prob_Grp: "cr_id_prob (BNF_Def.Grp A f) = BNF_Def.Grp {x. set_id x \ A} (return_pmf \ f \ extract)" apply(auto simp add: Grp_def fun_eq_iff simp add: cr_id_prob.simps intro: id.expand) subgoal for x by(cases x) auto done subsection \State and Reader\ text \When no state updates are needed, the operation @{term get} can be replaced by @{term ask}.\ named_theorems cr_envT_stateT_transfer definition cr_prod1 :: "'c \ ('a \ 'b \ bool) \ 'a \ 'b \ 'c \ bool" where "cr_prod1 c' A = (\a (b, c). A a b \ c' = c)" lemma cr_prod1_simps [simp]: "cr_prod1 c' A a (b, c) \ A a b \ c' = c" by(simp add: cr_prod1_def) lemma cr_prod1I: "A a b \ cr_prod1 c' A a (b, c')" by simp lemma cr_prod1_Pair_transfer [cr_envT_stateT_transfer]: "(A ===> eq_onp ((=) c) ===> cr_prod1 c A) (\a _. a) Pair" by(auto simp add: rel_fun_def eq_onp_def) lemma cr_prod1_fst_transfer [cr_envT_stateT_transfer]: "(cr_prod1 c A ===> A) (\a. a) fst" by(auto simp add: rel_fun_def) lemma cr_prod1_case_prod_transfer [cr_envT_stateT_transfer]: "((A ===> eq_onp ((=) c) ===> C) ===> cr_prod1 c A ===> C) (\f a. f a c) case_prod" by(simp add: rel_fun_def eq_onp_def) lemma cr_prod1_Grp: "cr_prod1 c (BNF_Def.Grp A f) = BNF_Def.Grp A (\b. (f b, c))" by(auto simp add: Grp_def fun_eq_iff) definition cr_envT_stateT :: "'s \ ('m1 \ 'm2 \ bool) \ ('s, 'm1) envT \ ('s, 'm2) stateT \ bool" where "cr_envT_stateT s M m1 m2 = (eq_onp ((=) s) ===> M) (run_env m1) (run_state m2)" lemma cr_envT_stateT_simps [simp]: "cr_envT_stateT s M (EnvT f) (StateT g) \ M (f s) (g s)" by(simp add: cr_envT_stateT_def rel_fun_def eq_onp_def) lemma cr_envT_stateTE: assumes "cr_envT_stateT s M m1 m2" obtains f g where "m1 = EnvT f" "m2 = StateT g" "(eq_onp ((=) s) ===> M) f g" using assms by(cases m1; cases m2; auto simp add: eq_onp_def) lemma cr_envT_stateTD: "cr_envT_stateT s M m1 m2 \ M (run_env m1 s) (run_state m2 s)" by(auto elim!: cr_envT_stateTE dest: rel_funD simp add: eq_onp_def) lemma cr_envT_stateT_run [cr_envT_stateT_transfer]: "(cr_envT_stateT s M ===> eq_onp ((=) s) ===> M) run_env run_state" by(rule rel_funI)(auto elim!: cr_envT_stateTE) lemma cr_envT_stateT_StateT_EnvT [cr_envT_stateT_transfer]: "((eq_onp ((=) s) ===> M) ===> cr_envT_stateT s M) EnvT StateT" by(auto 4 3 dest: rel_funD simp add: eq_onp_def) lemma cr_envT_stateT_rec [cr_envT_stateT_transfer]: "(((eq_onp ((=) s) ===> M) ===> C) ===> cr_envT_stateT s M ===> C) rec_envT rec_stateT" by(auto simp add: rel_fun_def elim!: cr_envT_stateTE) lemma cr_envT_stateT_return [cr_envT_stateT_transfer]: notes [transfer_rule] = cr_envT_stateT_transfer shows "((cr_prod1 s A ===> M) ===> A ===> cr_envT_stateT s M) return_env return_state" unfolding return_env_def return_state_def by transfer_prover lemma cr_envT_stateT_bind [cr_envT_stateT_transfer]: "((M ===> (cr_prod1 s A ===> M) ===> M) ===> cr_envT_stateT s M ===> (A ===> cr_envT_stateT s M) ===> cr_envT_stateT s M) bind_env bind_state" apply(rule rel_funI)+ apply(erule cr_envT_stateTE) apply(clarsimp simp add: split_def) apply(drule rel_funD) apply(erule rel_funD) apply(simp add: eq_onp_def) apply(erule rel_funD) apply(rule rel_funI) apply clarsimp apply(rule cr_envT_stateT_run[THEN rel_funD, THEN rel_funD, where B=M]) apply(erule (1) rel_funD) apply(simp add: eq_onp_def) done lemma cr_envT_stateT_ask_get [cr_envT_stateT_transfer]: "((eq_onp ((=) s) ===> cr_envT_stateT s M) ===> cr_envT_stateT s M) ask_env get_state" unfolding ask_env_def get_state_def apply(rule rel_funI)+ apply simp apply(rule cr_envT_stateT_run[THEN rel_funD, THEN rel_funD]) apply(erule rel_funD) apply(simp_all add: eq_onp_def) done lemma cr_envT_stateT_fail [cr_envT_stateT_transfer]: notes [transfer_rule] = cr_envT_stateT_transfer shows "(M ===> cr_envT_stateT s M) fail_env fail_state" unfolding fail_env_def fail_state_def by transfer_prover subsection \@{typ "_ spmf"} and @{typ "(_, _ prob) optionT"}\ text \ This section defines the mapping between the @{typ "_ spmf"} monad and the monad obtained by composing transforming @{typ "_ prob"} with @{typ "(_, _) optionT"}. \ definition cr_spmf_prob_optionT :: "('a \ 'b \ bool) \ ('a, 'a option prob) optionT \ 'b spmf \ bool" where "cr_spmf_prob_optionT A p q \ rel_spmf A (run_option p) q" lemma cr_spmf_prob_optionTI: "rel_spmf A (run_option p) q \ cr_spmf_prob_optionT A p q" by(simp add: cr_spmf_prob_optionT_def) lemma cr_spmf_prob_optionTD: "cr_spmf_prob_optionT A p q \ rel_spmf A (run_option p) q" by(simp add: cr_spmf_prob_optionT_def) lemma cr_spmf_prob_optionT_return_transfer: \ \Cannot be used as a transfer rule in @{method transfer_prover} because @{term return_spmf} is not a constant.\ "(A ===> cr_spmf_prob_optionT A) (return_option return_pmf) return_spmf" by(simp add: rel_fun_def cr_spmf_prob_optionTI) lemma cr_spmf_prob_optionT_bind_transfer: "(cr_spmf_prob_optionT A ===> (A ===> cr_spmf_prob_optionT A) ===> cr_spmf_prob_optionT A) (bind_option return_pmf bind_pmf) bind_spmf" by(rule rel_funI cr_spmf_prob_optionTI)+ (auto 4 4 simp add: run_bind_option bind_spmf_def dest!: cr_spmf_prob_optionTD dest: rel_funD intro: rel_pmf_bindI split: option.split) lemma cr_spmf_prob_optionT_fail_transfer: "cr_spmf_prob_optionT A (fail_option return_pmf) (return_pmf None)" by(rule cr_spmf_prob_optionTI) simp abbreviation (input) spmf_of_prob_optionT :: "('a, 'a option prob) optionT \ 'a spmf" where "spmf_of_prob_optionT \ run_option" abbreviation (input) prob_optionT_of_spmf :: "'a spmf \ ('a, 'a option prob) optionT" where "prob_optionT_of_spmf \ OptionT" lemma spmf_of_prob_optionT_transfer: "(cr_spmf_prob_optionT A ===> rel_spmf A) spmf_of_prob_optionT (\x. x)" by(auto simp add: rel_fun_def dest: cr_spmf_prob_optionTD) lemma prob_optionT_of_spmf_transfer: "(rel_spmf A ===> cr_spmf_prob_optionT A) prob_optionT_of_spmf (\x. x)" by(auto simp add: rel_fun_def intro: cr_spmf_prob_optionTI) subsection \Probabilities and countable non-determinism\ named_theorems cr_prob_ndi_transfer context includes cset.lifting begin interpretation cset_nondetM return_id bind_id merge_id merge_id .. lift_definition cset_pmf :: "'a pmf \ 'a cset" is set_pmf by simp inductive cr_pmf_cset :: "'a pmf \ 'a cset \ bool" for p where "cr_pmf_cset p (cset_pmf p)" lemma cr_pmf_cset_Grp: "cr_pmf_cset = BNF_Def.Grp UNIV cset_pmf" by(simp add: fun_eq_iff cr_pmf_cset.simps Grp_def) lemma cr_pmf_cset_return_pmf [cr_prob_ndi_transfer]: "((=) ===> cr_pmf_cset) return_pmf csingle" by(simp add: cr_pmf_cset.simps rel_fun_def)(transfer; simp) inductive cr_prob_ndi :: "('a \ 'b \ bool) \ 'a prob \ ('b, 'b cset id) nondetT \ bool" for A p B where "cr_prob_ndi A p B" if "rel_set A (set_pmf p) (rcset (extract (run_nondet B)))" lemma cr_prob_ndi_Grp: "cr_prob_ndi (BNF_Def.Grp UNIV f) = BNF_Def.Grp UNIV (NondetT \ return_id \ cimage f \ cset_pmf)" by(simp add: fun_eq_iff cr_prob_ndi.simps rel_set_Grp) (auto simp add: Grp_def cimage.rep_eq cset_pmf.rep_eq cin.rep_eq intro!: nondetT.expand id.expand) lemma cr_ndi_prob_return [cr_prob_ndi_transfer]: "(A ===> cr_prob_ndi A) return_pmf return_nondet" by(simp add: rel_fun_def cr_prob_ndi.simps)(transfer; simp add: rel_set_def) lemma cr_ndi_prob_bind [cr_prob_ndi_transfer]: "(cr_prob_ndi A ===> (A ===> cr_prob_ndi A) ===> cr_prob_ndi A) bind_pmf bind_nondet" apply (clarsimp simp add: cr_prob_ndi.simps cUnion.rep_eq cimage.rep_eq intro!: rel_funI) apply(rule Union_transfer[THEN rel_funD]) apply(rule image_transfer[THEN rel_funD, THEN rel_funD]) apply(rule rel_funI) apply(drule (1) rel_funD) apply(erule cr_prob_ndi.cases) apply assumption+ done lemma cr_ndi_prob_sample [cr_prob_ndi_transfer]: "(cr_pmf_cset ===> ((=) ===> cr_prob_ndi A) ===> cr_prob_ndi A) bind_pmf altc_nondet" apply(clarsimp intro!: rel_funI simp add: cr_pmf_cset.simps cr_prob_ndi.simps cUnion.rep_eq cimage.rep_eq cset_pmf.rep_eq) apply(rule Union_transfer[THEN rel_funD]) apply(rule image_transfer[THEN rel_funD, THEN rel_funD]) apply(rule rel_funI) apply(drule (1) rel_funD) apply(erule cr_prob_ndi.cases) apply assumption apply(simp add: rel_set_eq) done end end end diff --git a/thys/Regular_Tree_Relations/RR2_Infinite.thy b/thys/Regular_Tree_Relations/RR2_Infinite.thy --- a/thys/Regular_Tree_Relations/RR2_Infinite.thy +++ b/thys/Regular_Tree_Relations/RR2_Infinite.thy @@ -1,515 +1,516 @@ theory RR2_Infinite imports RRn_Automata Tree_Automata_Pumping begin lemma map_ta_rule_id [simp]: "map_ta_rule f id r = (r_root r) (map f (r_lhs_states r)) \ (f (r_rhs r))" for f r by (simp add: ta_rule.expand ta_rule.map_sel(1 - 3)) (* Finitness/Infinitness facts *) lemma no_upper_bound_infinite: assumes "\(n::nat). \t \ S. n < f t" shows "infinite S" proof (rule ccontr, simp) assume "finite S" then obtain n where "n = Max (f ` S)" "\ t \ S. f t \ n" by auto then show False using assms linorder_not_le by blast qed lemma set_constr_finite: assumes "finite F" shows "finite {h x | x. x \ F \ P x}" using assms by (induct) auto lemma bounded_depth_finite: assumes fin_F: "finite \" and "\ (funas_term ` S) \ \" and "\t \ S. depth t \ n" and "\t \ S. ground t" shows "finite S" using assms(2-) proof (induction n arbitrary: S) case 0 {fix t assume elem: "t \ S" from 0 have "depth t = 0" "ground t" "funas_term t \ \" using elem by auto then have "\ f. (f, 0) \ \ \ t = Fun f []" by (cases t rule: depth.cases) auto} then have "S \ {Fun f [] |f . (f, 0) \ \}" by (auto simp add: image_iff) from finite_subset[OF this] show ?case using set_constr_finite[OF fin_F, of "\ (f, n). Fun f []" "\ x. snd x = 0"] by auto next case (Suc n) from Suc obtain S' where S: "S' = {t :: ('a, 'b) term . ground t \ funas_term t \ \ \ depth t \ n}" "finite S'" by (auto simp add: SUP_le_iff) then obtain L F where L: "set L = S'" "set F = \" using fin_F by (meson finite_list) let ?Sn = "{Fun f ts | f ts. (f, length ts) \ \ \ set ts \ S'}" let ?Ln = "concat (map (\ (f, n). map (\ ts. Fun f ts) (List.n_lists n L)) F)" {fix t assume elem: "t \ S" from Suc have "depth t \ Suc n" "ground t" "funas_term t \ \" using elem by auto then have "funas_term t \ \ \ (\ x \ set (args t). depth x \ n) \ ground t" by (cases t rule: depth.cases) auto then have "t \ ?Sn \ S'" using S by (cases t) auto} note sub = this {fix t assume elem: "t \ ?Sn" then obtain f ts where [simp]: "t = Fun f ts" and invar: "(f, length ts) \ \" "set ts \ S'" by blast then have "Fun f ts \ set (map (\ ts. Fun f ts) (List.n_lists (length ts) L))" using L(1) by (auto simp: image_iff set_n_lists) then have "t \ set ?Ln" using invar(1) L(2) by auto} from this sub have sub: "?Sn \ set ?Ln" "S \ ?Sn \ S'" by blast+ from finite_subset[OF sub(1)] finite_subset[OF sub(2)] finite_UnI[of ?Sn, OF _ S(2)] show ?case by blast qed lemma infinite_imageD: "infinite (f ` S) \ inj_on f S \ infinite S" by blast lemma infinite_imageD2: "infinite (f ` S) \ inj f \ infinite S" by blast lemma infinite_inj_image_infinite: assumes "infinite S" and "inj_on f S" shows "infinite (f ` S)" using assms finite_image_iff by blast (*The following lemma tells us that number of terms greater than a certain depth are infinite*) lemma infinte_no_depth_limit: assumes "infinite S" and "finite \" and "\t \ S. funas_term t \ \" and "\t \ S. ground t" shows "\(n::nat). \t \ S. n < (depth t)" proof(rule allI, rule ccontr) fix n::nat assume "\ (\t \ S. (depth t) > n)" hence "\t \ S. depth t \ n" by auto from bounded_depth_finite[OF assms(2) _ this] show False using assms by auto qed lemma depth_gterm_conv: "depth (term_of_gterm t) = depth (term_of_gterm t)" by (metis leD nat_neq_iff poss_gposs_conv poss_length_bounded_by_depth poss_length_depth) lemma funs_term_ctxt [simp]: "funs_term C\s\ = funs_ctxt C \ funs_term s" by (induct C) auto lemma pigeonhole_ta_infinit_terms: fixes t ::"'f gterm" and \ :: "('q, 'f) ta" defines "t' \ term_of_gterm t :: ('f, 'q) term" assumes "fcard (\ \) < depth t'" and "q |\| gta_der \ t" and "P (funas_gterm t)" shows "infinite {t . q |\| gta_der \ t \ P (funas_gterm t)}" proof - from pigeonhole_tree_automata[OF _ assms(3)[unfolded gta_der_def]] assms(2,4) obtain C C2 s v p where ctxt: "C2 \ \" "C\s\ = t'" "C2\v\ = s" and loop: "p |\| ta_der \ v" "p |\| ta_der \ C2\Var p\" "q |\| ta_der \ C\Var p\" unfolding assms(1) by auto let ?terms = "\ n. C\(C2 ^n)\v\\" let ?inner = "\ n. (C2 ^n)\v\" have gr: "ground_ctxt C2" "ground_ctxt C" "ground v" using arg_cong[OF ctxt(2), of ground] unfolding ctxt(3)[symmetric] assms(1) by fastforce+ moreover have funas: "funas_term (?terms (Suc n)) = funas_term t'" for n unfolding ctxt(2, 3)[symmetric] using ctxt_comp_n_pres_funas by auto moreover have der: "q |\| ta_der \ (?terms (Suc n))" for n using loop by (meson ta_der_ctxt ta_der_ctxt_n_loop) moreover have "n < depth (?terms (Suc n))" for n by (meson ctxt(1) ctxt_comp_n_lower_bound depth_ctxt_less_eq less_le_trans) ultimately have "q |\| ta_der \ (?terms (Suc n)) \ ground (?terms (Suc n)) \ P (funas_term (?terms (Suc n))) \ n < depth (?terms (Suc n))" for n using assms(4) by (auto simp: assms(1) funas_term_of_gterm_conv) then have inf: "infinite {t. q |\| ta_der \ t \ ground t \ P (funas_term t)}" by (intro no_upper_bound_infinite[of _ depth]) blast have inj: "inj_on gterm_of_term {t. q |\| ta_der \ t \ ground t \ P (funas_term t)}" by (intro gterm_of_term_inj) simp show ?thesis by (intro infinite_super[OF _ infinite_inj_image_infinite[OF inf inj]]) (auto simp: image_def gta_der_def funas_gterm_gterm_of_term) qed lemma gterm_to_None_Some_funas [simp]: "funas_gterm (gterm_to_None_Some t) \ (\ (f, n). ((None, Some f), n)) ` \ \ funas_gterm t \ \" by (induct t) (auto simp: funas_gterm_def, blast) lemma funas_gterm_bot_some_decomp: assumes "funas_gterm s \ (\ (f, n). ((None, Some f), n)) ` \" shows "\ t. gterm_to_None_Some t = s \ funas_gterm t \ \" using assms proof (induct s) case (GFun f ts) from GFun(1)[OF nth_mem] obtain ss where l: "length ss = length ts \ (\i s i. gterm_to_None_Some s = ts ! i"] GFun(2-) by (auto simp: funas_gterm_def) (meson UN_subset_iff nth_mem) then have "i < length ss \ funas_gterm (ss ! i) \ \" for i using GFun(2) by (auto simp: UN_subset_iff) (smt (z3) gterm_to_None_Some_funas nth_mem subsetD) then show ?case using GFun(2-) l by (cases f) (force simp: map_nth_eq_conv UN_subset_iff dest!: in_set_idx intro!: exI[of _ "GFun (the (snd f)) ss"]) qed (* Definition INF, Q_infinity and automata construction *) definition "Inf_branching_terms \ \ = {t . infinite {u. (t, u) \ \ \ funas_gterm u \ fset \} \ funas_gterm t \ fset \}" definition "Q_infty \ \ = {|q | q. infinite {t | t. funas_gterm t \ fset \ \ q |\| ta_der \ (term_of_gterm (gterm_to_None_Some t))}|}" lemma Q_infty_fmember: "q |\| Q_infty \ \ \ infinite {t | t. funas_gterm t \ fset \ \ q |\| ta_der \ (term_of_gterm (gterm_to_None_Some t))}" proof - have "{q | q. infinite {t | t. funas_gterm t \ fset \ \ q |\| ta_der \ (term_of_gterm (gterm_to_None_Some t))}} \ fset (\ \)" using not_finite_existsD by fastforce from finite_subset[OF this] show ?thesis by (auto simp: Q_infty_def) qed abbreviation q_inf_dash_intro_rules where "q_inf_dash_intro_rules Q r \ if (r_rhs r) |\| Q \ fst (r_root r) = None then {|(r_root r) (map CInl (r_lhs_states r)) \ CInr (r_rhs r)|} else {||}" abbreviation args :: "'a list \ nat \ ('a + 'a) list" where "args \ \ qs i. map CInl (take i qs) @ CInr (qs ! i) # map CInl (drop (Suc i) qs)" abbreviation q_inf_dash_closure_rules :: "('q, 'f) ta_rule \ ('q + 'q, 'f) ta_rule list" where "q_inf_dash_closure_rules r \ (let (f, qs, q) = (r_root r, r_lhs_states r, r_rhs r) in (map (\ i. f (args qs i) \ CInr q) [0 ..< length qs]))" definition Inf_automata :: "('q, 'f option \ 'f option) ta \ 'q fset \ ('q + 'q, 'f option \ 'f option) ta" where "Inf_automata \ Q = TA (( |\| (q_inf_dash_intro_rules Q |`| rules \)) |\| ( |\| ((fset_of_list \ q_inf_dash_closure_rules) |`| rules \)) |\| map_ta_rule CInl id |`| rules \) (map_both Inl |`| eps \ |\| map_both CInr |`| eps \)" definition Inf_reg where "Inf_reg \ Q = Reg (CInr |`| fin \) (Inf_automata (ta \) Q)" lemma Inr_Inl_rel_comp: "map_both CInr |`| S |O| map_both CInl |`| S = {||}" by auto lemmas eps_split = ftrancl_Un2_separatorE[OF Inr_Inl_rel_comp] lemma Inf_automata_eps_simp [simp]: shows "(map_both Inl |`| eps \ |\| map_both CInr |`| eps \)|\<^sup>+| = (map_both CInl |`| eps \)|\<^sup>+| |\| (map_both CInr |`| eps \)|\<^sup>+|" proof - {fix x y z assume "(x, y) |\| (map_both CInl |`| eps \)|\<^sup>+|" "(y, z) |\| (map_both CInr |`| eps \)|\<^sup>+|" then have False by (metis Inl_Inr_False eps_statesI(1, 2) eps_states_image fimageE ftranclD ftranclD2)} then show ?thesis by (auto simp: Inf_automata_def eps_split) qed lemma map_both_CInl_ftrancl_conv: "(map_both CInl |`| eps \)|\<^sup>+| = map_both CInl |`| (eps \)|\<^sup>+|" by (intro ftrancl_map_both_fsubset) (auto simp: finj_CInl_CInr) lemma map_both_CInr_ftrancl_conv: "(map_both CInr |`| eps \)|\<^sup>+| = map_both CInr |`| (eps \)|\<^sup>+|" by (intro ftrancl_map_both_fsubset) (auto simp: finj_CInl_CInr) lemmas map_both_ftrancl_conv = map_both_CInl_ftrancl_conv map_both_CInr_ftrancl_conv lemma Inf_automata_Inl_to_eps [simp]: "(CInl p, CInl q) |\| (map_both CInl |`| eps \)|\<^sup>+| \ (p, q) |\| (eps \)|\<^sup>+|" "(CInr p, CInr q) |\| (map_both CInr |`| eps \)|\<^sup>+| \ (p, q) |\| (eps \)|\<^sup>+|" "(CInl q, CInl p) |\| (map_both CInr |`| eps \)|\<^sup>+| \ False" "(CInr q, CInr p) |\| (map_both CInl |`| eps \)|\<^sup>+| \ False" by (auto simp: map_both_ftrancl_conv dest: fmap_prod_fimageI) lemma Inl_eps_Inr: "(CInl q, CInl p) |\| (eps (Inf_automata \ Q))|\<^sup>+| \ (CInr q, CInr p) |\| (eps (Inf_automata \ Q))|\<^sup>+|" by (auto simp: Inf_automata_def) lemma Inr_rhs_eps_Inr_lhs: assumes "(q, CInr p) |\| (eps (Inf_automata \ Q))|\<^sup>+|" obtains q' where "q = CInr q'" using assms ftrancl_map_both_fsubset[OF finj_CInl_CInr(1)] by (cases q) (auto simp: Inf_automata_def map_both_ftrancl_conv) lemma Inl_rhs_eps_Inl_lhs: assumes "(q, CInl p) |\| (eps (Inf_automata \ Q))|\<^sup>+|" obtains q' where "q = CInl q'" using assms by (cases q) (auto simp: Inf_automata_def map_both_ftrancl_conv) lemma Inf_automata_eps [simp]: "(CInl q, CInr p) |\| (eps (Inf_automata \ Q))|\<^sup>+| \ False" "(CInr q, CInl p) |\| (eps (Inf_automata \ Q))|\<^sup>+| \ False" by (auto elim: Inr_rhs_eps_Inr_lhs Inl_rhs_eps_Inl_lhs) lemma Inl_A_res_Inf_automata: "ta_der (fmap_states_ta CInl \) t |\| ta_der (Inf_automata \ Q) t" proof (rule ta_der_mono) show "rules (fmap_states_ta CInl \) |\| rules (Inf_automata \ Q)" apply (rule fsubsetI) - apply (simp add: Inf_automata_def fmap_states_ta_def') - by force + by (auto simp: Inf_automata_def fmap_states_ta_def' image_iff Bex_def) next show "eps (fmap_states_ta CInl \) |\| eps (Inf_automata \ Q)" by (rule fsubsetI) (simp add: Inf_automata_def fmap_states_ta_def') qed lemma Inl_res_A_res_Inf_automata: "CInl |`| ta_der \ (term_of_gterm t) |\| ta_der (Inf_automata \ Q) (term_of_gterm t)" by (intro fsubset_trans[OF ta_der_fmap_states_ta_mono[of CInl \ t]]) (auto simp: Inl_A_res_Inf_automata) lemma r_rhs_CInl_args_A_rule: assumes "f qs \ CInl q |\| rules (Inf_automata \ Q)" obtains qs' where "qs = map CInl qs'" "f qs' \ q |\| rules \" using assms by (auto simp: Inf_automata_def split!: if_splits) lemma A_rule_to_dash_closure: assumes "f qs \ q |\| rules \" and "i < length qs" shows "f (args qs i) \ CInr q |\| rules (Inf_automata \ Q)" using assms by (auto simp add: Inf_automata_def fimage_iff fBall_def upt_fset intro!: fBexI[OF _ assms(1)]) lemma Inf_automata_reach_to_dash_reach: assumes "CInl p |\| ta_der (Inf_automata \ Q) C\Var (CInl q)\" shows "CInr p |\| ta_der (Inf_automata \ Q) C\Var (CInr q)\" (is "_ |\| ta_der ?A _") using assms proof (induct C arbitrary: p) case (More f ss C ts) from More(2) obtain qs q' where rule: "f qs \ q' |\| rules ?A" "length qs = Suc (length ss + length ts)" and eps: "q' = CInl p \ (q', CInl p) |\| (eps ?A)|\<^sup>+|" and reach: "\ i < Suc (length ss + length ts). qs ! i |\| ta_der ?A ((ss @ C\Var (CInl q)\ # ts) ! i)" by auto from eps obtain q'' where [simp]: "q' = CInl q''" by (cases q') (auto simp add: Inf_automata_def eps_split elim: ftranclE converse_ftranclE) from rule obtain qs' where args: "qs = map CInl qs'" "f qs' \ q'' |\| rules \" using r_rhs_CInl_args_A_rule by (metis \q' = CInl q''\) then have "CInl (qs' ! length ss) |\| ta_der (Inf_automata \ Q) C\Var (CInl q)\" using reach by (auto simp: all_Suc_conv nth_append_Cons) (metis length_map less_add_Suc1 local.rule(2) nth_append_length nth_map reach) from More(1)[OF this] More(2) show ?case using rule args eps reach A_rule_to_dash_closure[OF args(2), of "length ss" Q] by (auto simp: Inl_eps_Inr id_take_nth_drop all_Suc_conv intro!: exI[of _ "CInr q''"] exI[of _ "map CInl (take (length ss) qs') @ CInr (qs' ! length ss) # map CInl (drop (Suc (length ss)) qs')"]) (auto simp: nth_append_Cons min_def) qed (auto simp: Inf_automata_def) lemma Inf_automata_dashI: assumes "run \ r (gterm_to_None_Some t)" and "ex_rule_state r |\| Q" shows "CInr (ex_rule_state r) |\| gta_der (Inf_automata \ Q) (gterm_to_None_Some t)" proof (cases t) case [simp]: (GFun f ts) from run_root_rule[OF assms(1)] run_argsD[OF assms(1)] have rule: "TA_rule (None, Some f) (map ex_comp_state (gargs r)) (ex_rule_state r) |\| rules \" "length (gargs r) = length ts" and reach: "\ i < length ts. ex_comp_state (gargs r ! i) |\| ta_der \ (term_of_gterm (gterm_to_None_Some (ts ! i)))" by (auto intro!: run_to_comp_st_gta_der[unfolded gta_der_def comp_def]) from rule assms(2) have "(None, Some f) (map (CInl \ ex_comp_state) (gargs r)) \ CInr (ex_rule_state r) |\| rules (Inf_automata \ Q)" - by (auto simp: Inf_automata_def) force + apply (simp add: Inf_automata_def image_iff bex_Un) + apply (rule disjI1) + by force then show ?thesis using reach rule Inl_res_A_res_Inf_automata[of \ "gterm_to_None_Some (ts ! i)" Q for i] by (auto simp: gta_der_def intro!: exI[of _ "CInr (ex_rule_state r)"] exI[of _ "map (CInl \ ex_comp_state) (gargs r)"]) blast qed lemma Inf_automata_dash_reach_to_reach: assumes "p |\| ta_der (Inf_automata \ Q) t" (is "_ |\| ta_der ?A _") shows "remove_sum p |\| ta_der \ (map_vars_term remove_sum t)" using assms proof (induct t arbitrary: p) case (Var x) then show ?case by (cases p; cases x) (auto simp: Inf_automata_def ftrancl_map_both map_both_ftrancl_conv) next case (Fun f ss) from Fun(2) obtain qs q' where rule: "f qs \ q' |\| rules ?A" "length qs = length ss" and eps: "q' = p \ (q', p) |\| (eps ?A)|\<^sup>+|" and reach: "\ i < length ss. qs ! i |\| ta_der ?A (ss ! i)" by auto from rule have r: "f (map (remove_sum) qs) \ (remove_sum q') |\| rules \" by (auto simp: comp_def Inf_automata_def min_def id_take_nth_drop[symmetric] upt_fset simp flip: drop_map take_map split!: if_splits) moreover have "remove_sum q' = remove_sum p \ (remove_sum q', remove_sum p) |\| (eps \)|\<^sup>+|" using eps by (cases "is_Inl q'"; cases "is_Inl p") (auto elim!: is_InlE is_InrE, auto simp: Inf_automata_def) ultimately show ?case using reach rule(2) Fun(1)[OF nth_mem, of i "qs ! i" for i] by auto (metis (mono_tags, lifting) length_map map_nth_eq_conv)+ qed lemma depth_poss_split: assumes "Suc (depth (term_of_gterm t) + n) < depth (term_of_gterm u)" shows "\ p q. p @ q \ gposs u \ n < length q \ p \ gposs t" proof - from poss_length_depth obtain p m where p: "p \ gposs u" "length p = m" "depth (term_of_gterm u) = m" using poss_gposs_conv by blast then obtain m' where dt: "depth (term_of_gterm t) = m'" by blast from assms dt p(2, 3) have "length (take (Suc m') p) = Suc m'" by (metis Suc_leI depth_gterm_conv length_take less_add_Suc1 less_imp_le_nat less_le_trans min.absorb2) then have nt: "take (Suc m') p \ gposs t" using poss_length_bounded_by_depth dt depth_gterm_conv by (metis Suc_n_not_le_n gposs_to_poss) moreover have "n < length (drop (Suc m') p)" using assms depth_gterm_conv dt p(2-) by (metis add_Suc diff_diff_left length_drop zero_less_diff) ultimately show ?thesis by (metis append_take_drop_id p(1)) qed lemma Inf_to_automata: assumes "RR2_spec \ \" and "t \ Inf_branching_terms \ \" shows "\ u. gpair t u \ \ (Inf_reg \ (Q_infty (ta \) \))" (is "\ u. gpair t u \ \ ?B") proof - let ?A = "Inf_automata (ta \) (Q_infty (ta \) \)" let ?t_of_g = "\ t. term_of_gterm t :: ('b, 'a) term" obtain n where depth_card: "depth (?t_of_g t) + fcard (\ (ta \)) < n" by auto from assms(1, 2) have fin: "infinite {u. gpair t u \ \ \ \ funas_gterm u \ fset \}" by (auto simp: RR2_spec_def Inf_branching_terms_def) from infinte_no_depth_limit[of "?t_of_g ` {u. gpair t u \ \ \ \ funas_gterm u \ fset \}" "fset \"] this have "\ n. \t \ ?t_of_g ` {u. gpair t u \ \ \ \ funas_gterm u \ fset \}. n < depth t" by (simp add: infinite_inj_image_infinite[OF fin] funas_term_of_gterm_conv inj_term_of_gterm) from this depth_card obtain u where funas: "funas_gterm u \ fset \" and depth: "Suc n < depth (?t_of_g u)" and lang: "gpair t u \ \ \" by auto have "Suc (depth (term_of_gterm t) + fcard (\ (ta \))) < depth (term_of_gterm u)" using depth depth_card by (metis Suc_less_eq2 depth_gterm_conv less_trans) from depth_poss_split[OF this] obtain p q where pos: "p @ q \ gposs u" "p \ gposs t" and card: "fcard (\ (ta \)) < length q" by auto then have gp: "gsubt_at (gpair t u) p = gterm_to_None_Some (gsubt_at u p)" using subst_at_gpair_nt_poss_None_Some[of p] by force from lang obtain r where r: "run (ta \) r (gpair t u)" "ex_comp_state r |\| fin \" unfolding \_def gta_lang_def by (fastforce dest: gta_der_to_run) from pos have p_gtu: "p \ gposs (gpair t u)" and pu: "p \ gposs u" using not_gposs_append by auto have qinf: "ex_rule_state (gsubt_at r p) |\| Q_infty (ta \) \" using funas_gterm_gsubt_at_subseteq[OF pu] funas card unfolding Q_infty_fmember gta_der_def[symmetric] by (intro infinite_super[THEN infinite_imageD2[OF _ inj_gterm_to_None_Some], OF _ pigeonhole_ta_infinit_terms[of "ta \" "gsubt_at (gpair t u) p" _ "\ t. t \ (\(f, n). ((None, Some f), n)) ` fset \", OF _ run_to_gta_der_gsubt_at(1)[OF r(1) p_gtu]]]) (auto simp: poss_length_bounded_by_depth[of q] image_iff gp less_le_trans pos(1) poss_gposs_conv pu dest!: funas_gterm_bot_some_decomp) from Inf_automata_dashI[OF run_gsubt_cl[OF r(1) p_gtu, unfolded gp] qinf] have dashI: "CInr (ex_rule_state (gsubt_at r p)) |\| gta_der (Inf_automata (ta \) (Q_infty (ta \) \)) (gsubt_at (gpair t u) p)" unfolding gp[symmetric] . have "CInl (ex_comp_state r) |\| ta_der ?A (ctxt_at_pos (term_of_gterm (gpair t u)) p)\Var (CInl (ex_rule_state (gsubt_at r p)))\" using ta_der_fmap_states_ta[OF run_ta_der_ctxt_split2[OF r(1) p_gtu], of CInl, THEN fsubsetD[OF Inl_A_res_Inf_automata]] unfolding replace_term_at_replace_at_conv[OF gposs_to_poss[OF p_gtu]] by (auto simp: gterm.map_ident simp flip: map_term_replace_at_dist[OF gposs_to_poss[OF p_gtu]]) from ta_der_ctxt[OF dashI[unfolded gta_der_def] Inf_automata_reach_to_dash_reach[OF this]] have "CInr (ex_comp_state r) |\| gta_der (Inf_automata (ta \) (Q_infty (ta \) \)) (gpair t u)" unfolding replace_term_at_replace_at_conv[OF gposs_to_poss[OF p_gtu]] unfolding replace_gterm_conv[OF p_gtu] by (auto simp: gta_der_def) moreover from r(2) have "CInr (ex_comp_state r) |\| fin (Inf_reg \ (Q_infty (ta \) \))" by (auto simp: Inf_reg_def) ultimately show ?thesis using r(2) by (auto simp: \_def gta_der_def Inf_reg_def intro: exI[of _ u]) qed lemma CInr_Inf_automata_to_q_state: assumes "CInr p |\| ta_der (Inf_automata \ Q) t" and "ground t" shows "\ C s q. C\s\ = t \ CInr q |\| ta_der (Inf_automata \ Q) s \ q |\| Q \ CInr p |\| ta_der (Inf_automata \ Q) C\Var (CInr q)\ \ (fst \ fst \ the \ root) s = None" using assms proof (induct t arbitrary: p) case (Fun f ts) let ?A = "(Inf_automata \ Q)" from Fun(2) obtain qs q' where rule: "f qs \ CInr q' |\| rules ?A" "length qs = length ts" and eps: "q' = p \ (CInr q', CInr p) |\| (eps ?A)|\<^sup>+|" and reach: "\ i < length ts. qs ! i |\| ta_der ?A (ts ! i)" by auto (metis Inr_rhs_eps_Inr_lhs) consider (a) "\ i. i < length qs \ \ q''. qs ! i = CInl q''" | (b) "\ i < length qs. \ q''. qs ! i = CInr q''" by (meson remove_sum.cases) then show ?case proof cases case a then have "f qs \ CInr q' |\| |\| (q_inf_dash_intro_rules Q |`| rules \)" using rule by (auto simp: Inf_automata_def min_def upt_fset split!: if_splits) (metis (no_types, lifting) Inl_Inr_False Suc_pred append_eq_append_conv id_take_nth_drop length_Cons length_drop length_greater_0_conv length_map less_nat_zero_code list.size(3) nth_append_length rule(2)) then show ?thesis using reach eps rule by (intro exI[of _ Hole] exI[of _ "Fun f ts"] exI[of _ q']) (auto split!: if_splits) next case b then obtain i q'' where b: "i < length ts" "qs ! i = CInr q''" using rule(2) by auto then have "CInr q'' |\| ta_der ?A (ts ! i)" using rule(2) reach by auto from Fun(3) Fun(1)[OF nth_mem, OF b(1) this] b rule(2) obtain C s q''' where ctxt: "C\s\ = ts ! i" and qinf: "CInr q''' |\| ta_der (Inf_automata \ Q) s \ q''' |\| Q" and reach2: "CInr q'' |\| ta_der (Inf_automata \ Q) C\Var (CInr q''')\" and "(fst \ fst \ the \ root) s = None" by auto then show ?thesis using rule eps reach ctxt qinf reach2 b(1) b(2)[symmetric] by (auto simp: min_def nth_append_Cons simp flip: map_append id_take_nth_drop[OF b(1)] intro!: exI[of _ "More f (take i ts) C (drop (Suc i) ts)"] exI[of _ "s"] exI[of _ "q'''"] exI[of _ "CInr q'"] exI[of _ "qs"]) qed qed auto lemma aux_lemma: assumes "RR2_spec \ \" and "\ \ \\<^sub>G (fset \) \ \\<^sub>G (fset \)" and "infinite {u | u. gpair t u \ \ \}" shows "t \ Inf_branching_terms \ \" proof - from assms have [simp]: "gpair t u \ \ \ \ (t, u) \ \ \ u \ \\<^sub>G (fset \)" for u by (auto simp: RR2_spec_def) from assms have "t \ \\<^sub>G (fset \)" unfolding RR2_spec_def by (auto dest: not_finite_existsD) then show ?thesis using assms unfolding Inf_branching_terms_def by (auto simp: \\<^sub>G_equivalent_def) qed lemma Inf_automata_to_Inf: assumes "RR2_spec \ \" and "\ \ \\<^sub>G (fset \) \ \\<^sub>G (fset \)" and "gpair t u \ \ (Inf_reg \ (Q_infty (ta \) \))" shows "t \ Inf_branching_terms \ \" proof - let ?con = "\ t. term_of_gterm (gterm_to_None_Some t)" let ?A = "Inf_automata (ta \) (Q_infty (ta \) \)" from assms(3) obtain q where fin: "q |\| fin \" and reach_fin: "CInr q |\| ta_der ?A (term_of_gterm (gpair t u))" by (auto simp: Inf_reg_def \_def Inf_automata_def elim!: gta_langE) from CInr_Inf_automata_to_q_state[OF reach_fin] obtain C s p where ctxt: "C\s\ = term_of_gterm (gpair t u)" and q_inf_st: "CInr p |\| ta_der ?A s" "p |\| Q_infty (ta \) \" and reach: "CInr q |\| ta_der ?A C\Var (CInr p)\" and none: "(fst \ fst \ the \ root) s = None" by auto have gr: "ground s" "ground_ctxt C" using arg_cong[OF ctxt, of ground] by auto have reach: "q |\| ta_der (ta \) (adapt_vars_ctxt C)\Var p\" using gr Inf_automata_dash_reach_to_reach[OF reach] by (auto simp: map_vars_term_ctxt_apply) from q_inf_st(2) have inf: "infinite {v. funas_gterm v \ fset \ \ p |\| ta_der (ta \) (?con v)}" by (simp add: Q_infty_fmember) have inf: "infinite {v. funas_gterm v \ fset \ \ q |\| gta_der (ta \) (gctxt_of_ctxt C)\gterm_to_None_Some v\\<^sub>G}" using reach ground_ctxt_adapt_ground[OF gr(2)] gr by (intro infinite_super[OF _ inf], auto simp: gta_der_def) (smt (z3) adapt_vars_ctxt adapt_vars_term_of_gterm ground_gctxt_of_ctxt_apply_gterm ta_der_ctxt) have *: "gfun_at (gterm_of_term C\s\) (hole_pos C) = gfun_at (gterm_of_term s) []" by (induct C) (auto simp: nth_append_Cons) from arg_cong[OF ctxt, of "\ t. gfun_at (gterm_of_term t) (hole_pos C)"] none have hp_nt: "ghole_pos (gctxt_of_ctxt C) \ gposs t" unfolding ground_hole_pos_to_ghole[OF gr(2)] using gfun_at_gpair[of t u "hole_pos C"] gr * by (cases s) (auto simp flip: poss_gposs_mem_conv split: if_splits elim: gfun_at_possE) from gpair_ctxt_decomposition[OF hp_nt, of u "gsubt_at u (hole_pos C)"] have to_gpair: "gpair t (gctxt_at_pos u (hole_pos C))\v\\<^sub>G = (gctxt_of_ctxt C)\gterm_to_None_Some v\\<^sub>G" for v unfolding ground_hole_pos_to_ghole[OF gr(2)] using ctxt gr using subst_at_gpair_nt_poss_None_Some[OF _ hp_nt, of u] by (metis (no_types, lifting) UnE \ghole_pos (gctxt_of_ctxt C) = hole_pos C\ gposs_of_gpair gsubt_at_gctxt_apply_ghole hole_pos_in_ctxt_apply hp_nt poss_gposs_conv term_of_gterm_ctxt_apply) have inf: "infinite {v. gpair t ((gctxt_at_pos u (hole_pos C))\v\\<^sub>G) \ \ \}" using fin by (intro infinite_super[OF _ inf]) (auto simp: \_def gta_der_def simp flip: to_gpair) have "infinite {u |u. gpair t u \ \ \}" by (intro infinite_super[OF _ infinite_inj_image_infinite[OF inf gctxt_apply_inj_on_term[of "gctxt_at_pos u (hole_pos C)"]]]) (auto simp: image_def intro: infinite_super) then show ?thesis using assms(1, 2) by (intro aux_lemma[of \]) simp qed lemma Inf_automata_subseteq: "\ (Inf_reg \ (Q_infty (ta \) \)) \ \ \" (is "\ ?IA \ _") proof fix s assume l: "s \ \ ?IA" then obtain q where w: "q |\| fin ?IA" "q |\| ta_der (ta ?IA) (term_of_gterm s)" by (auto simp: \_def elim!: gta_langE) from w(1) have "remove_sum q |\| fin \" by (auto simp: Inf_reg_def Inf_automata_def) then show "s \ \ \" using Inf_automata_dash_reach_to_reach[of q "ta \"] w(2) by (auto simp: gterm.map_ident \_def Inf_reg_def) (metis gta_langI map_vars_term_term_of_gterm) qed lemma \_Inf_reg: assumes "RR2_spec \ \" and "\ \ \\<^sub>G (fset \) \ \\<^sub>G (fset \)" shows "gfst ` \ (Inf_reg \ (Q_infty (ta \) \)) = Inf_branching_terms \ \" proof - {fix s assume ass: "s \ \ (Inf_reg \ (Q_infty (ta \) \))" then have "\ t u. s = gpair t u" using Inf_automata_subseteq[of \ \] assms(1) by (auto simp: RR2_spec_def) then have "gfst s \ Inf_branching_terms \ \" using ass Inf_automata_to_Inf[OF assms] by (force simp: gfst_gpair)} then show ?thesis using Inf_to_automata[OF assms(1), of _ \] by (auto simp: gfst_gpair) (metis gfst_gpair image_iff) qed end \ No newline at end of file diff --git a/thys/Regular_Tree_Relations/RRn_Automata.thy b/thys/Regular_Tree_Relations/RRn_Automata.thy --- a/thys/Regular_Tree_Relations/RRn_Automata.thy +++ b/thys/Regular_Tree_Relations/RRn_Automata.thy @@ -1,1568 +1,1568 @@ theory RRn_Automata imports Tree_Automata_Complement Ground_Ctxt begin section \Regular relations\ subsection \Encoding pairs of terms\ text \The encoding of two terms $s$ and $t$ is given by its tree domain, which is the union of the domains of $s$ and $t$, and the labels, which arise from looking up each position in $s$ and $t$, respectively.\ definition gpair :: "'f gterm \ 'g gterm \ ('f option \ 'g option) gterm" where "gpair s t = glabel (\p. (gfun_at s p, gfun_at t p)) (gunion (gdomain s) (gdomain t))" text \We provide an efficient implementation of gpair.\ definition zip_fill :: "'a list \ 'b list \ ('a option \ 'b option) list" where "zip_fill xs ys = zip (map Some xs @ replicate (length ys - length xs) None) (map Some ys @ replicate (length xs - length ys) None)" lemma zip_fill_code [code]: "zip_fill xs [] = map (\x. (Some x, None)) xs" "zip_fill [] ys = map (\y. (None, Some y)) ys" "zip_fill (x # xs) (y # ys) = (Some x, Some y) # zip_fill xs ys" subgoal by (induct xs) (auto simp: zip_fill_def) subgoal by (induct ys) (auto simp: zip_fill_def) subgoal by (auto simp: zip_fill_def) done lemma length_zip_fill [simp]: "length (zip_fill xs ys) = max (length xs) (length ys)" by (auto simp: zip_fill_def) lemma nth_zip_fill: assumes "i < max (length xs) (length ys)" shows "zip_fill xs ys ! i = (if i < length xs then Some (xs ! i) else None, if i < length ys then Some (ys ! i) else None)" using assms by (auto simp: zip_fill_def nth_append) fun gpair_impl :: "'f gterm option \ 'g gterm option \ ('f option \ 'g option) gterm" where "gpair_impl (Some s) (Some t) = gpair s t" | "gpair_impl (Some s) None = map_gterm (\f. (Some f, None)) s" | "gpair_impl None (Some t) = map_gterm (\f. (None, Some f)) t" | "gpair_impl None None = GFun (None, None) []" declare gpair_impl.simps(2-4)[code] lemma gpair_impl_code [simp, code]: "gpair_impl (Some s) (Some t) = (case s of GFun f ss \ case t of GFun g ts \ GFun (Some f, Some g) (map (\(s, t). gpair_impl s t) (zip_fill ss ts)))" proof (induct "gdomain s" "gdomain t" arbitrary: s t rule: gunion.induct) case (1 f ss g ts) obtain f' ss' where [simp]: "s = GFun f' ss'" by (cases s) obtain g' ts' where [simp]: "t = GFun g' ts'" by (cases t) show ?case using 1(2,3) 1(1)[of i "ss' ! i" "ts' ! i" for i] by (auto simp: gpair_def comp_def nth_zip_fill intro: glabel_map_gterm_conv[unfolded comp_def] intro!: nth_equalityI) qed lemma gpair_code [code]: "gpair s t = gpair_impl (Some s) (Some t)" by simp (* export_code gpair in Haskell *) declare gpair_impl.simps(1)[simp del] text \We can easily prove some basic properties. I believe that proving them by induction with a definition along the lines of @{const gpair_impl} would be very cumbersome.\ lemma gpair_swap: "map_gterm prod.swap (gpair s t) = gpair t s" by (intro eq_gterm_by_gposs_gfun_at) (auto simp: gpair_def) lemma gpair_assoc: defines "f \ \(f, gh). (f, gh \ fst, gh \ snd)" defines "g \ \(fg, h). (fg \ fst, fg \ snd, h)" shows "map_gterm f (gpair s (gpair t u)) = map_gterm g (gpair (gpair s t) u)" by (intro eq_gterm_by_gposs_gfun_at) (auto simp: gpair_def f_def g_def) subsection \Decoding of pairs\ fun gcollapse :: "'f option gterm \ 'f gterm option" where "gcollapse (GFun None _) = None" | "gcollapse (GFun (Some f) ts) = Some (GFun f (map the (filter (\t. \ Option.is_none t) (map gcollapse ts))))" lemma gcollapse_groot_None [simp]: "groot_sym t = None \ gcollapse t = None" "fst (groot t) = None \ gcollapse t = None" by (cases t, simp)+ definition gfst :: "('f option \ 'g option) gterm \ 'f gterm" where "gfst = the \ gcollapse \ map_gterm fst" definition gsnd :: "('f option \ 'g option) gterm \ 'g gterm" where "gsnd = the \ gcollapse \ map_gterm snd" lemma filter_less_upt: "[i\[i.. m") case True then show ?thesis proof (induct rule: inc_induct) case (step n) then show ?case by (auto simp: upt_rec[of n]) qed simp qed simp lemma gcollapse_aux: assumes "gposs s = {p. p \ gposs t \ gfun_at t p \ Some None}" shows "gposs (the (gcollapse t)) = gposs s" "\p. p \ gposs s \ gfun_at (the (gcollapse t)) p = (gfun_at t p \ id)" proof (goal_cases) define s' t' where "s' \ gdomain s" and "t' \ gdomain t" have *: "gposs (the (gcollapse t)) = gposs s \ (\p. p \ gposs s \ gfun_at (the (gcollapse t)) p = (gfun_at t p \ id))" using assms s'_def t'_def proof (induct s' t' arbitrary: s t rule: gunion.induct) case (1 f' ss' g' ts') obtain f ss where s [simp]: "s = GFun f ss" by (cases s) obtain g ts where t [simp]: "t = GFun (Some g) ts" using arg_cong[OF 1(2), of "\P. [] \ P"] by (cases t) auto have *: "i < length ts \ \ Option.is_none (gcollapse (ts ! i)) \ i < length ss" for i using arg_cong[OF 1(2), of "\P. [i] \ P"] by (cases "ts ! i") auto have l: "length ss \ length ts" using arg_cong[OF 1(2), of "\P. [length ss-1] \ P"] by auto have [simp]: "[t\map gcollapse ts . \ Option.is_none t] = take (length ss) (map gcollapse ts)" by (subst (2) map_nth[symmetric]) (auto simp add: * filter_map comp_def filter_less_upt cong: filter_cong[OF refl, of "[0.. gposs (ss ! i) = gposs (the (gcollapse (ts ! i)))" for i using conjunct1[OF 1(1), of i "ss ! i" "ts ! i"] l arg_cong[OF 1(2), of "\P. {p. i # p \ P}"] 1(3,4) by simp show ?case proof (intro conjI allI, goal_cases A B) case A show ?case using l by (auto simp: comp_def split: if_splits) next case (B p) show ?case proof (cases p) case (Cons i q) then show ?thesis using arg_cong[OF 1(2), of "\P. {p. i # p \ P}"] spec[OF conjunct2[OF 1(1)], of i "ss ! i" "ts ! i" q] 1(3,4) by auto qed auto qed qed { case 1 then show ?case using * by blast next case 2 then show ?case using * by blast } qed lemma gfst_gpair: "gfst (gpair s t) = s" proof - have *: "gposs s = {p \ gposs (map_gterm fst (gpair s t)). gfun_at (map_gterm fst (gpair s t)) p \ Some None}" using gfun_at_nongposs by (fastforce simp: gpair_def elim: gfun_at_possE) show ?thesis unfolding gfst_def comp_def using gcollapse_aux[OF *] by (auto intro!: eq_gterm_by_gposs_gfun_at simp: gpair_def) qed lemma gsnd_gpair: "gsnd (gpair s t) = t" using gfst_gpair[of t s] gpair_swap[of t s, symmetric] by (simp add: gfst_def gsnd_def gpair_def gterm.map_comp comp_def) lemma gpair_impl_None_Inv: "map_gterm (the \ snd) (gpair_impl None (Some t)) = t" by (simp add: gterm.map_ident gterm.map_comp comp_def) subsection \Contexts to gpair\ lemma gpair_context1: assumes "length ts = length us" shows "gpair (GFun f ts) (GFun f us) = GFun (Some f, Some f) (map (case_prod gpair) (zip ts us))" using assms unfolding gpair_code by (auto intro!: nth_equalityI simp: zip_fill_def) lemma gpair_context2: assumes "\ i. i < length ts \ ts ! i = gpair (ss ! i) (us ! i)" and "length ss = length ts" and "length us = length ts" shows "GFun (Some f, Some h) ts = gpair (GFun f ss) (GFun h us)" using assms unfolding gpair_code gpair_impl_code by (auto simp: zip_fill_def intro!: nth_equalityI) lemma map_funs_term_some_gpair: shows "gpair t t = map_gterm (\f. (Some f, Some f)) t" proof (induct t) case (GFun f ts) then show ?case by (auto intro!: gpair_context2[symmetric]) qed lemma gpair_inject [simp]: "gpair s t = gpair s' t' \ s = s' \ t = t'" by (metis gfst_gpair gsnd_gpair) abbreviation gterm_to_None_Some :: "'f gterm \ ('f option \ 'f option) gterm" where "gterm_to_None_Some t \ map_gterm (\f. (None, Some f)) t" abbreviation "gterm_to_Some_None t \ map_gterm (\f. (Some f, None)) t" lemma inj_gterm_to_None_Some: "inj gterm_to_None_Some" by (meson Pair_inject gterm.inj_map inj_onI option.inject) lemma zip_fill1: assumes "length ss < length ts" shows "zip_fill ss ts = zip (map Some ss) (map Some (take (length ss) ts)) @ map (\ x. (None, Some x)) (drop (length ss) ts)" using assms by (auto simp: zip_fill_def list_eq_iff_nth_eq nth_append simp add: min.absorb2) lemma zip_fill2: assumes "length ts < length ss" shows "zip_fill ss ts = zip (map Some (take (length ts) ss)) (map Some ts) @ map (\ x. (Some x, None)) (drop (length ts) ss)" using assms by (auto simp: zip_fill_def list_eq_iff_nth_eq nth_append simp add: min.absorb2) (* GPair position lemmas *) (* MOVE me*) lemma not_gposs_append [simp]: assumes "p \ gposs t" shows "p @ q \ gposs t = False" using assms poss_gposs_conv using poss_append_poss by blast (*end Move *) lemma gfun_at_gpair: "gfun_at (gpair s t) p = (if p \ gposs s then (if p \ gposs t then Some (gfun_at s p, gfun_at t p) else Some (gfun_at s p, None)) else (if p \ gposs t then Some (None, gfun_at t p) else None))" using gfun_at_glabel by (auto simp: gpair_def) lemma gposs_of_gpair [simp]: shows "gposs (gpair s t) = gposs s \ gposs t" by (auto simp: gpair_def) lemma poss_to_gpair_poss: "p \ gposs s \ p \ gposs (gpair s t)" "p \ gposs t \ p \ gposs (gpair s t)" by auto lemma gsubt_at_gpair_poss: assumes "p \ gposs s" and "p \ gposs t" shows "gsubt_at (gpair s t) p = gpair (gsubt_at s p) (gsubt_at t p)" using assms by (auto simp: gunion_gsubt_at_poss gfun_at_gpair intro!: eq_gterm_by_gposs_gfun_at) lemma subst_at_gpair_nt_poss_Some_None: assumes "p \ gposs s" and "p \ gposs t" shows "gsubt_at (gpair s t) p = gterm_to_Some_None (gsubt_at s p)" using assms gfun_at_poss by (force simp: gunion_gsubt_at_poss gfun_at_gpair intro!: eq_gterm_by_gposs_gfun_at) lemma subst_at_gpair_nt_poss_None_Some: assumes "p \ gposs t" and "p \ gposs s" shows "gsubt_at (gpair s t) p = gterm_to_None_Some (gsubt_at t p)" using assms gfun_at_poss by (force simp: gunion_gsubt_at_poss gfun_at_gpair intro!: eq_gterm_by_gposs_gfun_at) lemma gpair_ctxt_decomposition: fixes C defines "p \ ghole_pos C" assumes "p \ gposs s" and "gpair s t = C\gterm_to_None_Some u\\<^sub>G" shows "gpair s (gctxt_at_pos t p)\v\\<^sub>G = C\gterm_to_None_Some v\\<^sub>G" using assms(2-) proof - note p[simp] = assms(1) have pt: "p \ gposs t" and pc: "p \ gposs C\gterm_to_None_Some v\\<^sub>G" and pu: "p \ gposs C\gterm_to_None_Some u\\<^sub>G" using arg_cong[OF assms(3), of gposs] assms(2) ghole_pos_in_apply by auto have *: "gctxt_at_pos (gpair s (gctxt_at_pos t (ghole_pos C))\v\\<^sub>G) (ghole_pos C) = gctxt_at_pos (gpair s t) (ghole_pos C)" using assms(2) pt by (intro eq_gctxt_at_pos) (auto simp: gposs_gctxt_at_pos gunion_gsubt_at_poss gfun_at_gpair gfun_at_gctxt_at_pos_not_after) have "gsubt_at (gpair s (gctxt_at_pos t p)\v\\<^sub>G) p = gsubt_at C\gterm_to_None_Some v\\<^sub>G p" using pt assms(2) subst_at_gpair_nt_poss_None_Some[OF _ assms(2), of "(gctxt_at_pos t p)\v\\<^sub>G"] using ghole_pos_gctxt_at_pos by (simp add: ghole_pos_in_apply) then show ?thesis using assms(2) ghole_pos_gctxt_at_pos using gsubst_at_gctxt_at_eq_gtermD[OF assms(3) pu] by (intro gsubst_at_gctxt_at_eq_gtermI[OF _ pc]) (auto simp: ghole_pos_in_apply * gposs_gctxt_at_pos[OF pt, unfolded p]) qed lemma groot_gpair [simp]: "fst (groot (gpair s t)) = (Some (fst (groot s)), Some (fst (groot t)))" by (cases s; cases t) (auto simp add: gpair_code) lemma ground_ctxt_adapt_ground [intro]: assumes "ground_ctxt C" shows "ground_ctxt (adapt_vars_ctxt C)" using assms by (induct C) auto lemma adapt_vars_ctxt2 : assumes "ground_ctxt C" shows "adapt_vars_ctxt (adapt_vars_ctxt C) = adapt_vars_ctxt C" using assms by (induct C) (auto simp: adapt_vars2) subsection \Encoding of lists of terms\ definition gencode :: "'f gterm list \ 'f option list gterm" where "gencode ts = glabel (\p. map (\t. gfun_at t p) ts) (gunions (map gdomain ts))" definition gdecode_nth :: "'f option list gterm \ nat \ 'f gterm" where "gdecode_nth t i = the (gcollapse (map_gterm (\f. f ! i) t))" lemma gdecode_nth_gencode: assumes "i < length ts" shows "gdecode_nth (gencode ts) i = ts ! i" proof - have *: "gposs (ts ! i) = {p \ gposs (map_gterm (\f. f ! i) (gencode ts)). gfun_at (map_gterm (\f. f ! i) (gencode ts)) p \ Some None}" using assms by (auto simp: gencode_def elim: gfun_at_possE dest: gfun_at_poss_gpossD) (force simp: fun_at_def' split: if_splits) show ?thesis unfolding gdecode_nth_def comp_def using assms gcollapse_aux[OF *] by (auto intro!: eq_gterm_by_gposs_gfun_at simp: gencode_def) (metis (no_types) gposs_map_gterm length_map list.set_map map_nth_eq_conv nth_mem) qed definition gdecode :: "'f option list gterm \ 'f gterm list" where "gdecode t = (case t of GFun f ts \ map (\i. gdecode_nth t i) [0..t. gfun_at t []"] by (auto simp: gencode_def) then show ?thesis using gdecode_nth_gencode[of _ ts] by (auto intro!: nth_equalityI simp: gdecode_def GFun) qed definition gencode_impl :: "'f gterm option list \ 'f option list gterm" where "gencode_impl ts = glabel (\p. map (\t. t \ (\t. gfun_at t p)) ts) (gunions (map (case_option (GFun () []) gdomain) ts))" lemma gencode_code [code]: "gencode ts = gencode_impl (map Some ts)" by (auto simp: gencode_def gencode_impl_def comp_def) lemma gencode_singleton: "gencode [t] = map_gterm (\f. [Some f]) t" using glabel_map_gterm_conv[unfolded comp_def, of "\t. [t]" t] by (simp add: gunions_def gencode_def) lemma gencode_pair: "gencode [t, u] = map_gterm (\(f, g). [f, g]) (gpair t u)" by (simp add: gunions_def gencode_def gpair_def map_gterm_glabel comp_def) subsection \RRn relations\ definition RR1_spec where "RR1_spec A T \ \ A = T" definition RR2_spec where "RR2_spec A T \ \ A = {gpair t u |t u. (t, u) \ T}" definition RRn_spec where "RRn_spec n A R \ \ A = gencode ` R \ (\ts \ R. length ts = n)" lemma RR1_to_RRn_spec: assumes "RR1_spec A T" shows "RRn_spec 1 (fmap_funs_reg (\f. [Some f]) A) ((\t. [t]) ` T)" proof - have [simp]: "inj_on (\f. [Some f]) X" for X by (auto simp: inj_on_def) show ?thesis using assms by (auto simp: RR1_spec_def RRn_spec_def fmap_funs_\ image_comp comp_def gencode_singleton) qed lemma RR2_to_RRn_spec: assumes "RR2_spec A T" shows "RRn_spec 2 (fmap_funs_reg (\(f, g). [f, g]) A) ((\(t, u). [t, u]) ` T)" proof - have [simp]: "inj_on (\(f, g). [f, g]) X" for X by (auto simp: inj_on_def) show ?thesis using assms by (auto simp: RR2_spec_def RRn_spec_def fmap_funs_\ image_comp comp_def prod.case_distrib gencode_pair) qed lemma RRn_to_RR2_spec: assumes "RRn_spec 2 A T" shows "RR2_spec (fmap_funs_reg (\ f. (f ! 0 , f ! 1)) A) ((\ f. (f ! 0, f ! 1)) ` T)" (is "RR2_spec ?A ?T") proof - {fix xs assume "xs \ T" then have "length xs = 2" using assms by (auto simp: RRn_spec_def) then obtain t u where *: "xs = [t, u]" by (metis (no_types, lifting) One_nat_def Suc_1 length_0_conv length_Suc_conv) have **: "(\f. (f ! 0, f ! Suc 0)) \ (\(f, g). [f, g]) = id" by auto have "map_gterm (\f. (f ! 0, f ! Suc 0)) (gencode xs) = gpair t u" unfolding * gencode_pair gterm.map_comp ** gterm.map_id .. then have "\ t u. xs = [t, u] \ map_gterm (\f. (f ! 0, f ! Suc 0)) (gencode xs) = gpair t u" using * by blast} then show ?thesis using assms by (force simp: RR2_spec_def RRn_spec_def fmap_funs_\ image_comp comp_def prod.case_distrib gencode_pair image_iff Bex_def) qed lemma relabel_RR1_spec [simp]: "RR1_spec (relabel_reg A) T \ RR1_spec A T" by (simp add: RR1_spec_def) lemma relabel_RR2_spec [simp]: "RR2_spec (relabel_reg A) T \ RR2_spec A T" by (simp add: RR2_spec_def) lemma relabel_RRn_spec [simp]: "RRn_spec n (relabel_reg A) T \ RRn_spec n A T" by (simp add: RRn_spec_def) lemma trim_RR1_spec [simp]: "RR1_spec (trim_reg A) T \ RR1_spec A T" by (simp add: RR1_spec_def \_trim) lemma trim_RR2_spec [simp]: "RR2_spec (trim_reg A) T \ RR2_spec A T" by (simp add: RR2_spec_def \_trim) lemma trim_RRn_spec [simp]: "RRn_spec n (trim_reg A) T \ RRn_spec n A T" by (simp add: RRn_spec_def \_trim) lemma swap_RR2_spec: assumes "RR2_spec A R" shows "RR2_spec (fmap_funs_reg prod.swap A) (prod.swap ` R)" using assms by (force simp add: RR2_spec_def fmap_funs_\ gpair_swap image_iff) subsection \Nullary automata\ lemma false_RRn_spec: "RRn_spec n empty_reg {}" by (auto simp: RRn_spec_def \_epmty) lemma true_RR0_spec: "RRn_spec 0 (Reg {|q|} (TA {|[] [] \ q|} {||})) {[]}" by (auto simp: RRn_spec_def \_def const_ta_lang gencode_def gunions_def) subsection \Pairing RR1 languages\ text \cf. @{const "gpair"}.\ abbreviation "lift_Some_None s \ (Some s, None)" abbreviation "lift_None_Some s \ (None, Some s)" abbreviation "pair_eps A B \ (\ (p, q). ((Some (fst p), q), (Some (snd p), q))) |`| (eps A |\| finsert None (Some |`| \ B))" abbreviation "pair_rule \ (\ (ra, rb). TA_rule (Some (r_root ra), Some (r_root rb)) (zip_fill (r_lhs_states ra) (r_lhs_states rb)) (Some (r_rhs ra), Some (r_rhs rb)))" lemma lift_Some_None_pord_swap [simp]: "prod.swap \ lift_Some_None = lift_None_Some" "prod.swap \ lift_None_Some = lift_Some_None" by auto lemma eps_to_pair_eps_Some_None: "(p, q) |\| eps \ \ (lift_Some_None p, lift_Some_None q) |\| pair_eps \ \" by force definition pair_automaton :: "('p, 'f) ta \ ('q, 'g) ta \ ('p option \ 'q option, 'f option \ 'g option) ta" where "pair_automaton A B = TA (map_ta_rule lift_Some_None lift_Some_None |`| rules A |\| map_ta_rule lift_None_Some lift_None_Some |`| rules B |\| pair_rule |`| (rules A |\| rules B)) (pair_eps A B |\| map_both prod.swap |`| (pair_eps B A))" definition pair_automaton_reg where "pair_automaton_reg R L = Reg (Some |`| fin R |\| Some |`| fin L) (pair_automaton (ta R) (ta L))" lemma pair_automaton_eps_simps: "(lift_Some_None p, p') |\| eps (pair_automaton A B) \ (lift_Some_None p, p') |\| pair_eps A B" "(q , lift_Some_None q') |\| eps (pair_automaton A B) \ (q , lift_Some_None q') |\| pair_eps A B" by (auto simp: pair_automaton_def eps_to_pair_eps_Some_None) lemma pair_automaton_eps_Some_SomeD: "((Some p, Some p'), r) |\| eps (pair_automaton A B) \ fst r \ None \ snd r \ None \ (Some p = fst r \ Some p' = snd r) \ (Some p \ fst r \ (p, the (fst r)) |\| (eps A)) \ (Some p' \ snd r \ (p', the (snd r)) |\| (eps B))" by (auto simp: pair_automaton_def) lemma pair_automaton_eps_Some_SomeD2: "(r, (Some p, Some p')) |\| eps (pair_automaton A B) \ fst r \ None \ snd r \ None \ (fst r = Some p \ snd r = Some p') \ (fst r \ Some p \ (the (fst r), p) |\| (eps A)) \ (snd r \ Some p' \ (the (snd r), p') |\| (eps B))" by (auto simp: pair_automaton_def) lemma pair_eps_Some_None: fixes p q q' defines "l \ (p, q)" and "r \ lift_Some_None q'" assumes "(l, r) |\| (eps (pair_automaton A B))|\<^sup>+|" shows "q = None \ p \ None \ (the p, q') |\| (eps A)|\<^sup>+|" using assms(3, 1, 2) proof (induct arbitrary: q' q rule: ftrancl_induct) case (Step b) then show ?case unfolding pair_automaton_eps_simps by (auto simp: pair_automaton_eps_simps) (meson not_ftrancl_into) qed (auto simp: pair_automaton_def) lemma pair_eps_Some_Some: fixes p q defines "l \ (Some p, Some q)" assumes "(l, r) |\| (eps (pair_automaton A B))|\<^sup>+|" shows "fst r \ None \ snd r \ None \ (fst l \ fst r \ (p, the (fst r)) |\| (eps A)|\<^sup>+|) \ (snd l \ snd r \ (q, the (snd r)) |\| (eps B)|\<^sup>+|)" using assms(2, 1) proof (induct arbitrary: p q rule: ftrancl_induct) case (Step b c) then obtain r r' where *: "b = (Some r, Some r')" by (cases b) auto show ?case using Step(2) using pair_automaton_eps_Some_SomeD[OF Step(3)[unfolded *]] by (auto simp: *) (meson not_ftrancl_into)+ qed (auto simp: pair_automaton_def) lemma pair_eps_Some_Some2: fixes p q defines "r \ (Some p, Some q)" assumes "(l, r) |\| (eps (pair_automaton A B))|\<^sup>+|" shows "fst l \ None \ snd l \ None \ (fst l \ fst r \ (the (fst l), p) |\| (eps A)|\<^sup>+|) \ (snd l \ snd r \ (the (snd l), q) |\| (eps B)|\<^sup>+|)" using assms(2, 1) proof (induct arbitrary: p q rule: ftrancl_induct) case (Step b c) from pair_automaton_eps_Some_SomeD2[OF Step(3)] obtain r r' where *: "c = (Some r, Some r')" by (cases c) auto from Step(2)[OF this] show ?case using pair_automaton_eps_Some_SomeD[OF Step(3)[unfolded *]] by (auto simp: *) (meson not_ftrancl_into)+ qed (auto simp: pair_automaton_def) lemma map_pair_automaton: "pair_automaton (fmap_funs_ta f A) (fmap_funs_ta g B) = fmap_funs_ta (\(a, b). (map_option f a, map_option g b)) (pair_automaton A B)" (is "?Ls = ?Rs") proof - let ?ls = "pair_rule \ map_prod (map_ta_rule id f) (map_ta_rule id g)" let ?rs = "map_ta_rule id (\(a, b). (map_option f a, map_option g b)) \ pair_rule" have *: "(\(a, b). (map_option f a, map_option g b)) \ lift_Some_None = lift_Some_None \ f" "(\(a, b). (map_option f a, map_option g b)) \ lift_None_Some = lift_None_Some \ g" by (auto simp: comp_def) have "?ls x = ?rs x" for x by (cases x) (auto simp: ta_rule.map_sel) then have [simp]: "?ls = ?rs" by blast then have "rules ?Ls = rules ?Rs" unfolding pair_automaton_def fmap_funs_ta_def by (simp add: fimage_funion map_ta_rule_comp * map_prod_ftimes) moreover have "eps ?Ls = eps ?Rs" unfolding pair_automaton_def fmap_funs_ta_def by (simp add: fimage_funion \_def) ultimately show ?thesis by (intro TA_equalityI) simp qed lemmas map_pair_automaton_12 = map_pair_automaton[of _ _ id, unfolded fmap_funs_ta_id option.map_id] map_pair_automaton[of id _ _, unfolded fmap_funs_ta_id option.map_id] lemma fmap_states_funs_ta_commute: "fmap_states_ta f (fmap_funs_ta g A) = fmap_funs_ta g (fmap_states_ta f A)" proof - have [simp]: "map_ta_rule f id (map_ta_rule id g r) = map_ta_rule id g (map_ta_rule f id r)" for r by (cases r) auto show ?thesis by (auto simp: ta_rule.case_distrib fmap_states_ta_def fmap_funs_ta_def fimage_iff fBex_def split: ta_rule.splits) qed lemma states_pair_automaton: "\ (pair_automaton A B) |\| (finsert None (Some |`| \ A) |\| (finsert None (Some |`| \ B)))" unfolding pair_automaton_def apply (intro \_subseteq_I) apply (auto simp: ta_rule.map_sel in_fset_conv_nth nth_zip_fill rule_statesD eps_statesD) apply (simp add: rule_statesD)+ done lemma swap_pair_automaton: assumes "(p, q) |\| ta_der (pair_automaton A B) (term_of_gterm t)" shows "(q, p) |\| ta_der (pair_automaton B A) (term_of_gterm (map_gterm prod.swap t))" proof - let ?m = "map_ta_rule prod.swap prod.swap" have [simp]: "map prod.swap (zip_fill xs ys) = zip_fill ys xs" for xs ys by (auto simp: zip_fill_def zip_nth_conv comp_def intro!: nth_equalityI) let ?swp = "\X. fmap_states_ta prod.swap (fmap_funs_ta prod.swap X)" { fix A B let ?AB = "?swp (pair_automaton A B)" and ?BA = "pair_automaton B A" have "eps ?AB |\| eps ?BA" "rules ?AB |\| rules ?BA" by (auto simp: fmap_states_ta_def fmap_funs_ta_def pair_automaton_def fimage_funion comp_def prod.map_comp ta_rule.map_comp) } note * = this let ?BA = "?swp (?swp (pair_automaton B A))" and ?AB = "?swp (pair_automaton A B)" have **: "r |\| rules (pair_automaton B A) \ ?m r |\| ?m |`| rules (pair_automaton B A)" for r by blast have "r |\| rules ?BA \ r |\| rules ?AB" "e |\| eps ?BA \ e |\| eps ?AB" for r e using *[of B A] map_ta_rule_prod_swap_id - unfolding fmap_funs_ta_def fmap_states_ta_def - by (auto simp: map_ta_rule_comp image_iff fBex_def ta_rule.map_id0 intro!: bexI[of _ "?m r"]) + unfolding fmap_funs_ta_def fmap_states_ta_def ta.sel + by (auto simp: map_ta_rule_comp image_iff ta_rule.map_id0 intro!: bexI[of _ "?m r"]) then have "eps ?BA |\| eps ?AB" "rules ?BA |\| rules ?AB" by blast+ then have "fmap_states_ta prod.swap (fmap_funs_ta prod.swap (pair_automaton A B)) = pair_automaton B A" using *[of A B] by (simp add: fmap_states_funs_ta_commute fmap_funs_ta_comp fmap_states_ta_comp TA_equalityI) then show ?thesis using ta_der_fmap_states_ta[OF ta_der_fmap_funs_ta[OF assms], of prod.swap prod.swap] by (simp add: gterm.map_comp comp_def) qed lemma to_ta_der_pair_automaton: "p |\| ta_der A (term_of_gterm t) \ (Some p, None) |\| ta_der (pair_automaton A B) (term_of_gterm (map_gterm (\f. (Some f, None)) t))" "q |\| ta_der B (term_of_gterm u) \ (None, Some q) |\| ta_der (pair_automaton A B) (term_of_gterm (map_gterm (\f. (None, Some f)) u))" "p |\| ta_der A (term_of_gterm t) \ q |\| ta_der B (term_of_gterm u) \ (Some p, Some q) |\| ta_der (pair_automaton A B) (term_of_gterm (gpair t u))" proof (goal_cases sn ns ss) let ?AB = "pair_automaton A B" have 1: "(Some p, None) |\| ta_der (pair_automaton A B) (term_of_gterm (map_gterm (\f. (Some f, None)) s))" if "p |\| ta_der A (term_of_gterm s)" for A B p s proof (rule fsubsetD[OF ta_der_mono]) show "rules (fmap_states_ta lift_Some_None (fmap_funs_ta lift_Some_None A)) |\| rules (pair_automaton A B)" by (auto simp: fmap_states_ta_def fmap_funs_ta_def comp_def ta_rule.map_comp pair_automaton_def) next show "eps (fmap_states_ta lift_Some_None (fmap_funs_ta lift_Some_None A)) |\| eps (pair_automaton A B)" by (rule fsubsetI) (auto simp: fmap_states_ta_def fmap_funs_ta_def pair_automaton_def comp_def fimage.rep_eq dest: eps_to_pair_eps_Some_None) next show "lift_Some_None p |\| ta_der (fmap_states_ta lift_Some_None (fmap_funs_ta lift_Some_None A)) (term_of_gterm (gterm_to_Some_None s))" using ta_der_fmap_states_ta[OF ta_der_fmap_funs_ta[OF that], of lift_Some_None] using ta_der_fmap_funs_ta ta_der_to_fmap_states_der that by fastforce qed have 2: "q |\| ta_der B (term_of_gterm t) \ (None, Some q) |\| ta_der ?AB (term_of_gterm (map_gterm (\g. (None, Some g)) t))" for q t using swap_pair_automaton[OF 1[of q B t A]] by (simp add: gterm.map_comp comp_def) { case sn then show ?case by (rule 1) next case ns then show ?case by (rule 2) next case ss then show ?case proof (induct t arbitrary: p q u) case (GFun f ts) obtain g us where u [simp]: "u = GFun g us" by (cases u) obtain p' ps where p': "f ps \ p' |\| rules A" "p' = p \ (p', p) |\| (eps A)|\<^sup>+|" "length ps = length ts" "\i. i < length ts \ ps ! i |\| ta_der A (term_of_gterm (ts ! i))" using GFun(2) by auto obtain q' qs where q': "g qs \ q' |\| rules B" "q' = q \ (q', q) |\| (eps B)|\<^sup>+|" "length qs = length us" "\i. i < length us \ qs ! i |\| ta_der B (term_of_gterm (us ! i))" using GFun(3) by auto have *: "Some p |\| Some |`| \ A" "Some q' |\| Some |`| \ B" using p'(2,1) q'(1) by (auto simp: rule_statesD eps_trancl_statesD) have eps: "p' = p \ q' = q \ ((Some p', Some q'), (Some p, Some q)) |\| (eps ?AB)|\<^sup>+|" proof (cases "p' = p") case True note p = this then show ?thesis proof (cases "q' = q") case False then have "(q', q) |\| (eps B)|\<^sup>+|" using q'(2) by auto hence "((Some p', Some q'), Some p', Some q) |\| (eps (pair_automaton A B))|\<^sup>+|" proof (rule ftrancl_map[rotated]) fix x y assume "(x, y) |\| eps B" thus "((Some p', Some x), Some p', Some y) |\| eps (pair_automaton A B)" using p'(1)[THEN rule_statesD(1), simplified] apply (simp add: pair_automaton_def image_iff fSigma.rep_eq) by fastforce qed thus ?thesis by (simp add: p) qed (simp add: p) next case False note p = this then have *: "(p', p) |\| (eps A)|\<^sup>+|" using p'(2) by auto then have eps: "((Some p', Some q'), Some p, Some q') |\| (eps (pair_automaton A B))|\<^sup>+|" proof (rule ftrancl_map[rotated]) fix x y assume "(x, y) |\| eps A" then show "((Some x, Some q'), Some y, Some q') |\| eps (pair_automaton A B)" using q'(1)[THEN rule_statesD(1), simplified] apply (simp add: pair_automaton_def image_iff fSigma.rep_eq) by fastforce qed then show ?thesis proof (cases "q' = q") case True then show ?thesis using eps by simp next case False then have "(q', q) |\| (eps B)|\<^sup>+|" using q'(2) by auto then have "((Some p, Some q'), Some p, Some q) |\| (eps (pair_automaton A B))|\<^sup>+|" apply (rule ftrancl_map[rotated]) using *[THEN eps_trancl_statesD] apply (simp add: p pair_automaton_def image_iff fSigma.rep_eq) by fastforce then show ?thesis using eps by (meson ftrancl_trans) qed qed have "(Some f, Some g) zip_fill ps qs \ (Some p', Some q') |\| rules ?AB" using p'(1) q'(1) by (force simp: pair_automaton_def) then show ?case using GFun(1)[OF nth_mem p'(4) q'(4)] p'(1 - 3) q'(1 - 3) eps using 1[OF p'(4), of _ B] 2[OF q'(4)] by (auto simp: gpair_code nth_zip_fill less_max_iff_disj intro!: exI[of _ "zip_fill ps qs"] exI[of _ "Some p'"] exI[of _ "Some q'"]) qed } qed lemma from_ta_der_pair_automaton: "(None, None) |\| ta_der (pair_automaton A B) (term_of_gterm s)" "(Some p, None) |\| ta_der (pair_automaton A B) (term_of_gterm s) \ \t. p |\| ta_der A (term_of_gterm t) \ s = map_gterm (\f. (Some f, None)) t" "(None, Some q) |\| ta_der (pair_automaton A B) (term_of_gterm s) \ \u. q |\| ta_der B (term_of_gterm u) \ s = map_gterm (\f. (None, Some f)) u" "(Some p, Some q) |\| ta_der (pair_automaton A B) (term_of_gterm s) \ \t u. p |\| ta_der A (term_of_gterm t) \ q |\| ta_der B (term_of_gterm u) \ s = gpair t u" proof (goal_cases nn sn ns ss) let ?AB = "pair_automaton A B" { fix p s A B assume "(Some p, None) |\| ta_der (pair_automaton A B) (term_of_gterm s)" then have "\t. p |\| ta_der A (term_of_gterm t) \ s = map_gterm (\f. (Some f, None)) t" proof (induct s arbitrary: p) case (GFun h ss) obtain rs sp nq where r: "h rs \ (sp, nq) |\| rules (pair_automaton A B)" "sp = Some p \ nq = None \ ((sp, nq), (Some p, None)) |\| (eps (pair_automaton A B))|\<^sup>+|" "length rs = length ss" "\i. i < length ss \ rs ! i |\| ta_der (pair_automaton A B) (term_of_gterm (ss ! i))" using GFun(2) by auto obtain p' where pq: "sp = Some p'" "nq = None" "p' = p \ (p', p) |\| (eps A)|\<^sup>+|" using r(2) pair_eps_Some_None[of sp nq p A B] by (cases sp) auto obtain f ps where h: "h = lift_Some_None f" "rs = map lift_Some_None ps" "f ps \ p' |\| rules A" using r(1) unfolding pq(1, 2) by (auto simp: pair_automaton_def map_ta_rule_cases) obtain ts where "\i. i < length ss \ ps ! i |\| ta_der A (term_of_gterm (ts i)) \ ss ! i = map_gterm (\f. (Some f, None)) (ts i)" using GFun(1)[OF nth_mem, of i "ps ! i" for i] r(4)[unfolded h(2)] r(3)[unfolded h(2) length_map] by auto metis then show ?case using h(3) pq(3) r(3)[unfolded h(2) length_map] by (intro exI[of _ "GFun f (map ts [0..u. q |\| ta_der B (term_of_gterm u) \ s = map_gterm (\g. (None, Some g)) u" if "(None, Some q) |\| ta_der ?AB (term_of_gterm s)" for q s using 1[OF swap_pair_automaton, OF that] by (auto simp: gterm.map_comp comp_def gterm.map_ident dest!: arg_cong[of "map_gterm prod.swap _" _ "map_gterm prod.swap", unfolded gterm.map_comp]) { case nn then show ?case by (intro ta_der_not_reach) (auto simp: pair_automaton_def map_ta_rule_cases) next case sn then show ?case by (rule 1) next case ns then show ?case by (rule 2) next case ss then show ?case proof (induct s arbitrary: p q) case (GFun h ss) obtain rs sp sq where r: "h rs \ (sp, sq) |\| rules ?AB" "sp = Some p \ sq = Some q \ ((sp, sq), (Some p, Some q)) |\| (eps ?AB)|\<^sup>+|" "length rs = length ss" "\i. i < length ss \ rs ! i |\| ta_der ?AB (term_of_gterm (ss ! i))" using GFun(2) by auto from r(2) have "sp \ None" "sq \ None" using pair_eps_Some_Some2[of "(sp, sq)" p q] by auto then obtain p' q' where pq: "sp = Some p'" "sq = Some q'" "p' = p \ (p', p) |\| (eps A)|\<^sup>+|" "q' = q \ (q', q) |\| (eps B)|\<^sup>+|" using r(2) pair_eps_Some_Some[where ?r = "(Some p, Some q)" and ?A = A and ?B = B] using pair_eps_Some_Some2[of "(sp, sq)" p q] by (cases sp; cases sq) auto obtain f g ps qs where h: "h = (Some f, Some g)" "rs = zip_fill ps qs" "f ps \ p' |\| rules A" "g qs \ q' |\| rules B" using r(1) unfolding pq(1,2) by (auto simp: pair_automaton_def map_ta_rule_cases) have *: "\t u. ps ! i |\| ta_der A (term_of_gterm t) \ qs ! i |\| ta_der B (term_of_gterm u) \ ss ! i = gpair t u" if "i < length ps" "i < length qs" for i using GFun(1)[OF nth_mem, of i "ps ! i" "qs ! i"] r(4)[unfolded pq(1,2) h(2), of i] that r(3)[symmetric] by (auto simp: nth_zip_fill h(2)) { fix i assume "i < length ss" then have "\t u. (i < length ps \ ps ! i |\| ta_der A (term_of_gterm t)) \ (i < length qs \ qs ! i |\| ta_der B (term_of_gterm u)) \ ss ! i = gpair_impl (if i < length ps then Some t else None) (if i < length qs then Some u else None)" using *[of i] 1[of "ps ! i" A B "ss ! i"] 2[of "qs ! i" "ss ! i"] r(4)[of i] r(3)[symmetric] by (cases "i < length ps"; cases "i < length qs") (auto simp add: h(2) nth_zip_fill less_max_iff_disj gpair_code split: gterm.splits) } then obtain ts us where "\i. i < length ss \ (i < length ps \ ps ! i |\| ta_der A (term_of_gterm (ts i))) \ (i < length qs \ qs ! i |\| ta_der B (term_of_gterm (us i))) \ ss ! i = gpair_impl (if i < length ps then Some (ts i) else None) (if i < length qs then Some (us i) else None)" by metis moreover then have "\i. i < length ps \ ps ! i |\| ta_der A (term_of_gterm (ts i))" "\i. i < length qs \ qs ! i |\| ta_der B (term_of_gterm (us i))" using r(3)[unfolded h(2)] by auto ultimately show ?case using h(3,4) pq(3,4) r(3)[symmetric] by (intro exI[of _ "GFun f (map ts [0..f. (Some f, Some f)) A) {(s, s) | s. s \ R}" using assms by (auto simp: RR2_spec_def RR1_spec_def fmap_funs_reg_def fmap_funs_gta_lang map_funs_term_some_gpair \_def) lemma pair_automaton: assumes "RR1_spec A T" "RR1_spec B U" shows "RR2_spec (pair_automaton_reg A B) (T \ U)" proof - let ?AB = "pair_automaton (ta A) (ta B)" { fix t u assume t: "t \ T" and u: "u \ U" obtain p where p: "p |\| fin A" "p |\| ta_der (ta A) (term_of_gterm t)" using t assms(1) by (auto simp: RR1_spec_def gta_lang_def \_def gta_der_def) obtain q where q: "q |\| fin B" "q |\| ta_der (ta B) (term_of_gterm u)" using u assms(2) by (auto simp: RR1_spec_def gta_lang_def \_def gta_der_def) have "gpair t u \ \ (pair_automaton_reg A B)" using p(1) q(1) to_ta_der_pair_automaton(3)[OF p(2) q(2)] by (auto simp: pair_automaton_reg_def \_def) } moreover { fix s assume "s \ \ (pair_automaton_reg A B)" from this[unfolded gta_lang_def \_def] obtain r where r: "r |\| fin (pair_automaton_reg A B)" "r |\| gta_der ?AB s" by (auto simp: pair_automaton_reg_def) obtain p q where pq: "r = (Some p, Some q)" "p |\| fin A" "q |\| fin B" using r(1) by (auto simp: pair_automaton_reg_def) from from_ta_der_pair_automaton(4)[OF r(2)[unfolded pq(1) gta_der_def]] obtain t u where "p |\| ta_der (ta A) (term_of_gterm t)" "q |\| ta_der (ta B) (term_of_gterm u)" "s = gpair t u" by (elim exE conjE) note * = this then have "t \ \ A" "u \ \ B" using pq(2,3) by (auto simp: \_def) then have "\t u. s = gpair t u \ t \ T \ u \ U" using assms by (auto simp: RR1_spec_def *(3) intro!: exI[of _ t, OF exI[of _ u]]) } ultimately show ?thesis by (auto simp: RR2_spec_def) qed lemma pair_automaton': shows "\ (pair_automaton_reg A B) = case_prod gpair ` (\ A \ \ B)" using pair_automaton[of A _ B] by (auto simp: RR1_spec_def RR2_spec_def) subsection \Collapsing\ text \cf. @{const "gcollapse"}.\ fun collapse_state_list where "collapse_state_list Qn Qs [] = [[]]" | "collapse_state_list Qn Qs (q # qs) = (let rec = collapse_state_list Qn Qs qs in (if q |\| Qn \ q |\| Qs then map (Cons None) rec @ map (Cons (Some q)) rec else if q |\| Qn then map (Cons None) rec else if q |\| Qs then map (Cons (Some q)) rec else [[]]))" lemma collapse_state_list_inner_length: assumes "qss = collapse_state_list Qn Qs qs" and "\ i < length qs. qs ! i |\| Qn \ qs ! i |\| Qs" and "i < length qss" shows "length (qss ! i) = length qs" using assms proof (induct qs arbitrary: qss i) case (Cons q qs) have "\i| Qn \ qs ! i |\| Qs" using Cons(3) by auto then show ?case using Cons(1)[of "collapse_state_list Qn Qs qs"] Cons(2-) by (auto simp: Let_def nth_append) qed auto lemma collapse_fset_inv_constr: assumes "\ i < length qs'. qs ! i |\| Qn \ qs' ! i = None \ qs ! i |\| Qs \ qs' ! i = Some (qs ! i)" and "length qs = length qs'" shows "qs' |\| fset_of_list (collapse_state_list Qn Qs qs)" using assms proof (induct qs arbitrary: qs') case (Cons q qs) have "(tl qs') |\| fset_of_list (collapse_state_list Qn Qs qs)" using Cons(2-) by (intro Cons(1)[of "tl qs'"]) (cases qs', auto) then show ?case using Cons(2-) by (cases qs') (auto simp: Let_def image_def) qed auto lemma collapse_fset_inv_constr2: assumes "\ i < length qs. qs ! i |\| Qn \ qs ! i |\| Qs" and "qs' |\| fset_of_list (collapse_state_list Qn Qs qs)" and "i < length qs'" shows "qs ! i |\| Qn \ qs' ! i = None \ qs ! i |\| Qs \ qs' ! i = Some (qs ! i)" using assms proof (induct qs arbitrary: qs' i) case (Cons a qs) have "i \ 0 \ qs ! (i - 1) |\| Qn \ tl qs' ! (i - 1) = None \ qs ! (i - 1) |\| Qs \ tl qs' ! (i - 1) = Some (qs ! (i - 1))" using Cons(2-) by (intro Cons(1)[of "tl qs'" "i - 1"]) (auto simp: Let_def split: if_splits) then show ?case using Cons(2-) by (cases i) (auto simp: Let_def split: if_splits) qed auto definition collapse_rule where "collapse_rule A Qn Qs = |\| ((\ r. fset_of_list (map (\ qs. TA_rule (r_root r) qs (Some (r_rhs r))) (collapse_state_list Qn Qs (r_lhs_states r)))) |`| ffilter (\ r. (\ i < length (r_lhs_states r). r_lhs_states r ! i |\| Qn \ r_lhs_states r ! i |\| Qs)) (ffilter (\ r. r_root r \ None) (rules A)))" definition collapse_rule_fset where "collapse_rule_fset A Qn Qs = (\ r. TA_rule (the (r_root r)) (map the (filter (\q. \ Option.is_none q) (r_lhs_states r))) (the (r_rhs r))) |`| collapse_rule A Qn Qs" lemma collapse_rule_set_conv: "fset (collapse_rule_fset A Qn Qs) = {TA_rule f (map the (filter (\q. \ Option.is_none q) qs')) q | f qs qs' q. TA_rule (Some f) qs q |\| rules A \ length qs = length qs' \ (\i < length qs. qs ! i |\| Qn \ qs' ! i = None \ qs ! i |\| Qs \ (qs' ! i) = Some (qs ! i))} " (is "?Ls = ?Rs") proof {fix f qs' q qs assume ass: "TA_rule (Some f) qs q |\| rules A" "length qs = length qs'" "\i < length qs. qs ! i |\| Qn \ qs' ! i = None \ qs ! i |\| Qs \ (qs' ! i) = Some (qs ! i)" then have "TA_rule (Some f) qs' (Some q) |\| collapse_rule A Qn Qs" using collapse_fset_inv_constr[of qs' qs Qn Qs] unfolding collapse_rule_def - by (auto simp: fset_of_list.rep_eq fimage_iff intro!: fBexI[of _ "TA_rule (Some f) qs q"]) + by (auto simp: fset_of_list.rep_eq fimage_iff intro!: bexI[of _ "TA_rule (Some f) qs q"]) then have "TA_rule f (map the (filter (\q. \ Option.is_none q) qs')) q \ ?Ls" unfolding collapse_rule_fset_def by (auto simp: image_iff Bex_def intro!: exI[of _"TA_rule (Some f) qs' (Some q)"])} then show "?Rs \ ?Ls" by blast next {fix f qs q assume ass: "TA_rule f qs q \ ?Ls" then obtain ps qs' where w: "TA_rule (Some f) ps q |\| rules A" "\ i < length ps. ps ! i |\| Qn \ ps ! i |\| Qs" "qs' |\| fset_of_list (collapse_state_list Qn Qs ps)" "qs = map the (filter (\q. \ Option.is_none q) qs')" unfolding collapse_rule_fset_def collapse_rule_def by (auto simp: ffUnion.rep_eq fset_of_list.rep_eq) (metis ta_rule.collapse) then have "\ i < length qs'. ps ! i |\| Qn \ qs' ! i = None \ ps ! i |\| Qs \ qs' ! i = Some (ps ! i)" using collapse_fset_inv_constr2 by metis moreover have "length ps = length qs'" using collapse_state_list_inner_length[of _ Qn Qs ps] using w(2, 3) calculation(1) by (auto simp: in_fset_conv_nth) ultimately have "TA_rule f qs q \ ?Rs" using w(1) unfolding w(4) by auto fastforce} then show "?Ls \ ?Rs" by (intro subsetI) (metis (no_types, lifting) ta_rule.collapse) qed lemma collapse_rule_fmember [simp]: "TA_rule f qs q |\| (collapse_rule_fset A Qn Qs) \ (\ qs' ps. qs = map the (filter (\q. \ Option.is_none q) qs') \ TA_rule (Some f) ps q |\| rules A \ length ps = length qs' \ (\i < length ps. ps ! i |\| Qn \ qs' ! i = None \ ps ! i |\| Qs \ (qs' ! i) = Some (ps ! i)))" unfolding collapse_rule_set_conv by auto definition "Qn A \ (let S = (r_rhs |`| ffilter (\ r. r_root r = None) (rules A)) in (eps A)|\<^sup>+| |``| S |\| S)" definition "Qs A \ (let S = (r_rhs |`| ffilter (\ r. r_root r \ None) (rules A)) in (eps A)|\<^sup>+| |``| S |\| S)" lemma Qn_member_iff [simp]: "q |\| Qn A \ (\ ps p. TA_rule None ps p |\| rules A \ (p = q \ (p, q) |\| (eps A)|\<^sup>+|))" (is "?Ls \ ?Rs") proof - {assume ass: "q |\| Qn A" then obtain r where "r_rhs r = q \ (r_rhs r, q) |\| (eps A)|\<^sup>+|" "r |\| rules A" "r_root r = None" by (force simp: Qn_def Image_def image_def Let_def fImage.rep_eq) then have "?Ls \ ?Rs" by (cases r) auto} moreover have "?Rs \ ?Ls" by (force simp: Qn_def Image_def image_def Let_def fImage.rep_eq) ultimately show ?thesis by blast qed lemma Qs_member_iff [simp]: "q |\| Qs A \ (\ f ps p. TA_rule (Some f) ps p |\| rules A \ (p = q \ (p, q) |\| (eps A)|\<^sup>+|))" (is "?Ls \ ?Rs") proof - {assume ass: "q |\| Qs A" then obtain f r where "r_rhs r = q \ (r_rhs r, q) |\| (eps A)|\<^sup>+|" "r |\| rules A" "r_root r = Some f" by (force simp: Qs_def Image_def image_def Let_def fImage.rep_eq) then have "?Ls \ ?Rs" by (cases r) auto} moreover have "?Rs \ ?Ls" by (force simp: Qs_def Image_def image_def Let_def fImage.rep_eq) ultimately show ?thesis by blast qed lemma collapse_Qn_Qs_set_conv: "fset (Qn A) = {q' |qs q q'. TA_rule None qs q |\| rules A \ (q = q' \ (q, q') |\| (eps A)|\<^sup>+|)}" (is "?Ls1 = ?Rs1") "fset (Qs A) = {q' |f qs q q'. TA_rule (Some f) qs q |\| rules A \ (q = q' \ (q, q') |\| (eps A)|\<^sup>+|)}" (is "?Ls2 = ?Rs2") by auto force+ definition collapse_automaton :: "('q, 'f option) ta \ ('q, 'f) ta" where "collapse_automaton A = TA (collapse_rule_fset A (Qn A) (Qs A)) (eps A)" definition collapse_automaton_reg where "collapse_automaton_reg R = Reg (fin R) (collapse_automaton (ta R))" lemma ta_states_collapse_automaton: "\ (collapse_automaton A) |\| \ A" apply (intro \_subseteq_I) apply (auto simp: collapse_automaton_def collapse_Qn_Qs_set_conv collapse_rule_set_conv fset_of_list.rep_eq in_set_conv_nth rule_statesD eps_statesD[unfolded]) apply (metis Option.is_none_def fnth_mem option.sel rule_statesD(3) ta_rule.sel(2)) done lemma last_nthI: assumes "i < length ts" "\ i < length ts - Suc 0" shows "ts ! i = last ts" using assms by (induct ts arbitrary: i) (auto, metis last_conv_nth length_0_conv less_antisym nth_Cons') lemma collapse_automaton': assumes "\ A |\| ta_reachable A" (* cf. ta_trim *) shows "gta_lang Q (collapse_automaton A) = the ` (gcollapse ` gta_lang Q A - {None})" proof - define Qn' where "Qn' = Qn A" define Qs' where "Qs' = Qs A" {fix t assume t: "t \ gta_lang Q (collapse_automaton A)" then obtain q where q: "q |\| Q" "q |\| ta_der (collapse_automaton A) (term_of_gterm t)" by auto have "\ t'. q |\| ta_der A (term_of_gterm t') \ gcollapse t' \ None \ t = the (gcollapse t')" using q(2) proof (induct rule: ta_der_gterm_induct) case (GFun f ts ps p q) from GFun(1 - 3) obtain qs rs where ps: "TA_rule (Some f) qs p |\| rules A" "length qs = length rs" "\i. i < length qs \ qs ! i |\| Qn' \ rs ! i = None \ qs ! i |\| Qs' \ rs ! i = Some (qs ! i)" "ps = map the (filter (\q. \ Option.is_none q) rs)" by (auto simp: collapse_automaton_def Qn'_def Qs'_def) obtain ts' where ts': "ps ! i |\| ta_der A (term_of_gterm (ts' i))" "gcollapse (ts' i) \ None" "ts ! i = the (gcollapse (ts' i))" if "i < length ts" for i using GFun(5) by metis from ps(2, 3, 4) have rs: "i < length qs \ rs ! i = None \ rs ! i = Some (qs ! i)" for i by auto {fix i assume "i < length qs" "rs ! i = None" then have "\ t'. groot_sym t' = None \ qs ! i |\| ta_der A (term_of_gterm t')" using ps(1, 2) ps(3)[of i] by (auto simp: ta_der_trancl_eps Qn'_def groot_sym_groot_conv elim!: ta_reachable_rule_gtermE[OF assms]) (force dest: ta_der_trancl_eps)+} note None = this {fix i assume i: "i < length qs" "rs ! i = Some (qs ! i)" have "map Some ps = filter (\q. \ Option.is_none q) rs" using ps(4) by (induct rs arbitrary: ps) (auto simp add: Option.is_none_def) from filter_rev_nth_idx[OF _ _ this, of i] have *: "rs ! i = map Some ps ! filter_rev_nth (\q. \ Option.is_none q) rs i" "filter_rev_nth (\q. \ Option.is_none q) rs i < length ps" using ps(2, 4) i by auto then have "the (rs ! i) = ps ! filter_rev_nth (\q. \ Option.is_none q) rs i" "filter_rev_nth (\q. \ Option.is_none q) rs i < length ps" by auto} note Some = this let ?P = "\ t i. qs ! i |\| ta_der A (term_of_gterm t) \ (rs ! i = None \ groot_sym t = None) \ (rs ! i = Some (qs ! i) \ t = ts' (filter_rev_nth (\q. \ Option.is_none q) rs i))" {fix i assume i: "i < length qs" then have "\ t. ?P t i" using Some[OF i] None[OF i] ts' ps(2, 4) GFun(2) rs[OF i] by (cases "rs ! i") auto} then obtain ts'' where ts'': "length ts'' = length qs" "i < length qs \ qs ! i |\| ta_der A (term_of_gterm (ts'' ! i))" "i < length qs \ rs ! i = None \ groot_sym (ts'' ! i) = None" "i < length qs \ rs ! i = Some (qs ! i) \ ts'' ! i = ts' (filter_rev_nth (\q. \ Option.is_none q) rs i)" for i using that Ex_list_of_length_P[of "length qs" ?P] by auto from ts''(1, 3, 4) Some ps(2, 4) GFun(2) rs ts'(2-) have "map Some ts = (filter (\q. \ Option.is_none q) (map gcollapse ts''))" proof (induct ts'' arbitrary: qs rs ps ts rule: rev_induct) case (snoc a us) from snoc(2, 7) obtain r rs' where [simp]: "rs = rs' @ [r]" by (metis append_butlast_last_id append_is_Nil_conv length_0_conv not_Cons_self2) have l: "length us = length (butlast qs)" "length (butlast qs) = length (butlast rs)" using snoc(2, 7) by auto have *: "i < length (butlast qs) \ butlast rs ! i = None \ groot_sym (us ! i) = None" for i using snoc(3)[of i] snoc(2, 7) by (auto simp:nth_append l(1) nth_butlast split!: if_splits) have **: "i < length (butlast qs) \ butlast rs ! i = None \ butlast rs ! i = Some (butlast qs ! i)" for i using snoc(10)[of i] snoc(2, 7) l by (auto simp: nth_append nth_butlast) have "i < length (butlast qs) \ butlast rs ! i = Some (butlast qs ! i) \ us ! i = ts' (filter_rev_nth (\q. \ Option.is_none q) (butlast rs) i)" for i using snoc(4)[of i] snoc(2, 7) l by (auto simp: nth_append nth_butlast filter_rev_nth_def take_butlast) note IH = snoc(1)[OF l(1) * this _ _ l(2) _ _ **] show ?case proof (cases "r = None") case True then have "map Some ts = filter (\q. \ Option.is_none q) (map gcollapse us)" using snoc(2, 5-) by (intro IH[of ps ts]) (auto simp: nth_append nth_butlast filter_rev_nth_butlast) then show ?thesis using True snoc(2, 7) snoc(3)[of "length (butlast rs)"] by (auto simp: nth_append l(1) last_nthI split!: if_splits) next case False then obtain t' ss where *: "ts = ss @ [t']" using snoc(2, 7, 8, 9) by (cases ts) (auto, metis append_butlast_last_id list.distinct(1)) let ?i = "filter_rev_nth (\q. \ Option.is_none q) rs (length us)" have "map Some (butlast ts) = filter (\q. \ Option.is_none q) (map gcollapse us)" using False snoc(2, 5-) l filter_rev_nth_idx by (intro IH[of "butlast ps" "butlast ts"]) (fastforce simp: nth_butlast filter_rev_nth_butlast)+ moreover have "a = ts' ?i" "?i < length ps" using False snoc(2, 9) l snoc(4, 6, 10)[of "length us"] by (auto simp: nth_append) moreover have "filter_rev_nth (\q. \ Option.is_none q) (rs' @ [r]) (length rs') = length ss" using l snoc(2, 7, 8, 9) False unfolding * by (auto simp: filter_rev_nth_def) ultimately show ?thesis using l snoc(2, 7, 9) snoc(11-)[of ?i] by (auto simp: nth_append *) qed qed simp then have "ts = map the (filter (\t. \ Option.is_none t) (map gcollapse ts''))" by (metis comp_the_Some list.map_id map_map) then show ?case using ps(1, 2) ts''(1, 2) GFun(3) by (auto simp: collapse_automaton_def intro!: exI[of _ "GFun (Some f) ts''"] exI[of _ qs] exI[of _ p]) qed then have "t \ the ` (gcollapse ` gta_lang Q A - {None})" by (meson Diff_iff gta_langI imageI q(1) singletonD) } moreover { fix t assume t: "t \ gta_lang Q A" "gcollapse t \ None" obtain q where q: "q |\| Q" "q |\| ta_der A (term_of_gterm t)" using t(1) by auto have "q |\| ta_der (collapse_automaton A) (term_of_gterm (the (gcollapse t)))" using q(2) t(2) proof (induct t arbitrary: q) case (GFun f ts) obtain qs q' where q: "TA_rule f qs q' |\| rules A" "q' = q \ (q', q) |\| (eps (collapse_automaton A))|\<^sup>+|" "length qs = length ts" "\i. i < length ts \ qs ! i |\| ta_der A (term_of_gterm (ts ! i))" using GFun(2) by (auto simp: collapse_automaton_def) obtain f' where f [simp]: "f = Some f'" using GFun(3) by (cases f) auto define qs' where "qs' = map (\i. if Option.is_none (gcollapse (ts ! i)) then None else Some (qs ! i)) [0.. qs ! i |\| Qn'" if "i < length qs" for i using q(4)[of i] that by (cases "ts ! i" rule: gcollapse.cases) (auto simp: q(3) Qn'_def collapse_Qn_Qs_set_conv) moreover have "\ Option.is_none (gcollapse (ts ! i)) \ qs ! i |\| Qs'" if "i < length qs" for i using q(4)[of i] that by (cases "ts ! i" rule: gcollapse.cases) (auto simp: q(3) Qs'_def collapse_Qn_Qs_set_conv) ultimately have "f' (map the (filter (\q. \ Option.is_none q) qs')) \ q' |\| rules (collapse_automaton A)" using q(1, 4) unfolding collapse_automaton_def Qn'_def[symmetric] Qs'_def[symmetric] by (fastforce simp: qs'_def q(3) intro: exI[of _ qs] exI[of _ qs'] split: if_splits) moreover have ***: "length (filter (\i.\ Option.is_none (gcollapse (ts ! i))) [0..t. \ Option.is_none (gcollapse t)) ts)" for ts by (subst length_map[of "(!) ts", symmetric] filter_map[unfolded comp_def, symmetric] map_nth)+ (rule refl) note this[of ts] moreover have "the (filter (\q. \ Option.is_none q) qs' ! i) |\| ta_der (collapse_automaton A) (term_of_gterm (the (filter (\t. \ Option.is_none t) (map gcollapse ts) ! i)))" if "i < length [x\ts . \ Option.is_none (gcollapse x)]" for i unfolding qs'_def using that q(3) GFun(1)[OF nth_mem q(4)] proof (induct ts arbitrary: qs rule: List.rev_induct) case (snoc t ts) have x1 [simp]: "i < length xs \ (xs @ ys) ! i = xs ! i" for xs ys i by (auto simp: nth_append) have x2: "i = length xs \ (xs @ ys) ! i = ys ! 0" for xs ys i by (auto simp: nth_append) obtain q qs' where qs [simp]: "qs = qs' @ [q]" using snoc(3) by (cases "rev qs") auto have [simp]: "map (\x. if Option.is_none (gcollapse ((ts @ [t]) ! x)) then None else Some ((qs' @ [q]) ! x)) [0..x. if Option.is_none (gcollapse (ts ! x)) then None else Some (qs' ! x)) [0..ts . \ Option.is_none (gcollapse x)]") case True then show ?thesis using snoc(3) snoc(4)[of "length ts"] unfolding qs length_append list.size add_0 add_0_right add_Suc_right upt_Suc_append[OF zero_le] filter_append map_append by (subst (5 6) x2) (auto simp: comp_def *** False' Option.is_none_def[symmetric]) next case False then show ?thesis using snoc(1)[of qs'] snoc(2,3) snoc(4)[unfolded length_append list.size add_0 add_0_right add_Suc_right, OF less_SucI] unfolding qs length_append list.size add_0 add_0_right add_Suc_right upt_Suc_append[OF zero_le] filter_append map_append by (subst (5 6) x1) (auto simp: comp_def *** False') qed qed qed auto ultimately show ?case using q(2) by (auto simp: qs'_def comp_def q(3) intro!: exI[of _ q'] exI[of _ "map the (filter (\q. \ Option.is_none q) qs')"]) qed then have "the (gcollapse t) \ gta_lang Q (collapse_automaton A)" by (metis q(1) gta_langI) } ultimately show ?thesis by blast qed lemma \_collapse_automaton': assumes "\\<^sub>r A |\| ta_reachable (ta A)" (* cf. ta_trim *) shows "\ (collapse_automaton_reg A) = the ` (gcollapse ` \ A - {None})" using assms by (auto simp: collapse_automaton_reg_def \_def collapse_automaton') lemma collapse_automaton: assumes "\\<^sub>r A |\| ta_reachable (ta A)" "RR1_spec A T" shows "RR1_spec (collapse_automaton_reg A) (the ` (gcollapse ` \ A - {None}))" using collapse_automaton'[OF assms(1)] assms(2) by (simp add: collapse_automaton_reg_def \_def RR1_spec_def) subsection \Cylindrification\ (* cylindrification is a product ("pairing") of a RR1 language accepting all terms, and an RRn language, modulo some fairly trivial renaming of symbols. *) definition pad_with_Nones where "pad_with_Nones n m = (\(f, g). case_option (replicate n None) id f @ case_option (replicate m None) id g)" lemma gencode_append: "gencode (ss @ ts) = map_gterm (pad_with_Nones (length ss) (length ts)) (gpair (gencode ss) (gencode ts))" proof - have [simp]: "p \ gposs (gunions (map gdomain ts)) \ map (\t. gfun_at t p) ts = replicate (length ts) None" for p ts by (intro nth_equalityI) (fastforce simp: poss_gposs_mem_conv fun_at_def' image_def all_set_conv_all_nth)+ note [simp] = glabel_map_gterm_conv[of "\(_ :: unit option). ()", unfolded comp_def gdomain_id] show ?thesis by (auto intro!: arg_cong[of _ _ "\x. glabel x _"] simp del: gposs_gunions simp: pad_with_Nones_def gencode_def gunions_append gpair_def map_gterm_glabel comp_def) qed lemma append_automaton: assumes "RRn_spec n A T" "RRn_spec m B U" shows "RRn_spec (n + m) (fmap_funs_reg (pad_with_Nones n m) (pair_automaton_reg A B)) {ts @ us |ts us. ts \ T \ us \ U}" using assms pair_automaton[of A "gencode ` T" B "gencode ` U"] unfolding RRn_spec_def proof (intro conjI set_eqI iffI, goal_cases) case (1 s) then obtain ts us where "ts \ T" "us \ U" "s = gencode (ts @ us)" by (fastforce simp: \_def fmap_funs_reg_def RR1_spec_def RR2_spec_def gencode_append fmap_funs_gta_lang) then show ?case by blast qed (fastforce simp: RR1_spec_def RR2_spec_def fmap_funs_reg_def \_def gencode_append fmap_funs_gta_lang)+ lemma cons_automaton: assumes "RR1_spec A T" "RRn_spec m B U" shows "RRn_spec (Suc m) (fmap_funs_reg (\(f, g). pad_with_Nones 1 m (map_option (\f. [Some f]) f, g)) (pair_automaton_reg A B)) {t # us |t us. t \ T \ us \ U}" proof - have [simp]: "{ts @ us |ts us. ts \ (\t. [t]) ` T \ us \ U} = {t # us |t us. t \ T \ us \ U}" by (auto intro: exI[of _ "[_]", OF exI]) show ?thesis using append_automaton[OF RR1_to_RRn_spec, OF assms] by (auto simp: \_def fmap_funs_reg_def pair_automaton_reg_def comp_def fmap_funs_gta_lang map_pair_automaton_12 fmap_funs_ta_comp prod.case_distrib gencode_append[of "[_]", unfolded gencode_singleton List.append.simps]) qed subsection \Projection\ (* projection is composed from fmap_funs_ta and collapse_automaton, corresponding to gsnd *) abbreviation "drop_none_rule m fs \ if list_all (Option.is_none) (drop m fs) then None else Some (drop m fs)" lemma drop_automaton_reg: assumes "\\<^sub>r A |\| ta_reachable (ta A)" "m < n" "RRn_spec n A T" defines "f \ \fs. drop_none_rule m fs" shows "RRn_spec (n - m) (collapse_automaton_reg (fmap_funs_reg f A)) (drop m ` T)" proof - have [simp]: "length ts = n - m ==> p \ gposs (gencode ts) \ \f. \t\set ts. Some f = gfun_at t p" for p ts proof (cases p, goal_cases Empty PCons) case Empty obtain t where "t \ set ts" using assms(2) Empty(1) by (cases ts) auto moreover then obtain f where "Some f = gfun_at t p" using Empty(3) by (cases t rule: gterm.exhaust) auto ultimately show ?thesis by auto next case (PCons i p') then have "p \ []" by auto then show ?thesis using PCons(2) by (auto 0 3 simp: gencode_def eq_commute[of "gfun_at _ _" "Some _"] elim!: gfun_at_possE) qed { fix p ts y assume that: "gfun_at (gencode ts) p = Some y" have "p \ gposs (gencode ts) \ gfun_at (gencode ts) p = Some (map (\t. gfun_at t p) ts)" by (auto simp: gencode_def) moreover have "gfun_at (gencode ts) p = Some y \ p \ gposs (gencode ts)" using gfun_at_nongposs by force ultimately have "y = map (\t. gfun_at t p) ts \ p \ gposs (gencode ts)" by (simp add: that) } note [dest!] = this have [simp]: "list_all f (replicate n x) \ n = 0 \ f x" for f n x by (induct n) auto have [dest]: "x \ set xs \ list_all f xs \ f x" for f x xs by (induct xs) auto have *: "f (pad_with_Nones m' n' z) = snd z" if "fst z = None \ fst z \ None \ length (the (fst z)) = m" "snd z = None \ snd z \ None \ length (the (snd z)) = n - m \ (\y. Some y \ set (the (snd z)))" "m' = m" "n' = n - m" for z m' n' using that by (auto simp: f_def pad_with_Nones_def split: option.splits prod.splits) { fix ts assume ts: "ts \ T" "length ts = n" have "gencode (drop m ts) = the (gcollapse (map_gterm f (gencode ts)))" "gcollapse (map_gterm f (gencode ts)) \ None" proof (goal_cases) case 1 show ?case using ts assms(2) apply (subst gsnd_gpair[of "gencode (take m ts)", symmetric]) apply (subst gencode_append[of "take m ts" "drop m ts", unfolded append_take_drop_id]) unfolding gsnd_def comp_def gterm.map_comp apply (intro arg_cong[where f = "\x. the (gcollapse x)"] gterm.map_cong refl) by (subst *) (auto simp: gpair_def set_gterm_gposs_conv image_def) next case 2 have [simp]: "gcollapse t = None \ gfun_at t [] = Some None" for t by (cases t rule: gcollapse.cases) auto obtain t where "t \ set (drop m ts)" using ts(2) assms(2) by (cases "drop m ts") auto moreover then have "\ Option.is_none (gfun_at t [])" by (cases t rule: gterm.exhaust) auto ultimately show ?case by (auto simp: f_def gencode_def list_all_def drop_map) qed } then show ?thesis using assms(3) by (fastforce simp: \_def collapse_automaton_reg_def fmap_funs_reg_def RRn_spec_def fmap_funs_gta_lang gsnd_def gpair_def gterm.map_comp comp_def glabel_map_gterm_conv[unfolded comp_def] assms(1) collapse_automaton') qed lemma gfst_collapse_simp: "the (gcollapse (map_gterm fst t)) = gfst t" by (simp add: gfst_def) lemma gsnd_collapse_simp: "the (gcollapse (map_gterm snd t)) = gsnd t" by (simp add: gsnd_def) definition proj_1_reg where "proj_1_reg A = collapse_automaton_reg (fmap_funs_reg fst (trim_reg A))" definition proj_2_reg where "proj_2_reg A = collapse_automaton_reg (fmap_funs_reg snd (trim_reg A))" lemmas proj_1_reg_simp = proj_1_reg_def collapse_automaton_reg_def fmap_funs_reg_def trim_reg_def lemmas proj_2_reg_simp = proj_2_reg_def collapse_automaton_reg_def fmap_funs_reg_def trim_reg_def lemma \_proj_1_reg_collapse: "\ (proj_1_reg \) = the ` (gcollapse ` map_gterm fst ` (\ \) - {None})" proof - have "\\<^sub>r (fmap_funs_reg fst (trim_reg \)) |\| ta_reachable (ta (fmap_funs_reg fst (trim_reg \)))" by (auto simp: fmap_funs_reg_def) note [simp] = \_collapse_automaton'[OF this] show ?thesis by (auto simp: proj_1_reg_def fmap_funs_\ \_trim) qed lemma \_proj_2_reg_collapse: "\ (proj_2_reg \) = the ` (gcollapse ` map_gterm snd ` (\ \) - {None})" proof - have "\\<^sub>r (fmap_funs_reg snd (trim_reg \)) |\| ta_reachable (ta (fmap_funs_reg snd (trim_reg \)))" by (auto simp: fmap_funs_reg_def) note [simp] = \_collapse_automaton'[OF this] show ?thesis by (auto simp: proj_2_reg_def fmap_funs_\ \_trim) qed lemma proj_1: assumes "RR2_spec A R" shows "RR1_spec (proj_1_reg A) (fst ` R)" proof - {fix s t assume ass: "(s, t) \ R" from ass have s: "s = the (gcollapse (map_gterm fst (gpair s t)))" by (auto simp: gfst_gpair gfst_collapse_simp) then have "Some s = gcollapse (map_gterm fst (gpair s t))" by (cases s; cases t) (auto simp: gpair_def) then have "s \ \ (proj_1_reg A)" using assms ass s by (auto simp: proj_1_reg_simp \_def trim_ta_reach trim_gta_lang image_def image_Collect RR2_spec_def fmap_funs_gta_lang collapse_automaton'[of "fmap_funs_ta fst (trim_ta (fin A) (ta A))"]) force} moreover {fix s assume "s \ \ (proj_1_reg A)" then have "s \ fst ` R" using assms by (auto simp: gfst_collapse_simp gfst_gpair rev_image_eqI RR2_spec_def trim_ta_reach trim_gta_lang \_def proj_1_reg_simp fmap_funs_gta_lang collapse_automaton'[of "fmap_funs_ta fst (trim_ta (fin A) (ta A))"])} ultimately show ?thesis using assms unfolding RR2_spec_def RR1_spec_def \_def proj_1_reg_simp by auto qed lemma proj_2: assumes "RR2_spec A R" shows "RR1_spec (proj_2_reg A) (snd ` R)" proof - {fix s t assume ass: "(s, t) \ R" then have s: "t = the (gcollapse (map_gterm snd (gpair s t)))" by (auto simp: gsnd_gpair gsnd_collapse_simp) then have "Some t = gcollapse (map_gterm snd (gpair s t))" by (cases s; cases t) (auto simp: gpair_def) then have "t \ \ (proj_2_reg A)" using assms ass s by (auto simp: \_def trim_ta_reach trim_gta_lang proj_2_reg_simp image_def image_Collect RR2_spec_def fmap_funs_gta_lang collapse_automaton'[of "fmap_funs_ta snd (trim_ta (fin A) (ta A))"]) fastforce} moreover {fix s assume "s \ \ (proj_2_reg A)" then have "s \ snd ` R" using assms by (auto simp: \_def gsnd_collapse_simp gsnd_gpair rev_image_eqI RR2_spec_def trim_ta_reach trim_gta_lang proj_2_reg_simp fmap_funs_gta_lang collapse_automaton'[of "fmap_funs_ta snd (trim_ta (fin A) (ta A))"])} ultimately show ?thesis using assms unfolding RR2_spec_def RR1_spec_def by auto qed lemma \_proj: assumes "RR2_spec A R" shows "\ (proj_1_reg A) = gfst ` \ A" "\ (proj_2_reg A) = gsnd ` \ A" proof - have [simp]: "gfst ` {gpair t u |t u. (t, u) \ R} = fst ` R" by (force simp: gfst_gpair image_def) have [simp]: "gsnd ` {gpair t u |t u. (t, u) \ R} = snd ` R" by (force simp: gsnd_gpair image_def) show "\ (proj_1_reg A) = gfst ` \ A" "\ (proj_2_reg A) = gsnd ` \ A" using proj_1[OF assms] proj_2[OF assms] assms gfst_gpair gsnd_gpair by (auto simp: RR1_spec_def RR2_spec_def) qed lemmas proj_automaton_gta_lang = proj_1 proj_2 subsection \Permutation\ (* permutations are a direct application of fmap_funs_ta *) lemma gencode_permute: assumes "set ps = {0..xs. map ((!) xs) ps) (gencode ts)" proof - have *: "(!) ts ` set ps = set ts" using assms by (auto simp: image_def set_conv_nth) show ?thesis using subsetD[OF equalityD1[OF assms]] apply (intro eq_gterm_by_gposs_gfun_at) unfolding gencode_def gposs_glabel gposs_map_gterm gposs_gunions gfun_at_map_gterm gfun_at_glabel set_map * by auto qed lemma permute_automaton: assumes "RRn_spec n A T" "set ps = {0..xs. map ((!) xs) ps) A) ((\xs. map ((!) xs) ps) ` T)" using assms by (auto simp: RRn_spec_def gencode_permute fmap_funs_reg_def \_def fmap_funs_gta_lang image_def) subsection \Intersection\ (* intersection is already defined in IsaFoR *) lemma intersect_automaton: assumes "RRn_spec n A T" "RRn_spec n B U" shows "RRn_spec n (reg_intersect A B) (T \ U)" using assms by (simp add: RRn_spec_def \_intersect) (metis gdecode_gencode image_Int inj_on_def) (* lemma swap_union_automaton: "fmap_states_ta (case_sum Inr Inl) (union_automaton B A) = union_automaton A B" by (simp add: fmap_states_ta_def' union_automaton_def image_Un image_comp case_sum_o_inj ta_rule.map_comp prod.map_comp comp_def id_def ac_simps) *) lemma union_automaton: assumes "RRn_spec n A T" "RRn_spec n B U" shows "RRn_spec n (reg_union A B) (T \ U)" using assms by (auto simp: RRn_spec_def \_union) subsection \Difference\ lemma RR1_difference: assumes "RR1_spec A T" "RR1_spec B U" shows "RR1_spec (difference_reg A B) (T - U)" using assms by (auto simp: RR1_spec_def \_difference_reg) lemma RR2_difference: assumes "RR2_spec A T" "RR2_spec B U" shows "RR2_spec (difference_reg A B) (T - U)" using assms by (auto simp: RR2_spec_def \_difference_reg) lemma RRn_difference: assumes "RRn_spec n A T" "RRn_spec n B U" shows "RRn_spec n (difference_reg A B) (T - U)" using assms by (auto simp: RRn_spec_def \_difference_reg) (metis gdecode_gencode)+ subsection \All terms over a signature\ definition term_automaton :: "('f \ nat) fset \ (unit, 'f) ta" where "term_automaton \ = TA ((\ (f, n). TA_rule f (replicate n ()) ()) |`| \) {||}" definition term_reg where "term_reg \ = Reg {|()|} (term_automaton \)" lemma term_automaton: "RR1_spec (term_reg \) (\\<^sub>G (fset \))" unfolding RR1_spec_def gta_lang_def term_reg_def \_def proof (intro set_eqI iffI, goal_cases lr rl) case (lr t) then have "() |\| ta_der (term_automaton \) (term_of_gterm t)" by (auto simp: gta_der_def) then show ?case by (induct t) (auto simp: term_automaton_def split: if_splits) next case (rl t) then have "() |\| ta_der (term_automaton \) (term_of_gterm t)" proof (induct t rule: \\<^sub>G.induct) case (const a) then show ?case by (auto simp: term_automaton_def image_iff intro: bexI[of _ "(a, 0)"]) next case (ind f n ss) then show ?case by (auto simp: term_automaton_def image_iff intro: bexI[of _ "(f, n)"]) qed then show ?case by (auto simp: gta_der_def) qed fun true_RRn :: "('f \ nat) fset \ nat \ (nat, 'f option list) reg" where "true_RRn \ 0 = Reg {|0|} (TA {|TA_rule [] [] 0|} {||})" | "true_RRn \ (Suc 0) = relabel_reg (fmap_funs_reg (\f. [Some f]) (term_reg \))" | "true_RRn \ (Suc n) = relabel_reg (trim_reg (fmap_funs_reg (pad_with_Nones 1 n) (pair_automaton_reg (true_RRn \ 1) (true_RRn \ n))))" lemma true_RRn_spec: "RRn_spec n (true_RRn \ n) {ts. length ts = n \ set ts \ \\<^sub>G (fset \)}" proof (induct \ n rule: true_RRn.induct) case (1 \) show ?case by (simp cong: conj_cong add: true_RR0_spec) next case (2 \) moreover have "{ts. length ts = 1 \ set ts \ \\<^sub>G (fset \)} = (\t. [t]) ` \\<^sub>G (fset \)" apply (intro equalityI subsetI) subgoal for ts by (cases ts) auto by auto ultimately show ?case using RR1_to_RRn_spec[OF term_automaton, of \] by auto next case (3 \ n) have [simp]: "{ts @ us |ts us. length ts = n \ set ts \ \\<^sub>G (fset \) \ length us = m \ set us \ \\<^sub>G (fset \)} = {ss. length ss = n + m \ set ss \ \\<^sub>G (fset \)}" for n m by (auto 0 4 intro!: exI[of _ "take n _", OF exI[of _ "drop n _"], of _ xs xs for xs] dest!: subsetD[OF set_take_subset] subsetD[OF set_drop_subset]) show ?case using append_automaton[OF 3] by simp qed subsection \RR2 composition\ abbreviation "RR2_to_RRn A \ fmap_funs_reg (\(f, g). [f, g]) A" abbreviation "RRn_to_RR2 A \ fmap_funs_reg (\f. (f ! 0, f ! 1)) A" definition rr2_compositon where "rr2_compositon \ A B = (let A' = RR2_to_RRn A in let B' = RR2_to_RRn B in let F = true_RRn \ 1 in let CA = trim_reg (fmap_funs_reg (pad_with_Nones 2 1) (pair_automaton_reg A' F)) in let CB = trim_reg (fmap_funs_reg (pad_with_Nones 1 2) (pair_automaton_reg F B')) in let PI = trim_reg (fmap_funs_reg (\xs. map ((!) xs) [1, 0, 2]) (reg_intersect CA CB)) in RRn_to_RR2 (collapse_automaton_reg (fmap_funs_reg (drop_none_rule 1) PI)) )" lemma list_length1E: assumes "length xs = Suc 0" obtains x where "xs = [x]" using assms by (cases xs) auto lemma rr2_compositon: assumes "\ \ \\<^sub>G (fset \) \ \\<^sub>G (fset \)" "\ \ \\<^sub>G (fset \) \ \\<^sub>G (fset \)" and "RR2_spec A \" and "RR2_spec B \" shows "RR2_spec (rr2_compositon \ A B) (\ O \)" proof - let ?R = "(\(t, u). [t, u]) ` \" let ?L = "(\(t, u). [t, u]) ` \" let ?A = "RR2_to_RRn A" let ?B = "RR2_to_RRn B" let ?F = "true_RRn \ 1" let ?CA = "trim_reg (fmap_funs_reg (pad_with_Nones 2 1) (pair_automaton_reg ?A ?F))" let ?CB = "trim_reg (fmap_funs_reg (pad_with_Nones 1 2) (pair_automaton_reg ?F ?B))" let ?PI = "trim_reg (fmap_funs_reg (\xs. map ((!) xs) [1, 0, 2]) (reg_intersect ?CA ?CB))" let ?DR = "collapse_automaton_reg (fmap_funs_reg (drop_none_rule 1) ?PI)" let ?Rs = "{ts @ us | ts us. ts \ ?R \ (\t. us = [t] \ t \ \\<^sub>G (fset \))}" let ?Ls = "{us @ ts | ts us. ts \ ?L \ (\t. us = [t] \ t \ \\<^sub>G (fset \))}" from RR2_to_RRn_spec assms(3, 4) have rr2: "RRn_spec 2 ?A ?R" "RRn_spec 2 ?B ?L" by auto have *: "{ts. length ts = 1 \ set ts \ \\<^sub>G (fset \)} = {[t] | t. t \ \\<^sub>G (fset \)}" by (auto elim!: list_length1E) have F: "RRn_spec 1 ?F {[t] | t. t \ \\<^sub>G (fset \)}" using true_RRn_spec[of 1 \] unfolding * . have "RRn_spec 3 ?CA ?Rs" "RRn_spec 3 ?CB ?Ls" using append_automaton[OF rr2(1) F] append_automaton[OF F rr2(2)] by (auto simp: numeral_3_eq_3) (smt Collect_cong) from permute_automaton[OF intersect_automaton[OF this], of "[1, 0, 2]"] have "RRn_spec 3 ?PI ((\xs. map ((!) xs) [1, 0, 2]) ` (?Rs \ ?Ls))" by (auto simp: atLeast0_lessThan_Suc insert_commute numeral_2_eq_2 numeral_3_eq_3) from drop_automaton_reg[OF _ _ this, of 1] have sp: "RRn_spec 2 ?DR (drop 1 ` (\xs. map ((!) xs) [1, 0, 2]) ` (?Rs \ ?Ls))" by auto {fix s assume "s \ (\(t, u). [t, u]) ` (\ O \)" then obtain t u v where comp: "s = [t, u]" "(t, v) \ \" "(v, u) \ \" by (auto simp: image_iff relcomp_unfold split!: prod.split) then have "[t, v] \ ?R" "[v , u] \ ?L" "u \ \\<^sub>G (fset \)" "v \ \\<^sub>G (fset \)" "t \ \\<^sub>G (fset \)" using assms(1, 2) by (auto simp: image_iff relcomp_unfold split!: prod.splits) then have "[t, v, u] \ ?Rs" "[t, v, u] \ ?Ls" apply (simp_all) subgoal apply (rule exI[of _ "[t, v]"], rule exI[of _ "[u]"]) apply simp done subgoal apply (rule exI[of _ "[v, u]"], rule exI[of _ "[t]"]) apply simp done done then have "s \ drop 1 ` (\xs. map ((!) xs) [1, 0, 2]) ` (?Rs \ ?Ls)" unfolding comp(1) apply (simp add: image_def Bex_def) apply (rule exI[of _ "[v, t, u]"]) apply simp apply (rule exI[of _ "[t, v, u]"]) apply simp done} moreover have "drop 1 ` (\xs. map ((!) xs) [1, 0, 2]) ` (?Rs \ ?Ls) \ (\(t, u). [t, u]) ` (\ O \)" by (auto simp: image_iff relcomp_unfold Bex_def split!: prod.splits) ultimately have *: "drop 1 ` (\xs. map ((!) xs) [1, 0, 2]) ` (?Rs \ ?Ls) = (\(t, u). [t, u]) ` (\ O \)" by (simp add: subsetI subset_antisym) have **: "(\f. (f ! 0, f ! 1)) ` (\(t, u). [t, u]) ` (\ O \) = \ O \" by (force simp: image_def relcomp_unfold split!: prod.splits) show ?thesis using sp unfolding * using RRn_to_RR2_spec[where ?T = "(\(t, u). [t, u]) ` (\ O \)" and ?A = ?DR] unfolding ** by (auto simp: rr2_compositon_def Let_def image_iff) qed end \ No newline at end of file diff --git a/thys/Regular_Tree_Relations/Tree_Automata/Tree_Automata.thy b/thys/Regular_Tree_Relations/Tree_Automata/Tree_Automata.thy --- a/thys/Regular_Tree_Relations/Tree_Automata/Tree_Automata.thy +++ b/thys/Regular_Tree_Relations/Tree_Automata/Tree_Automata.thy @@ -1,2191 +1,2192 @@ section \Tree automaton\ theory Tree_Automata imports FSet_Utils "HOL-Library.Product_Lexorder" "HOL-Library.Option_ord" begin subsection \Tree automaton definition and functionality\ datatype ('q, 'f) ta_rule = TA_rule (r_root: 'f) (r_lhs_states: "'q list") (r_rhs: 'q) ("_ _ \ _" [51, 51, 51] 52) datatype ('q, 'f) ta = TA (rules: "('q, 'f) ta_rule fset") (eps: "('q \ 'q) fset") text \In many application we are interested in specific subset of all terms. If these can be captured by a tree automaton (identified by a state) then we say the set is regular. This gives the motivation for the following definition\ datatype ('q, 'f) reg = Reg (fin: "'q fset") (ta: "('q, 'f) ta") text \The state set induced by a tree automaton is implicit in our representation. We compute it based on the rules and epsilon transitions of a given tree automaton\ abbreviation rule_arg_states where "rule_arg_states \ \ |\| ((fset_of_list \ r_lhs_states) |`| \)" abbreviation rule_target_states where "rule_target_states \ \ (r_rhs |`| \)" definition rule_states where "rule_states \ \ rule_arg_states \ |\| rule_target_states \" definition eps_states where "eps_states \\<^sub>\ \ (fst |`| \\<^sub>\) |\| (snd |`| \\<^sub>\)" definition "\ \ = rule_states (rules \) |\| eps_states (eps \)" abbreviation "\\<^sub>r \ \ \ (ta \)" definition ta_rhs_states :: "('q, 'f) ta \ 'q fset" where "ta_rhs_states \ \ {| q | p q. (p |\| rule_target_states (rules \)) \ (p = q \ (p, q) |\| (eps \)|\<^sup>+|)|}" definition "ta_sig \ = (\ r. (r_root r, length (r_lhs_states r))) |`| (rules \)" subsubsection \Rechability of a term induced by a tree automaton\ (* The reachable states of some term. *) fun ta_der :: "('q, 'f) ta \ ('f, 'q) term \ 'q fset" where "ta_der \ (Var q) = {|q' | q'. q = q' \ (q, q') |\| (eps \)|\<^sup>+| |}" | "ta_der \ (Fun f ts) = {|q' | q' q qs. TA_rule f qs q |\| (rules \) \ (q = q' \ (q, q') |\| (eps \)|\<^sup>+|) \ length qs = length ts \ (\ i < length ts. qs ! i |\| ta_der \ (ts ! i))|}" (* The reachable mixed terms of some term. *) fun ta_der' :: "('q,'f) ta \ ('f,'q) term \ ('f,'q) term fset" where "ta_der' \ (Var p) = {|Var q | q. p = q \ (p, q) |\| (eps \)|\<^sup>+| |}" | "ta_der' \ (Fun f ts) = {|Var q | q. q |\| ta_der \ (Fun f ts)|} |\| {|Fun f ss | ss. length ss = length ts \ (\i < length ts. ss ! i |\| ta_der' \ (ts ! i))|}" text \Sometimes it is useful to analyse a concrete computation done by a tree automaton. To do this we introduce the notion of run which keeps track which states are computed in each subterm to reach a certain state.\ abbreviation "ex_rule_state \ fst \ groot_sym" abbreviation "ex_comp_state \ snd \ groot_sym" inductive run for \ where step: "length qs = length ts \ (\ i < length ts. run \ (qs ! i) (ts ! i)) \ TA_rule f (map ex_comp_state qs) q |\| (rules \) \ (q = q' \ (q, q') |\| (eps \)|\<^sup>+|) \ run \ (GFun (q, q') qs) (GFun f ts)" subsubsection \Language acceptance\ definition ta_lang :: "'q fset \ ('q, 'f) ta \ ('f, 'v) terms" where [code del]: "ta_lang Q \ = {adapt_vars t | t. ground t \ Q |\| ta_der \ t \ {||}}" definition gta_der where "gta_der \ t = ta_der \ (term_of_gterm t)" definition gta_lang where "gta_lang Q \ = {t. Q |\| gta_der \ t \ {||}}" definition \ where "\ \ = gta_lang (fin \) (ta \)" definition reg_Restr_Q\<^sub>f where "reg_Restr_Q\<^sub>f R = Reg (fin R |\| \\<^sub>r R) (ta R)" subsubsection \Trimming\ definition ta_restrict where "ta_restrict \ Q = TA {| TA_rule f qs q| f qs q. TA_rule f qs q |\| rules \ \ fset_of_list qs |\| Q \ q |\| Q |} (fRestr (eps \) Q)" definition ta_reachable :: "('q, 'f) ta \ 'q fset" where "ta_reachable \ = {|q| q. \ t. ground t \ q |\| ta_der \ t |}" definition ta_productive :: "'q fset \ ('q,'f) ta \ 'q fset" where "ta_productive P \ \ {|q| q q' C. q' |\| ta_der \ (C\Var q\) \ q' |\| P |}" text \An automaton is trim if all its states are reachable and productive.\ definition ta_is_trim :: "'q fset \ ('q, 'f) ta \ bool" where "ta_is_trim P \ \ \ q. q |\| \ \ \ q |\| ta_reachable \ \ q |\| ta_productive P \" definition reg_is_trim :: "('q, 'f) reg \ bool" where "reg_is_trim R \ ta_is_trim (fin R) (ta R)" text \We obtain a trim automaton by restriction it to reachable and productive states.\ abbreviation ta_only_reach :: "('q, 'f) ta \ ('q, 'f) ta" where "ta_only_reach \ \ ta_restrict \ (ta_reachable \)" abbreviation ta_only_prod :: "'q fset \ ('q,'f) ta \ ('q,'f) ta" where "ta_only_prod P \ \ ta_restrict \ (ta_productive P \)" definition reg_reach where "reg_reach R = Reg (fin R) (ta_only_reach (ta R))" definition reg_prod where "reg_prod R = Reg (fin R) (ta_only_prod (fin R) (ta R))" definition trim_ta :: "'q fset \ ('q, 'f) ta \ ('q, 'f) ta" where "trim_ta P \ = ta_only_prod P (ta_only_reach \)" definition trim_reg where "trim_reg R = Reg (fin R) (trim_ta (fin R) (ta R))" subsubsection \Mapping over tree automata\ definition fmap_states_ta :: "('a \ 'b) \ ('a, 'f) ta \ ('b, 'f) ta" where "fmap_states_ta f \ = TA (map_ta_rule f id |`| rules \) (map_both f |`| eps \)" definition fmap_funs_ta :: "('f \ 'g) \ ('a, 'f) ta \ ('a, 'g) ta" where "fmap_funs_ta f \ = TA (map_ta_rule id f |`| rules \) (eps \)" definition fmap_states_reg :: "('a \ 'b) \ ('a, 'f) reg \ ('b, 'f) reg" where "fmap_states_reg f R = Reg (f |`| fin R) (fmap_states_ta f (ta R))" definition fmap_funs_reg :: "('f \ 'g) \ ('a, 'f) reg \ ('a, 'g) reg" where "fmap_funs_reg f R = Reg (fin R) (fmap_funs_ta f (ta R))" subsubsection \Product construction (language intersection)\ definition prod_ta_rules :: "('q1,'f) ta \ ('q2,'f) ta \ ('q1 \ 'q2, 'f) ta_rule fset" where "prod_ta_rules \ \ = {| TA_rule f qs q | f qs q. TA_rule f (map fst qs) (fst q) |\| rules \ \ TA_rule f (map snd qs) (snd q) |\| rules \|}" declare prod_ta_rules_def [simp] definition prod_epsLp where "prod_epsLp \ \ = (\ (p, q). (fst p, fst q) |\| eps \ \ snd p = snd q \ snd q |\| \ \)" definition prod_epsRp where "prod_epsRp \ \ = (\ (p, q). (snd p, snd q) |\| eps \ \ fst p = fst q \ fst q |\| \ \)" definition prod_ta :: "('q1,'f) ta \ ('q2,'f) ta \ ('q1 \ 'q2, 'f) ta" where "prod_ta \ \ = TA (prod_ta_rules \ \) (fCollect (prod_epsLp \ \) |\| fCollect (prod_epsRp \ \))" definition reg_intersect where "reg_intersect R L = Reg (fin R |\| fin L) (prod_ta (ta R) (ta L))" subsubsection \Union construction (language union)\ definition ta_union where "ta_union \ \ = TA (rules \ |\| rules \) (eps \ |\| eps \)" definition reg_union where "reg_union R L = Reg (Inl |`| (fin R |\| \\<^sub>r R) |\| Inr |`| (fin L |\| \\<^sub>r L)) (ta_union (fmap_states_ta Inl (ta R)) (fmap_states_ta Inr (ta L)))" subsubsection \Epsilon free and tree automaton accepting empty language\ definition eps_free_rulep where "eps_free_rulep \ = (\ r. \ f qs q q'. r = TA_rule f qs q' \ TA_rule f qs q |\| rules \ \ (q = q' \ (q, q') |\| (eps \)|\<^sup>+|))" definition eps_free :: "('q, 'f) ta \ ('q, 'f) ta" where "eps_free \ = TA (fCollect (eps_free_rulep \)) {||}" definition is_ta_eps_free :: "('q, 'f) ta \ bool" where "is_ta_eps_free \ \ eps \ = {||}" definition ta_empty :: "'q fset \ ('q,'f) ta \ bool" where "ta_empty Q \ \ ta_reachable \ |\| Q |\| {||}" definition eps_free_reg where "eps_free_reg R = Reg (fin R) (eps_free (ta R))" definition reg_empty where "reg_empty R = ta_empty (fin R) (ta R)" subsubsection \Relabeling tree automaton states to natural numbers\ definition map_fset_to_nat :: "('a :: linorder) fset \ 'a \ nat" where "map_fset_to_nat X = (\x. the (mem_idx x (sorted_list_of_fset X)))" definition map_fset_fset_to_nat :: "('a :: linorder) fset fset \ 'a fset \ nat" where "map_fset_fset_to_nat X = (\x. the (mem_idx (sorted_list_of_fset x) (sorted_list_of_fset (sorted_list_of_fset |`| X))))" definition relabel_ta :: "('q :: linorder, 'f) ta \ (nat, 'f) ta" where "relabel_ta \ = fmap_states_ta (map_fset_to_nat (\ \)) \" definition relabel_Q\<^sub>f :: "('q :: linorder) fset \ ('q :: linorder, 'f) ta \ nat fset" where "relabel_Q\<^sub>f Q \ = map_fset_to_nat (\ \) |`| (Q |\| \ \)" definition relabel_reg :: "('q :: linorder, 'f) reg \ (nat, 'f) reg" where "relabel_reg R = Reg (relabel_Q\<^sub>f (fin R) (ta R)) (relabel_ta (ta R))" \ \The instantiation of $<$ and $\leq$ for finite sets are $\mid \subset \mid$ and $\mid \subseteq \mid$ which don't give rise to a total order and therefore it cannot be an instance of the type class linorder. However taking the lexographic order of the sorted list of each finite set gives rise to a total order. Therefore we provide a relabeling for tree automata where the states are finite sets. This allows us to relabel the well known power set construction.\ definition relabel_fset_ta :: "(('q :: linorder) fset, 'f) ta \ (nat, 'f) ta" where "relabel_fset_ta \ = fmap_states_ta (map_fset_fset_to_nat (\ \)) \" definition relabel_fset_Q\<^sub>f :: "('q :: linorder) fset fset \ (('q :: linorder) fset, 'f) ta \ nat fset" where "relabel_fset_Q\<^sub>f Q \ = map_fset_fset_to_nat (\ \) |`| (Q |\| \ \)" definition relable_fset_reg :: "(('q :: linorder) fset, 'f) reg \ (nat, 'f) reg" where "relable_fset_reg R = Reg (relabel_fset_Q\<^sub>f (fin R) (ta R)) (relabel_fset_ta (ta R))" definition "srules \ = fset (rules \)" definition "seps \ = fset (eps \)" lemma rules_transfer [transfer_rule]: "rel_fun (=) (pcr_fset (=)) srules rules" unfolding rel_fun_def by (auto simp add: cr_fset_def fset.pcr_cr_eq srules_def) lemma eps_transfer [transfer_rule]: "rel_fun (=) (pcr_fset (=)) seps eps" unfolding rel_fun_def by (auto simp add: cr_fset_def fset.pcr_cr_eq seps_def) lemma TA_equalityI: "rules \ = rules \ \ eps \ = eps \ \ \ = \" using ta.expand by blast lemma rule_states_code [code]: "rule_states \ = |\| ((\ r. finsert (r_rhs r) (fset_of_list (r_lhs_states r))) |`| \)" unfolding rule_states_def by fastforce lemma eps_states_code [code]: "eps_states \\<^sub>\ = |\| ((\ (q,q'). {|q,q'|}) |`| \\<^sub>\)" (is "?Ls = ?Rs") unfolding eps_states_def by (force simp add: case_prod_beta') lemma rule_states_empty [simp]: "rule_states {||} = {||}" by (auto simp: rule_states_def) lemma eps_states_empty [simp]: "eps_states {||} = {||}" by (auto simp: eps_states_def) lemma rule_states_union [simp]: "rule_states (\ |\| \) = rule_states \ |\| rule_states \" unfolding rule_states_def by fastforce lemma rule_states_mono: "\ |\| \ \ rule_states \ |\| rule_states \" unfolding rule_states_def by force lemma eps_states_union [simp]: "eps_states (\ |\| \) = eps_states \ |\| eps_states \" unfolding eps_states_def by auto lemma eps_states_image [simp]: "eps_states (map_both f |`| \\<^sub>\) = f |`| eps_states \\<^sub>\" unfolding eps_states_def map_prod_def by (force simp: fimage_iff) lemma eps_states_mono: "\ |\| \ \ eps_states \ |\| eps_states \" unfolding eps_states_def by transfer auto lemma eps_statesI [intro]: "(p, q) |\| \ \ p |\| eps_states \" "(p, q) |\| \ \ q |\| eps_states \" unfolding eps_states_def by (auto simp add: rev_image_eqI) lemma eps_statesE [elim]: assumes "p |\| eps_states \" obtains q where "(p, q) |\| \ \ (q, p) |\| \" using assms unfolding eps_states_def by (transfer, auto)+ lemma rule_statesE [elim]: assumes "q |\| rule_states \" obtains f ps p where "TA_rule f ps p |\| \" "q |\| (fset_of_list ps) \ q = p" using assms proof - assume ass: "(\f ps p. f ps \ p |\| \ \ q |\| fset_of_list ps \ q = p \ thesis)" from assms obtain r where "r |\| \" "q |\| fset_of_list (r_lhs_states r) \ q = r_rhs r" by (auto simp: rule_states_def) then show thesis using ass by (cases r) auto qed lemma rule_statesI [intro]: assumes "r |\| \" "q |\| finsert (r_rhs r) (fset_of_list (r_lhs_states r))" shows "q |\| rule_states \" using assms by (auto simp: rule_states_def) text \Destruction rule for states\ lemma rule_statesD: "r |\| (rules \) \ r_rhs r |\| \ \" "f qs \ q |\| (rules \) \ q |\| \ \" "r |\| (rules \) \ p |\| fset_of_list (r_lhs_states r) \ p |\| \ \" "f qs \ q |\| (rules \) \ p |\| fset_of_list qs \ p |\| \ \" by (force simp: \_def rule_states_def fimage_iff)+ lemma eps_states [simp]: "(eps \) |\| \ \ |\| \ \" unfolding \_def eps_states_def rule_states_def by (auto simp add: rev_image_eqI) lemma eps_statesD: "(p, q) |\| (eps \) \ p |\| \ \ \ q |\| \ \" using eps_states by (auto simp add: \_def) lemma eps_trancl_statesD: "(p, q) |\| (eps \)|\<^sup>+| \ p |\| \ \ \ q |\| \ \" by (induct rule: ftrancl_induct) (auto dest: eps_statesD) lemmas eps_dest_all = eps_statesD eps_trancl_statesD text \Mapping over function symbols/states\ lemma finite_Collect_ta_rule: "finite {TA_rule f qs q | f qs q. TA_rule f qs q |\| rules \}" (is "finite ?S") proof - have "{f qs \ q |f qs q. f qs \ q |\| rules \} \ fset (rules \)" by auto from finite_subset[OF this] show ?thesis by simp qed lemma map_ta_rule_finite: "finite \ \ finite {TA_rule (g h) (map f qs) (f q) | h qs q. TA_rule h qs q \ \}" proof (induct rule: finite.induct) case (insertI A a) have union: "{TA_rule (g h) (map f qs) (f q) |h qs q. TA_rule h qs q \ insert a A} = {TA_rule (g h) (map f qs) (f q) | h qs q. TA_rule h qs q = a} \ {TA_rule (g h) (map f qs) (f q) |h qs q. TA_rule h qs q \ A}" by auto have "finite {g h map f qs \ f q |h qs q. h qs \ q = a}" by (cases a) auto from finite_UnI[OF this insertI(2)] show ?case unfolding union . qed auto lemmas map_ta_rule_fset_finite [simp] = map_ta_rule_finite[of "fset \" for \, simplified] lemmas map_ta_rule_states_finite [simp] = map_ta_rule_finite[of "fset \" id for \, simplified] lemmas map_ta_rule_funsym_finite [simp] = map_ta_rule_finite[of "fset \" _ id for \, simplified] lemma map_ta_rule_comp: "map_ta_rule f g \ map_ta_rule f' g' = map_ta_rule (f \ f') (g \ g')" using ta_rule.map_comp[of f g] by (auto simp: comp_def) lemma map_ta_rule_cases: "map_ta_rule f g r = TA_rule (g (r_root r)) (map f (r_lhs_states r)) (f (r_rhs r))" by (cases r) auto lemma map_ta_rule_prod_swap_id [simp]: "map_ta_rule prod.swap prod.swap (map_ta_rule prod.swap prod.swap r) = r" by (auto simp: map_ta_rule_cases) lemma rule_states_image [simp]: "rule_states (map_ta_rule f g |`| \) = f |`| rule_states \" (is "?Ls = ?Rs") proof - {fix q assume "q |\| ?Ls" then obtain r where "r |\| \" "q |\| finsert (r_rhs (map_ta_rule f g r)) (fset_of_list (r_lhs_states (map_ta_rule f g r)))" by (auto simp: rule_states_def) then have "q |\| ?Rs" by (cases r) (force simp: fimage_iff)} moreover {fix q assume "q |\| ?Rs" then obtain r p where "r |\| \" "f p = q" "p |\| finsert (r_rhs r) (fset_of_list (r_lhs_states r))" by (auto simp: rule_states_def) then have "q |\| ?Ls" by (cases r) (force simp: fimage_iff)} ultimately show ?thesis by blast qed lemma \_mono: "(rules \) |\| (rules \) \ (eps \) |\| (eps \) \ \ \ |\| \ \" using rule_states_mono eps_states_mono unfolding \_def by blast lemma \_subseteq_I: assumes "\ r. r |\| rules \ \ r_rhs r |\| S" and "\ r. r |\| rules \ \ fset_of_list (r_lhs_states r) |\| S" and "\ e. e |\| eps \ \ fst e |\| S \ snd e |\| S" shows "\ \ |\| S" using assms unfolding \_def by (auto simp: rule_states_def) blast lemma finite_states: "finite {q. \ f p ps. f ps \ p |\| rules \ \ (p = q \ (p, q) |\| (eps \)|\<^sup>+|)}" (is "finite ?set") proof - have "?set \ fset (\ \)" by (intro subsetI, drule CollectD) (metis eps_trancl_statesD rule_statesD(2)) from finite_subset[OF this] show ?thesis by auto qed text \Collecting all states reachable from target of rules\ lemma finite_ta_rhs_states [simp]: "finite {q. \p. p |\| rule_target_states (rules \) \ (p = q \ (p, q) |\| (eps \)|\<^sup>+|)}" (is "finite ?Set") proof - have "?Set \ fset (\ \)" by (auto dest: rule_statesD) (metis eps_trancl_statesD rule_statesD(1))+ from finite_subset[OF this] show ?thesis by auto qed text \Computing the signature induced by the rule set of given tree automaton\ lemma ta_sigI [intro]: "TA_rule f qs q |\| (rules \) \ length qs = n \ (f, n) |\| ta_sig \" unfolding ta_sig_def using mk_disjoint_finsert by fastforce lemma ta_sig_mono: "(rules \) |\| (rules \) \ ta_sig \ |\| ta_sig \" by (auto simp: ta_sig_def) lemma finite_eps: "finite {q. \ f ps p. f ps \ p |\| rules \ \ (p = q \ (p, q) |\| (eps \)|\<^sup>+|)}" (is "finite ?S") by (intro finite_subset[OF _ finite_ta_rhs_states[of \]]) (auto intro!: bexI) lemma collect_snd_trancl_fset: "{p. (q, p) |\| (eps \)|\<^sup>+|} = fset (snd |`| (ffilter (\ x. fst x = q) ((eps \)|\<^sup>+|)))" by (auto simp: image_iff) force lemma ta_der_Var: "q |\| ta_der \ (Var x) \ x = q \ (x, q) |\| (eps \)|\<^sup>+|" by (auto simp: collect_snd_trancl_fset) lemma ta_der_Fun: "q |\| ta_der \ (Fun f ts) \ (\ ps p. TA_rule f ps p |\| (rules \) \ (p = q \ (p, q) |\| (eps \)|\<^sup>+|) \ length ps = length ts \ (\ i < length ts. ps ! i |\| ta_der \ (ts ! i)))" (is "?Ls \ ?Rs") unfolding ta_der.simps by (intro iffI fCollect_memberI finite_Collect_less_eq[OF _ finite_eps[of \]]) auto declare ta_der.simps[simp del] declare ta_der.simps[code del] lemmas ta_der_simps [simp] = ta_der_Var ta_der_Fun lemma ta_der'_Var: "Var q |\| ta_der' \ (Var x) \ x = q \ (x, q) |\| (eps \)|\<^sup>+|" by (auto simp: collect_snd_trancl_fset) lemma ta_der'_Fun: "Var q |\| ta_der' \ (Fun f ts) \ q |\| ta_der \ (Fun f ts)" unfolding ta_der'.simps by (intro iffI funionI1 fCollect_memberI) (auto simp del: ta_der_Fun ta_der_Var simp: fset_image_conv) lemma ta_der'_Fun2: "Fun f ps |\| ta_der' \ (Fun g ts) \ f = g \ length ps = length ts \ (\i| ta_der' \ (ts ! i))" proof - have f: "finite {ss. set ss \ fset ( |\| (fset_of_list (map (ta_der' \) ts))) \ length ss = length ts}" by (intro finite_lists_length_eq) auto have "finite {ss. length ss = length ts \ (\i| ta_der' \ (ts ! i))}" by (intro finite_subset[OF _ f]) (force simp: in_fset_conv_nth simp flip: fset_of_list_elem) then show ?thesis unfolding ta_der'.simps by (intro iffI funionI2 fCollect_memberI) (auto simp del: ta_der_Fun ta_der_Var) qed declare ta_der'.simps[simp del] declare ta_der'.simps[code del] lemmas ta_der'_simps [simp] = ta_der'_Var ta_der'_Fun ta_der'_Fun2 text \Induction schemes for the most used cases\ lemma ta_der_induct[consumes 1, case_names Var Fun]: assumes reach: "q |\| ta_der \ t" and VarI: "\ q v. v = q \ (v, q) |\| (eps \)|\<^sup>+| \ P (Var v) q" and FunI: "\f ts ps p q. f ps \ p |\| rules \ \ length ts = length ps \ p = q \ (p, q) |\| (eps \)|\<^sup>+| \ (\i. i < length ts \ ps ! i |\| ta_der \ (ts ! i)) \ (\i. i < length ts \ P (ts ! i) (ps ! i)) \ P (Fun f ts) q" shows "P t q" using assms(1) by (induct t arbitrary: q) (auto simp: VarI FunI) lemma ta_der_gterm_induct[consumes 1, case_names GFun]: assumes reach: "q |\| ta_der \ (term_of_gterm t)" and Fun: "\f ts ps p q. TA_rule f ps p |\| rules \ \ length ts = length ps \ p = q \ (p, q) |\| (eps \)|\<^sup>+| \ (\i. i < length ts \ ps ! i |\| ta_der \ (term_of_gterm (ts ! i))) \ (\i. i < length ts \ P (ts ! i) (ps ! i)) \ P (GFun f ts) q" shows "P t q" using assms(1) by (induct t arbitrary: q) (auto simp: Fun) lemma ta_der_rule_empty: assumes "q |\| ta_der (TA {||} \\<^sub>\) t" obtains p where "t = Var p" "p = q \ (p, q) |\| \\<^sub>\|\<^sup>+|" using assms by (cases t) auto lemma ta_der_eps: assumes "(p, q) |\| (eps \)" and "p |\| ta_der \ t" shows "q |\| ta_der \ t" using assms by (cases t) (auto intro: ftrancl_into_trancl) lemma ta_der_trancl_eps: assumes "(p, q) |\| (eps \)|\<^sup>+|" and "p |\| ta_der \ t" shows "q |\| ta_der \ t" using assms by (induct rule: ftrancl_induct) (auto intro: ftrancl_into_trancl ta_der_eps) lemma ta_der_mono: "(rules \) |\| (rules \) \ (eps \) |\| (eps \) \ ta_der \ t |\| ta_der \ t" proof (induct t) case (Var x) then show ?case by (auto dest: ftrancl_mono[of _ "eps \" "eps \"]) next case (Fun f ts) show ?case using Fun(1)[OF nth_mem Fun(2, 3)] by (auto dest!: fsubsetD[OF Fun(2)] ftrancl_mono[OF _ Fun(3)]) blast+ qed lemma ta_der_el_mono: "(rules \) |\| (rules \) \ (eps \) |\| (eps \) \ q |\| ta_der \ t \ q |\| ta_der \ t" using ta_der_mono by blast lemma ta_der'_ta_der: assumes "t |\| ta_der' \ s" "p |\| ta_der \ t" shows "p |\| ta_der \ s" using assms proof (induction arbitrary: p t rule: ta_der'.induct) case (2 \ f ts) show ?case using 2(2-) proof (induction t) case (Var x) then show ?case by auto (meson ftrancl_trans) next case (Fun g ss) have ss_props: "g = f" "length ss = length ts" "\i < length ts. ss ! i |\| ta_der' \ (ts ! i)" using Fun(2) by auto then show ?thesis using Fun(1)[OF nth_mem] Fun(2-) by (auto simp: ss_props) (metis (no_types, lifting) "2.IH" ss_props(3))+ qed qed (auto dest: ftrancl_trans simp: ta_der'.simps) lemma ta_der'_empty: assumes "t |\| ta_der' (TA {||} {||}) s" shows "t = s" using assms by (induct s arbitrary: t) (auto simp add: ta_der'.simps nth_equalityI) lemma ta_der'_to_ta_der: "Var q |\| ta_der' \ s \ q |\| ta_der \ s" using ta_der'_ta_der by fastforce lemma ta_der_to_ta_der': "q |\| ta_der \ s \ Var q |\| ta_der' \ s " by (induct s arbitrary: q) auto lemma ta_der'_poss: assumes "t |\| ta_der' \ s" shows "poss t \ poss s" using assms proof (induct s arbitrary: t) case (Fun f ts) show ?case using Fun(2) Fun(1)[OF nth_mem, of i "args t ! i" for i] by (cases t) auto qed (auto simp: ta_der'.simps) lemma ta_der'_refl[simp]: "t |\| ta_der' \ t" by (induction t) fastforce+ lemma ta_der'_eps: assumes "Var p |\| ta_der' \ s" and "(p, q) |\| (eps \)|\<^sup>+|" shows "Var q |\| ta_der' \ s" using assms by (cases s, auto dest: ftrancl_trans) (meson ftrancl_trans) lemma ta_der'_trans: assumes "t |\| ta_der' \ s" and "u |\| ta_der' \ t" shows "u |\| ta_der' \ s" using assms proof (induct t arbitrary: u s) case (Fun f ts) note IS = Fun(2-) note IH = Fun(1)[OF nth_mem, of i "args s ! i" for i] show ?case proof (cases s) case (Var x1) then show ?thesis using IS by (auto simp: ta_der'.simps) next case [simp]: (Fun g ss) show ?thesis using IS IH by (cases u, auto) (metis ta_der_to_ta_der')+ qed qed (auto simp: ta_der'.simps ta_der'_eps) text \Connecting contexts to derivation definition\ lemma ta_der_ctxt: assumes p: "p |\| ta_der \ t" "q |\| ta_der \ C\Var p\" shows "q |\| ta_der \ C\t\" using assms(2) proof (induct C arbitrary: q) case Hole then show ?case using assms by (auto simp: ta_der_trancl_eps) next case (More f ss C ts) from More(2) obtain qs r where rule: "f qs \ r |\| rules \" "length qs = Suc (length ss + length ts)" and reach: "\ i < Suc (length ss + length ts). qs ! i |\| ta_der \ ((ss @ C\Var p\ # ts) ! i)" "r = q \ (r, q) |\| (eps \)|\<^sup>+|" by auto have "i < Suc (length ss + length ts) \ qs ! i |\| ta_der \ ((ss @ C\t\ # ts) ! i)" for i using More(1)[of "qs ! length ss"] assms rule(2) reach(1) unfolding nth_append_Cons by presburger then show ?case using rule reach(2) by auto qed lemma ta_der_eps_ctxt: assumes "p |\| ta_der A C\Var q'\" and "(q, q') |\| (eps A)|\<^sup>+|" shows "p |\| ta_der A C\Var q\" using assms by (meson ta_der_Var ta_der_ctxt) lemma rule_reachable_ctxt_exist: assumes rule: "f qs \ q |\| rules \" and "i < length qs" shows "\ C. q |\| ta_der \ (C \Var (qs ! i)\)" using assms by (intro exI[of _ "More f (map Var (take i qs)) \ (map Var (drop (Suc i) qs))"]) (auto simp: min_def nth_append_Cons intro!: exI[of _ q] exI[of _ qs]) lemma ta_der_ctxt_decompose: assumes "q |\| ta_der \ C\t\" shows "\ p . p |\| ta_der \ t \ q |\| ta_der \ C\Var p\" using assms proof (induct C arbitrary: q) case (More f ss C ts) from More(2) obtain qs r where rule: "f qs \ r |\| rules \" "length qs = Suc (length ss + length ts)" and reach: "\ i < Suc (length ss + length ts). qs ! i |\| ta_der \ ((ss @ C\t\ # ts) ! i)" "r = q \ (r, q) |\| (eps \)|\<^sup>+|" by auto obtain p where p: "p |\| ta_der \ t" "qs ! length ss |\| ta_der \ C\Var p\" using More(1)[of "qs ! length ss"] reach(1) rule(2) by (metis less_add_Suc1 nth_append_length) have "i < Suc (length ss + length ts) \ qs ! i |\| ta_der \ ((ss @ C\Var p\ # ts) ! i)" for i using reach rule(2) p by (auto simp: p(2) nth_append_Cons) then have "q |\| ta_der \ (More f ss C ts)\Var p\" using rule reach by auto then show ?case using p(1) by (intro exI[of _ p]) blast qed auto \ \Relation between reachable states and states of a tree automaton\ lemma ta_der_states: "ta_der \ t |\| \ \ |\| fvars_term t" proof (induct t) case (Var x) then show ?case by (auto simp: eq_onp_same_args) (metis eps_trancl_statesD) case (Fun f ts) then show ?case by (auto simp: rule_statesD(2) eps_trancl_statesD) qed lemma ground_ta_der_states: "ground t \ ta_der \ t |\| \ \" using ta_der_states[of \ t] by auto lemmas ground_ta_der_statesD = fsubsetD[OF ground_ta_der_states] lemma gterm_ta_der_states [simp]: "q |\| ta_der \ (term_of_gterm t) \ q |\| \ \" by (intro ground_ta_der_states[THEN fsubsetD, of "term_of_gterm t"]) simp lemma ta_der_states': "q |\| ta_der \ t \ q |\| \ \ \ fvars_term t |\| \ \" proof (induct rule: ta_der_induct) case (Fun f ts ps p r) then have "i < length ts \ fvars_term (ts ! i) |\| \ \" for i by (auto simp: in_fset_conv_nth dest!: rule_statesD(3)) then show ?case by (force simp: in_fset_conv_nth) qed (auto simp: eps_trancl_statesD) lemma ta_der_not_stateD: "q |\| ta_der \ t \ q |\| \ \ \ t = Var q" using fsubsetD[OF ta_der_states, of q \ t] by (cases t) (auto dest: rule_statesD eps_trancl_statesD) lemma ta_der_is_fun_stateD: "is_Fun t \ q |\| ta_der \ t \ q |\| \ \" using ta_der_not_stateD[of q \ t] by (cases t) auto lemma ta_der_is_fun_fvars_stateD: "is_Fun t \ q |\| ta_der \ t \ fvars_term t |\| \ \" using ta_der_is_fun_stateD[of t q \] using ta_der_states'[of q \ t] by (cases t) auto lemma ta_der_not_reach: assumes "\ r. r |\| rules \ \ r_rhs r \ q" and "\ e. e |\| eps \ \ snd e \ q" shows "q |\| ta_der \ (term_of_gterm t)" using assms by (cases t) (fastforce dest!: assms(1) ftranclD2[of _ q]) lemma ta_rhs_states_subset_states: "ta_rhs_states \ |\| \ \" by (auto simp: ta_rhs_states_def dest: rtranclD rule_statesD eps_trancl_statesD) (* a resulting state is always some rhs of a rule (or epsilon transition) *) lemma ta_rhs_states_res: assumes "is_Fun t" shows "ta_der \ t |\| ta_rhs_states \" proof fix q assume q: "q |\| ta_der \ t" from \is_Fun t\ obtain f ts where t: "t = Fun f ts" by (cases t, auto) from q[unfolded t] obtain q' qs where "TA_rule f qs q' |\| rules \" and q: "q' = q \ (q', q) |\| (eps \)|\<^sup>+|" by auto then show "q |\| ta_rhs_states \" unfolding ta_rhs_states_def by (auto intro!: bexI) qed text \Reachable states of ground terms are preserved over the @{const adapt_vars} function\ lemma ta_der_adapt_vars_ground [simp]: "ground t \ ta_der A (adapt_vars t) = ta_der A t" by (induct t) auto lemma gterm_of_term_inv': "ground t \ term_of_gterm (gterm_of_term t) = adapt_vars t" by (induct t) (auto 0 0 intro!: nth_equalityI) lemma map_vars_term_term_of_gterm: "map_vars_term f (term_of_gterm t) = term_of_gterm t" by (induct t) auto lemma adapt_vars_term_of_gterm: "adapt_vars (term_of_gterm t) = term_of_gterm t" by (induct t) auto (* a term can be reduced to a state, only if all symbols appear in the automaton *) lemma ta_der_term_sig: "q |\| ta_der \ t \ ffunas_term t |\| ta_sig \" proof (induct rule: ta_der_induct) case (Fun f ts ps p q) show ?case using Fun(1 - 4) Fun(5)[THEN fsubsetD] by (auto simp: in_fset_conv_nth) qed auto lemma ta_der_gterm_sig: "q |\| ta_der \ (term_of_gterm t) \ ffunas_gterm t |\| ta_sig \" using ta_der_term_sig ffunas_term_of_gterm_conv by fastforce text \@{const ta_lang} for terms with arbitrary variable type\ lemma ta_langE: assumes "t \ ta_lang Q \" obtains t' q where "ground t'" "q |\| Q" "q |\| ta_der \ t'" "t = adapt_vars t'" using assms unfolding ta_lang_def by blast lemma ta_langI: assumes "ground t'" "q |\| Q" "q |\| ta_der \ t'" "t = adapt_vars t'" shows "t \ ta_lang Q \" using assms unfolding ta_lang_def by blast lemma ta_lang_def2: "(ta_lang Q (\ :: ('q,'f)ta) :: ('f,'v)terms) = {t. ground t \ Q |\| ta_der \ (adapt_vars t) \ {||}}" by (auto elim!: ta_langE) (metis adapt_vars_adapt_vars ground_adapt_vars ta_langI) text \@{const ta_lang} for @{const gterms}\ lemma ta_lang_to_gta_lang [simp]: "ta_lang Q \ = term_of_gterm ` gta_lang Q \" (is "?Ls = ?Rs") proof - {fix t assume "t \ ?Ls" from ta_langE[OF this] obtain q t' where "ground t'" "q |\| Q" "q |\| ta_der \ t'" "t = adapt_vars t'" by blast then have "t \ ?Rs" unfolding gta_lang_def gta_der_def by (auto simp: image_iff gterm_of_term_inv' intro!: exI[of _ "gterm_of_term t'"])} moreover {fix t assume "t \ ?Rs" then have "t \ ?Ls" using ta_langI[OF ground_term_of_gterm _ _ gterm_of_term_inv'[OF ground_term_of_gterm]] by (force simp: gta_lang_def gta_der_def)} ultimately show ?thesis by blast qed lemma term_of_gterm_in_ta_lang_conv: "term_of_gterm t \ ta_lang Q \ \ t \ gta_lang Q \" by (metis (mono_tags, lifting) image_iff ta_lang_to_gta_lang term_of_gterm_inv) lemma gta_lang_def_sym: "gterm_of_term ` ta_lang Q \ = gta_lang Q \" (* this is nontrivial because the lhs has a more general type than the rhs of gta_lang_def *) unfolding gta_lang_def image_def by (intro Collect_cong) (simp add: gta_lang_def) lemma gta_langI [intro]: assumes "q |\| Q" and "q |\| ta_der \ (term_of_gterm t)" shows "t \ gta_lang Q \" using assms by (metis adapt_vars_term_of_gterm ground_term_of_gterm ta_langI term_of_gterm_in_ta_lang_conv) lemma gta_langE [elim]: assumes "t \ gta_lang Q \" obtains q where "q |\| Q" and "q |\| ta_der \ (term_of_gterm t)" using assms by (metis adapt_vars_adapt_vars adapt_vars_term_of_gterm ta_langE term_of_gterm_in_ta_lang_conv) lemma gta_lang_mono: assumes "\ t. ta_der \ t |\| ta_der \ t" and "Q\<^sub>\ |\| Q\<^sub>\" shows "gta_lang Q\<^sub>\ \ \ gta_lang Q\<^sub>\ \" using assms by (auto elim!: gta_langE intro!: gta_langI) lemma gta_lang_term_of_gterm [simp]: "term_of_gterm t \ term_of_gterm ` gta_lang Q \ \ t \ gta_lang Q \" by (auto elim!: gta_langE intro!: gta_langI) (metis term_of_gterm_inv) (* terms can be accepted, only if all their symbols appear in the automaton *) lemma gta_lang_subset_rules_funas: "gta_lang Q \ \ \\<^sub>G (fset (ta_sig \))" using ta_der_gterm_sig[THEN fsubsetD] by (force simp: \\<^sub>G_equivalent_def ffunas_gterm.rep_eq) lemma reg_funas: "\ \ \ \\<^sub>G (fset (ta_sig (ta \)))" using gta_lang_subset_rules_funas by (auto simp: \_def) lemma ta_syms_lang: "t \ ta_lang Q \ \ ffunas_term t |\| ta_sig \" using gta_lang_subset_rules_funas ffunas_gterm_gterm_of_term ta_der_gterm_sig ta_lang_def2 by fastforce lemma gta_lang_Rest_states_conv: "gta_lang Q \ = gta_lang (Q |\| \ \) \" by (auto elim!: gta_langE) lemma reg_Rest_fin_states [simp]: "\ (reg_Restr_Q\<^sub>f \) = \ \" using gta_lang_Rest_states_conv by (auto simp: \_def reg_Restr_Q\<^sub>f_def) text \Deterministic tree automatons\ definition ta_det :: "('q,'f) ta \ bool" where "ta_det \ \ eps \ = {||} \ (\ f qs q q'. TA_rule f qs q |\| rules \ \ TA_rule f qs q' |\| rules \ \ q = q')" definition "ta_subset \ \ \ rules \ |\| rules \ \ eps \ |\| eps \" (* determinism implies unique results *) lemma ta_detE[elim, consumes 1]: assumes det: "ta_det \" shows "q |\| ta_der \ t \ q' |\| ta_der \ t \ q = q'" using assms by (induct t arbitrary: q q') (auto simp: ta_det_def, metis nth_equalityI nth_mem) lemma ta_subset_states: "ta_subset \ \ \ \ \ |\| \ \" using \_mono by (auto simp: ta_subset_def) lemma ta_subset_refl[simp]: "ta_subset \ \" unfolding ta_subset_def by auto lemma ta_subset_trans: "ta_subset \ \ \ ta_subset \ \ \ ta_subset \ \" unfolding ta_subset_def by auto lemma ta_subset_det: "ta_subset \ \ \ ta_det \ \ ta_det \" unfolding ta_det_def ta_subset_def by blast lemma ta_der_mono': "ta_subset \ \ \ ta_der \ t |\| ta_der \ t" using ta_der_mono unfolding ta_subset_def by auto lemma ta_lang_mono': "ta_subset \ \ \ Q\<^sub>\ |\| Q\<^sub>\ \ ta_lang Q\<^sub>\ \ \ ta_lang Q\<^sub>\ \" using gta_lang_mono[of \ \] ta_der_mono'[of \ \] by auto blast (* the restriction of an automaton to a given set of states *) lemma ta_restrict_subset: "ta_subset (ta_restrict \ Q) \" unfolding ta_subset_def ta_restrict_def by auto lemma ta_restrict_states_Q: "\ (ta_restrict \ Q) |\| Q" by (auto simp: \_def ta_restrict_def rule_states_def eps_states_def dest!: fsubsetD) lemma ta_restrict_states: "\ (ta_restrict \ Q) |\| \ \" using ta_subset_states[OF ta_restrict_subset] by fastforce lemma ta_restrict_states_eq_imp_eq [simp]: assumes eq: "\ (ta_restrict \ Q) = \ \" shows "ta_restrict \ Q = \" using assms apply (auto simp: ta_restrict_def intro!: ta.expand finite_subset[OF _ finite_Collect_ta_rule, of _ \]) apply (metis (no_types, lifting) eq fsubsetD fsubsetI rule_statesD(1) rule_statesD(4) ta_restrict_states_Q ta_rule.collapse) apply (metis eps_statesD eq fin_mono ta_restrict_states_Q) by (metis eps_statesD eq fsubsetD ta_restrict_states_Q) lemma ta_der_ta_derict_states: "fvars_term t |\| Q \ q |\| ta_der (ta_restrict \ Q) t \ q |\| Q" by (induct t arbitrary: q) (auto simp: ta_restrict_def elim: ftranclE) lemma ta_derict_ruleI [intro]: "TA_rule f qs q |\| rules \ \ fset_of_list qs |\| Q \ q |\| Q \ TA_rule f qs q |\| rules (ta_restrict \ Q)" by (auto simp: ta_restrict_def intro!: ta.expand finite_subset[OF _ finite_Collect_ta_rule, of _ \]) text \Reachable and productive states: There always is a trim automaton\ lemma finite_ta_reachable [simp]: "finite {q. \t. ground t \ q |\| ta_der \ t}" proof - have "{q. \t. ground t \ q |\| ta_der \ t} \ fset (\ \)" using ground_ta_der_states[of _ \] by auto from finite_subset[OF this] show ?thesis by auto qed lemma ta_reachable_states: "ta_reachable \ |\| \ \" unfolding ta_reachable_def using ground_ta_der_states by force lemma ta_reachableE: assumes "q |\| ta_reachable \" obtains t where "ground t" "q |\| ta_der \ t" using assms[unfolded ta_reachable_def] by auto lemma ta_reachable_gtermE [elim]: assumes "q |\| ta_reachable \" obtains t where "q |\| ta_der \ (term_of_gterm t)" using ta_reachableE[OF assms] by (metis ground_term_to_gtermD) lemma ta_reachableI [intro]: assumes "ground t" and "q |\| ta_der \ t" shows "q |\| ta_reachable \" using assms finite_ta_reachable by (auto simp: ta_reachable_def) lemma ta_reachable_gtermI [intro]: "q |\| ta_der \ (term_of_gterm t) \ q |\| ta_reachable \" by (intro ta_reachableI[of "term_of_gterm t"]) simp lemma ta_reachableI_rule: assumes sub: "fset_of_list qs |\| ta_reachable \" and rule: "TA_rule f qs q |\| rules \" shows "q |\| ta_reachable \" "\ ts. length qs = length ts \ (\ i < length ts. ground (ts ! i)) \ (\ i < length ts. qs ! i |\| ta_der \ (ts ! i))" (is "?G") proof - { fix i assume i: "i < length qs" then have "qs ! i |\| fset_of_list qs" by auto with sub have "qs ! i |\| ta_reachable \" by auto from ta_reachableE[OF this] have "\ t. ground t \ qs ! i |\| ta_der \ t" by auto } then have "\ i. \ t. i < length qs \ ground t \ qs ! i |\| ta_der \ t" by auto from choice[OF this] obtain ts where ts: "\ i. i < length qs \ ground (ts i) \ qs ! i |\| ta_der \ (ts i)" by blast let ?t = "Fun f (map ts [0 ..< length qs])" have gt: "ground ?t" using ts by auto have r: "q |\| ta_der \ ?t" unfolding ta_der_Fun using rule ts by (intro exI[of _ qs] exI[of _ q]) simp with gt show "q |\| ta_reachable \" by blast from gt ts show ?G by (intro exI[of _ "map ts [0.. \ |\| ta_reachable \" and "TA_rule f qs q |\| rules \" obtains t where "groot t = (f, length qs)" "q |\| ta_der \ (term_of_gterm t)" proof - assume *: "\t. groot t = (f, length qs) \ q |\| ta_der \ (term_of_gterm t) \ thesis" from assms have "fset_of_list qs |\| ta_reachable \" by (auto dest: rule_statesD(3)) from ta_reachableI_rule[OF this assms(2)] obtain ts where args: "length qs = length ts" "\ i < length ts. ground (ts ! i)" "\ i < length ts. qs ! i |\| ta_der \ (ts ! i)" using assms by force then show ?thesis using assms(2) by (intro *[of "GFun f (map gterm_of_term ts)"]) auto qed lemma ta_reachableI_eps': assumes reach: "q |\| ta_reachable \" and eps: "(q, q') |\| (eps \)|\<^sup>+|" shows "q' |\| ta_reachable \" proof - from ta_reachableE[OF reach] obtain t where g: "ground t" and res: "q |\| ta_der \ t" by auto from ta_der_trancl_eps[OF eps res] g show ?thesis by blast qed lemma ta_reachableI_eps: assumes reach: "q |\| ta_reachable \" and eps: "(q, q') |\| eps \" shows "q' |\| ta_reachable \" by (rule ta_reachableI_eps'[OF reach], insert eps, auto) \ \Automata are productive on a set P if all states can reach a state in P\ lemma finite_ta_productive: "finite {p. \q q' C. p = q \ q' |\| ta_der \ C\Var q\ \ q' |\| P}" proof - {fix x q C assume ass: "x \ fset P" "q |\| P" "q |\| ta_der \ C\Var x\" then have "x \ fset (\ \)" proof (cases "is_Fun C\Var x\") case True then show ?thesis using ta_der_is_fun_fvars_stateD[OF _ ass(3)] by auto next case False then show ?thesis using ass by (cases C, auto, (metis eps_trancl_statesD)+) qed} then have "{q | q q' C. q' |\| ta_der \ (C\Var q\) \ q' |\| P} \ fset (\ \) \ fset P" by auto from finite_subset[OF this] show ?thesis by auto qed lemma ta_productiveE: assumes "q |\| ta_productive P \" obtains q' C where "q' |\| ta_der \ (C\Var q\)" "q' |\| P" using assms[unfolded ta_productive_def] by auto lemma ta_productiveI: assumes "q' |\| ta_der \ (C\Var q\)" "q' |\| P" shows "q |\| ta_productive P \" using assms unfolding ta_productive_def using finite_ta_productive by auto lemma ta_productiveI': assumes "q |\| ta_der \ (C\Var p\)" "q |\| ta_productive P \" shows "p |\| ta_productive P \" using assms unfolding ta_productive_def by auto (metis (mono_tags, lifting) ctxt_ctxt_compose ta_der_ctxt) lemma ta_productive_setI: "q |\| P \ q |\| ta_productive P \" using ta_productiveI[of q \ \ q] by simp lemma ta_reachable_empty_rules [simp]: "rules \ = {||} \ ta_reachable \ = {||}" by (auto simp: ta_reachable_def) (metis ground.simps(1) ta.exhaust_sel ta_der_rule_empty) lemma ta_reachable_mono: "ta_subset \ \ \ ta_reachable \ |\| ta_reachable \" using ta_der_mono' by (auto simp: ta_reachable_def) blast lemma ta_reachabe_rhs_states: "ta_reachable \ |\| ta_rhs_states \" proof - {fix q assume "q |\| ta_reachable \" then obtain t where "ground t" "q |\| ta_der \ t" by (auto simp: ta_reachable_def) then have "q |\| ta_rhs_states \" by (cases t) (auto simp: ta_rhs_states_def intro!: bexI)} then show ?thesis by blast qed lemma ta_reachable_eps: "(p, q) |\| (eps \)|\<^sup>+| \ p |\| ta_reachable \ \ (p, q) |\| (fRestr (eps \) (ta_reachable \))|\<^sup>+|" proof (induct rule: ftrancl_induct) case (Base a b) then show ?case by (metis fSigmaI finterI fr_into_trancl ta_reachableI_eps) next case (Step p q r) then have "q |\| ta_reachable \" "r |\| ta_reachable \" by (metis ta_reachableI_eps ta_reachableI_eps')+ then show ?case using Step by (metis fSigmaI finterI ftrancl_into_trancl) qed (* major lemma to show that one can restrict to reachable states *) lemma ta_der_only_reach: assumes "fvars_term t |\| ta_reachable \" shows "ta_der \ t = ta_der (ta_only_reach \) t" (is "?LS = ?RS") proof - have "?RS |\| ?LS" using ta_der_mono'[OF ta_restrict_subset] by fastforce moreover {fix q assume "q |\| ?LS" then have "q |\| ?RS" using assms proof (induct rule: ta_der_induct) case (Fun f ts ps p q) from Fun(2, 6) have ta_reach [simp]: "i < length ps \ fvars_term (ts ! i) |\| ta_reachable \" for i by auto (metis ffUnionI fimage_fset fnth_mem funionI2 length_map nth_map sup.orderE) from Fun have r: "i < length ts \ ps ! i |\| ta_der (ta_only_reach \) (ts ! i)" "i < length ts \ ps ! i |\| ta_reachable \" for i by (auto) (metis ta_reach ta_der_ta_derict_states)+ then have "f ps \ p |\| rules (ta_only_reach \)" using Fun(1, 2) by (intro ta_derict_ruleI) (fastforce simp: in_fset_conv_nth intro!: ta_reachableI_rule[OF _ Fun(1)])+ then show ?case using ta_reachable_eps[of p q] ta_reachableI_rule[OF _ Fun(1)] r Fun(2, 3) by (auto simp: ta_restrict_def intro!: exI[of _ p] exI[of _ ps]) qed (auto simp: ta_restrict_def intro: ta_reachable_eps)} ultimately show ?thesis by blast qed lemma ta_der_gterm_only_reach: "ta_der \ (term_of_gterm t) = ta_der (ta_only_reach \) (term_of_gterm t)" using ta_der_only_reach[of "term_of_gterm t" \] by simp lemma ta_reachable_ta_only_reach [simp]: "ta_reachable (ta_only_reach \) = ta_reachable \" (is "?LS = ?RS") proof - have "?LS |\| ?RS" using ta_der_mono'[OF ta_restrict_subset] by (auto simp: ta_reachable_def) fastforce moreover {fix t assume "ground (t :: ('b, 'a) term)" then have "ta_der \ t = ta_der (ta_only_reach \) t" using ta_der_only_reach[of t \] by simp} ultimately show ?thesis unfolding ta_reachable_def by auto qed lemma ta_only_reach_reachable: "\ (ta_only_reach \) |\| ta_reachable (ta_only_reach \)" using ta_restrict_states_Q[of \ "ta_reachable \"] by auto (* It is sound to restrict to reachable states. *) lemma gta_only_reach_lang: "gta_lang Q (ta_only_reach \) = gta_lang Q \" using ta_der_gterm_only_reach by (auto elim!: gta_langE intro!: gta_langI) force+ lemma \_only_reach: "\ (reg_reach R) = \ R" using gta_only_reach_lang by (auto simp: \_def reg_reach_def) lemma ta_only_reach_lang: "ta_lang Q (ta_only_reach \) = ta_lang Q \" using gta_only_reach_lang by (metis ta_lang_to_gta_lang) lemma ta_prod_epsD: "(p, q) |\| (eps \)|\<^sup>+| \ q |\| ta_productive P \ \ p |\| ta_productive P \" using ta_der_ctxt[of q \ "\\Var p\"] by (auto simp: ta_productive_def ta_der_trancl_eps) lemma ta_only_prod_eps: "(p, q) |\| (eps \)|\<^sup>+| \ q |\| ta_productive P \ \ (p, q) |\| (eps (ta_only_prod P \))|\<^sup>+|" proof (induct rule: ftrancl_induct) case (Base p q) then show ?case by (metis (no_types, lifting) fSigmaI finterI fr_into_trancl ta.sel(2) ta_prod_epsD ta_restrict_def) next case (Step p q r) note IS = this show ?case using IS(2 - 4) ta_prod_epsD[OF fr_into_trancl[OF IS(3)] IS(4)] by (auto simp: ta_restrict_def) (simp add: ftrancl_into_trancl) qed (* Major lemma to show that it is sound to restrict to productive states. *) lemma ta_der_only_prod: "q |\| ta_der \ t \ q |\| ta_productive P \ \ q |\| ta_der (ta_only_prod P \) t" proof (induct rule: ta_der_induct) case (Fun f ts ps p q) let ?\ = "ta_only_prod P \" have pr: "p |\| ta_productive P \" "i < length ts \ ps ! i |\| ta_productive P \" for i using Fun(2) ta_prod_epsD[of p q] Fun(3, 6) rule_reachable_ctxt_exist[OF Fun(1)] using ta_productiveI'[of p \ _ "ps ! i" P] by auto then have "f ps \ p |\| rules ?\" using Fun(1, 2) unfolding ta_restrict_def by (auto simp: in_fset_conv_nth intro: finite_subset[OF _ finite_Collect_ta_rule, of _ \]) then show ?case using pr Fun ta_only_prod_eps[of p q \ P] Fun(3, 6) by auto qed (auto intro: ta_only_prod_eps) lemma ta_der_ta_only_prod_ta_der: "q |\| ta_der (ta_only_prod P \) t \ q |\| ta_der \ t" by (meson ta_der_el_mono ta_restrict_subset ta_subset_def) (* It is sound to restrict to productive states. *) lemma gta_only_prod_lang: "gta_lang Q (ta_only_prod Q \) = gta_lang Q \" (is "gta_lang Q ?\ = _") proof show "gta_lang Q ?\ \ gta_lang Q \" using gta_lang_mono[OF ta_der_mono'[OF ta_restrict_subset]] by blast next {fix t assume "t \ gta_lang Q \" from gta_langE[OF this] obtain q where reach: "q |\| ta_der \ (term_of_gterm t)" "q |\| Q" . from ta_der_only_prod[OF reach(1) ta_productive_setI[OF reach(2)]] reach(2) have "t \ gta_lang Q ?\" by (auto intro: gta_langI)} then show "gta_lang Q \ \ gta_lang Q ?\" by blast qed lemma \_only_prod: "\ (reg_prod R) = \ R" using gta_only_prod_lang by (auto simp: \_def reg_prod_def) lemma ta_only_prod_lang: "ta_lang Q (ta_only_prod Q \) = ta_lang Q \" using gta_only_prod_lang by (metis ta_lang_to_gta_lang) (* the productive states are also productive w.r.t. the new automaton *) lemma ta_prodictive_ta_only_prod [simp]: "ta_productive P (ta_only_prod P \) = ta_productive P \" (is "?LS = ?RS") proof - have "?LS |\| ?RS" using ta_der_mono'[OF ta_restrict_subset] using finite_ta_productive[of \ P] by (auto simp: ta_productive_def) fastforce moreover have "?RS |\| ?LS" using ta_der_only_prod by (auto elim!: ta_productiveE) (smt (z3) ta_der_only_prod ta_productiveI ta_productive_setI) ultimately show ?thesis by blast qed lemma ta_only_prod_productive: "\ (ta_only_prod P \) |\| ta_productive P (ta_only_prod P \)" using ta_restrict_states_Q by force lemma ta_only_prod_reachable: assumes all_reach: "\ \ |\| ta_reachable \" shows "\ (ta_only_prod P \) |\| ta_reachable (ta_only_prod P \)" (is "?Ls |\| ?Rs") proof - {fix q assume "q |\| ?Ls" then obtain t where "ground t" "q |\| ta_der \ t" "q |\| ta_productive P \" using fsubsetD[OF ta_only_prod_productive[of \ P]] using fsubsetD[OF fsubset_trans[OF ta_restrict_states all_reach, of "ta_productive P \"]] by (auto elim!: ta_reachableE) then have "q |\| ?Rs" by (intro ta_reachableI[where ?\ = "ta_only_prod P \" and ?t = t]) (auto simp: ta_der_only_prod)} then show ?thesis by blast qed lemma ta_prod_reach_subset: "ta_subset (ta_only_prod P (ta_only_reach \)) \" by (rule ta_subset_trans, (rule ta_restrict_subset)+) lemma ta_prod_reach_states: "\ (ta_only_prod P (ta_only_reach \)) |\| \ \" by (rule ta_subset_states[OF ta_prod_reach_subset]) (* If all states are reachable then there exists a ground context for all productive states *) lemma ta_productive_aux: assumes "\ \ |\| ta_reachable \" "q |\| ta_der \ (C\t\)" shows "\C'. ground_ctxt C' \ q |\| ta_der \ (C'\t\)" using assms(2) proof (induct C arbitrary: q) case Hole then show ?case by (intro exI[of _ "\"]) auto next case (More f ts1 C ts2) from More(2) obtain qs q' where q': "f qs \ q' |\| rules \" "q' = q \ (q', q) |\| (eps \)|\<^sup>+|" "qs ! length ts1 |\| ta_der \ (C\t\)" "length qs = Suc (length ts1 + length ts2)" by simp (metis less_add_Suc1 nth_append_length) { fix i assume "i < length qs" then have "qs ! i |\| \ \" using q'(1) by (auto dest!: rule_statesD(4)) then have "\t. ground t \ qs ! i |\| ta_der \ t" using assms(1) by (simp add: ta_reachable_def) force} then obtain ts where ts: "i < length qs \ ground (ts i) \ qs ! i |\| ta_der \ (ts i)" for i by metis obtain C' where C: "ground_ctxt C'" "qs ! length ts1 |\| ta_der \ C'\t\" using More(1)[OF q'(3)] by blast define D where "D \ More f (map ts [0..| ta_der \ D\t\" using ts C(2) q' unfolding D_def by (auto simp: nth_append_Cons not_le not_less le_less_Suc_eq Suc_le_eq intro!: exI[of _ qs] exI[of _ q']) ultimately show ?case by blast qed lemma ta_productive_def': assumes "\ \ |\| ta_reachable \" shows "ta_productive Q \ = {| q| q q' C. ground_ctxt C \ q' |\| ta_der \ (C\Var q\) \ q' |\| Q |}" using ta_productive_aux[OF assms] by (auto simp: ta_productive_def intro!: finite_subset[OF _ finite_ta_productive, of _ \ Q]) force+ (* turn a finite automaton into a trim one, by removing first all unreachable and then all non-productive states *) lemma trim_gta_lang: "gta_lang Q (trim_ta Q \) = gta_lang Q \" unfolding trim_ta_def gta_only_reach_lang gta_only_prod_lang .. lemma trim_ta_subset: "ta_subset (trim_ta Q \) \" unfolding trim_ta_def by (rule ta_prod_reach_subset) theorem trim_ta: "ta_is_trim Q (trim_ta Q \)" unfolding ta_is_trim_def by (metis fin_mono ta_only_prod_reachable ta_only_reach_reachable ta_prodictive_ta_only_prod ta_restrict_states_Q trim_ta_def) lemma reg_is_trim_trim_reg [simp]: "reg_is_trim (trim_reg R)" unfolding reg_is_trim_def trim_reg_def by (simp add: trim_ta) lemma trim_reg_reach [simp]: "\\<^sub>r (trim_reg A) |\| ta_reachable (ta (trim_reg A))" by (auto simp: trim_reg_def) (meson ta_is_trim_def trim_ta) lemma trim_reg_prod [simp]: "\\<^sub>r (trim_reg A) |\| ta_productive (fin (trim_reg A)) (ta (trim_reg A))" by (auto simp: trim_reg_def) (meson ta_is_trim_def trim_ta) (* Proposition 7: every tree automaton can be turned into an equivalent trim one *) lemmas obtain_trimmed_ta = trim_ta trim_gta_lang ta_subset_det[OF trim_ta_subset] (* Trim tree automaton signature *) lemma \_trim_ta_sig: assumes "reg_is_trim R" "\ R \ \\<^sub>G (fset \)" shows "ta_sig (ta R) |\| \" proof - {fix r assume r: "r |\| rules (ta R)" then obtain f ps p where [simp]: "r = f ps \ p" by (cases r) auto from r assms(1) have "fset_of_list ps |\| ta_reachable (ta R)" by (auto simp add: rule_statesD(4) reg_is_trim_def ta_is_trim_def) from ta_reachableI_rule[OF this, of f p] r obtain ts where ts: "length ts = length ps" "\ i < length ps. ground (ts ! i)" "\ i < length ps. ps ! i |\| ta_der (ta R) (ts ! i)" by auto obtain C q where ctxt: "ground_ctxt C" "q |\| ta_der (ta R) (C\Var p\)" "q |\| fin R" using assms(1) unfolding reg_is_trim_def by (metis \r = f ps \ p\ fsubsetI r rule_statesD(2) ta_productiveE ta_productive_aux ta_is_trim_def) from ts ctxt r have reach: "q |\| ta_der (ta R) C\Fun f ts\" by auto (metis ta_der_Fun ta_der_ctxt) have gr: "ground C\Fun f ts\" using ts(1, 2) ctxt(1) by (auto simp: in_set_conv_nth) then have "C\Fun f ts\ \ ta_lang (fin R) (ta R)" using ctxt(1, 3) ts(1, 2) apply (intro ta_langI[OF _ _ reach, of "fin R" "C\Fun f ts\"]) apply (auto simp del: adapt_vars_ctxt) by (metis gr adapt_vars2 adapt_vars_adapt_vars) then have *: "gterm_of_term C\Fun f ts\ \ \ R" using gr by (auto simp: \_def) then have "funas_gterm (gterm_of_term C\Fun f ts\) \ fset \" using assms(2) gr by (auto simp: \\<^sub>G_equivalent_def) moreover have "(f, length ps) \ funas_gterm (gterm_of_term C\Fun f ts\)" using ts(1) by (auto simp: funas_gterm_gterm_of_term[OF gr]) ultimately have "(r_root r, length (r_lhs_states r)) |\| \" by auto} then show ?thesis by (auto simp: ta_sig_def) qed text \Map function over TA rules which change states/signature\ lemma map_ta_rule_iff: "map_ta_rule f g |`| \ = {|TA_rule (g h) (map f qs) (f q) | h qs q. TA_rule h qs q |\| \|}" apply (intro fequalityI fsubsetI) apply (auto simp add: rev_image_eqI) apply (metis map_ta_rule_cases ta_rule.collapse) done lemma \_trim: "\ (trim_reg R) = \ R" by (auto simp: trim_gta_lang \_def trim_reg_def) lemma fmap_funs_ta_def': "fmap_funs_ta h \ = TA {|(h f) qs \ q |f qs q. f qs \ q |\| rules \|} (eps \)" unfolding fmap_funs_ta_def map_ta_rule_iff by auto lemma fmap_states_ta_def': "fmap_states_ta h \ = TA {|f (map h qs) \ h q |f qs q. f qs \ q |\| rules \|} (map_both h |`| eps \)" unfolding fmap_states_ta_def map_ta_rule_iff by auto lemma fmap_states [simp]: "\ (fmap_states_ta h \) = h |`| \ \" unfolding fmap_states_ta_def \_def by auto lemma fmap_states_ta_sig [simp]: "ta_sig (fmap_states_ta f \) = ta_sig \" by (auto simp: fmap_states_ta_def ta_sig_def ta_rule.map_sel intro: fset.map_cong0) lemma fmap_states_ta_eps_wit: assumes "(h p, q) |\| (map_both h |`| eps \)|\<^sup>+|" "finj_on h (\ \)" "p |\| \ \" - obtains q' where "q = h q'" "(p, q') |\| (eps \)|\<^sup>+|" "q' |\| \ \" using assms - by (auto simp: fimage_iff finj_on_def' ftrancl_map_both_fsubset[OF assms(2), of "eps \"]) - (metis (mono_tags, lifting) assms(2) eps_trancl_statesD finj_on_eq_iff) + obtains q' where "q = h q'" "(p, q') |\| (eps \)|\<^sup>+|" "q' |\| \ \" + using assms(1)[unfolded ftrancl_map_both_fsubset[OF assms(2), of "eps \", simplified]] + using \finj_on h (\ \)\[unfolded finj_on_def', rule_format, OF \p |\| \ \\] + by (metis Pair_inject eps_trancl_statesD prod_fun_fimageE) lemma ta_der_fmap_states_inv_superset: assumes "\ \ |\| \" "finj_on h \" and "q |\| ta_der (fmap_states_ta h \) (term_of_gterm t)" shows "the_finv_into \ h q |\| ta_der \ (term_of_gterm t)" using assms(3) proof (induct rule: ta_der_gterm_induct) case (GFun f ts ps p q) from assms(1, 2) have inj: "finj_on h (\ \)" using fsubset_finj_on by blast have "x |\| \ \ \ the_finv_into (\ \) h (h x) = the_finv_into \ h (h x)" for x using assms(1, 2) by (metis fsubsetD inj the_finv_into_f_f) then show ?case using GFun the_finv_into_f_f[OF inj] assms(1) by (auto simp: fmap_states_ta_def' finj_on_def' rule_statesD eps_statesD elim!: fmap_states_ta_eps_wit[OF _ inj] intro!: exI[of _ "the_finv_into \ h p"]) qed lemma ta_der_fmap_states_inv: assumes "finj_on h (\ \)" "q |\| ta_der (fmap_states_ta h \) (term_of_gterm t)" shows "the_finv_into (\ \) h q |\| ta_der \ (term_of_gterm t)" by (simp add: ta_der_fmap_states_inv_superset assms) lemma ta_der_to_fmap_states_der: assumes "q |\| ta_der \ (term_of_gterm t)" shows "h q |\| ta_der (fmap_states_ta h \) (term_of_gterm t)" using assms proof (induct rule: ta_der_gterm_induct) case (GFun f ts ps p q) then show ?case using ftrancl_map_prod_mono[of h "eps \"] by (auto simp: fmap_states_ta_def' intro!: exI[of _ "h p"] exI[of _ "map h ps"]) qed lemma ta_der_fmap_states_conv: assumes "finj_on h (\ \)" shows "ta_der (fmap_states_ta h \) (term_of_gterm t) = h |`| ta_der \ (term_of_gterm t)" using ta_der_to_fmap_states_der[of _ \ t] ta_der_fmap_states_inv[OF assms] using f_the_finv_into_f[OF assms] finj_on_the_finv_into[OF assms] using gterm_ta_der_states by (auto intro!: rev_fimage_eqI) fastforce lemma fmap_states_ta_det: assumes "finj_on f (\ \)" shows "ta_det (fmap_states_ta f \) = ta_det \" (is "?Ls = ?Rs") proof {fix g ps p q assume ass: "?Ls" "TA_rule g ps p |\| rules \" "TA_rule g ps q |\| rules \" then have "TA_rule g (map f ps) (f p) |\| rules (fmap_states_ta f \)" "TA_rule g (map f ps) (f q) |\| rules (fmap_states_ta f \)" by (force simp: fmap_states_ta_def)+ then have "p = q" using ass finj_on_eq_iff[OF assms] by (auto simp: ta_det_def) (meson rule_statesD(2))} then show "?Ls \ ?Rs" by (auto simp: ta_det_def fmap_states_ta_def') next {fix g ps qs p q assume ass: "?Rs" "TA_rule g ps p |\| rules \" "TA_rule g qs q |\| rules \" then have "map f ps = map f qs \ ps = qs" using finj_on_eq_iff[OF assms] by (auto simp: map_eq_nth_conv in_fset_conv_nth dest!: rule_statesD(4) intro!: nth_equalityI)} then show "?Rs \ ?Ls" using finj_on_eq_iff[OF assms] by (auto simp: ta_det_def fmap_states_ta_def') blast qed lemma fmap_states_ta_lang: "finj_on f (\ \) \ Q |\| \ \ \ gta_lang (f |`| Q) (fmap_states_ta f \) = gta_lang Q \" using ta_der_fmap_states_conv[of f \] by (auto simp: finj_on_def' finj_on_eq_iff fsubsetD elim!: gta_langE intro!: gta_langI) lemma fmap_states_ta_lang2: "finj_on f (\ \ |\| Q) \ gta_lang (f |`| Q) (fmap_states_ta f \) = gta_lang Q \" using ta_der_fmap_states_conv[OF fsubset_finj_on[of f "\ \ |\| Q" "\ \"]] by (auto simp: finj_on_def' elim!: gta_langE intro!: gta_langI) fastforce definition funs_ta :: "('q, 'f) ta \ 'f fset" where "funs_ta \ = {|f |f qs q. TA_rule f qs q |\| rules \|}" lemma funs_ta[code]: "funs_ta \ = (\ r. case r of TA_rule f ps p \ f) |`| (rules \)" (is "?Ls = ?Rs") by (force simp: funs_ta_def rev_fimage_eqI simp flip: fset.set_map split!: ta_rule.splits intro!: finite_subset[of "{f. \qs q. TA_rule f qs q |\| rules \}" "fset ?Rs"]) lemma finite_funs_ta [simp]: "finite {f. \qs q. TA_rule f qs q |\| rules \}" by (intro finite_subset[of "{f. \qs q. TA_rule f qs q |\| rules \}" "fset (funs_ta \)"]) (auto simp: funs_ta rev_fimage_eqI simp flip: fset.set_map split!: ta_rule.splits) lemma funs_taE [elim]: assumes "f |\| funs_ta \" obtains ps p where "TA_rule f ps p |\| rules \" using assms by (auto simp: funs_ta_def) lemma funs_taI [intro]: "TA_rule f ps p |\| rules \ \ f |\| funs_ta \" by (auto simp: funs_ta_def) lemma fmap_funs_ta_cong: "(\x. x |\| funs_ta \ \ h x = k x) \ \ = \ \ fmap_funs_ta h \ = fmap_funs_ta k \" by (force simp: fmap_funs_ta_def') lemma [simp]: "{|TA_rule f qs q |f qs q. TA_rule f qs q |\| X|} = X" by (intro fset_eqI; case_tac x) auto lemma fmap_funs_ta_id [simp]: "fmap_funs_ta id \ = \" by (simp add: fmap_funs_ta_def') lemma fmap_states_ta_id [simp]: "fmap_states_ta id \ = \" by (auto simp: fmap_states_ta_def map_ta_rule_iff prod.map_id0) lemmas fmap_funs_ta_id' [simp] = fmap_funs_ta_id[unfolded id_def] lemma fmap_funs_ta_comp: "fmap_funs_ta h (fmap_funs_ta k A) = fmap_funs_ta (h \ k) A" proof - have "r |\| rules A \ map_ta_rule id h (map_ta_rule id k r) = map_ta_rule id (\x. h (k x)) r" for r by (cases r) (auto) then show ?thesis by (force simp: fmap_funs_ta_def fimage_iff cong: fmap_funs_ta_cong) qed lemma fmap_funs_reg_comp: "fmap_funs_reg h (fmap_funs_reg k A) = fmap_funs_reg (h \ k) A" using fmap_funs_ta_comp unfolding fmap_funs_reg_def by auto lemma fmap_states_ta_comp: "fmap_states_ta h (fmap_states_ta k A) = fmap_states_ta (h \ k) A" by (auto simp: fmap_states_ta_def ta_rule.map_comp comp_def id_def prod.map_comp) lemma funs_ta_fmap_funs_ta [simp]: "funs_ta (fmap_funs_ta f A) = f |`| funs_ta A" by (auto simp: funs_ta fmap_funs_ta_def' comp_def fimage_iff split!: ta_rule.splits) force+ lemma ta_der_funs_ta: "q |\| ta_der A t \ ffuns_term t |\| funs_ta A" proof (induct t arbitrary: q) case (Fun f ts) then have "f |\| funs_ta A" by (auto simp: funs_ta_def) then show ?case using Fun(1)[OF nth_mem, THEN fsubsetD] Fun(2) by (auto simp: in_fset_conv_nth) blast+ qed auto lemma ta_der_fmap_funs_ta: "q |\| ta_der A t \ q |\| ta_der (fmap_funs_ta f A) (map_funs_term f t)" by (induct t arbitrary: q) (auto 0 4 simp: fmap_funs_ta_def') lemma ta_der_fmap_states_ta: assumes "q |\| ta_der A t" shows "h q |\| ta_der (fmap_states_ta h A) (map_vars_term h t)" proof - have [intro]: "(q, q') |\| (eps A)|\<^sup>+| \ (h q, h q') |\| (eps (fmap_states_ta h A))|\<^sup>+|" for q q' by (force intro!: ftrancl_map[of "eps A"] simp: fmap_states_ta_def) show ?thesis using assms proof (induct rule: ta_der_induct) case (Fun f ts ps p q) have "f (map h ps) \ h p |\| rules (fmap_states_ta h A)" using Fun(1) by (force simp: fmap_states_ta_def') then show ?case using Fun by (auto 0 4) qed auto qed lemma ta_der_fmap_states_ta_mono: shows "f |`| ta_der A (term_of_gterm s) |\| ta_der (fmap_states_ta f A) (term_of_gterm s)" using ta_der_fmap_states_ta[of _ A "term_of_gterm s" f] by (simp add: fimage_fsubsetI ta_der_to_fmap_states_der) lemma ta_der_fmap_states_ta_mono2: assumes "finj_on f (\ A)" shows "ta_der (fmap_states_ta f A) (term_of_gterm s) |\| f |`| ta_der A (term_of_gterm s)" using ta_der_fmap_states_conv[OF assms] by auto lemma fmap_funs_ta_der': "q |\| ta_der (fmap_funs_ta h A) t \ \t'. q |\| ta_der A t' \ map_funs_term h t' = t" proof (induct rule: ta_der_induct) case (Var q v) then show ?case by (auto simp: fmap_funs_ta_def intro!: exI[of _ "Var v"]) next case (Fun f ts ps p q) obtain f' ts' where root: "f = h f'" "f' ps \ p |\| rules A" and "\i. i < length ts \ ps ! i |\| ta_der A (ts' i) \ map_funs_term h (ts' i) = ts ! i" using Fun(1, 5) unfolding fmap_funs_ta_def' by auto metis note [simp] = conjunct1[OF this(3)] conjunct2[OF this(3), unfolded id_def] have [simp]: "p = q \ f' ps \ q |\| rules A" using root(2) by auto show ?case using Fun(3) by (auto simp: comp_def Fun root fmap_funs_ta_def' intro!: exI[of _ "Fun f' (map ts' [0..) = map_gterm h ` gta_lang Q \" (is "?Ls = ?Rs") proof - {fix s assume "s \ ?Ls" then obtain q where lang: "q |\| Q" "q |\| ta_der (fmap_funs_ta h \) (term_of_gterm s)" by auto from fmap_funs_ta_der'[OF this(2)] obtain t where t: "q |\| ta_der \ t" "map_funs_term h t = term_of_gterm s" "ground t" by (metis ground_map_term ground_term_of_gterm) then have "s \ ?Rs" using map_gterm_of_term[OF t(3), of h id] lang by (auto simp: gta_lang_def gta_der_def image_iff) (metis fempty_iff finterI ground_term_to_gtermD map_term_of_gterm term_of_gterm_inv)} moreover have "?Rs \ ?Ls" using ta_der_fmap_funs_ta[of _ \ _ h] by (auto elim!: gta_langE intro!: gta_langI) fastforce ultimately show ?thesis by blast qed lemma fmap_funs_\: "\ (fmap_funs_reg h R) = map_gterm h ` \ R" using fmap_funs_gta_lang[of "fin R" h] by (auto simp: fmap_funs_reg_def \_def) lemma ta_states_fmap_funs_ta [simp]: "\ (fmap_funs_ta f A) = \ A" by (auto simp: fmap_funs_ta_def \_def) lemma ta_reachable_fmap_funs_ta [simp]: "ta_reachable (fmap_funs_ta f A) = ta_reachable A" unfolding ta_reachable_def by (metis (mono_tags, lifting) fmap_funs_ta_der' ta_der_fmap_funs_ta ground_map_term) lemma fin_in_states: "fin (reg_Restr_Q\<^sub>f R) |\| \\<^sub>r (reg_Restr_Q\<^sub>f R)" by (auto simp: reg_Restr_Q\<^sub>f_def) lemma fmap_states_reg_Restr_Q\<^sub>f_fin: "finj_on f (\ \) \ fin (fmap_states_reg f (reg_Restr_Q\<^sub>f R)) |\| \\<^sub>r (fmap_states_reg f (reg_Restr_Q\<^sub>f R))" by (auto simp: fmap_states_reg_def reg_Restr_Q\<^sub>f_def) lemma \_fmap_states_reg_Inl_Inr [simp]: "\ (fmap_states_reg Inl R) = \ R" "\ (fmap_states_reg Inr R) = \ R" unfolding \_def fmap_states_reg_def by (auto simp: finj_Inl_Inr intro!: fmap_states_ta_lang2) lemma finite_Collect_prod_ta_rules: "finite {f qs \ (a, b) |f qs a b. f map fst qs \ a |\| rules \ \ f map snd qs \ b |\| rules \}" (is "finite ?set") proof - have "?set \ (\ (ra, rb). case ra of f ps \ p \ case rb of g qs \ q \ f (zip ps qs) \ (p, q)) ` (srules \ \ srules \)" by (auto simp: srules_def image_iff split!: ta_rule.splits) (metis ta_rule.inject zip_map_fst_snd) from finite_imageI[of "srules \ \ srules \", THEN finite_subset[OF this]] show ?thesis by (auto simp: srules_def) qed \ \The product automaton of the automata A and B is constructed by applying the rules on pairs of states\ lemmas prod_eps_def = prod_epsLp_def prod_epsRp_def lemma finite_prod_epsLp: "finite (Collect (prod_epsLp \ \))" by (intro finite_subset[of "Collect (prod_epsLp \ \)" "fset ((\ \ |\| \ \) |\| \ \ |\| \ \)"]) (auto simp: prod_epsLp_def dest: eps_statesD) lemma finite_prod_epsRp: "finite (Collect (prod_epsRp \ \))" by (intro finite_subset[of "Collect (prod_epsRp \ \)" "fset ((\ \ |\| \ \) |\| \ \ |\| \ \)"]) (auto simp: prod_epsRp_def dest: eps_statesD) lemmas finite_prod_eps [simp] = finite_prod_epsLp[unfolded prod_epsLp_def] finite_prod_epsRp[unfolded prod_epsRp_def] lemma [simp]: "f qs \ q |\| rules (prod_ta \ \) \ f qs \ q |\| prod_ta_rules \ \" "r |\| rules (prod_ta \ \) \ r |\| prod_ta_rules \ \" by (auto simp: prod_ta_def) lemma prod_ta_states: "\ (prod_ta \ \) |\| \ \ |\| \ \" proof - {fix q assume "q |\| rule_states (rules (prod_ta \ \))" then obtain f ps p where "f ps \ p |\| rules (prod_ta \ \)" and "q |\| fset_of_list ps \ p = q" by (metis rule_statesE) then have "fst q |\| \ \ \ snd q |\| \ \" using rule_statesD(2, 4)[of f "map fst ps" "fst p" \] using rule_statesD(2, 4)[of f "map snd ps" "snd p" \] by auto} moreover {fix q assume "q |\| eps_states (eps (prod_ta \ \))" then have "fst q |\| \ \ \ snd q |\| \ \" by (auto simp: eps_states_def prod_ta_def prod_eps_def dest: eps_statesD)} ultimately show ?thesis by (auto simp: \_def) blast+ qed lemma prod_ta_det: assumes "ta_det \" and "ta_det \" shows "ta_det (prod_ta \ \)" using assms unfolding ta_det_def prod_ta_def prod_eps_def by auto lemma prod_ta_sig: "ta_sig (prod_ta \ \) |\| ta_sig \ |\| ta_sig \" proof (rule fsubsetI) fix x assume "x |\| ta_sig (prod_ta \ \)" hence "x |\| ta_sig \ \ x |\| ta_sig \" unfolding ta_sig_def prod_ta_def using image_iff by fastforce thus "x |\| ta_sig (prod_ta \ \) \ x |\| ta_sig \ |\| ta_sig \" by simp qed lemma from_prod_eps: "(p, q) |\| (eps (prod_ta \ \))|\<^sup>+| \ (snd p, snd q) |\| (eps \)|\<^sup>+| \ snd p = snd q \ (fst p, fst q) |\| (eps \)|\<^sup>+|" "(p, q) |\| (eps (prod_ta \ \))|\<^sup>+| \ (fst p, fst q) |\| (eps \)|\<^sup>+| \ fst p = fst q \ (snd p, snd q) |\| (eps \)|\<^sup>+|" apply (induct rule: ftrancl_induct) apply (auto simp: prod_ta_def prod_eps_def intro: ftrancl_into_trancl ) apply (simp add: fr_into_trancl not_ftrancl_into)+ done lemma to_prod_eps\: "(p, q) |\| (eps \)|\<^sup>+| \ r |\| \ \ \ ((p, r), (q, r)) |\| (eps (prod_ta \ \))|\<^sup>+|" by (induct rule: ftrancl_induct) (auto simp: prod_ta_def prod_eps_def intro: fr_into_trancl ftrancl_into_trancl) lemma to_prod_eps\: "(p, q) |\| (eps \)|\<^sup>+| \ r |\| \ \ \ ((r, p), (r, q)) |\| (eps (prod_ta \ \))|\<^sup>+|" by (induct rule: ftrancl_induct) (auto simp: prod_ta_def prod_eps_def intro: fr_into_trancl ftrancl_into_trancl) lemma to_prod_eps: "(p, q) |\| (eps \)|\<^sup>+| \ (p', q') |\| (eps \)|\<^sup>+| \ ((p, p'), (q, q')) |\| (eps (prod_ta \ \))|\<^sup>+|" proof (induct rule: ftrancl_induct) case (Base a b) show ?case using Base(2, 1) proof (induct rule: ftrancl_induct) case (Base c d) then have "((a, c), b, c) |\| (eps (prod_ta \ \))|\<^sup>+|" using finite_prod_eps by (auto simp: prod_ta_def prod_eps_def dest: eps_statesD intro!: fr_into_trancl ftrancl_into_trancl) moreover have "((b, c), b, d) |\| (eps (prod_ta \ \))|\<^sup>+|" using finite_prod_eps Base by (auto simp: prod_ta_def prod_eps_def dest: eps_statesD intro!: fr_into_trancl ftrancl_into_trancl) ultimately show ?case by (auto intro: ftrancl_trans) next case (Step p q r) then have "((b, q), b, r) |\| (eps (prod_ta \ \))|\<^sup>+|" using finite_prod_eps by (auto simp: prod_ta_def prod_eps_def dest: eps_statesD intro!: fr_into_trancl) then show ?case using Step by (auto intro: ftrancl_trans) qed next case (Step a b c) from Step have "q' |\| \ \" by (auto dest: eps_trancl_statesD) then have "((b, q'), (c,q')) |\| (eps (prod_ta \ \))|\<^sup>+|" using Step(3) finite_prod_eps by (auto simp: prod_ta_def prod_eps_def intro!: fr_into_trancl) then show ?case using ftrancl_trans Step by auto qed lemma prod_ta_der_to_\_\_der1: assumes "q |\| ta_der (prod_ta \ \) (term_of_gterm t)" shows "fst q |\| ta_der \ (term_of_gterm t)" using assms proof (induct rule: ta_der_gterm_induct) case (GFun f ts ps p q) then show ?case by (auto dest: from_prod_eps intro!: exI[of _ "map fst ps"] exI[of _ "fst p"]) qed lemma prod_ta_der_to_\_\_der2: assumes "q |\| ta_der (prod_ta \ \) (term_of_gterm t)" shows "snd q |\| ta_der \ (term_of_gterm t)" using assms proof (induct rule: ta_der_gterm_induct) case (GFun f ts ps p q) then show ?case by (auto dest: from_prod_eps intro!: exI[of _ "map snd ps"] exI[of _ "snd p"]) qed lemma \_\_der_to_prod_ta: assumes "fst q |\| ta_der \ (term_of_gterm t)" "snd q |\| ta_der \ (term_of_gterm t)" shows "q |\| ta_der (prod_ta \ \) (term_of_gterm t)" using assms proof (induct t arbitrary: q) case (GFun f ts) from GFun(2, 3) obtain ps qs p q' where rules: "f ps \ p |\| rules \" "f qs \ q' |\| rules \" "length ps = length ts" "length ps = length qs" and eps: "p = fst q \ (p, fst q) |\| (eps \)|\<^sup>+|" "q' = snd q \ (q', snd q) |\| (eps \)|\<^sup>+|" and steps: "\ i < length qs. ps ! i |\| ta_der \ (term_of_gterm (ts ! i))" "\ i < length qs. qs ! i |\| ta_der \ (term_of_gterm (ts ! i))" by auto from rules have st: "p |\| \ \" "q' |\| \ \" by (auto dest: rule_statesD) have "(p, snd q) = q \ ((p, q'), q) |\| (eps (prod_ta \ \))|\<^sup>+|" using eps st using to_prod_eps\[of q' "snd q" \ "fst q" \] using to_prod_eps\[of p "fst q" \ "snd q" \] using to_prod_eps[of p "fst q" \ q' "snd q" \] by (cases "p = fst q"; cases "q' = snd q") (auto simp: prod_ta_def) then show ?case using rules eps steps GFun(1) st by (cases "(p, snd q) = q") (auto simp: finite_Collect_prod_ta_rules dest: to_prod_eps\ intro!: exI[of _ p] exI[of _ q'] exI[of _ "zip ps qs"]) qed lemma prod_ta_der: "q |\| ta_der (prod_ta \ \) (term_of_gterm t) \ fst q |\| ta_der \ (term_of_gterm t) \ snd q |\| ta_der \ (term_of_gterm t)" using prod_ta_der_to_\_\_der1 prod_ta_der_to_\_\_der2 \_\_der_to_prod_ta by blast lemma intersect_ta_gta_lang: "gta_lang (Q\<^sub>\ |\| Q\<^sub>\) (prod_ta \ \) = gta_lang Q\<^sub>\ \ \ gta_lang Q\<^sub>\ \" by (auto simp: prod_ta_der elim!: gta_langE intro: gta_langI) lemma \_intersect: "\ (reg_intersect R L) = \ R \ \ L" by (auto simp: intersect_ta_gta_lang \_def reg_intersect_def) lemma intersect_ta_ta_lang: "ta_lang (Q\<^sub>\ |\| Q\<^sub>\) (prod_ta \ \) = ta_lang Q\<^sub>\ \ \ ta_lang Q\<^sub>\ \" using intersect_ta_gta_lang[of Q\<^sub>\ Q\<^sub>\ \ \] by auto (metis IntI imageI term_of_gterm_inv) \ \Union of tree automata\ lemma ta_union_ta_subset: "ta_subset \ (ta_union \ \)" "ta_subset \ (ta_union \ \)" unfolding ta_subset_def ta_union_def by auto lemma ta_union_states [simp]: "\ (ta_union \ \) = \ \ |\| \ \" by (auto simp: ta_union_def \_def) lemma ta_union_sig [simp]: "ta_sig (ta_union \ \) = ta_sig \ |\| ta_sig \" by (auto simp: ta_union_def ta_sig_def) lemma ta_union_eps_disj_states: assumes "\ \ |\| \ \ = {||}" and "(p, q) |\| (eps (ta_union \ \))|\<^sup>+|" shows "(p, q) |\| (eps \)|\<^sup>+| \ (p, q) |\| (eps \)|\<^sup>+|" using assms(2, 1) by (induct rule: ftrancl_induct) (auto simp: ta_union_def ftrancl_into_trancl dest: eps_statesD eps_trancl_statesD) lemma eps_ta_union_eps [simp]: "(p, q) |\| (eps \)|\<^sup>+| \ (p, q) |\| (eps (ta_union \ \))|\<^sup>+|" "(p, q) |\| (eps \)|\<^sup>+| \ (p, q) |\| (eps (ta_union \ \))|\<^sup>+|" by (auto simp add: in_ftrancl_UnI ta_union_def) lemma disj_states_eps [simp]: "\ \ |\| \ \ = {||} \ f ps \ p |\| rules \ \ (p, q) |\| (eps \)|\<^sup>+| \ False" "\ \ |\| \ \ = {||} \ f ps \ p |\| rules \ \ (p, q) |\| (eps \)|\<^sup>+| \ False" by (auto simp: rtrancl_eq_or_trancl dest: rule_statesD eps_trancl_statesD) lemma ta_union_der_disj_states: assumes "\ \ |\| \ \ = {||}" and "q |\| ta_der (ta_union \ \) t" shows "q |\| ta_der \ t \ q |\| ta_der \ t" using assms(2) proof (induct rule: ta_der_induct) case (Var q v) then show ?case using ta_union_eps_disj_states[OF assms(1)] by auto next case (Fun f ts ps p q) have dist: "fset_of_list ps |\| \ \ \ i < length ts \ ps ! i |\| ta_der \ (ts ! i)" "fset_of_list ps |\| \ \ \ i < length ts \ ps ! i |\| ta_der \ (ts ! i)" for i using Fun(2) Fun(5)[of i] assms(1) by (auto dest!: ta_der_not_stateD fsubsetD) from Fun(1) consider (a) "fset_of_list ps |\| \ \" | (b) "fset_of_list ps |\| \ \" by (auto simp: ta_union_def dest: rule_statesD) then show ?case using dist Fun(1, 2) assms(1) ta_union_eps_disj_states[OF assms(1), of p q] Fun(3) by (cases) (auto simp: fsubsetI rule_statesD ta_union_def intro!: exI[of _ p] exI[of _ ps]) qed lemma ta_union_der_disj_states': assumes "\ \ |\| \ \ = {||}" shows "ta_der (ta_union \ \) t = ta_der \ t |\| ta_der \ t" using ta_union_der_disj_states[OF assms] ta_der_mono' ta_union_ta_subset by (auto, fastforce) blast lemma ta_union_gta_lang: assumes "\ \ |\| \ \ = {||}" and "Q\<^sub>\ |\| \ \" and "Q\<^sub>\ |\| \ \" shows"gta_lang (Q\<^sub>\ |\| Q\<^sub>\) (ta_union \ \) = gta_lang Q\<^sub>\ \ \ gta_lang Q\<^sub>\ \" (is "?Ls = ?Rs") proof - {fix s assume "s \ ?Ls" then obtain q where w: "q |\| Q\<^sub>\ |\| Q\<^sub>\" "q |\| ta_der (ta_union \ \) (term_of_gterm s)" by (auto elim: gta_langE) from ta_union_der_disj_states[OF assms(1) w(2)] consider (a) "q |\| ta_der \ (term_of_gterm s)" | "q |\| ta_der \ (term_of_gterm s)" by blast then have "s \ ?Rs" using w(1) assms by (cases, auto simp: gta_langI) (metis fempty_iff finterI funion_iff gterm_ta_der_states sup.orderE)} moreover have "?Rs \ ?Ls" using ta_union_der_disj_states'[OF assms(1)] by (auto elim!: gta_langE intro!: gta_langI) ultimately show ?thesis by blast qed lemma \_union: "\ (reg_union R L) = \ R \ \ L" proof - let ?inl = "Inl :: 'b \ 'b + 'c" let ?inr = "Inr :: 'c \ 'b + 'c" let ?fr = "?inl |`| (fin R |\| \\<^sub>r R)" let ?fl = "?inr |`| (fin L |\| \\<^sub>r L)" have [simp]:"gta_lang (?fr |\| ?fl) (ta_union (fmap_states_ta ?inl (ta R)) (fmap_states_ta ?inr (ta L))) = gta_lang ?fr (fmap_states_ta ?inl (ta R)) \ gta_lang ?fl (fmap_states_ta ?inr (ta L))" by (intro ta_union_gta_lang) (auto simp: fimage_iff) have [simp]: "gta_lang ?fr (fmap_states_ta ?inl (ta R)) = gta_lang (fin R |\| \\<^sub>r R) (ta R)" by (intro fmap_states_ta_lang) (auto simp: finj_Inl_Inr) have [simp]: "gta_lang ?fl (fmap_states_ta ?inr (ta L)) = gta_lang (fin L |\| \\<^sub>r L) (ta L)" by (intro fmap_states_ta_lang) (auto simp: finj_Inl_Inr) show ?thesis using gta_lang_Rest_states_conv by (auto simp: \_def reg_union_def ta_union_gta_lang) fastforce qed lemma reg_union_states: "\\<^sub>r (reg_union A B) = (Inl |`| \\<^sub>r A) |\| (Inr |`| \\<^sub>r B)" by (auto simp: reg_union_def) \ \Deciding emptiness\ lemma ta_empty [simp]: "ta_empty Q \ = (gta_lang Q \ = {})" by (auto simp: ta_empty_def elim!: gta_langE ta_reachable_gtermE intro: ta_reachable_gtermI gta_langI) lemma reg_empty [simp]: "reg_empty R = (\ R = {})" by (simp add: \_def reg_empty_def) text \Epsilon free automaton\ lemma finite_eps_free_rulep [simp]: "finite (Collect (eps_free_rulep \))" proof - let ?par = "(\ r. case r of f qs \ q \ (f, qs)) |`| (rules \)" let ?st = "(\ (r, q). case r of (f, qs) \ TA_rule f qs q) |`| (?par |\| \ \)" show ?thesis using rule_statesD eps_trancl_statesD by (intro finite_subset[of "Collect (eps_free_rulep \)" "fset ?st"]) (auto simp: eps_free_rulep_def fimage_iff simp flip: fset.set_map split!: ta_rule.splits, fastforce+) qed lemmas finite_eps_free_rule [simp] = finite_eps_free_rulep[unfolded eps_free_rulep_def] lemma ta_res_eps_free: "ta_der (eps_free \) (term_of_gterm t) = ta_der \ (term_of_gterm t)" (is "?Ls = ?Rs") proof - {fix q assume "q |\| ?Ls" then have "q |\| ?Rs" by (induct rule: ta_der_gterm_induct) (auto simp: eps_free_def eps_free_rulep_def)} moreover {fix q assume "q |\| ?Rs" then have "q |\| ?Ls" proof (induct rule: ta_der_gterm_induct) case (GFun f ts ps p q) then show ?case by (auto simp: eps_free_def eps_free_rulep_def intro!: exI[of _ ps]) qed} ultimately show ?thesis by blast qed lemma ta_lang_eps_free [simp]: "gta_lang Q (eps_free \) = gta_lang Q \" by (auto simp add: ta_res_eps_free elim!: gta_langE intro: gta_langI) lemma \_eps_free: "\ (eps_free_reg R) = \ R" by (auto simp: \_def eps_free_reg_def) text \Sufficient criterion for containment\ (* sufficient criterion to check whether automaton accepts at least T_g(F) where F is a subset of the signature *) definition ta_contains_aux :: "('f \ nat) set \ 'q fset \ ('q, 'f) ta \ 'q fset \ bool" where "ta_contains_aux \ Q\<^sub>1 \ Q\<^sub>2 \ (\ f qs. (f, length qs) \ \ \ fset_of_list qs |\| Q\<^sub>1 \ (\ q q'. TA_rule f qs q |\| rules \ \ q' |\| Q\<^sub>2 \ (q = q' \ (q, q') |\| (eps \)|\<^sup>+|)))" lemma ta_contains_aux_state_set: assumes "ta_contains_aux \ Q \ Q" "t \ \\<^sub>G \" shows "\ q. q |\| Q \ q |\| ta_der \ (term_of_gterm t)" using assms(2) proof (induct rule: \\<^sub>G.induct) case (const a) then show ?case using assms(1) by (force simp: ta_contains_aux_def) next case (ind f n ss) obtain qs where "fset_of_list qs |\| Q" "length ss = length qs" "\ i < length qs. qs ! i |\| ta_der \ (term_of_gterm (ss ! i))" using ind(4) Ex_list_of_length_P[of "length ss" "\ q i. q |\| Q \ q |\| ta_der \ (term_of_gterm (ss ! i))"] by (auto simp: fset_list_fsubset_eq_nth_conv) metis then show ?case using ind(1 - 3) assms(1) by (auto simp: ta_contains_aux_def) blast qed lemma ta_contains_aux_mono: assumes "ta_subset \ \" and "Q\<^sub>2 |\| Q\<^sub>2'" shows "ta_contains_aux \ Q\<^sub>1 \ Q\<^sub>2 \ ta_contains_aux \ Q\<^sub>1 \ Q\<^sub>2'" using assms unfolding ta_contains_aux_def ta_subset_def by (meson fin_mono ftrancl_mono) definition ta_contains :: "('f \ nat) set \ ('f \ nat) set \ ('q, 'f) ta \ 'q fset \ 'q fset \ bool" where "ta_contains \ \ \ Q Q\<^sub>f \ ta_contains_aux \ Q \ Q \ ta_contains_aux \ Q \ Q\<^sub>f" lemma ta_contains_mono: assumes "ta_subset \ \" and "Q\<^sub>f |\| Q\<^sub>f'" shows "ta_contains \ \ \ Q Q\<^sub>f \ ta_contains \ \ \ Q Q\<^sub>f'" unfolding ta_contains_def using ta_contains_aux_mono[OF assms(1) fsubset_refl] using ta_contains_aux_mono[OF assms] by blast lemma ta_contains_both: assumes contain: "ta_contains \ \ \ Q Q\<^sub>f" shows "\ t. groot t \ \ \ \ (funas_gterm ` set (gargs t)) \ \ \ t \ gta_lang Q\<^sub>f \" proof - fix t :: "'a gterm" assume F: "\ (funas_gterm ` set (gargs t)) \ \" and G: "groot t \ \" obtain g ss where t[simp]: "t = GFun g ss" by (cases t, auto) then have "\ qs. length qs = length ss \ (\ i < length qs. qs ! i |\| ta_der \ (term_of_gterm (ss ! i)) \ qs ! i |\| Q)" using contain ta_contains_aux_state_set[of \ Q \ "ss ! i" for i] F unfolding ta_contains_def \\<^sub>G_funas_gterm_conv using Ex_list_of_length_P[of "length ss" "\ q i. q |\| Q \ q |\| ta_der \ (term_of_gterm (ss ! i))"] by auto (metis SUP_le_iff nth_mem) then obtain qs where " length qs = length ss" "\ i < length qs. qs ! i |\| ta_der \ (term_of_gterm (ss ! i))" "\ i < length qs. qs ! i |\| Q" by blast then obtain q where "q |\| Q\<^sub>f" "q |\| ta_der \ (term_of_gterm t)" using conjunct2[OF contain[unfolded ta_contains_def]] G by (auto simp: ta_contains_def ta_contains_aux_def fset_list_fsubset_eq_nth_conv) metis then show "t \ gta_lang Q\<^sub>f \" by (intro gta_langI) simp qed lemma ta_contains: assumes contain: "ta_contains \ \ \ Q Q\<^sub>f" shows "\\<^sub>G \ \ gta_lang Q\<^sub>f \" (is "?A \ _") proof - have [simp]: "funas_gterm t \ \ \ groot t \ \" for t by (cases t) auto have [simp]: "funas_gterm t \ \ \ \ (funas_gterm ` set (gargs t)) \ \" for t by (cases t) auto show ?thesis using ta_contains_both[OF contain] by (auto simp: \\<^sub>G_equivalent_def) qed text \Relabeling, map finite set to natural numbers\ lemma map_fset_to_nat_inj: assumes "Y |\| X" shows "finj_on (map_fset_to_nat X) Y" proof - { fix x y assume "x |\| X" "y |\| X" then have "x |\| fset_of_list (sorted_list_of_fset X)" "y |\| fset_of_list (sorted_list_of_fset X)" by simp_all note this[unfolded mem_idx_fset_sound] then have "x = y" if "map_fset_to_nat X x = map_fset_to_nat X y" using that nth_eq_iff_index_eq[OF distinct_sorted_list_of_fset[of X]] by (force dest: mem_idx_sound_output simp: map_fset_to_nat_def) } then show ?thesis using assms by (auto simp add: finj_on_def' fBall_def) qed lemma map_fset_fset_to_nat_inj: assumes "Y |\| X" shows "finj_on (map_fset_fset_to_nat X) Y" using assms proof - let ?f = "map_fset_fset_to_nat X" { fix x y assume "x |\| X" "y |\| X" then have "sorted_list_of_fset x |\| fset_of_list (sorted_list_of_fset (sorted_list_of_fset |`| X))" "sorted_list_of_fset y |\| fset_of_list (sorted_list_of_fset (sorted_list_of_fset |`| X))" unfolding map_fset_fset_to_nat_def by auto note this[unfolded mem_idx_fset_sound] then have "x = y" if "?f x = ?f y" using that nth_eq_iff_index_eq[OF distinct_sorted_list_of_fset[of "sorted_list_of_fset |`| X"]] by (auto simp: map_fset_fset_to_nat_def) (metis mem_idx_sound_output sorted_list_of_fset_simps(1))+} then show ?thesis using assms by (auto simp add: finj_on_def' fBall_def) qed lemma relabel_gta_lang [simp]: "gta_lang (relabel_Q\<^sub>f Q \) (relabel_ta \) = gta_lang Q \" proof - have "gta_lang (relabel_Q\<^sub>f Q \) (relabel_ta \) = gta_lang (Q |\| \ \) \" unfolding relabel_ta_def relabel_Q\<^sub>f_def by (intro fmap_states_ta_lang2 map_fset_to_nat_inj) simp then show ?thesis by fastforce qed lemma \_relable [simp]: "\ (relabel_reg R) = \ R" by (auto simp: \_def relabel_reg_def) lemma relabel_ta_lang [simp]: "ta_lang (relabel_Q\<^sub>f Q \) (relabel_ta \) = ta_lang Q \" unfolding ta_lang_to_gta_lang using relabel_gta_lang by simp lemma relabel_fset_gta_lang [simp]: "gta_lang (relabel_fset_Q\<^sub>f Q \) (relabel_fset_ta \) = gta_lang Q \" proof - have "gta_lang (relabel_fset_Q\<^sub>f Q \) (relabel_fset_ta \) = gta_lang (Q |\| \ \) \" unfolding relabel_fset_Q\<^sub>f_def relabel_fset_ta_def by (intro fmap_states_ta_lang2 map_fset_fset_to_nat_inj) simp then show ?thesis by fastforce qed lemma \_relable_fset [simp]: "\ (relable_fset_reg R) = \ R" by (auto simp: \_def relable_fset_reg_def) lemma ta_states_trim_ta: "\ (trim_ta Q \) |\| \ \" unfolding trim_ta_def using ta_prod_reach_states . lemma trim_ta_reach: "\ (trim_ta Q \) |\| ta_reachable (trim_ta Q \)" unfolding trim_ta_def using ta_only_prod_reachable ta_only_reach_reachable by metis lemma trim_ta_prod: "\ (trim_ta Q A) |\| ta_productive Q (trim_ta Q A)" unfolding trim_ta_def using ta_only_prod_productive by metis lemma empty_gta_lang: "gta_lang Q (TA {||} {||}) = {}" using ta_reachable_gtermI by (force simp: gta_lang_def gta_der_def elim!: ta_langE) abbreviation empty_reg where "empty_reg \ Reg {||} (TA {||} {||})" lemma \_epmty: "\ empty_reg = {}" by (auto simp: \_def empty_gta_lang) lemma const_ta_lang: "gta_lang {|q|} (TA {| TA_rule f [] q |} {||}) = {GFun f []}" proof - have [dest!]: "q' |\| ta_der (TA {| TA_rule f [] q |} {||}) t' \ ground t' \ t' = Fun f []" for t' q' by (induct t') auto show ?thesis by (auto simp: gta_lang_def gta_der_def elim!: gta_langE) (metis gterm_of_term.simps list.simps(8) term_of_gterm_inv) qed lemma run_argsD: "run \ s t \ length (gargs s) = length (gargs t) \ (\ i < length (gargs t). run \ (gargs s ! i) (gargs t ! i))" using run.cases by fastforce lemma run_root_rule: "run \ s t \ TA_rule (groot_sym t) (map ex_comp_state (gargs s)) (ex_rule_state s) |\| (rules \) \ (ex_rule_state s = ex_comp_state s \ (ex_rule_state s, ex_comp_state s) |\| (eps \)|\<^sup>+|)" by (cases s; cases t) (auto elim: run.cases) lemma run_poss_eq: "run \ s t \ gposs s = gposs t" by (induct rule: run.induct) auto lemma run_gsubt_cl: assumes "run \ s t" and "p \ gposs t" shows "run \ (gsubt_at s p) (gsubt_at t p)" using assms proof (induct p arbitrary: s t) case (Cons i p) show ?case using Cons(1) Cons(2-) by (cases s; cases t) (auto dest: run_argsD) qed auto lemma run_replace_at: assumes "run \ s t" and "run \ u v" and "p \ gposs s" and "ex_comp_state (gsubt_at s p) = ex_comp_state u" shows "run \ s[p \ u]\<^sub>G t[p \ v]\<^sub>G" using assms proof (induct p arbitrary: s t) case (Cons i p) obtain r q qs f ts where [simp]: "s = GFun (r, q) qs" "t = GFun f ts" by (cases s, cases t) auto have *: "j < length qs \ ex_comp_state (qs[i := (qs ! i)[p \ u]\<^sub>G] ! j) = ex_comp_state (qs ! j)" for j using Cons(5) by (cases "i = j", cases p) auto have [simp]: "map ex_comp_state (qs[i := (qs ! i)[p \ u]\<^sub>G]) = map ex_comp_state qs" using Cons(5) by (auto simp: *[unfolded comp_def] intro!: nth_equalityI) have "run \ (qs ! i)[p \ u]\<^sub>G (ts ! i)[p \ v]\<^sub>G" using Cons(2-) by (intro Cons(1)) (auto dest: run_argsD) moreover have "i < length qs" "i < length ts" using Cons(4) run_poss_eq[OF Cons(2)] by force+ ultimately show ?case using Cons(2) run_root_rule[OF Cons(2)] by (fastforce simp: nth_list_update dest: run_argsD intro!: run.intros) qed simp text \relating runs to derivation definition\ lemma run_to_comp_st_gta_der: "run \ s t \ ex_comp_state s |\| gta_der \ t" proof (induct s arbitrary: t) case (GFun q qs) show ?case using GFun(1)[OF nth_mem, of i "gargs t ! i" for i] using run_argsD[OF GFun(2)] run_root_rule[OF GFun(2-)] by (cases t) (auto simp: gta_der_def intro!: exI[of _ "map ex_comp_state qs"] exI[of _ "fst q"]) qed lemma run_to_rule_st_gta_der: assumes "run \ s t" shows "ex_rule_state s |\| gta_der \ t" proof (cases s) case [simp]: (GFun q qs) have "i < length qs \ ex_comp_state (qs ! i) |\| gta_der \ (gargs t ! i)" for i using run_to_comp_st_gta_der[of \] run_argsD[OF assms] by force then show ?thesis using conjunct1[OF run_argsD[OF assms]] run_root_rule[OF assms] by (cases t) (auto simp: gta_der_def intro!: exI[of _ "map ex_comp_state qs"] exI[of _ "fst q"]) qed lemma run_to_gta_der_gsubt_at: assumes "run \ s t" and "p \ gposs t" shows "ex_rule_state (gsubt_at s p) |\| gta_der \ (gsubt_at t p)" "ex_comp_state (gsubt_at s p) |\| gta_der \ (gsubt_at t p)" using assms run_gsubt_cl[THEN run_to_comp_st_gta_der] run_gsubt_cl[THEN run_to_rule_st_gta_der] by blast+ lemma gta_der_to_run: "q |\| gta_der \ t \ (\ p qs. run \ (GFun (p, q) qs) t)" unfolding gta_der_def proof (induct rule: ta_der_gterm_induct) case (GFun f ts ps p q) from GFun(5) Ex_list_of_length_P[of "length ts" "\ qs i. run \ (GFun (fst qs, ps ! i) (snd qs)) (ts ! i)"] obtain qss where mid: "length qss = length ts" "\ i < length ts .run \ (GFun (fst (qss ! i), ps ! i) (snd (qss ! i))) (ts ! i)" by auto have [simp]: "map (ex_comp_state \ (\(qs, y). GFun (fst y, qs) (snd y))) (zip ps qss) = ps" using GFun(2) mid(1) by (intro nth_equalityI) auto show ?case using mid GFun(1 - 4) by (intro exI[of _ p] exI[of _ "map2 (\ f args. GFun (fst args, f) (snd args)) ps qss"]) (auto intro: run.intros) qed lemma run_ta_der_ctxt_split1: assumes "run \ s t" "p \ gposs t" shows "ex_comp_state s |\| ta_der \ (ctxt_at_pos (term_of_gterm t) p)\Var (ex_comp_state (gsubt_at s p))\" using assms proof (induct p arbitrary: s t) case (Cons i p) obtain q f qs ts where [simp]: "s = GFun q qs" "t = GFun f ts" and l: "length qs = length ts" using run_argsD[OF Cons(2)] by (cases s, cases t) auto from Cons(2, 3) l have "ex_comp_state (qs ! i) |\| ta_der \ (ctxt_at_pos (term_of_gterm (ts ! i)) p)\Var (ex_comp_state (gsubt_at (qs ! i) p))\" by (intro Cons(1)) (auto dest: run_argsD) then show ?case using Cons(2-) l by (fastforce simp: nth_append_Cons min_def dest: run_root_rule run_argsD intro!: exI[of _ "map ex_comp_state (gargs s)"] exI[of _ "ex_rule_state s"] run_to_comp_st_gta_der[of \ "qs ! i" "ts ! i" for i, unfolded comp_def gta_der_def]) qed auto lemma run_ta_der_ctxt_split2: assumes "run \ s t" "p \ gposs t" shows "ex_comp_state s |\| ta_der \ (ctxt_at_pos (term_of_gterm t) p)\Var (ex_rule_state (gsubt_at s p))\" proof (cases "ex_rule_state (gsubt_at s p) = ex_comp_state (gsubt_at s p)") case False then show ?thesis using run_root_rule[OF run_gsubt_cl[OF assms]] by (intro ta_der_eps_ctxt[OF run_ta_der_ctxt_split1[OF assms]]) auto qed (auto simp: run_ta_der_ctxt_split1[OF assms, unfolded comp_def]) end \ No newline at end of file diff --git a/thys/Ribbon_Proofs/More_Finite_Map.thy b/thys/Ribbon_Proofs/More_Finite_Map.thy --- a/thys/Ribbon_Proofs/More_Finite_Map.thy +++ b/thys/Ribbon_Proofs/More_Finite_Map.thy @@ -1,106 +1,109 @@ section \Finite partial functions\ theory More_Finite_Map imports "HOL-Library.Finite_Map" begin +lemma fdisjoint_iff: "A |\| B = {||} \ (\x. x |\| A \ x |\| B)" + by (smt (verit) disjoint_iff_fnot_equal) + unbundle lifting_syntax unbundle fmap.lifting type_notation fmap (infix "\\<^sub>f" 9) subsection \Difference\ definition map_diff :: "('k \ 'v) \ 'k fset \ ('k \ 'v)" where "map_diff f ks = restrict_map f (- fset ks)" lift_definition fmap_diff :: "('k \\<^sub>f 'v) \ 'k fset \ ('k \\<^sub>f 'v)" (infix "\" 110) is "map_diff" by (auto simp add: map_diff_def) subsection \Comprehension\ definition make_map :: "'k fset \ 'v \ ('k \ 'v)" where "make_map ks v \ \k. if k \ fset ks then Some v else None" lemma make_map_transfer[transfer_rule]: "(rel_fset (=) ===> A ===> rel_map A) make_map make_map" unfolding make_map_def by transfer_prover lemma dom_make_map: "dom (make_map ks v) = fset ks" by (metis domIff make_map_def not_Some_eq set_eqI) lift_definition make_fmap :: "'k fset \ 'v \ ('k \\<^sub>f 'v)" ("[ _ |=> _ ]") is "make_map" parametric make_map_transfer by (unfold make_map_def dom_def, auto) lemma make_fmap_empty[simp]: "[ {||} |=> f ] = fmempty" by transfer (simp add: make_map_def) subsection \Domain\ lemma fmap_add_commute: assumes "fmdom A |\| fmdom B = {||}" shows "A ++\<^sub>f B = B ++\<^sub>f A" using assms including fset.lifting apply (transfer) apply (rule ext) apply (auto simp: dom_def map_add_def split: option.splits) done lemma make_fmap_union: "[ xs |=> v ] ++\<^sub>f [ ys |=> v] = [ xs |\| ys |=> v ]" by (transfer, auto simp add: make_map_def map_add_def) lemma fdom_make_fmap: "fmdom [ ks |=> v ] = ks" (* FIXME proper transfer proof *) apply (subst fmdom_def) apply transfer apply (auto simp: dom_def make_map_def fset_inverse) done subsection \Lookup\ lift_definition lookup :: "('k \\<^sub>f 'v) \ 'k \ 'v" is "(\) the" . lemma lookup_make_fmap: assumes "k \ fset ks" shows "lookup [ ks |=> v ] k = v" using assms by (transfer, simp add: make_map_def) lemma lookup_make_fmap1: "lookup [ {|k|} |=> v ] k = v" by (metis finsert.rep_eq insert_iff lookup_make_fmap) lemma lookup_union1: assumes "k |\| fmdom ys" shows "lookup (xs ++\<^sub>f ys) k = lookup ys k" using assms including fset.lifting by transfer auto lemma lookup_union2: assumes "k |\| fmdom ys" shows "lookup (xs ++\<^sub>f ys) k = lookup xs k" using assms including fset.lifting by transfer (auto simp: map_add_def split: option.splits) lemma lookup_union3: assumes "k |\| fmdom xs" shows "lookup (xs ++\<^sub>f ys) k = lookup ys k" using assms including fset.lifting by transfer (auto simp: map_add_def split: option.splits) end diff --git a/thys/Ribbon_Proofs/Ribbons_Graphical_Soundness.thy b/thys/Ribbon_Proofs/Ribbons_Graphical_Soundness.thy --- a/thys/Ribbon_Proofs/Ribbons_Graphical_Soundness.thy +++ b/thys/Ribbon_Proofs/Ribbons_Graphical_Soundness.thy @@ -1,1125 +1,1133 @@ section \Soundness for graphical diagrams\ theory Ribbons_Graphical_Soundness imports Ribbons_Graphical More_Finite_Map begin text \We prove that the proof rules for graphical ribbon proofs are sound with respect to the rules of separation logic. We impose an additional assumption to achieve soundness: that the Frame rule has no side-condition. This assumption is reasonable because there are several separation logics that lack such a side-condition, such as ``variables-as-resource''. We first describe how to extract proofchains from a diagram. This process is similar to the process of extracting commands from a diagram, which was described in @{theory Ribbon_Proofs.Ribbons_Graphical}. When we extract a proofchain, we don't just include the commands, but the assertions in between them. Our main lemma for proving soundness says that each of these proofchains corresponds to a valid separation logic proof. \ subsection \Proofstate chains\ text \When extracting a proofchain from a diagram, we need to keep track of which nodes we have processed and which ones we haven't. A proofstate, defined below, maps a node to ``Top'' if it hasn't been processed and ``Bot'' if it has.\ datatype topbot = Top | Bot type_synonym proofstate = "node \\<^sub>f topbot" text \A proofstate chain contains all the nodes and edges of a graphical diagram, interspersed with proofstates that track which nodes have been processed at each point.\ type_synonym ps_chain = "(proofstate, node + edge) chain" text \The @{term "next_ps \"} function processes one node or one edge in a diagram, given the current proofstate @{term \}. It processes a node @{term v} by replacing the mapping from @{term v} to @{term Top} with a mapping from @{term v} to @{term Bot}. It processes an edge @{term e} (whose source and target nodes are @{term vs} and @{term ws} respectively) by removing all the mappings from @{term vs} to @{term Bot}, and adding mappings from @{term ws} to @{term Top}.\ fun next_ps :: "proofstate \ node + edge \ proofstate" where "next_ps \ (Inl v) = \ \ {|v|} ++\<^sub>f [{|v|} |=> Bot]" | "next_ps \ (Inr e) = \ \ fst3 e ++\<^sub>f [thd3 e |=> Top]" text \The function @{term "mk_ps_chain \ \"} generates from @{term \}, which is a list of nodes and edges, a proofstate chain, by interspersing the elements of @{term \} with the appropriate proofstates. The first argument @{term \} is the part of the chain that has already been converted.\ definition mk_ps_chain :: "[ps_chain, (node + edge) list] \ ps_chain" where "mk_ps_chain \ foldl (\\ x. cSnoc \ x (next_ps (post \) x))" lemma mk_ps_chain_preserves_length: fixes \ \ shows "chainlen (mk_ps_chain \ \) = chainlen \ + length \" proof (induct \ arbitrary: \) case Nil show ?case by (unfold mk_ps_chain_def, auto) next case (Cons x \) show ?case apply (unfold mk_ps_chain_def list.size foldl.simps) apply (fold mk_ps_chain_def) apply (auto simp add: Cons len_snoc) done qed text \Distributing @{term mk_ps_chain} over @{term Cons}.\ lemma mk_ps_chain_cons: "mk_ps_chain \ (x # \) = mk_ps_chain (cSnoc \ x (next_ps (post \) x)) \" by (auto simp add: mk_ps_chain_def) text \Distributing @{term mk_ps_chain} over @{term snoc}.\ lemma mk_ps_chain_snoc: "mk_ps_chain \ (\ @ [x]) = cSnoc (mk_ps_chain \ \) x (next_ps (post (mk_ps_chain \ \)) x)" by (unfold mk_ps_chain_def, auto) text \Distributing @{term mk_ps_chain} over @{term cCons}.\ lemma mk_ps_chain_ccons: fixes \ \ shows "mk_ps_chain (\ \ \ \ x \ \) \ = \ \ \ \ x \ mk_ps_chain \ \ " by (induct \ arbitrary: \, auto simp add: mk_ps_chain_cons mk_ps_chain_def) lemma pre_mk_ps_chain: fixes \ \ shows "pre (mk_ps_chain \ \) = pre \" apply (induct \ arbitrary: \) apply (auto simp add: mk_ps_chain_def mk_ps_chain_cons pre_snoc) done text \A chain which is obtained from the list @{term \}, has @{term \} as its list of commands. The following lemma states this in a slightly more general form, that allows for part of the chain to have already been processed.\ lemma comlist_mk_ps_chain: "comlist (mk_ps_chain \ \) = comlist \ @ \" proof (induct \ arbitrary: \) case Nil thus ?case by (auto simp add: mk_ps_chain_def) next case (Cons x \') show ?case apply (unfold mk_ps_chain_def foldl.simps, fold mk_ps_chain_def) apply (auto simp add: Cons comlist_snoc) done qed text \In order to perform induction over our diagrams, we shall wish to obtain ``smaller'' diagrams, by removing nodes or edges. However, the syntax and well-formedness constraints for diagrams are such that although we can always remove an edge from a diagram, we cannot (in general) remove a node -- the resultant diagram would not be a well-formed if an edge connected to that node. Hence, we consider ``partially-processed diagrams'' @{term "(G,S)"}, which comprise a diagram @{term G} and a set @{term S} of nodes. @{term S} denotes the subset of @{term G}'s initial nodes that have already been processed, and can be thought of as having been removed from @{term G}. We now give an updated version of the @{term "lins G"} function. This was originally defined in @{theory Ribbon_Proofs.Ribbons_Graphical}. We provide an extra parameter, @{term S}, which denotes the subset of @{term G}'s initial nodes that shouldn't be included in the linear extensions.\ definition lins2 :: "[node fset, diagram] \ lin set" where "lins2 S G \ {\ :: lin . (distinct \) \ (set \ = (fset G^V - fset S) <+> set G^E) \ (\i j v e. i < length \ \ j < length \ \ \!i = Inl v \ \!j = Inr e \ v |\| fst3 e \ i (\j k w e. j < length \ \ k < length \ \ \!j = Inr e \ \!k = Inl w \ w |\| thd3 e \ j \ lins2 S G" shows "distinct \" and "set \ = (fset G^V - fset S) <+> set G^E" and "\i j v e. \ i < length \ ; j < length \ ; \!i = Inl v ; \!j = Inr e ; v |\| fst3 e \ \ ii k w e. \ j < length \ ; k < length \ ; \!j = Inr e ; \!k = Inl w ; w |\| thd3 e \ \ j" and "set \ = (fset G^V - fset S) <+> set G^E" and "\i j v e. \ i < length \ ; j < length \ ; \!i = Inl v ; \!j = Inr e ; v |\| fst3 e \ \ ij k w e. \ j < length \ ; k < length \ ; \!j = Inr e ; \!k = Inl w ; w |\| thd3 e \ \ j \ lins2 S G" using assms apply (unfold lins2_def Collect_iff, intro conjI) apply assumption+ apply blast+ done text \When @{term S} is empty, the two definitions coincide.\ lemma lins_is_lins2_with_empty_S: "lins G = lins2 {||} G" by (unfold lins_def lins2_def, auto) text \The first proofstate for a diagram @{term G} is obtained by mapping each of its initial nodes to @{term Top}.\ definition initial_ps :: "diagram \ proofstate" where "initial_ps G \ [ initials G |=> Top ]" text \The first proofstate for the partially-processed diagram @{term G} is obtained by mapping each of its initial nodes to @{term Top}, except those in @{term S}, which are mapped to @{term Bot}.\ definition initial_ps2 :: "[node fset, diagram] \ proofstate" where "initial_ps2 S G \ [ initials G - S |=> Top ] ++\<^sub>f [ S |=> Bot ]" text \When @{term S} is empty, the above two definitions coincide.\ lemma initial_ps_is_initial_ps2_with_empty_S: "initial_ps = initial_ps2 {||}" apply (unfold fun_eq_iff, intro allI) apply (unfold initial_ps_def initial_ps2_def) apply simp done text \The following function extracts the set of proofstate chains from a diagram.\ definition ps_chains :: "diagram \ ps_chain set" where "ps_chains G \ mk_ps_chain (cNil (initial_ps G)) ` lins G" text \The following function extracts the set of proofstate chains from a partially-processed diagram. Nodes in @{term S} are excluded from the resulting chains.\ definition ps_chains2 :: "[node fset, diagram] \ ps_chain set" where "ps_chains2 S G \ mk_ps_chain (cNil (initial_ps2 S G)) ` lins2 S G" text \When @{term S} is empty, the above two definitions coincide.\ lemma ps_chains_is_ps_chains2_with_empty_S: "ps_chains = ps_chains2 {||}" apply (unfold fun_eq_iff, intro allI) apply (unfold ps_chains_def ps_chains2_def) apply (fold initial_ps_is_initial_ps2_with_empty_S) apply (fold lins_is_lins2_with_empty_S) apply auto done text \We now wish to describe proofstates chain that are well-formed. First, let us say that @{term "f ++\<^sub>fdisjoint g"} is defined, when @{term f} and @{term g} have disjoint domains, as @{term "f ++\<^sub>f g"}. Then, a well-formed proofstate chain consists of triples of the form @{term "(\ ++\<^sub>fdisjoint [{| v |} |=> Top], Inl v, \ ++\<^sub>fdisjoint [{| v |} |=> Bot])"}, where @{term v} is a node, or of the form @{term "(\ ++\<^sub>fdisjoint [{| vs |} |=> Bot], Inr e, \ ++\<^sub>fdisjoint [{| ws |} |=> Top])"}, where @{term e} is an edge with source and target nodes @{term vs} and @{term ws} respectively. The definition below describes a well-formed triple; we then lift this to complete chains shortly.\ definition wf_ps_triple :: "proofstate \ (node + edge) \ proofstate \ bool" where "wf_ps_triple T = (case snd3 T of Inl v \ (\\. v |\| fmdom \ \ fst3 T = [ {|v|} |=> Top ] ++\<^sub>f \ \ thd3 T = [ {|v|} |=> Bot ] ++\<^sub>f \) | Inr e \ (\\. (fst3 e |\| thd3 e) |\| fmdom \ = {||} \ fst3 T = [ fst3 e |=> Bot ] ++\<^sub>f \ \ thd3 T = [ thd3 e |=> Top ] ++\<^sub>f \))" lemma wf_ps_triple_nodeI: assumes "\\. v |\| fmdom \ \ \1 = [ {|v|} |=> Top ] ++\<^sub>f \ \ \2 = [ {|v|} |=> Bot ] ++\<^sub>f \" shows "wf_ps_triple (\1, Inl v, \2)" using assms unfolding wf_ps_triple_def by (auto simp add: fst3_simp snd3_simp thd3_simp) lemma wf_ps_triple_edgeI: assumes "\\. (fst3 e |\| thd3 e) |\| fmdom \ = {||} \ \1 = [ fst3 e |=> Bot ] ++\<^sub>f \ \ \2 = [ thd3 e |=> Top ] ++\<^sub>f \" shows "wf_ps_triple (\1, Inr e, \2)" using assms unfolding wf_ps_triple_def by (auto simp add: fst3_simp snd3_simp thd3_simp) definition wf_ps_chain :: "ps_chain \ bool" where "wf_ps_chain \ chain_all wf_ps_triple" lemma next_initial_ps2_vertex: "initial_ps2 ({|v|} |\| S) G = initial_ps2 S G \ {|v|} ++\<^sub>f [ {|v|} |=> Bot ]" apply (unfold initial_ps2_def) apply transfer apply (auto simp add: make_map_def map_diff_def map_add_def restrict_map_def) done lemma next_initial_ps2_edge: assumes "G = Graph V \ E" and "G' = Graph V' \ E'" and "V' = V - fst3 e" and "E' = removeAll e E" and "e \ set E" and "fst3 e |\| S" and "S |\| initials G" and "wf_dia G" shows "initial_ps2 (S - fst3 e) G' = initial_ps2 S G \ fst3 e ++\<^sub>f [ thd3 e |=> Top ]" proof (insert assms, unfold initial_ps2_def, transfer) fix G V \ E G' V' E' e S assume G_def: "G = Graph V \ E" and G'_def: "G' = Graph V' \ E'" and V'_def: "V' = V - fst3 e" and E'_def: "E' = removeAll e E" and e_in_E: "e \ set E" and fst_e_in_S: "fst3 e |\| S" and S_initials: "S |\| initials G" and wf_G: "wf_dia G" have "thd3 e |\| initials G = {||}" by (auto simp add: initials_def G_def e_in_E) show "make_map (initials G' - (S - fst3 e)) Top ++ make_map (S - fst3 e) Bot = map_diff (make_map (initials G - S) Top ++ make_map S Bot) (fst3 e) ++ make_map (thd3 e) Top" apply (unfold make_map_def map_diff_def) apply (unfold map_add_def restrict_map_def) apply (unfold minus_fset) apply (unfold fun_eq_iff initials_def) apply (unfold G_def G'_def V'_def E'_def) apply (unfold edges.simps vertices.simps) apply (simp add: less_eq_fset.rep_eq e_in_E) apply safe apply (insert \thd3 e |\| initials G = {||}\)[1] apply (insert S_initials, fold fset_cong)[2] apply (unfold less_eq_fset.rep_eq initials_def filter_fset) apply (auto simp add: G_def e_in_E)[1] apply (auto simp add: G_def e_in_E)[1] apply (auto simp add: G_def e_in_E)[1] apply (insert wf_G)[1] apply (unfold G_def vertices.simps edges.simps) apply (drule wf_dia_inv(3)) apply (unfold acyclicity_def) apply (metis fst_e_in_S inter_fset le_iff_inf subsetD) apply (insert wf_G)[1] apply (unfold G_def vertices.simps edges.simps) apply (drule wf_dia_inv(4)) apply (drule linearityD2) apply (fold fset_cong, unfold inter_fset fset_simps) apply (insert e_in_E, blast)[1] apply (insert wf_G)[1] apply (unfold G_def vertices.simps edges.simps) apply (drule wf_dia_inv(3)) apply (metis (lifting) e_in_E G_def empty_iff fset_simps(1) finter_iff linearityD(2) wf_G wf_dia_inv(4)) apply (insert wf_G)[1] apply (unfold G_def vertices.simps edges.simps) apply (drule wf_dia_inv(4)) apply (drule linearityD2) apply (fold fset_cong, unfold inter_fset fset_simps) apply (insert e_in_E, blast)[1] apply (insert wf_G)[1] apply (unfold G_def vertices.simps edges.simps) apply (drule wf_dia_inv(3)) apply (metis (lifting) e_in_E G_def empty_iff fset_simps(1) finter_iff linearityD(2) wf_G wf_dia_inv(4)) apply (insert wf_G) apply (unfold G_def vertices.simps edges.simps) apply (drule wf_dia_inv(5)) apply (unfold less_eq_fset.rep_eq union_fset) apply auto[1] apply (drule wf_dia_inv(5)) apply (unfold less_eq_fset.rep_eq union_fset) apply auto[1] apply (drule wf_dia_inv(5)) apply (unfold less_eq_fset.rep_eq union_fset) apply (auto simp add: e_in_E)[1] apply (drule wf_dia_inv(5)) apply (unfold less_eq_fset.rep_eq union_fset) apply (auto simp add: e_in_E)[1] done qed lemma next_lins2_vertex: assumes "Inl v # \ \ lins2 S G" assumes "v |\| S" shows "\ \ lins2 ({|v|} |\| S) G" proof - note lins2D = lins2D[OF assms(1)] show ?thesis proof (intro lins2I) show "distinct \" using lins2D(1) by auto next have "set \ = set (Inl v # \) - {Inl v}" using lins2D(1) by auto also have "... = (fset G^V - fset ({|v|} |\| S)) <+> set G^E" using lins2D(2) by auto finally show "set \ = (fset G^V - fset ({|v|} |\| S)) <+> set G^E" by auto next fix i j v e assume "i < length \" "j < length \" "\ ! i = Inl v" "\ ! j = Inr e" "v |\| fst3 e" thus "i < j" using lins2D(3)[of "i+1" "j+1"] by auto next fix j k w e assume "j < length \" "k < length \" "\ ! j = Inr e" "\ ! k = Inl w" "w |\| thd3 e" thus "j < k" using lins2D(4)[of "j+1" "k+1"] by auto qed qed lemma next_lins2_edge: assumes "Inr e # \ \ lins2 S (Graph V \ E)" and "vs |\| S" and "e = (vs,c,ws)" shows "\ \ lins2 (S - vs) (Graph (V - vs) \ (removeAll e E))" proof - note lins2D = lins2D[OF assms(1)] show ?thesis proof (intro lins2I, unfold vertices.simps edges.simps) show "distinct \" using lins2D(1) by auto next show "set \ = (fset (V - vs) - fset (S - vs)) <+> set (removeAll e E)" apply (insert lins2D(1) lins2D(2) assms(2)) apply (unfold assms(3) vertices.simps edges.simps less_eq_fset.rep_eq, simp) apply (unfold diff_diff_eq) proof - have "\a aa b. insert (Inr (vs, c, ws)) (set \) = (fset V - fset S) <+> set E \ fset vs \ fset S \ Inr (vs, c, ws) \ set \ \ distinct \ \ (a, aa, b) \ set E \ Inr (a, aa, b) \ set \ \ b = ws" by (metis (lifting) InrI List.set_simps(2) prod.inject set_ConsD sum.simps(2)) moreover have "\a aa b. insert (Inr (vs, c, ws)) (set \) = (fset V - fset S) <+> set E \ fset vs \ fset S \ Inr (vs, c, ws) \ set \ \ distinct \ \ (a, aa, b) \ set E \ Inr (a, aa, b) \ set \ \ aa = c" by (metis (lifting) InrI List.set_simps(2) prod.inject set_ConsD sum.simps(2)) moreover have "\x. insert (Inr (vs, c, ws)) (set \) = (fset V - fset S) <+> set E \ fset vs \ fset S \ Inr (vs, c, ws) \ set \ \ distinct \ \ x \ set \ \ x \ (fset V - fset S) <+> set E - {(vs, c, ws)}" apply (unfold insert_is_Un[of _ "set \"]) apply (fold assms(3)) apply clarify apply (subgoal_tac "set \ = ((fset V - fset S) <+> set E) - {Inr e}") by auto ultimately show "Inr (vs, c, ws) \ set \ \ distinct \ \ insert (Inr (vs, c, ws)) (set \) = (fset V - fset S) <+> set E \ fset vs \ fset S \ set \ = (fset V - fset S) <+> set E - {(vs, c, ws)}" by blast qed next fix i j v e assume "i < length \" "j < length \" "\ ! i = Inl v" "\ ! j = Inr e" "v |\| fst3 e" thus "i < j" using lins2D(3)[of "i+1" "j+1"] by auto next fix j k w e assume "j < length \" "k < length \" "\ ! j = Inr e" "\ ! k = Inl w" "w |\| thd3 e" thus "j < k" using lins2D(4)[of "j+1" "k+1"] by auto qed qed text \We wish to prove that every proofstate chain that can be obtained from a linear extension of @{term G} is well-formed and has as its final proofstate that state in which every terminal node in @{term G} is mapped to @{term Bot}. We first prove this for partially-processed diagrams, for then the result for ordinary diagrams follows as an easy corollary. We use induction on the size of the partially-processed diagram. The size of a partially-processed diagram @{term "(G,S)"} is defined as the number of nodes in @{term G}, plus the number of edges, minus the number of nodes in @{term S}.\ lemma wf_chains2: fixes k assumes "S |\| initials G" and "wf_dia G" and "\ \ ps_chains2 S G" and "fcard G^V + length G^E = k + fcard S" shows "wf_ps_chain \ \ (post \ = [ terminals G |=> Bot ])" using assms proof (induct k arbitrary: S G \) case 0 obtain V \ E where G_def: "G = Graph V \ E" by (metis diagram.exhaust) have "S |\| V" using "0.prems"(1) initials_in_vertices[of "G"] by (auto simp add: G_def) have "fcard V \ fcard S" using "0.prems"(4) by (unfold G_def, auto) from fcard_seteq[OF \S |\| V\ this] have "S = V" by auto hence "E = []" using "0.prems"(4) by (unfold G_def, auto) have "initials G = V" by (unfold G_def \E=[]\, rule no_edges_imp_all_nodes_initial) have "terminals G = V" by (unfold G_def \E=[]\, rule no_edges_imp_all_nodes_terminal) have "{} <+> {} = {}" by auto have "lins2 S G = { [] }" apply (unfold G_def \S=V\ \E=[]\) apply (unfold lins2_def, auto simp add: \{} <+> {} = {}\) done hence \_def: "\ = \ initial_ps2 S G \" using "0.prems"(3) by (auto simp add: ps_chains2_def mk_ps_chain_def) show ?case apply (intro conjI) apply (unfold \_def wf_ps_chain_def, auto) apply (unfold post.simps initial_ps2_def \initials G = V\ \terminals G = V\) apply (unfold \S=V\) apply (subgoal_tac "V - V = {||}", simp_all) done next case (Suc k) obtain V \ E where G_def: "G = Graph V \ E" by (metis diagram.exhaust) from Suc.prems(3) obtain \ where \_def: "\ = mk_ps_chain \ initial_ps2 S G \ \" and \_in: "\ \ lins2 S G" by (auto simp add: ps_chains2_def) note lins2 = lins2D[OF \_in] have "S |\| V" using Suc.prems(1) initials_in_vertices[of "G"] by (auto simp add: G_def) show ?case proof (cases \) case Nil from \_in have "V = S" "E = []" apply (-, unfold \\ = []\ lins2_def, simp_all) apply (unfold empty_eq_Plus_conv) apply (unfold G_def vertices.simps edges.simps, auto) by (metis \S |\| V\ less_eq_fset.rep_eq subset_antisym) with Suc.prems(4) have False by (simp add: G_def) thus ?thesis by auto next case (Cons x \') note \_def = this show ?thesis proof (cases x) case (Inl v) note x_def = this have "v |\| S \ v |\| V" apply (subgoal_tac "v \ fset V - fset S") apply (simp) apply (subgoal_tac "Inl v \ (fset V - fset S) <+> set E") apply (metis Inl_inject Inr_not_Inl PlusE) apply (metis lins2(1) lins2(2) Cons G_def Inl distinct.simps(2) distinct_length_2_or_more edges.simps vertices.simps) done hence v_notin_S: "v |\| S" and v_in_V: "v |\| V" by auto have v_initial_not_S: "v |\| initials G - S" apply (simp only: G_def initials_def vertices.simps edges.simps) apply (simp only: fminus_iff) apply (simp only: conj_commute, intro conjI, rule v_notin_S) apply (subgoal_tac "v \ fset (ffilter (\v. \e\set E. v |\| thd3 e) V)") apply simp apply (simp only: filter_fset, simp, simp only: conj_commute) apply (intro conjI ballI notI) apply (insert v_in_V, simp) proof - fix e :: edge assume "v \ fset (thd3 e)" then have "v |\| (thd3 e)" by auto assume "e \ set E" hence "Inr e \ set \" using lins2(2) by (auto simp add: G_def) then obtain j where "j < length \" "0 < length \" "\!j = Inr e" "\!0 = Inl v" by (metis \_def x_def in_set_conv_nth length_pos_if_in_set nth_Cons_0) with lins2(4)[OF this \v |\| (thd3 e)\] show False by auto qed define S' where "S' = {|v|} |\| S" define \' where "\' = mk_ps_chain \ initial_ps2 S' G \ \'" hence pre_\': "pre \' = initial_ps2 S' G" by (metis pre.simps(1) pre_mk_ps_chain) define \ where "\ = [ initials G - ({|v|} |\| S) |=> Top ] ++\<^sub>f [ S |=> Bot ]" have "wf_ps_chain \' \ (post \' = [terminals G |=> Bot])" proof (intro Suc.hyps[of "S'"]) show "S' |\| initials G" apply (unfold S'_def, auto) apply (metis fminus_iff v_initial_not_S) by (metis Suc.prems(1) fset_rev_mp) next show "wf_dia G" by (rule Suc.prems(2)) next show "\' \ ps_chains2 S' G" apply (unfold ps_chains2_def \'_def) apply (intro imageI) apply (unfold S'_def) apply (intro next_lins2_vertex) apply (fold x_def, fold \_def) apply (rule \_in) by (metis v_notin_S) next show "fcard G^V + length G^E = k + fcard S'" apply (unfold S'_def) by (auto simp add: Suc.prems(4) fcard_finsert_disjoint[OF v_notin_S]) qed hence wf_\': "wf_ps_chain \'" and post_\': "post \' = [terminals G |=> Bot]" by auto show ?thesis proof (intro conjI) have 1: "fmdom [ {|v|} |=> Bot ] |\| fmdom ([ initials G - ({|v|} |\| S) |=> Top ] ++\<^sub>f [ S |=> Bot ]) = {||}" by (metis (no_types) fdom_make_fmap fmdom_add bot_least funion_iff finter_finsert_left le_iff_inf fminus_iff finsert_fsubset sup_ge1 v_initial_not_S) show "wf_ps_chain \" using [[unfold_abs_def = false]] apply (simp only: \_def \_def x_def mk_ps_chain_cons) apply simp apply (unfold mk_ps_chain_ccons) apply (fold next_initial_ps2_vertex S'_def) apply (fold \'_def) apply (unfold wf_ps_chain_def chain_all.simps conj_commute) apply (intro conjI) apply (fold wf_ps_chain_def, rule wf_\') apply (intro wf_ps_triple_nodeI exI[of _ "\"] conjI) apply (unfold \_def fmdom_add fdom_make_fmap) apply (metis finsertI1 fminus_iff funion_iff v_notin_S) apply (unfold pre_\' initial_ps2_def S'_def) apply (unfold fmap_add_commute[OF 1]) apply (unfold fmadd_assoc) apply (fold fmadd_assoc[of _ "[ S |=> Bot ]"]) apply (unfold make_fmap_union sup.commute[of "{|v|}"]) apply (unfold fminus_funion) using v_initial_not_S apply auto by (metis (opaque_lifting, no_types) finsert_absorb finsert_fminus_single finter_fminus inf_commute inf_idem v_initial_not_S) next show "post \ = [ terminals G |=> Bot ]" apply (unfold \_def \_def x_def mk_ps_chain_cons, simp) apply (unfold mk_ps_chain_ccons post.simps) apply (fold next_initial_ps2_vertex S'_def) apply (fold \'_def, rule post_\') done qed next case (Inr e) note x_def = this define vs where "vs = fst3 e" define ws where "ws = thd3 e" obtain c where e_def: "e = (vs, c, ws)" by (metis vs_def ws_def fst3_simp thd3_simp prod_cases3) have "linearity E" and "acyclicity E" and e_in_V: "\e. e \ set E \ fst3 e |\| thd3 e |\| V" by (insert Suc.prems(2) wf_dia_inv, unfold G_def, blast)+ note lin = linearityD[OF this(1)] have acy: "\e. e \ set E \ fst3 e |\| thd3 e = {||}" apply (fold fset_cong, insert \acyclicity E\) apply (unfold acyclicity_def acyclic_def, auto) done note lins = lins2D[OF \_in] have e_in_E: "e \ set E" apply (subgoal_tac "set \ = (fset G^V - fset S) <+> set G^E") apply (unfold \_def x_def G_def edges.simps, auto)[1] apply (simp add: lins(2)) done have vs_in_S: "vs |\| S" apply (insert e_in_V[OF e_in_E]) apply (unfold less_eq_fset.rep_eq) apply (intro subsetI) apply (unfold vs_def) apply (rule ccontr) apply (subgoal_tac "x \ fset V") prefer 2 apply (auto) proof - fix v assume a: "v \ fset (fst3 e)" assume "v \ fset S" and "v \ fset V" hence "Inl v \ set \" by (metis (lifting) DiffI G_def InlI lins(2) vertices.simps) then obtain i where "i < length \" "0 < length \" "\!i = Inl v" "\!0 = Inr e" by (metis Cons Inr in_set_conv_nth length_pos_if_in_set nth_Cons_0) from lins(3)[OF this] show "False" by (auto simp add: a) qed have "ws |\| (initials G) = {||}" apply (insert e_in_V[OF e_in_E]) apply (unfold initials_def less_eq_fset.rep_eq, fold fset_cong) apply (unfold ws_def G_def, auto simp add: e_in_E) done define S' where "S' = S - vs" define V' where "V' = V - vs" define E' where "E' = removeAll e E" define G' where "G' = Graph V' \ E'" define \' where "\' = mk_ps_chain \ initial_ps2 S' G' \ \'" hence pre_\': "pre \' = initial_ps2 S' G'" by (metis pre.simps(1) pre_mk_ps_chain) define \ where "\ = [ initials G - S |=> Top ] ++\<^sub>f [ S - vs |=> Bot ]" have next_initial_ps2: "initial_ps2 S' G' = initial_ps2 S G \ vs ++\<^sub>f [ws |=> Top]" using next_initial_ps2_edge[OF G_def _ _ _ e_in_E _ Suc.prems(1) Suc.prems(2)] G'_def E'_def vs_def ws_def V'_def vs_in_S S'_def by auto have "wf_ps_chain \' \ post \' = [ terminals G' |=> Bot ]" proof (intro Suc.hyps[of "S'"]) show "S' |\| initials G'" apply (insert Suc.prems(1)) apply (unfold G'_def G_def initials_def) apply (unfold less_eq_fset.rep_eq S'_def E'_def V'_def) apply auto done next from Suc.prems(2) have "wf_dia (Graph V \ E)" by (unfold G_def) note wf_G = wf_dia_inv[OF this] show "wf_dia G'" apply (unfold G'_def V'_def E'_def) apply (insert wf_G e_in_E vs_in_S Suc.prems(1)) apply (unfold vs_def) apply (intro wf_dia) apply (unfold linearity_def initials_def G_def) apply (fold fset_cong, unfold less_eq_fset.rep_eq) apply (simp, simp) apply (unfold acyclicity_def, rule acyclic_subset) apply (auto simp add: distinct_removeAll) apply (metis (lifting) IntI empty_iff) done next show "\' \ ps_chains2 S' G'" apply (unfold \_def \'_def ps_chains2_def) apply (intro imageI) apply (unfold S'_def G'_def V'_def E'_def) apply (intro next_lins2_edge) apply (metis \_def G_def x_def \_in) by (simp only: vs_in_S e_def)+ next have "vs |\| V" by (metis (lifting) \S |\| V\ order_trans vs_in_S) have "distinct E" using \linearity E\ linearity_def by auto show "fcard G'^V + length G'^E = k + fcard S'" apply (insert Suc.prems(4)) apply (unfold G_def G'_def vertices.simps edges.simps) apply (unfold V'_def E'_def S'_def) apply (unfold fcard_funion_fsubset[OF \vs |\| V\]) apply (unfold fcard_funion_fsubset[OF \vs |\| S\]) apply (fold distinct_remove1_removeAll[OF \distinct E\]) apply (unfold length_remove1) apply (simp add: e_in_E) apply (drule arg_cong[of _ _ "\x. x - fcard vs - 1"]) apply (subst (asm) add_diff_assoc2[symmetric]) apply (simp add: fcard_mono[OF \vs |\| V\]) apply (subst add_diff_assoc, insert length_pos_if_in_set[OF e_in_E], arith, auto) apply (subst add_diff_assoc, auto simp add: fcard_mono[OF \vs |\| S\]) done qed hence wf_\': "wf_ps_chain \'" and post_\': "post \' = [ terminals G' |=> Bot ]" by auto have terms_same: "terminals G = terminals G'" apply (unfold G'_def G_def terminals_def edges.simps vertices.simps) apply (unfold E'_def V'_def) apply (fold fset_cong, auto simp add: e_in_E vs_def) done have 1: "fmdom [ fst3 e |=> Bot ] |\| fmdom([ ffilter (\v. \e\set E. v |\| thd3 e) V - S |=> Top ] ++\<^sub>f [ S - fst3 e |=> Bot ]) = {||}" apply (unfold fmdom_add fdom_make_fmap) apply (fold fset_cong) apply auto apply (metis in_mono less_eq_fset.rep_eq vs_def vs_in_S) done show ?thesis proof (intro conjI) show "wf_ps_chain \" using [[unfold_abs_def = false]] apply (unfold \_def \_def x_def mk_ps_chain_cons) apply simp apply (unfold mk_ps_chain_ccons) apply (fold vs_def ws_def) apply (fold next_initial_ps2) apply (fold \'_def) apply (unfold wf_ps_chain_def chain_all.simps conj_commute) apply (intro conjI) apply (fold wf_ps_chain_def) apply (rule wf_\') apply (intro wf_ps_triple_edgeI exI[of _ "\"]) apply (unfold e_def fst3_simp thd3_simp \_def, intro conjI) apply (insert Suc.prems(1)) apply (unfold pre_\' initial_ps2_def initials_def) apply (insert vs_in_S acy[OF e_in_E]) apply (fold fset_cong) apply (unfold less_eq_fset.rep_eq)[1] apply (unfold G_def G'_def vs_def ws_def V'_def E'_def S'_def) apply (unfold vertices.simps edges.simps) apply (unfold fmap_add_commute[OF 1]) apply (fold fmadd_assoc) apply (unfold make_fmap_union) apply (auto simp add: fdom_make_fmap e_in_E)[1] apply simp apply (unfold fmadd_assoc) apply (unfold make_fmap_union) apply (metis (lifting) funion_absorb2 vs_def vs_in_S) apply (intro arg_cong2[of _ _ "[ S - fst3 e |=> Bot ]" "[ S - fst3 e |=> Bot ]" "(++\<^sub>f)"]) apply (intro arg_cong2[of _ _ "Top" "Top" "make_fmap"]) defer 1 apply (simp, simp) apply (fold fset_cong) apply (unfold less_eq_fset.rep_eq, simp) apply (elim conjE) apply (intro set_eqI iffI, simp_all) apply (elim conjE, intro disjI conjI ballI, simp) apply (case_tac "ea=e", simp_all) apply (elim disjE conjE, intro conjI ballI impI, simp_all) apply (insert e_in_E lin(2))[1] apply (subst (asm) (2) fset_cong[symmetric]) apply (elim conjE) apply (subst (asm) inter_fset) apply (subst (asm) fset_simps) apply (insert disjoint_iff_not_equal)[1] apply blast apply (metis G_def Suc(3) e_in_E subsetD less_eq_fset.rep_eq wf_dia_inv') prefer 2 apply (metis (lifting) IntI Suc(2) \ws |\| initials G = {||}\ empty_iff fset_simps(1) in_mono inter_fset less_eq_fset.rep_eq ws_def) apply auto done next show "post \ = [terminals G |=> Bot]" apply (unfold \_def \_def x_def mk_ps_chain_cons) apply simp apply (unfold mk_ps_chain_ccons post.simps) apply (fold vs_def ws_def) apply (fold next_initial_ps2) apply (fold \'_def) apply (unfold terms_same) apply (rule post_\') done qed qed qed qed corollary wf_chains: assumes "wf_dia G" assumes "\ \ ps_chains G" shows "wf_ps_chain \ \ post \ = [ terminals G |=> Bot ]" apply (intro wf_chains2[of "{||}"], insert assms(2)) by (auto simp add: assms(1) ps_chains_is_ps_chains2_with_empty_S fcard_fempty) subsection \Interface chains\ type_synonym int_chain = "(interface, assertion_gadget + command_gadget) chain" text \An interface chain is similar to a proofstate chain. However, where a proofstate chain talks about nodes and edges, an interface chain talks about the assertion-gadgets and command-gadgets that label those nodes and edges in a diagram. And where a proofstate chain talks about proofstates, an interface chain talks about the interfaces obtained from those proofstates. The following functions convert a proofstate chain into an interface chain.\ definition ps_to_int :: "[diagram, proofstate] \ interface" where "ps_to_int G \ \ \v |\| fmdom \. case_topbot top_ass bot_ass (lookup \ v) (G^\ v)" definition ps_chain_to_int_chain :: "[diagram, ps_chain] \ int_chain" where "ps_chain_to_int_chain G \ \ chainmap (ps_to_int G) ((case_sum (Inl \ G^\) (Inr \ snd3))) \" lemma ps_chain_to_int_chain_simp: "ps_chain_to_int_chain (Graph V \ E) \ = chainmap (ps_to_int (Graph V \ E)) ((case_sum (Inl \ \) (Inr \ snd3))) \" by (simp add: ps_chain_to_int_chain_def) subsection \Soundness proof\ text \We assume that @{term wr_com} always returns @{term "{}"}. This is equivalent to changing our axiomatization of separation logic such that the frame rule has no side-condition. One way to obtain a separation logic lacking a side-condition on its frame rule is to use variables-as- resource. We proceed by induction on the proof rules for graphical diagrams. We show that: (1) if a diagram @{term G} is provable w.r.t. interfaces @{term P} and @{term Q}, then @{term P} and @{term Q} are the top and bottom interfaces of @{term G}, and that the Hoare triple @{term "(asn P, c, asn Q)"} is provable for each command @{term c} that can be extracted from @{term G}; (2) if a command-gadget @{term C} is provable w.r.t. interfaces @{term P} and @{term Q}, then the Hoare triple @{term "(asn P, c, asn Q)"} is provable for each command @{term c} that can be extracted from @{term C}; and (3) if an assertion-gadget @{term A} is provable, and if the top and bottom interfaces of @{term A} are @{term P} and @{term Q} respectively, then the Hoare triple @{term "(asn P, c, asn Q)"} is provable for each command @{term c} that can be extracted from @{term A}.\ lemma soundness_graphical_helper: assumes no_var_interference: "\c. wr_com c = {}" shows "(prov_dia G P Q \ (P = top_dia G \ Q = bot_dia G \ (\c. coms_dia G c \ prov_triple (asn P, c, asn Q)))) \ (prov_com C P Q \ (\c. coms_com C c \ prov_triple (asn P, c, asn Q))) \ (prov_ass A \ (\c. coms_ass A c \ prov_triple (asn (top_ass A), c, asn (bot_ass A))))" proof (induct rule: prov_dia_prov_com_prov_ass.induct) case (Skip p) thus ?case apply (intro allI impI, elim conjE coms_skip_inv) apply (auto simp add: prov_triple.skip) done next case (Exists G P Q x) thus ?case apply (intro allI impI, elim conjE coms_exists_inv) apply (auto simp add: prov_triple.exists) done next case (Basic P c Q) thus ?case by (intro allI impI, elim conjE coms_basic_inv, auto) next case (Choice G P Q H) thus ?case apply (intro allI impI, elim conjE coms_choice_inv) apply (auto simp add: prov_triple.choose) done next case (Loop G P) thus ?case apply (intro allI impI, elim conjE coms_loop_inv) apply (auto simp add: prov_triple.loop) done next case (Main G) thus ?case apply (intro conjI) apply (simp, simp) apply (intro allI impI) apply (elim coms_main_inv, simp) proof - fix c V \ E fix \::"lin" fix cs::"command list" assume wf_G: "wf_dia (Graph V \ E)" assume "\v. v \ fset V \ \c. coms_ass (\ v) c \ prov_triple (asn (top_ass (\ v)), c, asn (bot_ass (\ v)))" hence prov_vertex: "\v c P Q F. \ coms_ass (\ v) c; v \ fset V; P = (top_ass (\ v) \ F) ; Q = (bot_ass (\ v) \ F) \ \ prov_triple (asn P, c, asn Q)" by (auto simp add: prov_triple.frame no_var_interference) assume "\e. e \ set E \ \c. coms_com (snd3 e) c \ prov_triple (asn (\v|\|fst3 e. bot_ass (\ v)),c,asn (\v|\|thd3 e. top_ass (\ v)))" hence prov_edge: "\e c P Q F. \ e \ set E ; coms_com (snd3 e) c ; P = ((\v|\|fst3 e. bot_ass (\ v)) \ F) ; Q = ((\v|\|thd3 e. top_ass (\ v)) \ F) \ \ prov_triple (asn P, c, asn Q)" by (auto simp add: prov_triple.frame no_var_interference) assume len_cs: "length cs = length \" assume "\i. case_sum (coms_ass \ \) (coms_com \ snd3) (\ ! i) (cs ! i)" hence \_cs: "\i. i < length \ \ case_sum (coms_ass \ \) (coms_com \ snd3) (\ ! i) (cs ! i)" by auto assume G_def: "G = Graph V \ E" assume c_def: "c = foldr (;;) cs Skip" assume \_lin: "\ \ lins (Graph V \ E)" note lins = linsD[OF \_lin] define \ where "\ = mk_ps_chain \ initial_ps G \ \" have "\ \ ps_chains G" by (simp add: \_lin \_def ps_chains_def G_def) hence 1: "post \ = [ terminals G |=> Bot ]" and 2: "chain_all wf_ps_triple \" by (insert wf_chains G_def wf_G, auto simp add: wf_ps_chain_def) show "prov_triple (asn (\v|\|initials (Graph V \ E). top_ass (\ v)), foldr (;;) cs Skip, asn (\v|\|terminals (Graph V \ E). bot_ass (\ v)))" using [[unfold_abs_def = false]] apply (intro seq_fold[of _ "ps_chain_to_int_chain G \"]) apply (unfold len_cs) apply (unfold ps_chain_to_int_chain_def chainmap_preserves_length \_def) apply (unfold mk_ps_chain_preserves_length, simp) apply (unfold pre_chainmap post_chainmap) apply (unfold pre_mk_ps_chain pre.simps) apply (fold \_def, unfold 1) apply (unfold initial_ps_def) apply (unfold ps_to_int_def) apply (unfold fdom_make_fmap) apply (unfold G_def labelling.simps, fold G_def) apply (subgoal_tac "\v \ fset (initials G). top_ass (\ v) = case_topbot top_ass bot_ass (lookup [ initials G |=> Top ] v) (\ v)") apply (unfold iter_hcomp_cong, simp) apply (metis lookup_make_fmap topbot.simps(3)) apply (subgoal_tac "\v \ fset (terminals G). bot_ass (\ v) = case_topbot top_ass bot_ass (lookup [ terminals G |=> Bot ] v) (\ v)") apply (unfold iter_hcomp_cong, simp) apply (metis lookup_make_fmap topbot.simps(4), simp) apply (unfold G_def, fold ps_chain_to_int_chain_simp G_def) proof - fix i assume "i < length \" hence "i < chainlen \" by (metis \_def add_0_left chainlen.simps(1) mk_ps_chain_preserves_length) hence wf_\i: "wf_ps_triple (nthtriple \ i)" by (insert 2, simp add: chain_all_nthtriple) show "prov_triple (asn (fst3 (nthtriple (ps_chain_to_int_chain G \) i)), cs ! i, asn (thd3 (nthtriple (ps_chain_to_int_chain G \) i)))" apply (unfold ps_chain_to_int_chain_def) apply (unfold nthtriple_chainmap[OF \i < chainlen \\]) apply (unfold fst3_simp thd3_simp) proof (cases "\!i") case (Inl v) have "snd3 (nthtriple \ i) = Inl v" apply (unfold snds_of_triples_form_comlist[OF \i < chainlen \\]) apply (auto simp add: \_def comlist_mk_ps_chain Inl) done with wf_\i wf_ps_triple_def obtain \ where v_notin_\: "v |\| fmdom \" and fst_\i: "fst3 (nthtriple \ i) = [ {|v|} |=> Top ] ++\<^sub>f \" and thd_\i: "thd3 (nthtriple \ i) = [ {|v|} |=> Bot ] ++\<^sub>f \" by auto show "prov_triple (asn (ps_to_int G (fst3 (nthtriple \ i))), cs ! i, asn (ps_to_int G (thd3 (nthtriple \ i))))" apply (intro prov_vertex[where v=v]) apply (metis (no_types) Inl \i < length \\ \_cs o_def sum.simps(5)) apply (metis (lifting) Inl lins(2) Inl_not_Inr PlusE \i < length \\ nth_mem sum.simps(1) vertices.simps) apply (unfold fst_\i thd_\i) apply (unfold ps_to_int_def) apply (unfold fmdom_add fdom_make_fmap) apply (unfold finsert_is_funion[symmetric]) apply (insert v_notin_\) apply (unfold iter_hcomp_insert) apply (unfold lookup_union2 lookup_make_fmap1) apply (unfold G_def labelling.simps) apply (subgoal_tac "\va \ fset (fmdom \). case_topbot top_ass bot_ass (lookup ([ {|v|} |=> Top ] ++\<^sub>f \) va) (\ va) = case_topbot top_ass bot_ass (lookup ([{|v|} |=> Bot] ++\<^sub>f \) va)(\ va)") apply (unfold iter_hcomp_cong, simp) apply (metis lookup_union1, simp) done next case (Inr e) have "snd3 (nthtriple \ i) = Inr e" apply (unfold snds_of_triples_form_comlist[OF \i < chainlen \\]) apply (auto simp add: \_def comlist_mk_ps_chain Inr) done with wf_\i wf_ps_triple_def obtain \ where fst_e_disjoint_\: "fst3 e |\| fmdom \ = {||}" and thd_e_disjoint_\: "thd3 e |\| fmdom \ = {||}" and fst_\i: "fst3 (nthtriple \ i) = [ fst3 e |=> Bot ] ++\<^sub>f \" and thd_\i: "thd3 (nthtriple \ i) = [ thd3 e |=> Top ] ++\<^sub>f \" by (auto simp add: inf_sup_distrib2) + + let ?F = "\v|\|fmdom \. + case_topbot top_ass bot_ass (lookup ([ fst3 e |=> Bot ] ++\<^sub>f \) v) (\ v)" + show "prov_triple (asn (ps_to_int G (fst3 (nthtriple \ i))), cs ! i, asn (ps_to_int G (thd3 (nthtriple \ i))))" - apply (intro prov_edge[where e=e]) - apply (subgoal_tac "Inr e \ set \") - apply (metis Inr_not_Inl PlusE edges.simps lins(2) sum.simps(2)) - apply (metis Inr \i < length \\ nth_mem) - apply (metis (no_types) Inr \i < length \\ \_cs o_def sum.simps(6)) - apply (unfold fst_\i thd_\i) - apply (unfold ps_to_int_def) - apply (unfold G_def labelling.simps) - apply (unfold fmdom_add fdom_make_fmap) - apply (insert fst_e_disjoint_\) - apply (unfold iter_hcomp_union) - apply (subgoal_tac "\v \ fset (fst3 e). case_topbot top_ass bot_ass + proof (intro prov_edge[where e=e]) + show "e \ set E" + apply (subgoal_tac "Inr e \ set \") + apply (metis Inr_not_Inl PlusE edges.simps lins(2) sum.simps(2)) + by (metis Inr \i < length \\ nth_mem) + next + show "coms_com (snd3 e) (cs ! i)" + by (metis (no_types) Inr \i < length \\ \_cs o_def sum.simps(6)) + next + show "ps_to_int G (fst3 (nthtriple \ i)) = ((\v|\|fst3 e. bot_ass (\ v)) \ ?F)" + unfolding fst_\i ps_to_int_def G_def labelling.simps fmdom_add fdom_make_fmap + apply (insert fst_e_disjoint_\) + apply (unfold iter_hcomp_union) + apply (subgoal_tac "\v \ fset (fst3 e). case_topbot top_ass bot_ass (lookup ([ fst3 e |=> Bot ] ++\<^sub>f \) v) (\ v) = bot_ass (\ v)") - apply (unfold iter_hcomp_cong) - apply (simp) - apply (intro ballI) - apply (subgoal_tac "v |\| fmdom \") - apply (unfold lookup_union2) - apply (metis lookup_make_fmap topbot.simps(4)) - apply (metis fempty_iff finterI) - apply (insert thd_e_disjoint_\) - apply (unfold iter_hcomp_union) - apply (subgoal_tac "\v \ fset (thd3 e). case_topbot top_ass bot_ass + apply (unfold iter_hcomp_cong) + apply (simp) + apply (intro ballI) + apply (subgoal_tac "v |\| fmdom \") + apply (unfold lookup_union2) + apply (metis lookup_make_fmap topbot.simps(4)) + by (metis fempty_iff finterI) + next + show "ps_to_int G (thd3 (nthtriple \ i)) = ((\v|\|thd3 e. top_ass (\ v)) \ ?F)" + unfolding thd_\i ps_to_int_def G_def labelling.simps fmdom_add fdom_make_fmap + apply (insert thd_e_disjoint_\) + apply (unfold iter_hcomp_union) + apply (subgoal_tac "\v \ fset (thd3 e). case_topbot top_ass bot_ass (lookup ([ thd3 e |=> Top ] ++\<^sub>f \) v) (\ v) = top_ass (\ v)") - apply (unfold iter_hcomp_cong) - apply (subgoal_tac "\v \ fset (fmdom \). case_topbot top_ass bot_ass + apply (unfold iter_hcomp_cong) + apply (subgoal_tac "\v \ fset (fmdom \). case_topbot top_ass bot_ass (lookup ([ thd3 e |=> Top ] ++\<^sub>f \) v) (\ v) = case_topbot top_ass bot_ass (lookup ([fst3 e |=> Bot] ++\<^sub>f \) v) (\ v)") - apply (unfold iter_hcomp_cong) - apply simp - apply (intro ballI) - apply (subgoal_tac "v |\| fmdom \") - apply (unfold lookup_union1, auto) - apply (subgoal_tac "v |\| fmdom \") - apply (unfold lookup_union2) - apply (metis lookup_make_fmap topbot.simps(3)) - by (metis fempty_iff finterI) + apply (unfold iter_hcomp_cong) + apply simp + subgoal + by (simp add: lookup_union1) + subgoal + by (simp add: fdisjoint_iff lookup_union2 lookup_make_fmap) + done + qed qed qed qed qed text \The soundness theorem states that any diagram provable using the proof rules for ribbons can be recreated as a valid proof in separation logic.\ corollary soundness_graphical: assumes "\c. wr_com c = {}" assumes "prov_dia G P Q" shows "\c. coms_dia G c \ prov_triple (asn P, c, asn Q)" using soundness_graphical_helper[OF assms(1)] and assms(2) by auto 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,10454 +1,10447 @@ (*********************************************************************************** * 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 (metis empty_iff insertE new_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t) apply (metis 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" by (auto simp add: get_shadow_root_safe_def bind_pure_I option.case_eq_if) 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 (verit, best) 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 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 by (metis get_host_ptr_in_heap local.a_shadow_root_valid_def local.get_shadow_root_ok local.heap_is_wellformed_def 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 using \|h \ get_tag_name host'|\<^sub>r \ safe_shadow_root_element_types\ assms(4) by auto 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 by (auto simp add: 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'" + using disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 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, opaque_lifting) children_eq2_h2 fin_mono 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) + object_ptr_kinds_eq_h2)[1] + by (metis (no_types, opaque_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'\ + children_eq2_h2 fsubsetD in_mono object_ptr_kinds_eq_h2) 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 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 (metis (no_types, lifting) IntI \known_ptrs h'\ \type_wf h'\ children_eq2_h2 empty_iff finite_fset is_OK_returns_result_I l_get_child_nodes.get_child_nodes_ptr_in_heap local.get_child_nodes_ok local.known_ptrs_known_ptr local.l_get_child_nodes_axioms returns_result_select_result sorted_list_of_set.set_sorted_key_list_of_set) apply(case_tac "xa = cast shadow_root_ptr") using \cast shadow_root_ptr |\| object_ptr_kinds h'\ apply simp by (metis (mono_tags, lifting) CD.distinct_lists_no_parent \known_ptrs h'\ \local.CD.a_distinct_lists h2\ \type_wf h'\ children_eq2_h2 children_eq_h2 disconnected_nodes_eq_h2 is_OK_returns_result_I l_get_child_nodes.get_child_nodes_ptr_in_heap local.get_child_nodes_ok local.get_disconnected_nodes_ok local.known_ptrs_known_ptr local.l_get_child_nodes_axioms returns_result_select_result) 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, opaque_lifting) \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, opaque_lifting)) 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, opaque_lifting) 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, opaque_lifting) 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'" by (smt (verit) \h \ get_shadow_root ptr \\<^sub>r Some shadow_root_ptr\ assms(2) 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 element_ptr_kinds_eq_h element_ptr_kinds_eq_h2 finsert_iff funion_finsert_right get_shadow_root_eq2_h2 get_shadow_root_eq_h h' local.a_shadow_root_valid_def local.get_shadow_root_ok object_ptr_kinds_eq2_h2 object_ptr_kinds_eq_h option.sel returns_result_select_result select_result_I2 shadow_root_ptr_kinds_commutes sup_bot.right_neutral tag_name_eq2_h tag_name_eq2_h2) 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 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 (verit) assms(1) assms(2) bind_pure_returns_result_I2 finite_fset 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_set(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 (verit, best) 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_I return_returns_result returns_result_select_result) 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 (verit) 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" by (meson DocumentMonad.ptr_kinds_ptr_kinds_M assms(1) disjoint_iff local.get_disconnected_nodes_ok local.heap_is_wellformed_one_disc_parent returns_result_select_result type_wf) have "h \ get_parent node_ptr \\<^sub>r None" by (metis (mono_tags, lifting) "2" "4" known_ptrs local.get_parent_child_dual local.get_parent_ok local.get_parent_parent_in_heap returns_result_select_result select_result_I2 split_option_ex type_wf) 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 (smt (verit) DocumentMonad.ptr_kinds_ptr_kinds_M bind_pure_returns_result_I2 local.get_disconnected_nodes_ok local.get_disconnected_nodes_pure node_ptr_inclusion return_returns_result select_result_I2 type_wf) 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 (metis (no_types, lifting) assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) 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_eq returns_result_select_result) 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 option.distinct(1) returns_result_eq returns_result_select_result root by (metis (no_types, opaque_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 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, opaque_lifting) disc_nodes_old_document_h2 disc_nodes_old_document_h3 disconnected_nodes_eq2_h2 document_ptr_kinds_eq3_h2 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, opaque_lifting) children_eq2_h3 object_ptr_kinds_h3_eq3 subsetD) by (metis (no_types, opaque_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 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, opaque_lifting) \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 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) 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, opaque_lifting) children_eq2_h2 in_mono object_ptr_kinds_M_eq3_h2) by (metis (no_types, opaque_lifting) NodeMonad.ptr_kinds_ptr_kinds_M disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3 document_ptr_kinds_commutes 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, opaque_lifting) contra_subsetD 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 select_result_I2 subsetD) by (metis (no_types, lifting) DocumentMonad.ptr_kinds_ptr_kinds_M disconnected_nodes_eq2_h3 document_ptr_kinds_eq2_h3 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 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, opaque_lifting) 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, opaque_lifting) 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 notin_set_remove1 select_result_I2 by (metis (no_types, opaque_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, opaque_lifting) "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, opaque_lifting) "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, opaque_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 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, opaque_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 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 (verit) children_eq2_h3 disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 finsertCI funion_finsert_right h' insert_iff list.simps(15) local.CD.a_all_ptrs_in_heap_def local.set_disconnected_nodes_get_disconnected_nodes node_ptr_kinds_eq_h node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h3 object_ptr_kinds_eq_h3 select_result_I2 subsetD subsetI) 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_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" by (smt (verit) 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.CD.a_all_ptrs_in_heap h\ disc_nodes_document_ptr_h disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disjoint_iff document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 h' l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes local.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, opaque_lifting) "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, opaque_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 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 (verit) Diff_empty Diff_insert0 ElementMonad.ptr_kinds_M_ptr_kinds ElementMonad.ptr_kinds_ptr_kinds_M assms(1) assms(3) 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 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 (metis \previous_host \ fset (element_ptr_kinds h2)\ \previous_host \ new_element_ptr\ element_ptr_kinds_eq_h2 select_result_eq shadow_root_eq_h2 tag_name_eq2_h2) qed then have "a_shadow_root_valid h'" by (smt (verit) \type_wf h3\ element_ptr_kinds_eq_h3 local.a_shadow_root_valid_def local.get_shadow_root_ok returns_result_select_result select_result_I2 shadow_root_eq_h3 shadow_root_ptr_kinds_eq_h3 tag_name_eq2_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 (verit, del_insts) assms(3) case_prodI h2 local.get_shadow_root_ok mem_Collect_eq new_element_ptr new_element_ptr_not_in_heap pair_imageI returns_result_select_result select_result_I2 shadow_root_eq_h) using shadow_root_none apply auto[1] by (metis (no_types, lifting) assms(3) 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 (metis (mono_tags, lifting) \type_wf h2\ case_prodI local.get_shadow_root_ok mem_Collect_eq pair_imageI returns_result_select_result select_result_I2 shadow_root_eq_h2) apply (metis (mono_tags, lifting) case_prodI mem_Collect_eq pair_imageI select_result_eq shadow_root_eq_h2) 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 (mono_tags, lifting) \type_wf h'\ case_prodI element_ptr_kinds_eq_h2 element_ptr_kinds_eq_h3 local.get_shadow_root_ok mem_Collect_eq pair_imageI returns_result_select_result select_result_I2 shadow_root_eq_h3) apply (metis (mono_tags, lifting) \type_wf h'\ case_prodI element_ptr_kinds_eq_h2 element_ptr_kinds_eq_h3 local.get_shadow_root_ok mem_Collect_eq pair_imageI returns_result_select_result select_result_I2 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, opaque_lifting) NodeMonad.ptr_kinds_ptr_kinds_M \parent_child_rel h = parent_child_rel h2\ children_eq2_h 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, opaque_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 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 (verit) children_eq2_h3 disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 finsertCI funion_finsert_right h' local.CD.a_all_ptrs_in_heap_def local.set_disconnected_nodes_get_disconnected_nodes node_ptr_kinds_eq_h node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h3 object_ptr_kinds_eq_h3 select_result_I2 set_ConsD subsetD subsetI) 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_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, opaque_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>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 (verit) \local.CD.a_all_ptrs_in_heap h\ disc_nodes_document_ptr_h disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disjoint_iff document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 h' l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes local.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(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 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 by (auto simp add: 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 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, opaque_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 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_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 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_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 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 (verit, best) ShadowRootMonad.ptr_kinds_ptr_kinds_M \new_shadow_root_ptr \ set |h \ shadow_root_ptr_kinds_M|\<^sub>r\ \type_wf h'\ assms(1) assms(2) element_ptr_kinds_eq_h3 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_shadow_root_ok local.get_shadow_root_shadow_root_ptr_in_heap 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 returns_result_select_result select_result_I2 shadow_root_eq_h shadow_root_eq_h2 shadow_root_eq_h3 shadow_root_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 (verit, best) \\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_h2 tag_name_eq2_h3 tag_name_eq_h) 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 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 (metis (no_types, opaque_lifting) assms(2) 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_ok local.l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms returns_result_select_result select_result_I2 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 by (smt (verit)) 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,12903 +1,12890 @@ (*********************************************************************************** * 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 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 (z3) assms(1) 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 by (auto simp add: document_ptr_kinds_def delete_shadow_root_pointers[OF h']) 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 document_ptr_kinds_eq_h2 by (auto simp add: 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. 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 subsetD) - by (metis (no_types, lifting) disconnected_nodes_eq2_h document_ptr_kinds_eq_h - 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) + hence "CD.a_all_ptrs_in_heap h2" + by (simp add: children_eq2_h disconnected_nodes_eq2_h document_ptr_kinds_eq_h + 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_eq_h + object_ptr_kinds_eq_h) + hence "CD.a_all_ptrs_in_heap h'" + by (metis (no_types, opaque_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'\ children_eq2_h2 + 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 disconnected_nodes_eq2_h2 document_ptr_kinds_eq2_h2 + fsubsetD h' CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 object_ptr_kinds_eq_h2 + shadow_root_ptr_kinds_commutes) 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 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 (verit) IntI children_eq2_h2 empty_iff finite_fset 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 insert_iff set_sorted_list_of_set 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, opaque_lifting) \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 (z3) \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 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, opaque_lifting)) 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, opaque_lifting) 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, opaque_lifting) 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 (z3) \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 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 (z3) 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 (verit) assms(1) assms(2) bind_pure_returns_result_I2 bind_returns_result_E 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 (verit) 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 image_iff Bex_def split: option.splits)[1] by (metis (mono_tags) 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 (verit) 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 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 image_iff Bex_def split: option.splits)[1] by (metis (mono_tags, lifting) 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 (verit) 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 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 (verit) assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) 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 option.distinct(1) returns_result_eq returns_result_select_result root by (metis (no_types, opaque_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 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, opaque_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 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 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, opaque_lifting) disc_nodes_old_document_h2 disc_nodes_old_document_h3 disconnected_nodes_eq2_h2 document_ptr_kinds_eq3_h2 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 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, opaque_lifting) \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 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) 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, opaque_lifting) disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3 document_ptr_kinds_eq_h2 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, opaque_lifting) contra_subsetD 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' select_result_I2 subsetD) by (metis (no_types, lifting) disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 in_mono) 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 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, opaque_lifting) 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, opaque_lifting) 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 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, opaque_lifting) "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, opaque_lifting) "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 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, opaque_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 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 (verit) 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 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_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 (verit) 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 h' set_disconnected_nodes_get_disconnected_nodes CD.a_all_ptrs_in_heap_def select_result_I2 set_ConsD subsetD) by (smt (verit) 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 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, opaque_lifting) "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, opaque_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 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 (verit) Diff_empty Diff_insert0 ElementMonad.ptr_kinds_M_ptr_kinds ElementMonad.ptr_kinds_ptr_kinds_M assms(1) assms(3) 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 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 (verit) \previous_host \ fset (element_ptr_kinds h2)\ \previous_host \ new_element_ptr\ \type_wf h2\ \type_wf h3\ element_ptr_kinds_eq_h2 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 (z3) \type_wf h3\ 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, opaque_lifting) NodeMonad.ptr_kinds_ptr_kinds_M \parent_child_rel h = parent_child_rel h2\ children_eq2_h 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, opaque_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 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 (verit) 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 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_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, opaque_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>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 (verit) 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 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, opaque_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 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 by (auto simp add: document_ptr_kinds_def) 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, opaque_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 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_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 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 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_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, opaque_lifting) \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 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_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 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\ by (auto simp add: shadow_root_ptr_kinds_eq_h3 shadow_root_ptr_kinds_eq_h2 shadow_root_ptr_kinds_eq_h) 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 (verit) \type_wf h'\ assms(2) 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